Page MenuHomeIsabelle/Phabricator

No OneTemporary

This file is larger than 256 KB, so syntax highlighting was skipped.
diff --git a/thys/Aggregation_Algebras/Aggregation_Algebras.thy b/thys/Aggregation_Algebras/Aggregation_Algebras.thy
--- a/thys/Aggregation_Algebras/Aggregation_Algebras.thy
+++ b/thys/Aggregation_Algebras/Aggregation_Algebras.thy
@@ -1,184 +1,184 @@
(* Title: Algebras for Aggregation and Minimisation
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Algebras for Aggregation and Minimisation\<close>
text \<open>
This theory gives algebras with operations for aggregation and minimisation.
In the weighted-graph model of matrices over (extended) numbers, the operations have the following meaning.
The binary operation $+$ adds the weights of corresponding edges of two graphs.
Addition does not have to be the standard addition on numbers, but can be any aggregation satisfying certain basic properties as demonstrated by various models of the algebras in another theory.
The unary operation \<open>sum\<close> adds the weights of all edges of a graph.
The result is a single aggregated weight using the same aggregation as $+$ but applied internally to the edges of a single graph.
The unary operation \<open>minarc\<close> finds an edge with a minimal weight in a graph.
It yields the position of such an edge as a regular element of a Stone relation algebra.
We give axioms for these operations which are sufficient to prove the correctness of Prim's and Kruskal's minimum spanning tree algorithms.
The operations have been proposed and axiomatised first in \cite{Guttmann2016c} with simplified axioms given in \cite{Guttmann2018a}.
The present version adds two axioms to prove total correctness of the spanning tree algorithms as discussed in \cite{Guttmann2018b}.
\<close>
theory Aggregation_Algebras
imports Stone_Kleene_Relation_Algebras.Kleene_Relation_Algebras
begin
context sup
begin
no_notation
sup (infixl "+" 65)
end
context plus
begin
notation
plus (infixl "+" 65)
end
text \<open>
We first introduce s-algebras as a class with the operations $+$ and \<open>sum\<close>.
Axiom \<open>sum_plus_right_isotone\<close> states that for non-empty graphs, the operation $+$ is $\leq$-isotone in its second argument on the image of the aggregation operation \<open>sum\<close>.
Axiom \<open>sum_bot\<close> expresses that the empty graph contributes no weight.
Axiom \<open>sum_plus\<close> generalises the inclusion-exclusion principle to sets of weights.
Axiom \<open>sum_conv\<close> specifies that reversing edge directions does not change the aggregated weight.
In instances of \<open>s_algebra\<close>, aggregated weights can be partially ordered.
\<close>
class sum =
fixes sum :: "'a \<Rightarrow> 'a"
class s_algebra = stone_relation_algebra + plus + sum +
assumes sum_plus_right_isotone: "x \<noteq> bot \<and> sum x \<le> sum y \<longrightarrow> sum z + sum x \<le> sum z + sum y"
assumes sum_bot: "sum x + sum bot = sum x"
assumes sum_plus: "sum x + sum y = sum (x \<squnion> y) + sum (x \<sqinter> y)"
assumes sum_conv: "sum (x\<^sup>T) = sum x"
begin
lemma sum_disjoint:
assumes "x \<sqinter> y = bot"
shows "sum ((x \<squnion> y) \<sqinter> z) = sum (x \<sqinter> z) + sum (y \<sqinter> z)"
by (subst sum_plus) (metis assms inf.sup_monoid.add_assoc inf.sup_monoid.add_commute inf_bot_left inf_sup_distrib2 sum_bot)
lemma sum_disjoint_3:
assumes "w \<sqinter> x = bot"
and "w \<sqinter> y = bot"
and "x \<sqinter> y = bot"
shows "sum ((w \<squnion> x \<squnion> y) \<sqinter> z) = sum (w \<sqinter> z) + sum (x \<sqinter> z) + sum (y \<sqinter> z)"
by (metis assms inf_sup_distrib2 sup_idem sum_disjoint)
lemma sum_symmetric:
assumes "y = y\<^sup>T"
shows "sum (x\<^sup>T \<sqinter> y) = sum (x \<sqinter> y)"
by (metis assms sum_conv conv_dist_inf)
lemma sum_commute:
"sum x + sum y = sum y + sum x"
by (metis inf_commute sum_plus sup_commute)
end
text \<open>
We next introduce the operation \<open>minarc\<close>.
Axiom \<open>minarc_below\<close> expresses that the result of \<open>minarc\<close> is contained in the graph ignoring the weights.
Axiom \<open>minarc_arc\<close> states that the result of \<open>minarc\<close> is a single unweighted edge if the graph is not empty.
Axiom \<open>minarc_min\<close> specifies that any edge in the graph weighs at least as much as the edge at the position indicated by the result of \<open>minarc\<close>, where weights of edges between different nodes are compared by applying the operation \<open>sum\<close> to single-edge graphs.
Axiom \<open>sum_linear\<close> requires that aggregated weights are linearly ordered, which is necessary for both Prim's and Kruskal's minimum spanning tree algorithms.
Axiom \<open>finite_regular\<close> ensures that there are only finitely many unweighted graphs, and therefore only finitely many edges and nodes in a graph; again this is necessary for the minimum spanning tree algorithms we consider.
\<close>
class minarc =
fixes minarc :: "'a \<Rightarrow> 'a"
class m_algebra = s_algebra + minarc +
assumes minarc_below: "minarc x \<le> --x"
assumes minarc_arc: "x \<noteq> bot \<longrightarrow> arc (minarc x)"
assumes minarc_min: "arc y \<and> y \<sqinter> x \<noteq> bot \<longrightarrow> sum (minarc x \<sqinter> x) \<le> sum (y \<sqinter> x)"
assumes sum_linear: "sum x \<le> sum y \<or> sum y \<le> sum x"
assumes finite_regular: "finite { x . regular x }"
begin
text \<open>
Axioms \<open>minarc_below\<close> and \<open>minarc_arc\<close> suffice to derive the Tarski rule in Stone relation algebras.
\<close>
subclass stone_relation_algebra_tarski
proof unfold_locales
fix x
let ?a = "minarc x"
assume 1: "regular x"
assume "x \<noteq> bot"
hence "arc ?a"
by (simp add: minarc_arc)
hence "top = top * ?a * top"
by (simp add: comp_associative)
also have "... \<le> top * --x * top"
by (simp add: minarc_below mult_isotone)
finally show "top * x * top = top"
- using 1 antisym by simp
+ using 1 order.antisym by simp
qed
lemma minarc_bot:
"minarc bot = bot"
by (metis bot_unique minarc_below regular_closed_bot)
lemma minarc_bot_iff:
"minarc x = bot \<longleftrightarrow> x = bot"
using covector_bot_closed inf_bot_right minarc_arc vector_bot_closed minarc_bot by fastforce
lemma minarc_meet_bot:
assumes "minarc x \<sqinter> x = bot"
shows "minarc x = bot"
proof -
have "minarc x \<le> -x"
using assms pseudo_complement by auto
thus ?thesis
by (metis minarc_below inf_absorb1 inf_import_p inf_p)
qed
lemma minarc_meet_bot_minarc_iff:
"minarc x \<sqinter> x = bot \<longleftrightarrow> minarc x = bot"
using comp_inf.semiring.mult_not_zero minarc_meet_bot by blast
lemma minarc_meet_bot_iff:
"minarc x \<sqinter> x = bot \<longleftrightarrow> x = bot"
using inf_bot_right minarc_bot_iff minarc_meet_bot by blast
lemma minarc_regular:
"regular (minarc x)"
proof (cases "x = bot")
assume "x = bot"
thus ?thesis
by (simp add: minarc_bot)
next
assume "x \<noteq> bot"
thus ?thesis
by (simp add: arc_regular minarc_arc)
qed
lemma minarc_selection:
"selection (minarc x \<sqinter> y) y"
using inf_assoc minarc_regular selection_closed_id by auto
lemma minarc_below_regular:
"regular x \<Longrightarrow> minarc x \<le> x"
by (metis minarc_below)
(*
lemma sum_bot: "sum bot = bot" nitpick [expect=genuine] oops
lemma plus_bot: "x + bot = x" nitpick [expect=genuine] oops
lemma "sum x = bot \<longrightarrow> x = bot" nitpick [expect=genuine] oops
*)
end
class m_kleene_algebra = m_algebra + stone_kleene_relation_algebra
end
diff --git a/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy b/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy
--- a/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy
+++ b/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy
@@ -1,1017 +1,1017 @@
(* Title: Matrix Algebras for Aggregation and Minimisation
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Matrix Algebras for Aggregation and Minimisation\<close>
text \<open>
This theory formalises aggregation orders and lattices as introduced in \cite{Guttmann2018a}.
Aggregation in these algebras is an associative and commutative operation satisfying additional properties related to the partial order and its least element.
We apply the aggregation operation to finite matrices over the aggregation algebras, which shows that they form an s-algebra.
By requiring the aggregation algebras to be linearly ordered, we also obtain that the matrices form an m-algebra.
This is an intermediate step in demonstrating that weighted graphs form an s-algebra and an m-algebra.
The present theory specifies abstract properties for the edge weights and shows that matrices over such weights form an instance of s-algebras and m-algebras.
A second step taken in a separate theory gives concrete instances of edge weights satisfying the abstract properties introduced here.
Lifting the aggregation to matrices requires finite sums over elements taken from commutative semigroups with an element that is a unit on the image of the semigroup operation.
Because standard sums assume a commutative monoid we have to derive a number of properties of these generalised sums as their statements or proofs are different.
\<close>
theory Matrix_Aggregation_Algebras
imports Stone_Kleene_Relation_Algebras.Matrix_Kleene_Algebras Aggregation_Algebras Semigroups_Big
begin
no_notation
inf (infixl "\<sqinter>" 70)
and uminus ("- _" [81] 80)
subsection \<open>Aggregation Orders and Finite Sums\<close>
text \<open>
An aggregation order is a partial order with a least element and an associative commutative operation satisfying certain properties.
Axiom \<open>add_add_bot\<close> introduces almost a commutative monoid; we require that \<open>bot\<close> is a unit only on the image of the aggregation operation.
This is necessary since it is not a unit of a number of aggregation operations we are interested in.
Axiom \<open>add_right_isotone\<close> states that aggregation is $\leq$-isotone on the image of the aggregation operation.
Its assumption $x \neq bot$ is necessary because the introduction of new edges can decrease the aggregated value.
Axiom \<open>add_bot\<close> expresses that aggregation is zero-sum-free.
\<close>
class aggregation_order = order_bot + ab_semigroup_add +
assumes add_right_isotone: "x \<noteq> bot \<and> x + bot \<le> y + bot \<longrightarrow> x + z \<le> y + z"
assumes add_add_bot [simp]: "x + y + bot = x + y"
assumes add_bot: "x + y = bot \<longrightarrow> x = bot"
begin
abbreviation "zero \<equiv> bot + bot"
sublocale aggregation: ab_semigroup_add_0 where plus = plus and zero = zero
apply unfold_locales
using add_assoc add_add_bot by auto
lemma add_bot_bot_bot:
"x + bot + bot + bot = x + bot"
by simp
end
text \<open>
We introduce notation for finite sums over aggregation orders.
The index variable of the summation ranges over the finite universe of its type.
Finite sums are defined recursively using the binary aggregation and $\bot + \bot$ for the empty sum.
\<close>
syntax (xsymbols)
"_sum_ab_semigroup_add_0" :: "idt \<Rightarrow> 'a::bounded_semilattice_sup_bot \<Rightarrow> 'a" ("(\<Sum>\<^sub>_ _)" [0,10] 10)
translations
"\<Sum>\<^sub>x t" => "XCONST ab_semigroup_add_0.sum_0 XCONST plus (XCONST plus XCONST bot XCONST bot) (\<lambda>x . t) { x . CONST True }"
text \<open>
The following are basic properties of such sums.
\<close>
lemma agg_sum_bot:
"(\<Sum>\<^sub>k bot::'a::aggregation_order) = bot + bot"
proof (induct rule: infinite_finite_induct)
case (infinite A)
thus ?case
by simp
next
case empty
thus ?case
by simp
next
case (insert x F)
thus ?case
by (metis add.commute add_add_bot aggregation.sum_0.insert)
qed
lemma agg_sum_bot_bot:
"(\<Sum>\<^sub>k bot + bot::'a::aggregation_order) = bot + bot"
by (rule aggregation.sum_0.neutral_const)
lemma agg_sum_not_bot_1:
fixes f :: "'a::finite \<Rightarrow> 'b::aggregation_order"
assumes "f i \<noteq> bot"
shows "(\<Sum>\<^sub>k f k) \<noteq> bot"
by (metis assms add_bot aggregation.sum_0.remove finite_code mem_Collect_eq)
lemma agg_sum_not_bot:
fixes f :: "('a::finite,'b::aggregation_order) square"
assumes "f (i,j) \<noteq> bot"
shows "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) \<noteq> bot"
by (metis assms agg_sum_not_bot_1)
lemma agg_sum_distrib:
fixes f g :: "'a \<Rightarrow> 'b::aggregation_order"
shows "(\<Sum>\<^sub>k f k + g k) = (\<Sum>\<^sub>k f k) + (\<Sum>\<^sub>k g k)"
by (rule aggregation.sum_0.distrib)
lemma agg_sum_distrib_2:
fixes f g :: "('a,'b::aggregation_order) square"
shows "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) + g (k,l)) = (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
proof -
have "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) + g (k,l)) = (\<Sum>\<^sub>k (\<Sum>\<^sub>l f (k,l)) + (\<Sum>\<^sub>l g (k,l)))"
by (metis (no_types) aggregation.sum_0.distrib)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
by (metis (no_types) aggregation.sum_0.distrib)
finally show ?thesis
.
qed
lemma agg_sum_add_bot:
fixes f :: "'a \<Rightarrow> 'b::aggregation_order"
shows "(\<Sum>\<^sub>k f k) = (\<Sum>\<^sub>k f k) + bot"
by (metis (no_types) add_add_bot aggregation.sum_0.F_one)
lemma agg_sum_add_bot_2:
fixes f :: "'a \<Rightarrow> 'b::aggregation_order"
shows "(\<Sum>\<^sub>k f k + bot) = (\<Sum>\<^sub>k f k)"
proof -
have "(\<Sum>\<^sub>k f k + bot) = (\<Sum>\<^sub>k f k) + (\<Sum>\<^sub>k::'a bot::'b)"
using agg_sum_distrib by simp
also have "... = (\<Sum>\<^sub>k f k) + (bot + bot)"
by (metis agg_sum_bot)
also have "... = (\<Sum>\<^sub>k f k)"
by simp
finally show ?thesis
by simp
qed
lemma agg_sum_commute:
fixes f :: "('a,'b::aggregation_order) square"
shows "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) = (\<Sum>\<^sub>l \<Sum>\<^sub>k f (k,l))"
by (rule aggregation.sum_0.swap)
lemma agg_delta:
fixes f :: "'a::finite \<Rightarrow> 'b::aggregation_order"
shows "(\<Sum>\<^sub>l if l = j then f l else zero) = f j + bot"
apply (subst aggregation.sum_0.delta)
apply simp
by (metis add.commute add.left_commute add_add_bot mem_Collect_eq)
lemma agg_delta_1:
fixes f :: "'a::finite \<Rightarrow> 'b::aggregation_order"
shows "(\<Sum>\<^sub>l if l = j then f l else bot) = f j + bot"
proof -
let ?f = "(\<lambda>l . if l = j then f l else bot)"
let ?S = "{l::'a . True}"
show ?thesis
proof (cases "j \<in> ?S")
case False
thus ?thesis by simp
next
case True
let ?A = "?S - {j}"
let ?B = "{j}"
from True have eq: "?S = ?A \<union> ?B"
by blast
have dj: "?A \<inter> ?B = {}"
by simp
have fAB: "finite ?A" "finite ?B"
by auto
have "aggregation.sum_0 ?f ?S = aggregation.sum_0 ?f ?A + aggregation.sum_0 ?f ?B"
using aggregation.sum_0.union_disjoint[OF fAB dj, of ?f, unfolded eq [symmetric]] by simp
also have "... = aggregation.sum_0 (\<lambda>l . bot) ?A + aggregation.sum_0 ?f ?B"
by (subst aggregation.sum_0.cong[where ?B="?A"]) simp_all
also have "... = zero + aggregation.sum_0 ?f ?B"
by (metis (no_types, lifting) add.commute add_add_bot aggregation.sum_0.F_g_one aggregation.sum_0.neutral)
also have "... = zero + (f j + zero)"
by simp
also have "... = f j + bot"
by (metis add.commute add.left_commute add_add_bot)
finally show ?thesis
.
qed
qed
lemma agg_delta_2:
fixes f :: "('a::finite,'b::aggregation_order) square"
shows "(\<Sum>\<^sub>k \<Sum>\<^sub>l if k = i \<and> l = j then f (k,l) else bot) = f (i,j) + bot"
proof -
have "\<forall>k . (\<Sum>\<^sub>l if k = i \<and> l = j then f (k,l) else bot) = (if k = i then f (k,j) + bot else zero)"
proof
fix k
have "(\<Sum>\<^sub>l if k = i \<and> l = j then f (k,l) else bot) = (\<Sum>\<^sub>l if l = j then if k = i then f (k,l) else bot else bot)"
by meson
also have "... = (if k = i then f (k,j) else bot) + bot"
by (rule agg_delta_1)
finally show "(\<Sum>\<^sub>l if k = i \<and> l = j then f (k,l) else bot) = (if k = i then f (k,j) + bot else zero)"
by simp
qed
hence "(\<Sum>\<^sub>k \<Sum>\<^sub>l if k = i \<and> l = j then f (k,l) else bot) = (\<Sum>\<^sub>k if k = i then f (k,j) + bot else zero)"
using aggregation.sum_0.cong by auto
also have "... = f (i,j) + bot"
apply (subst agg_delta)
by simp
finally show ?thesis
.
qed
subsection \<open>Matrix Aggregation\<close>
text \<open>
The following definitions introduce the matrix of unit elements, componentwise aggregation and aggregation on matrices.
The aggregation of a matrix is a single value, but because s-algebras are single-sorted the result has to be encoded as a matrix of the same type (size) as the input.
We store the aggregated matrix value in the `first' entry of a matrix, setting all other entries to the unit value.
The first entry is determined by requiring an enumeration of indices.
It does not have to be the first entry; any fixed location in the matrix would work as well.
\<close>
definition zero_matrix :: "('a,'b::{plus,bot}) square" ("mzero") where "mzero = (\<lambda>e . bot + bot)"
definition plus_matrix :: "('a,'b::plus) square \<Rightarrow> ('a,'b) square \<Rightarrow> ('a,'b) square" (infixl "\<oplus>\<^sub>M" 65) where "plus_matrix f g = (\<lambda>e . f e + g e)"
definition sum_matrix :: "('a::enum,'b::{plus,bot}) square \<Rightarrow> ('a,'b) square" ("sum\<^sub>M _" [80] 80) where "sum_matrix f = (\<lambda>(i,j) . if i = hd enum_class.enum \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) else bot + bot)"
text \<open>
Basic properties of these operations are given in the following.
\<close>
lemma bot_plus_bot:
"mbot \<oplus>\<^sub>M mbot = mzero"
by (simp add: plus_matrix_def bot_matrix_def zero_matrix_def)
lemma sum_bot:
"sum\<^sub>M (mbot :: ('a::enum,'b::aggregation_order) square) = mzero"
proof (rule ext, rule prod_cases)
fix i j :: "'a"
have "(sum\<^sub>M mbot :: ('a,'b) square) (i,j) = (if i = hd enum_class.enum \<and> j = i then \<Sum>\<^sub>(k::'a) \<Sum>\<^sub>(l::'a) bot else bot + bot)"
by (unfold sum_matrix_def bot_matrix_def) simp
also have "... = bot + bot"
using agg_sum_bot aggregation.sum_0.neutral by fastforce
also have "... = mzero (i,j)"
by (simp add: zero_matrix_def)
finally show "(sum\<^sub>M mbot :: ('a,'b) square) (i,j) = mzero (i,j)"
.
qed
lemma sum_plus_bot:
fixes f :: "('a::enum,'b::aggregation_order) square"
shows "sum\<^sub>M f \<oplus>\<^sub>M mbot = sum\<^sub>M f"
proof (rule ext, rule prod_cases)
let ?h = "hd enum_class.enum"
fix i j
have "(sum\<^sub>M f \<oplus>\<^sub>M mbot) (i,j) = (if i = ?h \<and> j = i then (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + bot else zero + bot)"
by (simp add: plus_matrix_def bot_matrix_def sum_matrix_def)
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) else zero)"
by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one)
also have "... = (sum\<^sub>M f) (i,j)"
by (simp add: sum_matrix_def)
finally show "(sum\<^sub>M f \<oplus>\<^sub>M mbot) (i,j) = (sum\<^sub>M f) (i,j)"
by simp
qed
lemma sum_plus_zero:
fixes f :: "('a::enum,'b::aggregation_order) square"
shows "sum\<^sub>M f \<oplus>\<^sub>M mzero = sum\<^sub>M f"
by (rule ext, rule prod_cases) (simp add: plus_matrix_def zero_matrix_def sum_matrix_def)
lemma agg_matrix_bot:
fixes f :: "('a,'b::aggregation_order) square"
assumes "\<forall>i j . f (i,j) = bot"
shows "f = mbot"
apply (unfold bot_matrix_def)
using assms by auto
text \<open>
We consider a different implementation of matrix aggregation which stores the aggregated value in all entries of the matrix instead of a particular one.
This does not require an enumeration of the indices.
All results continue to hold using this alternative implementation.
\<close>
definition sum_matrix_2 :: "('a,'b::{plus,bot}) square \<Rightarrow> ('a,'b) square" ("sum2\<^sub>M _" [80] 80) where "sum_matrix_2 f = (\<lambda>e . \<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l))"
lemma sum_bot_2:
"sum2\<^sub>M (mbot :: ('a,'b::aggregation_order) square) = mzero"
proof
fix e
have "(sum2\<^sub>M mbot :: ('a,'b) square) e = (\<Sum>\<^sub>(k::'a) \<Sum>\<^sub>(l::'a) bot)"
by (unfold sum_matrix_2_def bot_matrix_def) simp
also have "... = bot + bot"
using agg_sum_bot aggregation.sum_0.neutral by fastforce
also have "... = mzero e"
by (simp add: zero_matrix_def)
finally show "(sum2\<^sub>M mbot :: ('a,'b) square) e = mzero e"
.
qed
lemma sum_plus_bot_2:
fixes f :: "('a,'b::aggregation_order) square"
shows "sum2\<^sub>M f \<oplus>\<^sub>M mbot = sum2\<^sub>M f"
proof
fix e
have "(sum2\<^sub>M f \<oplus>\<^sub>M mbot) e = (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + bot"
by (simp add: plus_matrix_def bot_matrix_def sum_matrix_2_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l))"
by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one)
also have "... = (sum2\<^sub>M f) e"
by (simp add: sum_matrix_2_def)
finally show "(sum2\<^sub>M f \<oplus>\<^sub>M mbot) e = (sum2\<^sub>M f) e"
by simp
qed
lemma sum_plus_zero_2:
fixes f :: "('a,'b::aggregation_order) square"
shows "sum2\<^sub>M f \<oplus>\<^sub>M mzero = sum2\<^sub>M f"
by (simp add: plus_matrix_def zero_matrix_def sum_matrix_2_def)
subsection \<open>Aggregation Lattices\<close>
text \<open>
We extend aggregation orders to dense bounded distributive lattices.
Axiom \<open>add_lattice\<close> implements the inclusion-exclusion principle at the level of edge weights.
\<close>
class aggregation_lattice = bounded_distrib_lattice + dense_lattice + aggregation_order +
assumes add_lattice: "x + y = (x \<squnion> y) + (x \<sqinter> y)"
text \<open>
Aggregation lattices form a Stone relation algebra by reusing the meet operation as composition, using identity as converse and a standard implementation of pseudocomplement.
\<close>
class aggregation_algebra = aggregation_lattice + uminus + one + times + conv +
assumes uminus_def [simp]: "-x = (if x = bot then top else bot)"
assumes one_def [simp]: "1 = top"
assumes times_def [simp]: "x * y = x \<sqinter> y"
assumes conv_def [simp]: "x\<^sup>T = x"
begin
subclass stone_algebra
apply unfold_locales
using bot_meet_irreducible bot_unique by auto
subclass stone_relation_algebra
apply unfold_locales
prefer 9 using bot_meet_irreducible apply auto[1]
by (simp_all add: inf.assoc le_infI2 inf_sup_distrib1 inf_sup_distrib2 inf.commute inf.left_commute)
end
text \<open>
We show that matrices over aggregation lattices form an s-algebra using the above operations.
\<close>
interpretation agg_square_s_algebra: s_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix
proof
fix f g h :: "('a,'b) square"
show "f \<noteq> mbot \<and> sum\<^sub>M f \<preceq> sum\<^sub>M g \<longrightarrow> h \<oplus>\<^sub>M sum\<^sub>M f \<preceq> h \<oplus>\<^sub>M sum\<^sub>M g"
proof
let ?h = "hd enum_class.enum"
assume 1: "f \<noteq> mbot \<and> sum\<^sub>M f \<preceq> sum\<^sub>M g"
hence "\<exists>k l . f (k,l) \<noteq> bot"
by (meson agg_matrix_bot)
hence 2: "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) \<noteq> bot"
using agg_sum_not_bot by blast
have "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) = (sum\<^sub>M f) (?h,?h)"
by (simp add: sum_matrix_def)
also have "... \<le> (sum\<^sub>M g) (?h,?h)"
using 1 by (simp add: less_eq_matrix_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
by (simp add: sum_matrix_def)
finally have "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) \<le> (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
by simp
hence 3: "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + bot \<le> (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l)) + bot"
by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one)
show "h \<oplus>\<^sub>M sum\<^sub>M f \<preceq> h \<oplus>\<^sub>M sum\<^sub>M g"
proof (unfold less_eq_matrix_def, rule allI, rule prod_cases, unfold plus_matrix_def)
fix i j
have 4: "h (i,j) + (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) \<le> h (i,j) + (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
using 2 3 by (metis (no_types, lifting) add_right_isotone add.commute)
have "h (i,j) + (sum\<^sub>M f) (i,j) = h (i,j) + (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) else zero)"
by (simp add: sum_matrix_def)
also have "... = (if i = ?h \<and> j = i then h (i,j) + (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) else h (i,j) + zero)"
by simp
also have "... \<le> (if i = ?h \<and> j = i then h (i,j) + (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l)) else h (i,j) + zero)"
- using 4 inf.eq_iff by auto
+ using 4 order.eq_iff by auto
also have "... = h (i,j) + (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l) else zero)"
by simp
finally show "h (i,j) + (sum\<^sub>M f) (i,j) \<le> h (i,j) + (sum\<^sub>M g) (i,j)"
by (simp add: sum_matrix_def)
qed
qed
next
fix f :: "('a,'b) square"
show "sum\<^sub>M f \<oplus>\<^sub>M sum\<^sub>M mbot = sum\<^sub>M f"
by (simp add: sum_bot sum_plus_zero)
next
fix f g :: "('a,'b) square"
show "sum\<^sub>M f \<oplus>\<^sub>M sum\<^sub>M g = sum\<^sub>M (f \<oplus> g) \<oplus>\<^sub>M sum\<^sub>M (f \<otimes> g)"
proof (rule ext, rule prod_cases)
fix i j
let ?h = "hd enum_class.enum"
have "(sum\<^sub>M f \<oplus>\<^sub>M sum\<^sub>M g) (i,j) = (sum\<^sub>M f) (i,j) + (sum\<^sub>M g) (i,j)"
by (simp add: plus_matrix_def)
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) else zero) + (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l) else zero)"
by (simp add: sum_matrix_def)
also have "... = (if i = ?h \<and> j = i then (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l)) else zero)"
by simp
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) + g (k,l) else zero)"
using agg_sum_distrib_2 by (metis (no_types))
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l (f (k,l) \<squnion> g (k,l)) + (f (k,l) \<sqinter> g (k,l)) else zero)"
using add_lattice aggregation.sum_0.cong by (metis (no_types, lifting))
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l (f \<oplus> g) (k,l) + (f \<otimes> g) (k,l) else zero)"
by (simp add: sup_matrix_def inf_matrix_def)
also have "... = (if i = ?h \<and> j = i then (\<Sum>\<^sub>k \<Sum>\<^sub>l (f \<oplus> g) (k,l)) + (\<Sum>\<^sub>k \<Sum>\<^sub>l (f \<otimes> g) (k,l)) else zero)"
using agg_sum_distrib_2 by (metis (no_types))
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l (f \<oplus> g) (k,l) else zero) + (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l (f \<otimes> g) (k,l) else zero)"
by simp
also have "... = (sum\<^sub>M (f \<oplus> g)) (i,j) + (sum\<^sub>M (f \<otimes> g)) (i,j)"
by (simp add: sum_matrix_def)
also have "... = (sum\<^sub>M (f \<oplus> g) \<oplus>\<^sub>M sum\<^sub>M (f \<otimes> g)) (i,j)"
by (simp add: plus_matrix_def)
finally show "(sum\<^sub>M f \<oplus>\<^sub>M sum\<^sub>M g) (i,j) = (sum\<^sub>M (f \<oplus> g) \<oplus>\<^sub>M sum\<^sub>M (f \<otimes> g)) (i,j)"
.
qed
next
fix f :: "('a,'b) square"
show "sum\<^sub>M (f\<^sup>t) = sum\<^sub>M f"
proof (rule ext, rule prod_cases)
fix i j
let ?h = "hd enum_class.enum"
have "(sum\<^sub>M (f\<^sup>t)) (i,j) = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l (f\<^sup>t) (k,l) else zero)"
by (simp add: sum_matrix_def)
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l (f (l,k))\<^sup>T else zero)"
by (simp add: conv_matrix_def)
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>k \<Sum>\<^sub>l f (l,k) else zero)"
by simp
also have "... = (if i = ?h \<and> j = i then \<Sum>\<^sub>l \<Sum>\<^sub>k f (l,k) else zero)"
by (metis agg_sum_commute)
also have "... = (sum\<^sub>M f) (i,j)"
by (simp add: sum_matrix_def)
finally show "(sum\<^sub>M (f\<^sup>t)) (i,j) = (sum\<^sub>M f) (i,j)"
.
qed
qed
text \<open>
We show the same for the alternative implementation that stores the result of aggregation in all elements of the matrix.
\<close>
interpretation agg_square_s_algebra_2: s_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::finite,'b::aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix_2
proof
fix f g h :: "('a,'b) square"
show "f \<noteq> mbot \<and> sum2\<^sub>M f \<preceq> sum2\<^sub>M g \<longrightarrow> h \<oplus>\<^sub>M sum2\<^sub>M f \<preceq> h \<oplus>\<^sub>M sum2\<^sub>M g"
proof
assume 1: "f \<noteq> mbot \<and> sum2\<^sub>M f \<preceq> sum2\<^sub>M g"
hence "\<exists>k l . f (k,l) \<noteq> bot"
by (meson agg_matrix_bot)
hence 2: "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) \<noteq> bot"
using agg_sum_not_bot by blast
obtain c :: 'a where True
by simp
have "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) = (sum2\<^sub>M f) (c,c)"
by (simp add: sum_matrix_2_def)
also have "... \<le> (sum2\<^sub>M g) (c,c)"
using 1 by (simp add: less_eq_matrix_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
by (simp add: sum_matrix_2_def)
finally have "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) \<le> (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
by simp
hence 3: "(\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + bot \<le> (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l)) + bot"
by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one)
show "h \<oplus>\<^sub>M sum2\<^sub>M f \<preceq> h \<oplus>\<^sub>M sum2\<^sub>M g"
proof (unfold less_eq_matrix_def, rule allI, unfold plus_matrix_def)
fix e
have "h e + (sum2\<^sub>M f) e = h e + (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l))"
by (simp add: sum_matrix_2_def)
also have "... \<le> h e + (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
using 2 3 by (metis (no_types, lifting) add_right_isotone add.commute)
finally show "h e + (sum2\<^sub>M f) e \<le> h e + (sum2\<^sub>M g) e"
by (simp add: sum_matrix_2_def)
qed
qed
next
fix f :: "('a,'b) square"
show "sum2\<^sub>M f \<oplus>\<^sub>M sum2\<^sub>M mbot = sum2\<^sub>M f"
by (simp add: sum_bot_2 sum_plus_zero_2)
next
fix f g :: "('a,'b) square"
show "sum2\<^sub>M f \<oplus>\<^sub>M sum2\<^sub>M g = sum2\<^sub>M (f \<oplus> g) \<oplus>\<^sub>M sum2\<^sub>M (f \<otimes> g)"
proof
fix e
have "(sum2\<^sub>M f \<oplus>\<^sub>M sum2\<^sub>M g) e = (sum2\<^sub>M f) e + (sum2\<^sub>M g) e"
by (simp add: plus_matrix_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l)) + (\<Sum>\<^sub>k \<Sum>\<^sub>l g (k,l))"
by (simp add: sum_matrix_2_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l f (k,l) + g (k,l))"
using agg_sum_distrib_2 by (metis (no_types))
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l (f (k,l) \<squnion> g (k,l)) + (f (k,l) \<sqinter> g (k,l)))"
using add_lattice aggregation.sum_0.cong by (metis (no_types, lifting))
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l (f \<oplus> g) (k,l) + (f \<otimes> g) (k,l))"
by (simp add: sup_matrix_def inf_matrix_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l (f \<oplus> g) (k,l)) + (\<Sum>\<^sub>k \<Sum>\<^sub>l (f \<otimes> g) (k,l))"
using agg_sum_distrib_2 by (metis (no_types))
also have "... = (sum2\<^sub>M (f \<oplus> g)) e + (sum2\<^sub>M (f \<otimes> g)) e"
by (simp add: sum_matrix_2_def)
also have "... = (sum2\<^sub>M (f \<oplus> g) \<oplus>\<^sub>M sum2\<^sub>M (f \<otimes> g)) e"
by (simp add: plus_matrix_def)
finally show "(sum2\<^sub>M f \<oplus>\<^sub>M sum2\<^sub>M g) e = (sum2\<^sub>M (f \<oplus> g) \<oplus>\<^sub>M sum2\<^sub>M (f \<otimes> g)) e"
.
qed
next
fix f :: "('a,'b) square"
show "sum2\<^sub>M (f\<^sup>t) = sum2\<^sub>M f"
proof
fix e
have "(sum2\<^sub>M (f\<^sup>t)) e = (\<Sum>\<^sub>k \<Sum>\<^sub>l (f\<^sup>t) (k,l))"
by (simp add: sum_matrix_2_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l (f (l,k))\<^sup>T)"
by (simp add: conv_matrix_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l f (l,k))"
by simp
also have "... = (\<Sum>\<^sub>l \<Sum>\<^sub>k f (l,k))"
by (metis agg_sum_commute)
also have "... = (sum2\<^sub>M f) e"
by (simp add: sum_matrix_2_def)
finally show "(sum2\<^sub>M (f\<^sup>t)) e = (sum2\<^sub>M f) e"
.
qed
qed
subsection \<open>Matrix Minimisation\<close>
text \<open>
We construct an operation that finds the minimum entry of a matrix.
Because a matrix can have several entries with the same minimum value, we introduce a lexicographic order on the indices to make the operation deterministic.
The order is obtained by enumerating the universe of the index.
\<close>
primrec enum_pos' :: "'a list \<Rightarrow> 'a::enum \<Rightarrow> nat" where
"enum_pos' Nil x = 0"
| "enum_pos' (y#ys) x = (if x = y then 0 else 1 + enum_pos' ys x)"
lemma enum_pos'_inverse:
"List.member xs x \<Longrightarrow> xs!(enum_pos' xs x) = x"
apply (induct xs)
apply (simp add: member_rec(2))
by (metis diff_add_inverse enum_pos'.simps(2) less_one member_rec(1) not_add_less1 nth_Cons')
text \<open>
The following function finds the position of an index in the enumerated universe.
\<close>
fun enum_pos :: "'a::enum \<Rightarrow> nat" where "enum_pos x = enum_pos' (enum_class.enum::'a list) x"
lemma enum_pos_inverse [simp]:
"enum_class.enum!(enum_pos x) = x"
apply (unfold enum_pos.simps)
apply (rule enum_pos'_inverse)
by (metis in_enum List.member_def)
lemma enum_pos_injective [simp]:
"enum_pos x = enum_pos y \<Longrightarrow> x = y"
by (metis enum_pos_inverse)
text \<open>
The position in the enumerated universe determines the order.
\<close>
abbreviation enum_pos_less_eq :: "'a::enum \<Rightarrow> 'a \<Rightarrow> bool" where "enum_pos_less_eq x y \<equiv> (enum_pos x \<le> enum_pos y)"
abbreviation enum_pos_less :: "'a::enum \<Rightarrow> 'a \<Rightarrow> bool" where "enum_pos_less x y \<equiv> (enum_pos x < enum_pos y)"
sublocale enum < enum_order: order where less_eq = "\<lambda>x y . enum_pos_less_eq x y" and less = "\<lambda>x y . enum_pos x < enum_pos y"
apply unfold_locales
by auto
text \<open>
Based on this, a lexicographic order is defined on pairs, which represent locations in a matrix.
\<close>
abbreviation enum_lex_less :: "'a::enum \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> bool" where "enum_lex_less \<equiv> (\<lambda>(i,j) (k,l) . enum_pos_less i k \<or> (i = k \<and> enum_pos_less j l))"
abbreviation enum_lex_less_eq :: "'a::enum \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> bool" where "enum_lex_less_eq \<equiv> (\<lambda>(i,j) (k,l) . enum_pos_less i k \<or> (i = k \<and> enum_pos_less_eq j l))"
text \<open>
The $m$-operation determines the location of the non-$\bot$ minimum element which is first in the lexicographic order.
The result is returned as a regular matrix with $\top$ at that location and $\bot$ everywhere else.
In the weighted-graph model, this represents a single unweighted edge of the graph.
\<close>
definition minarc_matrix :: "('a::enum,'b::{bot,ord,plus,top}) square \<Rightarrow> ('a,'b) square" ("minarc\<^sub>M _" [80] 80) where "minarc_matrix f = (\<lambda>e . if f e \<noteq> bot \<and> (\<forall>d . (f d \<noteq> bot \<longrightarrow> (f e + bot \<le> f d + bot \<and> (enum_lex_less d e \<longrightarrow> f e + bot \<noteq> f d + bot)))) then top else bot)"
lemma minarc_at_most_one:
fixes f :: "('a::enum,'b::{aggregation_order,top}) square"
assumes "(minarc\<^sub>M f) e \<noteq> bot"
and "(minarc\<^sub>M f) d \<noteq> bot"
shows "e = d"
proof -
have 1: "f e + bot \<le> f d + bot"
by (metis assms minarc_matrix_def)
have "f d + bot \<le> f e + bot"
by (metis assms minarc_matrix_def)
hence "f e + bot = f d + bot"
using 1 by simp
hence "\<not> enum_lex_less d e \<and> \<not> enum_lex_less e d"
using assms by (unfold minarc_matrix_def) (metis (lifting))
thus ?thesis
using enum_pos_injective less_linear by auto
qed
subsection \<open>Linear Aggregation Lattices\<close>
text \<open>
We now assume that the aggregation order is linear and forms a bounded lattice.
It follows that these structures are aggregation lattices.
A linear order on matrix entries is necessary to obtain a unique minimum entry.
\<close>
class linear_aggregation_lattice = linear_bounded_lattice + aggregation_order
begin
subclass aggregation_lattice
apply unfold_locales
by (metis add_commute sup_inf_selective)
sublocale heyting: bounded_heyting_lattice where implies = "\<lambda>x y . if x \<le> y then top else y"
apply unfold_locales
by (simp add: inf_less_eq)
end
text \<open>
Every non-empty set with a transitive total relation has a least element with respect to this relation.
\<close>
lemma least_order:
assumes transitive: "\<forall>x y z . le x y \<and> le y z \<longrightarrow> le x z"
and total: "\<forall>x y . le x y \<or> le y x"
shows "finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> \<exists>x . x \<in> A \<and> (\<forall>y . y \<in> A \<longrightarrow> le x y)"
proof (induct A rule: finite_ne_induct)
case singleton
thus ?case
using total by auto
next
case insert
thus ?case
by (metis insert_iff transitive total)
qed
lemma minarc_at_least_one:
fixes f :: "('a::enum,'b::linear_aggregation_lattice) square"
assumes "f \<noteq> mbot"
shows "\<exists>e . (minarc\<^sub>M f) e = top"
proof -
let ?nbe = "{ (e,f e) | e . f e \<noteq> bot }"
have 1: "finite ?nbe"
using finite_code finite_image_set by blast
have 2: "?nbe \<noteq> {}"
using assms agg_matrix_bot by fastforce
let ?le = "\<lambda>(e::'a \<times> 'a,fe::'b) (d::'a \<times> 'a,fd) . fe + bot \<le> fd + bot"
have 3: "\<forall>x y z . ?le x y \<and> ?le y z \<longrightarrow> ?le x z"
by auto
have 4: "\<forall>x y . ?le x y \<or> ?le y x"
by (simp add: linear)
have "\<exists>x . x \<in> ?nbe \<and> (\<forall>y . y \<in> ?nbe \<longrightarrow> ?le x y)"
by (rule least_order, rule 3, rule 4, rule 1, rule 2)
then obtain e fe where 5: "(e,fe) \<in> ?nbe \<and> (\<forall>y . y \<in> ?nbe \<longrightarrow> ?le (e,fe) y)"
by auto
let ?me = "{ e . f e \<noteq> bot \<and> f e + bot = fe + bot }"
have 6: "finite ?me"
using finite_code finite_image_set by blast
have 7: "?me \<noteq> {}"
using 5 by auto
have 8: "\<forall>x y z . enum_lex_less_eq x y \<and> enum_lex_less_eq y z \<longrightarrow> enum_lex_less_eq x z"
by auto
have 9: "\<forall>x y . enum_lex_less_eq x y \<or> enum_lex_less_eq y x"
by auto
have "\<exists>x . x \<in> ?me \<and> (\<forall>y . y \<in> ?me \<longrightarrow> enum_lex_less_eq x y)"
by (rule least_order, rule 8, rule 9, rule 6, rule 7)
then obtain m where 10: "m \<in> ?me \<and> (\<forall>y . y \<in> ?me \<longrightarrow> enum_lex_less_eq m y)"
by auto
have 11: "f m \<noteq> bot"
using 10 5 by auto
have 12: "\<forall>d. f d \<noteq> bot \<longrightarrow> f m + bot \<le> f d + bot"
using 10 5 by simp
have "\<forall>d. f d \<noteq> bot \<and> enum_lex_less d m \<longrightarrow> f m + bot \<noteq> f d + bot"
using 10 by fastforce
hence "(minarc\<^sub>M f) m = top"
using 11 12 by (simp add: minarc_matrix_def)
thus ?thesis
by blast
qed
text \<open>
Linear aggregation lattices form a Stone relation algebra by reusing the meet operation as composition, using identity as converse and a standard implementation of pseudocomplement.
\<close>
class linear_aggregation_algebra = linear_aggregation_lattice + uminus + one + times + conv +
assumes uminus_def_2 [simp]: "-x = (if x = bot then top else bot)"
assumes one_def_2 [simp]: "1 = top"
assumes times_def_2 [simp]: "x * y = x \<sqinter> y"
assumes conv_def_2 [simp]: "x\<^sup>T = x"
begin
subclass aggregation_algebra
apply unfold_locales
using inf_dense by auto
lemma regular_bot_top_2:
"regular x \<longleftrightarrow> x = bot \<or> x = top"
by simp
sublocale heyting: heyting_stone_algebra where implies = "\<lambda>x y . if x \<le> y then top else y"
apply unfold_locales
- apply (simp add: antisym)
+ apply (simp add: order.antisym)
by auto
end
text \<open>
We show that matrices over linear aggregation lattices form an m-algebra using the above operations.
\<close>
interpretation agg_square_m_algebra: m_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix and minarc = minarc_matrix
proof
fix f :: "('a,'b) square"
show "minarc\<^sub>M f \<preceq> \<ominus>\<ominus>f"
proof (unfold less_eq_matrix_def, rule allI)
fix e :: "'a \<times> 'a"
have "(minarc\<^sub>M f) e \<le> (if f e \<noteq> bot then top else --(f e))"
by (simp add: minarc_matrix_def)
also have "... = --(f e)"
by simp
also have "... = (\<ominus>\<ominus>f) e"
by (simp add: uminus_matrix_def)
finally show "(minarc\<^sub>M f) e \<le> (\<ominus>\<ominus>f) e"
.
qed
next
fix f :: "('a,'b) square"
let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix"
show "f \<noteq> mbot \<longrightarrow> ?at (minarc\<^sub>M f)"
proof
assume 1: "f \<noteq> mbot"
have "minarc\<^sub>M f \<odot> mtop \<odot> (minarc\<^sub>M f \<odot> mtop)\<^sup>t = minarc\<^sub>M f \<odot> mtop \<odot> (minarc\<^sub>M f)\<^sup>t"
by (metis matrix_bounded_idempotent_semiring.surjective_top_closed matrix_monoid.mult_assoc matrix_stone_relation_algebra.conv_dist_comp matrix_stone_relation_algebra.conv_top)
also have "... \<preceq> mone"
proof (unfold less_eq_matrix_def, rule allI, rule prod_cases)
fix i j
have "(minarc\<^sub>M f \<odot> mtop \<odot> (minarc\<^sub>M f)\<^sup>t) (i,j) = (\<Squnion>\<^sub>l (\<Squnion>\<^sub>k (minarc\<^sub>M f) (i,k) * mtop (k,l)) * ((minarc\<^sub>M f)\<^sup>t) (l,j))"
by (simp add: times_matrix_def)
also have "... = (\<Squnion>\<^sub>l (\<Squnion>\<^sub>k (minarc\<^sub>M f) (i,k) * top) * ((minarc\<^sub>M f) (j,l))\<^sup>T)"
by (simp add: top_matrix_def conv_matrix_def)
also have "... = (\<Squnion>\<^sub>l \<Squnion>\<^sub>k (minarc\<^sub>M f) (i,k) * top * ((minarc\<^sub>M f) (j,l))\<^sup>T)"
by (metis comp_right_dist_sum)
also have "... = (\<Squnion>\<^sub>l \<Squnion>\<^sub>k if i = j \<and> l = k then (minarc\<^sub>M f) (i,k) * top * ((minarc\<^sub>M f) (j,l))\<^sup>T else bot)"
apply (rule sup_monoid.sum.cong)
apply simp
by (metis (no_types, lifting) comp_left_zero comp_right_zero conv_bot prod.inject minarc_at_most_one)
also have "... = (if i = j then (\<Squnion>\<^sub>l \<Squnion>\<^sub>k if l = k then (minarc\<^sub>M f) (i,k) * top * ((minarc\<^sub>M f) (j,l))\<^sup>T else bot) else bot)"
by auto
also have "... \<le> (if i = j then top else bot)"
by simp
also have "... = mone (i,j)"
by (simp add: one_matrix_def)
finally show "(minarc\<^sub>M f \<odot> mtop \<odot> (minarc\<^sub>M f)\<^sup>t) (i,j) \<le> mone (i,j)"
.
qed
finally have 2: "minarc\<^sub>M f \<odot> mtop \<odot> (minarc\<^sub>M f \<odot> mtop)\<^sup>t \<preceq> mone"
.
have 3: "mtop \<odot> (minarc\<^sub>M f \<odot> mtop) = mtop"
proof (rule ext, rule prod_cases)
fix i j
from minarc_at_least_one obtain ei ej where "(minarc\<^sub>M f) (ei,ej) = top"
using 1 by force
hence 4: "top * top \<le> (\<Squnion>\<^sub>l (minarc\<^sub>M f) (ei,l) * top)"
by (metis comp_inf.ub_sum)
have "top * (\<Squnion>\<^sub>l (minarc\<^sub>M f) (ei,l) * top) \<le> (\<Squnion>\<^sub>k top * (\<Squnion>\<^sub>l (minarc\<^sub>M f) (k,l) * top))"
by (rule comp_inf.ub_sum)
hence "top \<le> (\<Squnion>\<^sub>k top * (\<Squnion>\<^sub>l (minarc\<^sub>M f) (k,l) * top))"
using 4 by auto
also have "... = (\<Squnion>\<^sub>k mtop (i,k) * (\<Squnion>\<^sub>l (minarc\<^sub>M f) (k,l) * mtop (l,j)))"
by (simp add: top_matrix_def)
also have "... = (mtop \<odot> (minarc\<^sub>M f \<odot> mtop)) (i,j)"
by (simp add: times_matrix_def)
finally show "(mtop \<odot> (minarc\<^sub>M f \<odot> mtop)) (i,j) = mtop (i,j)"
by (simp add: eq_iff top_matrix_def)
qed
have "(minarc\<^sub>M f)\<^sup>t \<odot> mtop \<odot> ((minarc\<^sub>M f)\<^sup>t \<odot> mtop)\<^sup>t = (minarc\<^sub>M f)\<^sup>t \<odot> mtop \<odot> (minarc\<^sub>M f)"
by (metis matrix_stone_relation_algebra.comp_associative matrix_stone_relation_algebra.conv_dist_comp matrix_stone_relation_algebra.conv_involutive matrix_stone_relation_algebra.conv_top matrix_bounded_idempotent_semiring.surjective_top_closed)
also have "... \<preceq> mone"
proof (unfold less_eq_matrix_def, rule allI, rule prod_cases)
fix i j
have "((minarc\<^sub>M f)\<^sup>t \<odot> mtop \<odot> minarc\<^sub>M f) (i,j) = (\<Squnion>\<^sub>l (\<Squnion>\<^sub>k ((minarc\<^sub>M f)\<^sup>t) (i,k) * mtop (k,l)) * (minarc\<^sub>M f) (l,j))"
by (simp add: times_matrix_def)
also have "... = (\<Squnion>\<^sub>l (\<Squnion>\<^sub>k ((minarc\<^sub>M f) (k,i))\<^sup>T * top) * (minarc\<^sub>M f) (l,j))"
by (simp add: top_matrix_def conv_matrix_def)
also have "... = (\<Squnion>\<^sub>l \<Squnion>\<^sub>k ((minarc\<^sub>M f) (k,i))\<^sup>T * top * (minarc\<^sub>M f) (l,j))"
by (metis comp_right_dist_sum)
also have "... = (\<Squnion>\<^sub>l \<Squnion>\<^sub>k if i = j \<and> l = k then ((minarc\<^sub>M f) (k,i))\<^sup>T * top * (minarc\<^sub>M f) (l,j) else bot)"
apply (rule sup_monoid.sum.cong)
apply simp
by (metis (no_types, lifting) comp_left_zero comp_right_zero conv_bot prod.inject minarc_at_most_one)
also have "... = (if i = j then (\<Squnion>\<^sub>l \<Squnion>\<^sub>k if l = k then ((minarc\<^sub>M f) (k,i))\<^sup>T * top * (minarc\<^sub>M f) (l,j) else bot) else bot)"
by auto
also have "... \<le> (if i = j then top else bot)"
by simp
also have "... = mone (i,j)"
by (simp add: one_matrix_def)
finally show "((minarc\<^sub>M f)\<^sup>t \<odot> mtop \<odot> (minarc\<^sub>M f)) (i,j) \<le> mone (i,j)"
.
qed
finally have 5: "(minarc\<^sub>M f)\<^sup>t \<odot> mtop \<odot> ((minarc\<^sub>M f)\<^sup>t \<odot> mtop)\<^sup>t \<preceq> mone"
.
have "mtop \<odot> ((minarc\<^sub>M f)\<^sup>t \<odot> mtop) = mtop"
using 3 by (metis matrix_monoid.mult_assoc matrix_stone_relation_algebra.conv_dist_comp matrix_stone_relation_algebra.conv_top)
thus "?at (minarc\<^sub>M f)"
using 2 3 5 by blast
qed
next
fix f g :: "('a,'b) square"
let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix"
show "?at g \<and> g \<otimes> f \<noteq> mbot \<longrightarrow> sum\<^sub>M (minarc\<^sub>M f \<otimes> f) \<preceq> sum\<^sub>M (g \<otimes> f)"
proof
assume 1: "?at g \<and> g \<otimes> f \<noteq> mbot"
hence 2: "g = \<ominus>\<ominus>g"
using matrix_stone_relation_algebra.arc_regular by blast
show "sum\<^sub>M (minarc\<^sub>M f \<otimes> f) \<preceq> sum\<^sub>M (g \<otimes> f)"
proof (unfold less_eq_matrix_def, rule allI, rule prod_cases)
fix i j
from minarc_at_least_one obtain ei ej where 3: "(minarc\<^sub>M f) (ei,ej) = top"
using 1 by force
hence 4: "\<forall>k l . \<not>(k = ei \<and> l = ej) \<longrightarrow> (minarc\<^sub>M f) (k,l) = bot"
by (metis (mono_tags, hide_lams) bot.extremum inf.bot_unique prod.inject minarc_at_most_one)
from agg_matrix_bot obtain di dj where 5: "(g \<otimes> f) (di,dj) \<noteq> bot"
using 1 by force
hence 6: "g (di,dj) \<noteq> bot"
by (metis inf_bot_left inf_matrix_def)
hence 7: "g (di,dj) = top"
using 2 by (metis uminus_matrix_def uminus_def)
hence 8: "(g \<otimes> f) (di,dj) = f (di,dj)"
by (metis inf_matrix_def inf_top.left_neutral)
have 9: "\<forall>k l . k \<noteq> di \<longrightarrow> g (k,l) = bot"
proof (intro allI, rule impI)
fix k l
assume 10: "k \<noteq> di"
have "top * (g (k,l))\<^sup>T = g (di,dj) * top * (g\<^sup>t) (l,k)"
using 7 by (simp add: conv_matrix_def)
also have "... \<le> (\<Squnion>\<^sub>n g (di,n) * top) * (g\<^sup>t) (l,k)"
by (metis comp_inf.ub_sum comp_right_dist_sum)
also have "... \<le> (\<Squnion>\<^sub>m (\<Squnion>\<^sub>n g (di,n) * top) * (g\<^sup>t) (m,k))"
by (metis comp_inf.ub_sum)
also have "... = (g \<odot> mtop \<odot> g\<^sup>t) (di,k)"
by (simp add: times_matrix_def top_matrix_def)
also have "... \<le> mone (di,k)"
using 1 by (metis matrix_stone_relation_algebra.arc_expanded less_eq_matrix_def)
also have "... = bot"
apply (unfold one_matrix_def)
using 10 by auto
finally have "g (k,l) \<noteq> top"
using 5 by (metis bot.extremum conv_def inf.bot_unique mult.left_neutral one_def)
thus "g (k,l) = bot"
using 2 by (metis uminus_def uminus_matrix_def)
qed
have "\<forall>k l . l \<noteq> dj \<longrightarrow> g (k,l) = bot"
proof (intro allI, rule impI)
fix k l
assume 11: "l \<noteq> dj"
have "(g (k,l))\<^sup>T * top = (g\<^sup>t) (l,k) * top * g (di,dj)"
using 7 by (simp add: comp_associative conv_matrix_def)
also have "... \<le> (\<Squnion>\<^sub>n (g\<^sup>t) (l,n) * top) * g (di,dj)"
by (metis comp_inf.ub_sum comp_right_dist_sum)
also have "... \<le> (\<Squnion>\<^sub>m (\<Squnion>\<^sub>n (g\<^sup>t) (l,n) * top) * g (m,dj))"
by (metis comp_inf.ub_sum)
also have "... = (g\<^sup>t \<odot> mtop \<odot> g) (l,dj)"
by (simp add: times_matrix_def top_matrix_def)
also have "... \<le> mone (l,dj)"
using 1 by (metis matrix_stone_relation_algebra.arc_expanded less_eq_matrix_def)
also have "... = bot"
apply (unfold one_matrix_def)
using 11 by auto
finally have "g (k,l) \<noteq> top"
using 5 by (metis bot.extremum comp_right_one conv_def one_def top.extremum_unique)
thus "g (k,l) = bot"
using 2 by (metis uminus_def uminus_matrix_def)
qed
hence 12: "\<forall>k l . \<not>(k = di \<and> l = dj) \<longrightarrow> (g \<otimes> f) (k,l) = bot"
using 9 by (metis inf_bot_left inf_matrix_def)
have "(\<Sum>\<^sub>k \<Sum>\<^sub>l (minarc\<^sub>M f \<otimes> f) (k,l)) = (\<Sum>\<^sub>k \<Sum>\<^sub>l if k = ei \<and> l = ej then (minarc\<^sub>M f \<otimes> f) (k,l) else (minarc\<^sub>M f \<otimes> f) (k,l))"
by simp
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l if k = ei \<and> l = ej then (minarc\<^sub>M f \<otimes> f) (k,l) else (minarc\<^sub>M f) (k,l) \<sqinter> f (k,l))"
by (unfold inf_matrix_def) simp
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l if k = ei \<and> l = ej then (minarc\<^sub>M f \<otimes> f) (k,l) else bot)"
apply (rule aggregation.sum_0.cong)
apply simp
using 4 by (metis inf_bot_left)
also have "... = (minarc\<^sub>M f \<otimes> f) (ei,ej) + bot"
by (unfold agg_delta_2) simp
also have "... = f (ei,ej) + bot"
using 3 by (simp add: inf_matrix_def)
also have "... \<le> (g \<otimes> f) (di,dj) + bot"
using 3 5 6 7 8 by (metis minarc_matrix_def)
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l if k = di \<and> l = dj then (g \<otimes> f) (k,l) else bot)"
by (unfold agg_delta_2) simp
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l if k = di \<and> l = dj then (g \<otimes> f) (k,l) else (g \<otimes> f) (k,l))"
apply (rule aggregation.sum_0.cong)
apply simp
using 12 by metis
also have "... = (\<Sum>\<^sub>k \<Sum>\<^sub>l (g \<otimes> f) (k,l))"
by simp
finally show "(sum\<^sub>M (minarc\<^sub>M f \<otimes> f)) (i,j) \<le> (sum\<^sub>M (g \<otimes> f)) (i,j)"
by (simp add: sum_matrix_def)
qed
qed
next
fix f g :: "('a,'b) square"
let ?h = "hd enum_class.enum"
show "sum\<^sub>M f \<preceq> sum\<^sub>M g \<or> sum\<^sub>M g \<preceq> sum\<^sub>M f"
proof (cases "(sum\<^sub>M f) (?h,?h) \<le> (sum\<^sub>M g) (?h,?h)")
case 1: True
have "sum\<^sub>M f \<preceq> sum\<^sub>M g"
apply (unfold less_eq_matrix_def, rule allI, rule prod_cases)
using 1 by (unfold sum_matrix_def) auto
thus ?thesis
by simp
next
case False
hence 2: "(sum\<^sub>M g) (?h,?h) \<le> (sum\<^sub>M f) (?h,?h)"
by (simp add: linear)
have "sum\<^sub>M g \<preceq> sum\<^sub>M f"
apply (unfold less_eq_matrix_def, rule allI, rule prod_cases)
using 2 by (unfold sum_matrix_def) auto
thus ?thesis
by simp
qed
next
have "finite { f :: ('a,'b) square . (\<forall>e . regular (f e)) }"
by (unfold regular_bot_top_2, rule finite_set_of_finite_funs_pred) auto
thus "finite { f :: ('a,'b) square . matrix_p_algebra.regular f }"
by (unfold uminus_matrix_def) meson
qed
text \<open>
We show the same for the alternative implementation that stores the result of aggregation in all elements of the matrix.
\<close>
interpretation agg_square_m_algebra_2: m_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix_2 and minarc = minarc_matrix
proof
fix f :: "('a,'b) square"
show "minarc\<^sub>M f \<preceq> \<ominus>\<ominus>f"
by (simp add: agg_square_m_algebra.minarc_below)
next
fix f :: "('a,'b) square"
let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix"
show "f \<noteq> mbot \<longrightarrow> ?at (minarc\<^sub>M f)"
by (simp add: agg_square_m_algebra.minarc_arc)
next
fix f g :: "('a,'b) square"
let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix"
show "?at g \<and> g \<otimes> f \<noteq> mbot \<longrightarrow> sum2\<^sub>M (minarc\<^sub>M f \<otimes> f) \<preceq> sum2\<^sub>M (g \<otimes> f)"
proof
let ?h = "hd enum_class.enum"
assume "?at g \<and> g \<otimes> f \<noteq> mbot"
hence "sum\<^sub>M (minarc\<^sub>M f \<otimes> f) \<preceq> sum\<^sub>M (g \<otimes> f)"
by (simp add: agg_square_m_algebra.minarc_min)
hence "(sum\<^sub>M (minarc\<^sub>M f \<otimes> f)) (?h,?h) \<le> (sum\<^sub>M (g \<otimes> f)) (?h,?h)"
by (simp add: less_eq_matrix_def)
thus "sum2\<^sub>M (minarc\<^sub>M f \<otimes> f) \<preceq> sum2\<^sub>M (g \<otimes> f)"
by (simp add: sum_matrix_def sum_matrix_2_def less_eq_matrix_def)
qed
next
fix f g :: "('a,'b) square"
let ?h = "hd enum_class.enum"
have "sum\<^sub>M f \<preceq> sum\<^sub>M g \<or> sum\<^sub>M g \<preceq> sum\<^sub>M f"
by (simp add: agg_square_m_algebra.sum_linear)
hence "(sum\<^sub>M f) (?h,?h) \<le> (sum\<^sub>M g) (?h,?h) \<or> (sum\<^sub>M g) (?h,?h) \<le> (sum\<^sub>M f) (?h,?h)"
using less_eq_matrix_def by auto
thus "sum2\<^sub>M f \<preceq> sum2\<^sub>M g \<or> sum2\<^sub>M g \<preceq> sum2\<^sub>M f"
by (simp add: sum_matrix_def sum_matrix_2_def less_eq_matrix_def)
next
show "finite { f :: ('a,'b) square . matrix_p_algebra.regular f }"
by (simp add: agg_square_m_algebra.finite_regular)
qed
text \<open>
By defining the Kleene star as $\top$ aggregation lattices form a Kleene algebra.
\<close>
class aggregation_kleene_algebra = aggregation_algebra + star +
assumes star_def [simp]: "x\<^sup>\<star> = top"
begin
subclass stone_kleene_relation_algebra
apply unfold_locales
by (simp_all add: inf.assoc le_infI2 inf_sup_distrib1 inf_sup_distrib2)
end
class linear_aggregation_kleene_algebra = linear_aggregation_algebra + star +
assumes star_def_2 [simp]: "x\<^sup>\<star> = top"
begin
subclass aggregation_kleene_algebra
apply unfold_locales
by simp
end
interpretation agg_square_m_kleene_algebra: m_kleene_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_kleene_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and plus = plus_matrix and sum = sum_matrix and minarc = minarc_matrix ..
interpretation agg_square_m_kleene_algebra_2: m_kleene_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_kleene_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and plus = plus_matrix and sum = sum_matrix_2 and minarc = minarc_matrix ..
end
diff --git a/thys/Algebraic_Numbers/Compare_Complex.thy b/thys/Algebraic_Numbers/Compare_Complex.thy
--- a/thys/Algebraic_Numbers/Compare_Complex.thy
+++ b/thys/Algebraic_Numbers/Compare_Complex.thy
@@ -1,89 +1,89 @@
(*
Author: René Thiemann
Akihisa Yamada
License: BSD
*)
subsection \<open>Compare Instance for Complex Numbers\<close>
text \<open>We define some code equations for complex numbers, provide a comparator for complex
numbers, and register complex numbers for the container framework.\<close>
theory Compare_Complex
imports
HOL.Complex
Polynomial_Interpolation.Missing_Unsorted
Deriving.Compare_Real
Containers.Set_Impl
begin
declare [[code drop: Gcd_fin]]
declare [[code drop: Lcm_fin]]
definition gcds :: "'a::semiring_gcd list \<Rightarrow> 'a"
where [simp, code_abbrev]: "gcds xs = gcd_list xs"
lemma [code]:
"gcds xs = fold gcd xs 0"
by (simp add: Gcd_fin.set_eq_fold)
definition lcms :: "'a::semiring_gcd list \<Rightarrow> 'a"
where [simp, code_abbrev]: "lcms xs = lcm_list xs"
lemma [code]:
"lcms xs = fold lcm xs 1"
by (simp add: Lcm_fin.set_eq_fold)
lemma in_reals_code [code_unfold]:
"x \<in> \<real> \<longleftrightarrow> Im x = 0"
by (fact complex_is_Real_iff)
definition is_norm_1 :: "complex \<Rightarrow> bool" where
"is_norm_1 z = ((Re z)\<^sup>2 + (Im z)\<^sup>2 = 1)"
lemma is_norm_1[simp]: "is_norm_1 x = (norm x = 1)"
unfolding is_norm_1_def norm_complex_def by simp
definition is_norm_le_1 :: "complex \<Rightarrow> bool" where
"is_norm_le_1 z = ((Re z)\<^sup>2 + (Im z)\<^sup>2 \<le> 1)"
lemma is_norm_le_1[simp]: "is_norm_le_1 x = (norm x \<le> 1)"
unfolding is_norm_le_1_def norm_complex_def by simp
instantiation complex :: finite_UNIV
begin
definition "finite_UNIV = Phantom(complex) False"
instance
by (intro_classes, unfold finite_UNIV_complex_def, simp add: infinite_UNIV_char_0)
end
instantiation complex :: compare
begin
definition compare_complex :: "complex \<Rightarrow> complex \<Rightarrow> order" where
"compare_complex x y = compare (Re x, Im x) (Re y, Im y)"
instance
proof (intro_classes, unfold_locales; unfold compare_complex_def)
fix x y z :: complex
let ?c = "compare :: (real \<times> real) comparator"
interpret comparator ?c by (rule comparator_compare)
show "invert_order (?c (Re x, Im x) (Re y, Im y)) = ?c (Re y, Im y) (Re x, Im x)"
by (rule sym)
{
assume "?c (Re x, Im x) (Re y, Im y) = Lt"
"?c (Re y, Im y) (Re z, Im z) = Lt"
thus "?c (Re x, Im x) (Re z, Im z) = Lt"
- by (rule trans)
+ by (rule comp_trans)
}
{
assume "?c (Re x, Im x) (Re y, Im y) = Eq"
from weak_eq[OF this] show "x = y" unfolding complex_eq_iff by auto
}
qed
end
derive (eq) ceq complex real
derive (compare) ccompare complex
derive (compare) ccompare real
derive (dlist) set_impl complex real
end
diff --git a/thys/Algebraic_VCs/AVC_KAD/VC_KAD_scratch.thy b/thys/Algebraic_VCs/AVC_KAD/VC_KAD_scratch.thy
--- a/thys/Algebraic_VCs/AVC_KAD/VC_KAD_scratch.thy
+++ b/thys/Algebraic_VCs/AVC_KAD/VC_KAD_scratch.thy
@@ -1,532 +1,532 @@
(* Title: Program Correctness Component Based on Kleene Algebra with Domain
Author: Victor Gomes, Georg Struth
Maintainer: Victor Gomes <victor.gomes@cl.cam.ac.uk>
Georg Struth <g.struth@sheffield.ac.uk>
*)
subsection\<open>Component Based on Kleene Algebra with Domain\<close>
text \<open>This component supports the verification and step-wise refinement of simple while programs
in a partial correctness setting.\<close>
theory VC_KAD_scratch
imports Main
begin
subsubsection \<open>KAD: Definitions and Basic Properties\<close>
notation times (infixl "\<cdot>" 70)
class plus_ord = plus + ord +
assumes less_eq_def: "x \<le> y \<longleftrightarrow> x + y = y"
and less_def: "x < y \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
class dioid = semiring + one + zero + plus_ord +
assumes add_idem [simp]: "x + x = x"
and mult_onel [simp]: "1 \<cdot> x = x"
and mult_oner [simp]: "x \<cdot> 1 = x"
and add_zerol [simp]: "0 + x = x"
and annil [simp]: "0 \<cdot> x = 0"
and annir [simp]: "x \<cdot> 0 = 0"
begin
subclass monoid_mult
by (standard, simp_all)
subclass order
by (standard, simp_all add: less_def less_eq_def add_commute, auto, metis add_assoc)
lemma mult_isor: "x \<le> y \<Longrightarrow> x \<cdot> z \<le> y \<cdot> z"
by (metis distrib_right less_eq_def)
lemma mult_isol: "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
by (metis distrib_left less_eq_def)
lemma add_iso: "x \<le> y \<Longrightarrow> z + x \<le> z + y"
by (metis add.semigroup_axioms add_idem less_eq_def semigroup.assoc)
lemma add_ub: "x \<le> x + y"
by (metis add_assoc add_idem less_eq_def)
lemma add_lub: "x + y \<le> z \<longleftrightarrow> x \<le> z \<and> y \<le> z"
by (metis add_assoc add_ub add.left_commute less_eq_def)
end
class kleene_algebra = dioid +
fixes star :: "'a \<Rightarrow> 'a" ("_\<^sup>\<star>" [101] 100)
assumes star_unfoldl: "1 + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
and star_unfoldr: "1 + x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
and star_inductl: "z + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> y"
and star_inductr: "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"
begin
lemma star_sim: "x \<cdot> y \<le> z \<cdot> x \<Longrightarrow> x \<cdot> y\<^sup>\<star> \<le> z\<^sup>\<star> \<cdot> x"
proof -
assume "x \<cdot> y \<le> z \<cdot> x"
hence "x + z\<^sup>\<star> \<cdot> x \<cdot> y \<le> x + z\<^sup>\<star> \<cdot> z \<cdot> x"
by (metis add_lub distrib_left eq_refl less_eq_def mult_assoc)
also have "... \<le> z\<^sup>\<star> \<cdot> x"
using add_lub mult_isor star_unfoldr by fastforce
finally show ?thesis
by (simp add: star_inductr)
qed
end
class antidomain_kleene_algebra = kleene_algebra +
fixes ad :: "'a \<Rightarrow> 'a" ("ad")
assumes as1 [simp]: "ad x \<cdot> x = 0"
and as2 [simp]: "ad (x \<cdot> y) + ad (x \<cdot> ad (ad y)) = ad (x \<cdot> ad (ad y))"
and as3 [simp]: "ad (ad x) + ad x = 1"
begin
definition dom_op :: "'a \<Rightarrow> 'a" ("d") where
"d x = ad (ad x)"
lemma a_subid_aux: "ad x \<cdot> y \<le> y"
by (metis add_commute add_ub as3 mult_1_left mult_isor)
lemma d1_a [simp]: "d x \<cdot> x = x"
by (metis add_commute dom_op_def add_zerol as1 as3 distrib_right mult_1_left)
lemma a_mul_d [simp]: "ad x \<cdot> d x = 0"
by (metis add_commute dom_op_def add_zerol as1 as2 distrib_right mult_1_left)
lemma a_d_closed [simp]: "d (ad x) = ad x"
by (metis d1_a dom_op_def add_zerol as1 as3 distrib_left mult_1_right)
lemma a_idem [simp]: "ad x \<cdot> ad x = ad x"
by (metis a_d_closed d1_a)
lemma meet_ord: "ad x \<le> ad y \<longleftrightarrow> ad x \<cdot> ad y = ad x"
- by (metis a_d_closed a_subid_aux d1_a antisym mult_1_right mult_isol)
+ by (metis a_d_closed a_subid_aux d1_a order.antisym mult_1_right mult_isol)
lemma d_wloc: "x \<cdot> y = 0 \<longleftrightarrow> x \<cdot> d y = 0"
- by (metis a_subid_aux d1_a dom_op_def add_ub antisym as1 as2 mult_1_right mult_assoc)
+ by (metis a_subid_aux d1_a dom_op_def add_ub order.antisym as1 as2 mult_1_right mult_assoc)
lemma gla_1: "ad x \<cdot> y = 0 \<Longrightarrow> ad x \<le> ad y"
by (metis a_subid_aux d_wloc dom_op_def add_zerol as3 distrib_left mult_1_right)
lemma a2_eq [simp]: "ad (x \<cdot> d y) = ad (x \<cdot> y)"
- by (metis a_mul_d d1_a dom_op_def gla_1 add_ub antisym as1 as2 mult_assoc)
+ by (metis a_mul_d d1_a dom_op_def gla_1 add_ub order.antisym as1 as2 mult_assoc)
lemma a_supdist: "ad (x + y) \<le> ad x"
by (metis add_commute gla_1 add_ub add_zerol as1 distrib_left less_eq_def)
lemma a_antitone: "x \<le> y \<Longrightarrow> ad y \<le> ad x"
by (metis a_supdist less_eq_def)
lemma a_comm: "ad x \<cdot> ad y = ad y \<cdot> ad x"
proof -
{ fix x y
have "ad x \<cdot> ad y = d (ad x \<cdot> ad y) \<cdot> ad x \<cdot> ad y"
by (simp add: mult_assoc)
also have "... \<le> d (ad y) \<cdot> ad x"
by (metis a_antitone a_d_closed a_subid_aux mult_oner a_subid_aux dom_op_def mult_isol mult_isor meet_ord)
finally have "ad x \<cdot> ad y \<le> ad y \<cdot> ad x"
by simp }
thus ?thesis
- by (simp add: antisym)
+ by (simp add: order.antisym)
qed
lemma a_closed [simp]: "d (ad x \<cdot> ad y) = ad x \<cdot> ad y"
proof -
have f1: "\<And>x y. ad x \<le> ad (ad y \<cdot> x)"
by (simp add: a_antitone a_subid_aux)
have "\<And>x y. d (ad x \<cdot> y) \<le> ad x"
by (metis a2_eq a_antitone a_comm a_d_closed dom_op_def f1)
hence "\<And>x y. d (ad x \<cdot> y) \<cdot> y = ad x \<cdot> y"
by (metis d1_a dom_op_def meet_ord mult_assoc)
thus ?thesis
by (metis a_comm a_idem dom_op_def)
qed
lemma a_exp [simp]: "ad (ad x \<cdot> y) = d x + ad y"
-proof (rule antisym)
+proof (rule order.antisym)
have "ad (ad x \<cdot> y) \<cdot> ad x \<cdot> d y = 0"
using d_wloc mult_assoc by fastforce
hence a: "ad (ad x \<cdot> y) \<cdot> d y \<le> d x"
by (metis a_closed a_comm dom_op_def gla_1 mult_assoc)
have "ad (ad x \<cdot> y) = ad (ad x \<cdot> y) \<cdot> d y + ad (ad x \<cdot> y) \<cdot> ad y"
by (metis dom_op_def as3 distrib_left mult_oner)
also have "... \<le> d x + ad (ad x \<cdot> y) \<cdot> ad y"
using a add_lub dual_order.trans by blast
finally show "ad (ad x \<cdot> y) \<le> d x + ad y"
by (metis a_antitone a_comm a_subid_aux meet_ord)
next
have "ad y \<le> ad (ad x \<cdot> y)"
by (simp add: a_antitone a_subid_aux)
thus "d x + ad y \<le> ad (ad x \<cdot> y)"
by (metis a2_eq a_antitone a_comm a_subid_aux dom_op_def add_lub)
qed
lemma d1_sum_var: "x + y \<le> (d x + d y) \<cdot> (x + y)"
proof -
have "x + y = d x \<cdot> x + d y \<cdot> y"
by simp
also have "... \<le> (d x + d y) \<cdot> x + (d x + d y) \<cdot> y"
by (metis add_commute add_lub add_ub combine_common_factor)
finally show ?thesis
by (simp add: distrib_left)
qed
lemma a4: "ad (x + y) = ad x \<cdot> ad y"
-proof (rule antisym)
+proof (rule order.antisym)
show "ad (x + y) \<le> ad x \<cdot> ad y"
by (metis a_supdist add_commute mult_isor meet_ord)
hence "ad x \<cdot> ad y = ad x \<cdot> ad y + ad (x + y)"
using less_eq_def add_commute by simp
also have "... = ad (ad (ad x \<cdot> ad y) \<cdot> (x + y))"
by (metis a_closed a_exp)
finally show "ad x \<cdot> ad y \<le> ad (x + y)"
using a_antitone d1_sum_var dom_op_def by auto
qed
lemma kat_prop: "d x \<cdot> y \<le> y \<cdot> d z \<longleftrightarrow> d x \<cdot> y \<cdot> ad z = 0"
proof
show "d x \<cdot> y \<le> y \<cdot> d z \<Longrightarrow> d x \<cdot> y \<cdot> ad z = 0"
by (metis add_commute dom_op_def add_zerol annir as1 less_eq_def mult_isor mult_assoc)
next
assume h: "d x \<cdot> y \<cdot> ad z = 0"
hence "d x \<cdot> y = d x \<cdot> y \<cdot> d z + d x \<cdot> y \<cdot> ad z"
by (metis dom_op_def as3 distrib_left mult_1_right)
thus "d x \<cdot> y \<le> y \<cdot> d z"
by (metis a_subid_aux add_commute dom_op_def h add_zerol mult_assoc)
qed
lemma shunt: "ad x \<le> ad y + ad z \<longleftrightarrow> ad x \<cdot> d y \<le> ad z"
proof
assume "ad x \<le> ad y + ad z"
hence "ad x \<cdot> d y \<le> ad y \<cdot> d y + ad z \<cdot> d y"
by (metis distrib_right mult_isor)
thus " ad x \<cdot> d y \<le> ad z"
by (metis a_closed a_d_closed a_exp a_mul_d a_supdist dom_op_def dual_order.trans less_eq_def)
next
assume h: "ad x \<cdot> d y \<le> ad z"
have "ad x = ad x \<cdot> ad y + ad x \<cdot> d y"
by (metis add_commute dom_op_def as3 distrib_left mult_1_right)
also have "... \<le> ad x \<cdot> ad y + ad z"
using h add_lub dual_order.trans by blast
also have "... \<le> ad y + ad z"
by (metis a_subid_aux add_commute add_lub add_ub dual_order.trans)
finally show "ad x \<le> ad y + ad z"
by simp
qed
subsubsection \<open>wp Calculus\<close>
definition if_then_else :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("if _ then _ else _ fi" [64,64,64] 63) where
"if p then x else y fi = d p \<cdot> x + ad p \<cdot> y"
definition while :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("while _ do _ od" [64,64] 63) where
"while p do x od = (d p \<cdot> x)\<^sup>\<star> \<cdot> ad p"
definition while_inv :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("while _ inv _ do _ od" [64,64,64] 63) where
"while p inv i do x od = while p do x od"
definition wp :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"wp x p = ad (x \<cdot> ad p)"
lemma demod: " d p \<le> wp x q \<longleftrightarrow> d p \<cdot> x \<le> x \<cdot> d q"
by (metis as1 dom_op_def gla_1 kat_prop meet_ord mult_assoc wp_def)
lemma wp_weaken: "wp x p \<le> wp (x \<cdot> ad q) (d p \<cdot> ad q)"
by (metis a4 a_antitone a_d_closed a_mul_d dom_op_def gla_1 mult_isol mult_assoc wp_def)
lemma wp_seq [simp]: "wp (x \<cdot> y) q = wp x (wp y q)"
using a2_eq dom_op_def mult_assoc wp_def by auto
lemma wp_seq_var: "p \<le> wp x r \<Longrightarrow> r \<le> wp y q \<Longrightarrow> p \<le> wp (x \<cdot> y) q"
proof -
assume a1: "p \<le> wp x r"
assume a2: "r \<le> wp y q"
have "\<forall>z. \<not> wp x r \<le> z \<or> p \<le> z"
using a1 dual_order.trans by blast
then show ?thesis
using a2 a_antitone mult_isol wp_def wp_seq by auto
qed
lemma wp_cond_var [simp]: "wp (if p then x else y fi) q = (ad p + wp x q) \<cdot> (d p + wp y q)"
using a4 a_d_closed dom_op_def if_then_else_def distrib_right mult_assoc wp_def by auto
lemma wp_cond_aux1 [simp]: "d p \<cdot> wp (if p then x else y fi) q = d p \<cdot> wp x q"
proof -
have "d p \<cdot> wp (if p then x else y fi) q = ad (ad p) \<cdot> (ad p + wp x q) \<cdot> (d p + wp y q)"
using dom_op_def mult.semigroup_axioms semigroup.assoc wp_cond_var by fastforce
also have "... = wp x q \<cdot> d p \<cdot> (d p + d (wp y q))"
using a_comm a_d_closed dom_op_def distrib_left wp_def by auto
also have "... = wp x q \<cdot> d p"
by (metis a_exp dom_op_def add_ub meet_ord mult_assoc)
finally show ?thesis
by (simp add: a_comm dom_op_def wp_def)
qed
lemma wp_cond_aux2 [simp]: "ad p \<cdot> wp (if p then x else y fi) q = ad p \<cdot> wp y q"
by (metis (no_types) abel_semigroup.commute if_then_else_def a_d_closed add.abel_semigroup_axioms dom_op_def wp_cond_aux1)
lemma wp_cond [simp]: "wp (if p then x else y fi) q = (d p \<cdot> wp x q) + (ad p \<cdot> wp y q)"
by (metis as3 distrib_right dom_op_def mult_1_left wp_cond_aux2 wp_cond_aux1)
lemma wp_star_induct_var: "d q \<le> wp x q \<Longrightarrow> d q \<le> wp (x\<^sup>\<star>) q"
using demod star_sim by blast
lemma wp_while: "d p \<cdot> d r \<le> wp x p \<Longrightarrow> d p \<le> wp (while r do x od) (d p \<cdot> ad r)"
proof -
assume "d p \<cdot> d r \<le> wp x p"
hence "d p \<le> wp (d r \<cdot> x) p"
using dom_op_def mult.semigroup_axioms semigroup.assoc shunt wp_def by fastforce
hence "d p \<le> wp ((d r \<cdot> x)\<^sup>\<star>) p"
using wp_star_induct_var by blast
thus ?thesis
by (simp add: while_def) (use local.dual_order.trans wp_weaken in fastforce)
qed
lemma wp_while_inv: "d p \<le> d i \<Longrightarrow> d i \<cdot> ad r \<le> d q \<Longrightarrow> d i \<cdot> d r \<le> wp x i \<Longrightarrow> d p \<le> wp (while r inv i do x od) q"
proof -
assume a1: "d p \<le> d i" and a2: "d i \<cdot> ad r \<le> d q" and "d i \<cdot> d r \<le> wp x i"
hence "d i \<le> wp (while r inv i do x od) (d i \<cdot> ad r)"
by (simp add: while_inv_def wp_while)
also have "... \<le> wp (while r inv i do x od) q"
by (metis a2 a_antitone a_d_closed dom_op_def mult_isol wp_def)
finally show ?thesis
using a1 dual_order.trans by blast
qed
lemma wp_while_inv_break: "d p \<le> wp y i \<Longrightarrow> d i \<cdot> ad r \<le> d q \<Longrightarrow> d i \<cdot> d r \<le> wp x i \<Longrightarrow> d p \<le> wp (y \<cdot> (while r inv i do x od)) q"
by (metis dom_op_def eq_refl mult_1_left mult_1_right wp_def wp_seq wp_seq_var wp_while_inv)
end
subsubsection \<open>Soundness and Relation KAD\<close>
notation relcomp (infixl ";" 70)
interpretation rel_d: dioid Id "{}" "(\<union>)" "(;)" "(\<subseteq>)" "(\<subset>)"
by (standard, auto)
lemma (in dioid) pow_inductl: "z + x \<cdot> y \<le> y \<Longrightarrow> x ^ i \<cdot> z \<le> y"
apply (induct i; clarsimp simp add: add_lub)
by (metis local.dual_order.trans local.mult_isol mult_assoc)
lemma (in dioid) pow_inductr: "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x ^ i \<le> y"
apply (induct i; clarsimp simp add: add_lub)
proof -
fix i
assume "z \<cdot> x ^ i \<le> y" "z \<le> y" "y \<cdot> x \<le> y"
hence "(z \<cdot> x ^ i) \<cdot> x \<le> y"
using local.dual_order.trans local.mult_isor by blast
thus "z \<cdot> (x \<cdot> x ^ i) \<le> y"
by (simp add: mult_assoc local.power_commutes)
qed
lemma power_is_relpow: "rel_d.power X i = X ^^ i"
by (induct i, simp_all add: relpow_commute)
lemma rel_star_def: "X^* = (\<Union>i. rel_d.power X i)"
by (simp add: power_is_relpow rtrancl_is_UN_relpow)
lemma rel_star_contl: "X ; Y^* = (\<Union>i. X ; rel_d.power Y i)"
by (simp add: rel_star_def relcomp_UNION_distrib)
lemma rel_star_contr: "X^* ; Y = (\<Union>i. (rel_d.power X i) ; Y)"
by (simp add: rel_star_def relcomp_UNION_distrib2)
definition rel_ad :: "'a rel \<Rightarrow> 'a rel" where
"rel_ad R = {(x,x) | x. \<not> (\<exists>y. (x,y) \<in> R)}"
interpretation rel_aka: antidomain_kleene_algebra Id "{}" "(\<union>)" "(;)" "(\<subseteq>)" "(\<subset>)" rtrancl rel_ad
apply standard
apply auto[2]
by (auto simp: rel_star_contr rel_d.pow_inductl rel_star_contl SUP_least rel_d.pow_inductr rel_ad_def)
subsubsection \<open>Embedding Predicates in Relations\<close>
type_synonym 'a pred = "'a \<Rightarrow> bool"
abbreviation p2r :: "'a pred \<Rightarrow> 'a rel" ("\<lceil>_\<rceil>") where
"\<lceil>P\<rceil> \<equiv> {(s,s) |s. P s}"
lemma d_p2r [simp]: "rel_aka.dom_op \<lceil>P\<rceil> = \<lceil>P\<rceil>"
by (auto simp: rel_aka.dom_op_def rel_ad_def)
lemma p2r_neg_hom [simp]: "rel_ad \<lceil>P\<rceil> = \<lceil>\<lambda>s. \<not>P s\<rceil>"
by (auto simp: rel_ad_def)
lemma p2r_conj_hom [simp]: "\<lceil>P\<rceil> \<inter> \<lceil>Q\<rceil> = \<lceil>\<lambda>s. P s \<and> Q s\<rceil>"
by auto
lemma p2r_conj_hom_var [simp]: "\<lceil>P\<rceil> ; \<lceil>Q\<rceil> = \<lceil>\<lambda>s. P s \<and> Q s\<rceil>"
by auto
lemma p2r_disj_hom [simp]: "\<lceil>P\<rceil> \<union> \<lceil>Q\<rceil> = \<lceil>\<lambda>s. P s \<or> Q s\<rceil>"
by auto
subsubsection \<open>Store and Assignment\<close>
type_synonym 'a store = "string \<Rightarrow> 'a"
definition gets :: "string \<Rightarrow> ('a store \<Rightarrow> 'a) \<Rightarrow> 'a store rel" ("_ ::= _" [70, 65] 61) where
"v ::= e = {(s,s (v := e s)) |s. True}"
lemma wp_assign [simp]: "rel_aka.wp (v ::= e) \<lceil>Q\<rceil> = \<lceil>\<lambda>s. Q (s (v := e s))\<rceil>"
by (auto simp: rel_aka.wp_def gets_def rel_ad_def)
abbreviation spec_sugar :: "'a pred \<Rightarrow> 'a rel \<Rightarrow> 'a pred \<Rightarrow> bool" ("PRE _ _ POST _" [64,64,64] 63) where
"PRE P X POST Q \<equiv> rel_aka.dom_op \<lceil>P\<rceil> \<subseteq> rel_aka.wp X \<lceil>Q\<rceil>"
abbreviation if_then_else_sugar :: "'a pred \<Rightarrow> 'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" ("IF _ THEN _ ELSE _ FI" [64,64,64] 63) where
"IF P THEN X ELSE Y FI \<equiv> rel_aka.if_then_else \<lceil>P\<rceil> X Y"
abbreviation while_inv_sugar :: "'a pred \<Rightarrow> 'a pred \<Rightarrow> 'a rel \<Rightarrow> 'a rel" ("WHILE _ INV _ DO _ OD" [64,64,64] 63) where
"WHILE P INV I DO X OD \<equiv> rel_aka.while_inv \<lceil>P\<rceil> \<lceil>I\<rceil> X"
subsubsection \<open>Verification Example\<close>
lemma euclid:
"PRE (\<lambda>s::nat store. s ''x'' = x \<and> s ''y'' = y)
(WHILE (\<lambda>s. s ''y'' \<noteq> 0) INV (\<lambda>s. gcd (s ''x'') (s ''y'') = gcd x y)
DO
(''z'' ::= (\<lambda>s. s ''y''));
(''y'' ::= (\<lambda>s. s ''x'' mod s ''y''));
(''x'' ::= (\<lambda>s. s ''z''))
OD)
POST (\<lambda>s. s ''x'' = gcd x y)"
apply (rule rel_aka.wp_while_inv, simp_all) using gcd_red_nat by auto
context antidomain_kleene_algebra
begin
subsubsection\<open>Propositional Hoare Logic\<close>
definition H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"H p x q \<longleftrightarrow> d p \<le> wp x q"
lemma H_skip: "H p 1 p"
by (simp add: H_def dom_op_def wp_def)
lemma H_cons: "d p \<le> d p' \<Longrightarrow> d q' \<le> d q \<Longrightarrow> H p' x q' \<Longrightarrow> H p x q"
by (meson H_def demod mult_isol order_trans)
lemma H_seq: "H p x r \<Longrightarrow> H r y q \<Longrightarrow> H p (x \<cdot> y) q"
by (metis H_def a_d_closed demod dom_op_def wp_seq_var)
lemma H_cond: "H (d p \<cdot> d r) x q \<Longrightarrow> H (d p \<cdot> ad r) y q \<Longrightarrow> H p (if r then x else y fi) q"
proof -
assume h1: "H (d p \<cdot> d r) x q" and h2: "H (d p \<cdot> ad r) y q"
hence h3: "d p \<cdot> d r \<le> wp x q" and h4: "d p \<cdot> ad r \<le> wp y q"
using H_def a_closed dom_op_def apply auto[1]
using H_def h2 a_closed dom_op_def by auto
hence h5: "d p \<le> ad r + wp x q" and h6: "d p \<le> d r + wp y q"
apply (simp add: dom_op_def shunt wp_def)
using h4 a_d_closed dom_op_def shunt wp_def by auto
hence "d p \<le> d p \<cdot> (d r + wp y q)"
by (metis a_idem distrib_left dom_op_def less_eq_def)
also have "... \<le> (ad r + wp x q) \<cdot> (d r + wp y q)"
by (simp add: h5 mult_isor)
finally show ?thesis
by (simp add: H_def)
qed
lemma H_loop: "H (d p \<cdot> d r) x p \<Longrightarrow> H p (while r do x od) (d p \<cdot> ad r)"
by (metis (full_types) H_def a_closed dom_op_def wp_while)
lemma H_while_inv: "d p \<le> d i \<Longrightarrow> d i \<cdot> ad r \<le> d q \<Longrightarrow> H (d i \<cdot> d r) x i \<Longrightarrow> H p (while r inv i do x od) q"
using H_def a_closed dom_op_def wp_while_inv by auto
end
subsubsection\<open>Definition of Refinement KAD\<close>
class rkad = antidomain_kleene_algebra +
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
assumes R_def: "x \<le> R p q \<longleftrightarrow> d p \<le> wp x q"
begin
subsubsection \<open>Propositional Refinement Calculus\<close>
lemma HR: "H p x q \<longleftrightarrow> x \<le> R p q"
by (simp add: H_def R_def)
lemma wp_R1: "d p \<le> wp (R p q) q"
using R_def by blast
lemma wp_R2: "x \<le> R (wp x q) q"
by (simp add: R_def wp_def)
lemma wp_R3: "d p \<le> wp x q \<Longrightarrow> x \<le> R p q"
by (simp add: R_def)
lemma H_R1: "H p (R p q) q"
by (simp add: HR)
lemma H_R2: "H p x q \<Longrightarrow> x \<le> R p q"
by (simp add: HR)
lemma R_skip: "1 \<le> R p p"
by (simp add: H_R2 H_skip)
lemma R_cons: "d p \<le> d p' \<Longrightarrow> d q' \<le> d q \<Longrightarrow> R p' q' \<le> R p q"
by (simp add: H_R1 H_R2 H_cons)
lemma R_seq: "(R p r) \<cdot> (R r q) \<le> R p q"
using H_R1 H_R2 H_seq by blast
lemma R_cond: "if v then (R (d v \<cdot> d p) q) else (R (ad v \<cdot> d p) q) fi \<le> R p q"
by (simp add: H_R1 H_R2 H_cond a_comm dom_op_def)
lemma R_loop: "while q do (R (d p \<cdot> d q) p) od \<le> R p (d p \<cdot> ad q)"
by (simp add: H_R1 H_R2 H_loop)
end
subsubsection \<open>Soundness and Relation RKAD\<close>
definition rel_R :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" where
"rel_R P Q = \<Union>{X. rel_aka.dom_op P \<subseteq> rel_aka.wp X Q}"
interpretation rel_rkad: rkad Id "{}" "(\<union>)" "(;)" "(\<subseteq>)" "(\<subset>)" rtrancl rel_ad rel_R
by (standard, auto simp: rel_R_def rel_aka.dom_op_def rel_ad_def rel_aka.wp_def, blast)
subsubsection \<open>Assignment Laws\<close>
lemma R_assign: "(\<forall>s. P s \<longrightarrow> Q (s (v := e s))) \<Longrightarrow> (v ::= e) \<subseteq> rel_R \<lceil>P\<rceil> \<lceil>Q\<rceil>"
by (auto simp: rel_rkad.R_def)
lemma H_assign_var: "(\<forall>s. P s \<longrightarrow> Q (s (v := e s))) \<Longrightarrow> rel_aka.H \<lceil>P\<rceil> (v ::= e) \<lceil>Q\<rceil>"
by (auto simp: rel_aka.H_def rel_aka.dom_op_def rel_ad_def gets_def rel_aka.wp_def)
lemma R_assignr: "(\<forall>s. Q' s \<longrightarrow> Q (s (v := e s))) \<Longrightarrow> (rel_R \<lceil>P\<rceil> \<lceil>Q'\<rceil>) ; (v ::= e) \<subseteq> rel_R \<lceil>P\<rceil> \<lceil>Q\<rceil>"
apply (subst rel_rkad.HR[symmetric], rule rel_aka.H_seq) defer
by (erule H_assign_var, simp add: rel_rkad.H_R1)
lemma R_assignl: "(\<forall>s. P s \<longrightarrow> P' (s (v := e s))) \<Longrightarrow> (v ::= e) ; (rel_R \<lceil>P'\<rceil> \<lceil>Q\<rceil>) \<subseteq> rel_R \<lceil>P\<rceil> \<lceil>Q\<rceil>"
by (subst rel_rkad.HR[symmetric], rule rel_aka.H_seq, erule H_assign_var, simp add: rel_rkad.H_R1)
subsubsection \<open>Refinement Example\<close>
lemma var_swap_ref1:
"rel_R \<lceil>\<lambda>s. s ''x'' = a \<and> s ''y'' = b\<rceil> \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil>
\<supseteq> (''z'' ::= (\<lambda>s. s ''x'')); rel_R \<lceil>\<lambda>s. s ''z'' = a \<and> s ''y'' = b\<rceil> \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil>"
by (rule R_assignl, auto)
lemma var_swap_ref2:
"rel_R \<lceil>\<lambda>s. s ''z'' = a \<and> s ''y'' = b\<rceil> \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil>
\<supseteq> (''x'' ::= (\<lambda>s. s ''y'')); rel_R \<lceil>\<lambda>s. s ''z'' = a \<and> s ''x'' = b\<rceil> \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil>"
by (rule R_assignl, auto)
lemma var_swap_ref3:
"rel_R \<lceil>\<lambda>s. s ''z'' = a \<and> s ''x'' = b\<rceil> \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil>
\<supseteq> (''y'' ::= (\<lambda>s. s ''z'')); rel_R \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil> \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil>"
by (rule R_assignl, auto)
lemma var_swap_ref_var:
"rel_R \<lceil>\<lambda>s. s ''x'' = a \<and> s ''y'' = b\<rceil> \<lceil>\<lambda>s. s ''x'' = b \<and> s ''y'' = a\<rceil>
\<supseteq> (''z'' ::= (\<lambda>s. s ''x'')); (''x'' ::= (\<lambda>s. s ''y'')); (''y'' ::= (\<lambda>s. s ''z''))"
using var_swap_ref1 var_swap_ref2 var_swap_ref3 rel_rkad.R_skip by fastforce
end
diff --git a/thys/Algebraic_VCs/AVC_KAD/VC_KAD_wf.thy b/thys/Algebraic_VCs/AVC_KAD/VC_KAD_wf.thy
--- a/thys/Algebraic_VCs/AVC_KAD/VC_KAD_wf.thy
+++ b/thys/Algebraic_VCs/AVC_KAD/VC_KAD_wf.thy
@@ -1,190 +1,190 @@
(* Title: Verification Component Based on Divergence Kleene Algebra for Total Correctness
Author: Victor Gomes, Georg Struth
Maintainer: Victor Gomes <victor.gomes@cl.cam.ac.uk>
Georg Struth <g.struth@sheffield.ac.uk>
*)
subsection\<open>Verification Component for Total Correctness\<close>
theory VC_KAD_wf
imports VC_KAD KAD.Modal_Kleene_Algebra_Applications
begin
text \<open>This component supports the verification of simple while programs
in a total correctness setting.\<close>
subsubsection \<open>Relation Divergence Kleene Algebras\<close>
text\<open>Divergence Kleene algebras have been formalised in the AFP entry for Kleene algebra with domain.
The nabla or divergence operation models those states of a relation from which infinitely ascending chains
may start.\<close>
definition rel_nabla :: "'a rel \<Rightarrow> 'a rel" where
"rel_nabla X = \<Union> {P. P \<subseteq> relfdia X P}"
definition rel_nabla_bin :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" where
"rel_nabla_bin X Q = \<Union> {P. P \<subseteq> relfdia X P \<union> rdom Q}"
lemma rel_nabla_d_closed [simp]: "rdom (rel_nabla x) = rel_nabla x"
by (auto simp: rel_nabla_def rel_antidomain_kleene_algebra.fdia_def rel_antidomain_kleene_algebra.ads_d_def rel_ad_def)
lemma rel_nabla_bin_d_closed [simp]: "rdom (rel_nabla_bin x q) = rel_nabla_bin x q"
by (auto simp: rel_nabla_bin_def rel_antidomain_kleene_algebra.fdia_def rel_antidomain_kleene_algebra.ads_d_def rel_ad_def)
lemma rel_nabla_unfold: "rel_nabla X \<subseteq> relfdia X (rel_nabla X)"
by (simp add: rel_nabla_def rel_ad_def rel_antidomain_kleene_algebra.fdia_def, blast)
lemma rel_nabla_bin_unfold: "rel_nabla_bin X Q \<subseteq> relfdia X (rel_nabla_bin X Q) \<union> rdom Q"
by (simp add: rel_nabla_bin_def rel_ad_def rel_antidomain_kleene_algebra.fdia_def, blast)
lemma rel_nabla_coinduct_var: "P \<subseteq> relfdia X P \<Longrightarrow> P \<subseteq> rel_nabla X"
by (simp add: rel_nabla_def rel_antidomain_kleene_algebra.fdia_def rel_ad_def, blast)
lemma rel_nabla_bin_coinduct: "P \<subseteq> relfdia X P \<union> rdom Q \<Longrightarrow> P \<subseteq> rel_nabla_bin X Q"
by (simp add: rel_nabla_bin_def rel_antidomain_kleene_algebra.fdia_def rel_ad_def, blast)
text \<open>The two fusion lemmas are, in fact, hard-coded fixpoint fusion proofs. They might be replaced
by more generic fusion proofs eventually.\<close>
lemma nabla_fusion1: "rel_nabla X \<union> relfdia (X\<^sup>*) Q \<subseteq> rel_nabla_bin X Q"
proof -
have "rel_nabla X \<union> relfdia (X\<^sup>*) Q \<subseteq> relfdia X (rel_nabla X) \<union> relfdia X (relfdia (X\<^sup>*) Q) \<union> rdom Q"
by (metis (no_types, lifting) Un_mono inf_sup_aci(6) order_refl rel_antidomain_kleene_algebra.dka.fdia_star_unfold_var rel_nabla_unfold sup.commute)
also have "... = relfdia X (rel_nabla X \<union> relfdia (X\<^sup>*) Q) \<union> rdom Q"
by (simp add: rel_antidomain_kleene_algebra.dka.fdia_add1)
finally show ?thesis
using rel_nabla_bin_coinduct by blast
qed
lemma rel_ad_inter_seq: "rel_ad X \<inter> rel_ad Y = rel_ad X ; rel_ad Y"
by (auto simp: rel_ad_def)
lemma fusion2_aux2: "rdom (rel_nabla_bin X Q) \<subseteq> rdom (rel_nabla_bin X Q \<inter> rel_ad (relfdia (X\<^sup>*) Q) \<union> relfdia (X\<^sup>*) Q)"
apply (auto simp: rel_antidomain_kleene_algebra.ads_d_def rel_ad_def)
by (metis pair_in_Id_conv r_into_rtrancl rel_antidomain_kleene_algebra.a_one rel_antidomain_kleene_algebra.a_star rel_antidomain_kleene_algebra.addual.ars_r_def rel_antidomain_kleene_algebra.dka.dns1'' rel_antidomain_kleene_algebra.dpdz.dom_one rel_antidomain_kleene_algebra.ds.ddual.rsr5 rel_antidomain_kleene_algebra.dual.conway.dagger_unfoldr_eq rel_antidomain_kleene_algebra.dual.tc_eq rel_nabla_bin_d_closed)
lemma nabla_fusion2: "rel_nabla_bin X Q \<subseteq> rel_nabla X \<union> relfdia (X\<^sup>*) Q"
proof -
have "rel_nabla_bin X Q \<inter> rel_ad (relfdia (X\<^sup>*) Q) \<subseteq> (relfdia X (rel_nabla_bin X Q) \<union> rdom Q) \<inter> rel_ad (relfdia (X\<^sup>*) Q)"
by (meson Int_mono equalityD1 rel_nabla_bin_unfold)
also have "... \<subseteq> (relfdia X (rel_nabla_bin X Q \<inter> rel_ad (relfdia (X\<^sup>*) Q) \<union> relfdia (X\<^sup>*) Q) \<union> rdom Q) \<inter> rel_ad (relfdia (X\<^sup>*) Q)"
using fusion2_aux2 rel_antidomain_kleene_algebra.dka.fd_iso1 by blast
also have "... = (relfdia X (rel_nabla_bin X Q \<inter> rel_ad (relfdia (X\<^sup>*) Q)) \<union> relfdia X (relfdia (X\<^sup>*) Q) \<union> rdom Q) \<inter> rel_ad (relfdia (X\<^sup>*) Q)"
by (simp add: rel_antidomain_kleene_algebra.dka.fdia_add1)
also have "... = (relfdia X (rel_nabla_bin X Q \<inter> rel_ad (relfdia (X\<^sup>*) Q)) \<union> relfdia (X\<^sup>*) Q) \<inter> rel_ad (relfdia (X\<^sup>*) Q)"
using rel_antidomain_kleene_algebra.dka.fdia_star_unfold_var by blast
finally have "rel_nabla_bin X Q \<inter> rel_ad (relfdia (X\<^sup>*) Q) \<subseteq> relfdia X ((rel_nabla_bin X Q) \<inter> rel_ad (relfdia (X\<^sup>*) Q))"
by (metis (no_types, lifting) inf_commute order_trans_rules(23) rel_ad_inter_seq rel_antidomain_kleene_algebra.a_mult_add rel_antidomain_kleene_algebra.a_subid_aux1' rel_antidomain_kleene_algebra.addual.bdia_def rel_antidomain_kleene_algebra.ds.ddual.rsr5)
hence "rdom (rel_nabla_bin X Q) ; rel_ad (relfdia (X\<^sup>*) Q) \<subseteq> rdom (rel_nabla X)"
by (metis rel_ad_inter_seq rel_antidomain_kleene_algebra.addual.ars_r_def rel_nabla_bin_d_closed rel_nabla_coinduct_var rel_nabla_d_closed)
thus ?thesis
by (metis rel_antidomain_kleene_algebra.addual.ars_r_def rel_antidomain_kleene_algebra.addual.bdia_def rel_antidomain_kleene_algebra.d_a_galois1 rel_antidomain_kleene_algebra.dpdz.domain_invol rel_nabla_bin_d_closed rel_nabla_d_closed)
qed
lemma rel_nabla_coinduct: "P \<subseteq> relfdia X P \<union> rdom Q \<Longrightarrow> P \<subseteq> rel_nabla X \<union> relfdia (rtrancl X) Q"
by (meson nabla_fusion2 order_trans rel_nabla_bin_coinduct)
interpretation rel_fdivka: fdivergence_kleene_algebra rel_ad "(\<union>)" "(;) " Id "{}" "(\<subseteq>)" "(\<subset>)" rtrancl rel_nabla
proof
fix x y z:: "'a rel"
show "rdom (rel_nabla x) = rel_nabla x"
by simp
show "rel_nabla x \<subseteq> relfdia x (rel_nabla x)"
by (simp add: rel_nabla_unfold)
show "rdom y \<subseteq> relfdia x y \<union> rdom z \<Longrightarrow> rdom y \<subseteq> rel_nabla x \<union> relfdia (x\<^sup>*) z"
by (simp add: rel_nabla_coinduct)
qed
subsubsection \<open>Meta-Equational Loop Rule\<close>
context fdivergence_kleene_algebra
begin
text \<open>The rule below is inspired by Arden' rule from language theory. It can be used in total correctness proofs.\<close>
lemma fdia_arden: "\<nabla>x = 0 \<Longrightarrow> d p \<le> d q + |x\<rangle> p \<Longrightarrow> d p \<le> |x\<^sup>\<star>\<rangle> q"
proof -
assume a1: "\<nabla>x = zero_class.zero"
assume "d p \<le> d q + |x\<rangle> p"
then have "ad (ad p) \<le> zero_class.zero + ad (ad (x\<^sup>\<star> \<cdot> q))"
using a1 add_commute ads_d_def dka.fd_def nabla_coinduction by force
then show ?thesis
by (simp add: ads_d_def dka.fd_def)
qed
lemma fdia_arden_eq: "\<nabla>x = 0 \<Longrightarrow> d p = d q + |x\<rangle> p \<Longrightarrow> d p = |x\<^sup>\<star>\<rangle> q"
- by (simp add: fdia_arden dka.fdia_star_induct_eq eq_iff)
+ by (simp add: fdia_arden dka.fdia_star_induct_eq order.eq_iff)
lemma fdia_arden_iff: "\<nabla>x = 0 \<Longrightarrow> (d p = d q + |x\<rangle> p \<longleftrightarrow> d p = |x\<^sup>\<star>\<rangle> q)"
by (metis fdia_arden_eq dka.fdia_d_simp dka.fdia_star_unfold_var)
lemma "|x\<^sup>\<star>] p \<le> |x] p"
by (simp add: fbox_antitone_var)
lemma fbox_arden: "\<nabla>x = 0 \<Longrightarrow> d q \<cdot> |x] p \<le> d p \<Longrightarrow> |x\<^sup>\<star>] q \<le> d p"
proof -
assume h1: "\<nabla>x = 0" and "d q \<cdot> |x] p \<le> d p"
hence "ad p \<le> ad (d q \<cdot> |x] p)"
by (metis a_antitone' a_subid addual.ars_r_def dpdz.domain_subid dual_order.trans)
hence "ad p \<le> ad q + |x\<rangle> ad p"
by (simp add: a_6 addual.bbox_def ds.fd_def)
hence "ad p \<le> |x\<^sup>\<star>\<rangle> ad q"
by (metis fdia_arden h1 a_4 ads_d_def dpdz.dsg1 fdia_def meet_ord_def)
thus ?thesis
by (metis a_antitone' ads_d_def fbox_simp fdia_fbox_de_morgan_2)
qed
lemma fbox_arden_eq: "\<nabla>x = 0 \<Longrightarrow> d q \<cdot> |x] p = d p \<Longrightarrow> |x\<^sup>\<star>] q = d p"
- by (simp add: fbox_arden antisym fbox_star_induct_eq)
+ by (simp add: fbox_arden order.antisym fbox_star_induct_eq)
lemma fbox_arden_iff: "\<nabla>x = 0 \<Longrightarrow> (d p = d q \<cdot> |x] p \<longleftrightarrow> d p = |x\<^sup>\<star>] q)"
by (metis fbox_arden_eq fbox_simp fbox_star_unfold_var)
lemma fbox_arden_while_iff: "\<nabla> (d t \<cdot> x) = 0 \<Longrightarrow> (d p = (d t + d q) \<cdot> |d t \<cdot> x] p \<longleftrightarrow> d p = |while t do x od] q)"
by (metis fbox_arden_iff dka.dom_add_closed fbox_export3 while_def)
lemma fbox_arden_whilei: "\<nabla> (d t \<cdot> x) = 0 \<Longrightarrow> (d i = (d t + d q) \<cdot> |d t \<cdot> x] i \<Longrightarrow> d i = |while t inv i do x od] q)"
using fbox_arden_while_iff whilei_def by auto
lemma fbox_arden_whilei_iff: "\<nabla> (d t \<cdot> x) = 0 \<Longrightarrow> (d i = (d t + d q) \<cdot> |d t \<cdot> x] i \<longleftrightarrow> d i = |while t inv i do x od] q)"
using fbox_arden_while_iff whilei_def by auto
subsubsection \<open>Noethericity and Absence of Divergence\<close>
text \<open>Noetherian elements have been defined in the AFP entry for Kleene algebra with domain. First we show
that noethericity and absence of divergence coincide. Then we turn to the relational model and
show that noetherian relations model terminating programs.\<close>
lemma noether_nabla: "Noetherian x \<Longrightarrow> \<nabla> x = 0"
by (metis nabla_closure nabla_unfold noetherian_alt)
lemma nabla_noether_iff: "Noetherian x \<longleftrightarrow> \<nabla> x = 0"
using nabla_noether noether_nabla by blast
lemma nabla_preloeb_iff: "\<nabla> x = 0 \<longleftrightarrow> PreLoebian x"
using Noetherian_iff_PreLoebian nabla_noether noether_nabla by blast
end
lemma rel_nabla_prop: "rel_nabla R = {} \<longleftrightarrow> (\<forall>P. P \<subseteq> relfdia R P \<longrightarrow> P = {})"
by (metis bot.extremum_uniqueI rel_nabla_coinduct_var rel_nabla_unfold)
lemma fdia_rel_im1: "s2r ((converse R) `` P) = relfdia R (s2r P)"
by (auto simp: Id_on_def rel_antidomain_kleene_algebra.ads_d_def rel_ad_def rel_antidomain_kleene_algebra.fdia_def Image_def converse_def)
lemma fdia_rel_im2: "s2r ((converse R) `` (r2s (rdom P))) = relfdia R P"
by (simp add: fdia_rel_im1 rsr)
lemma wf_nabla_aux: "(P \<subseteq> (converse R) `` P \<longrightarrow> P = {}) \<longleftrightarrow> (s2r P \<subseteq> relfdia R (s2r P) \<longrightarrow> s2r P = {})"
apply (standard, metis Domain_Id_on Domain_mono Id_on_empty fdia_rel_im1)
using fdia_rel_im1 by fastforce
text \<open>A relation is noeterian if its converse is wellfounded. Hence a relation is noetherian if and only if its
divergence is empty. In the relational program semantics, noetherian programs terminate.\<close>
lemma wf_nabla: "wf (converse R) \<longleftrightarrow> rel_nabla R = {}"
by (metis (no_types, lifting) fdia_rel_im2 rel_fdivka.nabla_unfold_eq rel_nabla_prop rel_nabla_unfold wfE_pf wfI_pf wf_nabla_aux)
end
diff --git a/thys/Algebraic_VCs/Domain_Quantale.thy b/thys/Algebraic_VCs/Domain_Quantale.thy
--- a/thys/Algebraic_VCs/Domain_Quantale.thy
+++ b/thys/Algebraic_VCs/Domain_Quantale.thy
@@ -1,523 +1,523 @@
(* Title: Domain Quantales
Author: Victor Gomes, Georg Struth
Maintainer: Victor Gomes <victor.gomes@cl.cam.ac.uk>
Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Component for Recursive Programs\<close>
theory Domain_Quantale
imports KAD.Modal_Kleene_Algebra
begin
text \<open>This component supports the verification and step-wise refinement of recursive programs
in a partial correctness setting.\<close>
notation
times (infixl "\<cdot>" 70) and
bot ("\<bottom>") and
top ("\<top>") and
inf (infixl "\<sqinter>" 65) and
sup (infixl "\<squnion>" 65)
(*
Inf ("\<Sqinter>_" [900] 900) and
Sup ("\<Squnion>_" [900] 900)
*)
(*
syntax
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10)
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10)
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
*)
subsection \<open>Lattice-Ordered Monoids with Domain\<close>
class bd_lattice_ordered_monoid = bounded_lattice + distrib_lattice + monoid_mult +
assumes left_distrib: "x \<cdot> (y \<squnion> z) = x \<cdot> y \<squnion> x \<cdot> z"
and right_distrib: "(x \<squnion> y) \<cdot> z = x \<cdot> z \<squnion> y \<cdot> z"
and bot_annil [simp]: "\<bottom> \<cdot> x = \<bottom>"
and bot_annir [simp]: "x \<cdot> \<bottom> = \<bottom>"
begin
sublocale semiring_one_zero "(\<squnion>)" "(\<cdot>)" "1" "bot"
by (standard, auto simp: sup.assoc sup.commute sup_left_commute left_distrib right_distrib sup_absorb1)
sublocale dioid_one_zero "(\<squnion>)" "(\<cdot>)" "1" bot "(\<le>)" "(<)"
by (standard, simp add: le_iff_sup, auto)
end
no_notation ads_d ("d")
and ars_r ("r")
and antirange_op ("ar _" [999] 1000)
class domain_bdlo_monoid = bd_lattice_ordered_monoid +
assumes rdv: "(z \<sqinter> x \<cdot> top) \<cdot> y = z \<cdot> y \<sqinter> x \<cdot> top"
begin
definition "d x = 1 \<sqinter> x \<cdot> \<top>"
sublocale ds: domain_semiring "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" "d" "(\<le>)" "(<)"
proof standard
fix x y
show "x \<squnion> d x \<cdot> x = d x \<cdot> x"
by (metis d_def inf_sup_absorb left_distrib mult_1_left mult_1_right rdv sup.absorb_iff1 sup.idem sup.left_commute top_greatest)
show "d (x \<cdot> d y) = d (x \<cdot> y)"
by (simp add: d_def inf_absorb2 rdv mult_assoc)
show "d x \<squnion> 1 = 1"
by (simp add: d_def sup.commute)
show "d bot = bot"
by (simp add: d_def inf.absorb1 inf.commute)
show "d (x \<squnion> y) = d x \<squnion> d y"
by (simp add: d_def inf_sup_distrib1)
qed
end
subsection\<open>Boolean Monoids with Domain\<close>
class boolean_monoid = boolean_algebra + monoid_mult +
assumes left_distrib': "x \<cdot> (y \<squnion> z) = x \<cdot> y \<squnion> x \<cdot> z"
and right_distrib': "(x \<squnion> y) \<cdot> z = x \<cdot> z \<squnion> y \<cdot> z"
and bot_annil' [simp]: "\<bottom> \<cdot> x = \<bottom>"
and bot_annir' [simp]: "x \<cdot> \<bottom> = \<bottom>"
begin
subclass bd_lattice_ordered_monoid
by (standard, simp_all add: left_distrib' right_distrib')
lemma inf_bot_iff_le: "x \<sqinter> y = \<bottom> \<longleftrightarrow> x \<le> -y"
by (metis le_iff_inf inf_sup_distrib1 inf_top_right sup_bot.left_neutral sup_compl_top compl_inf_bot inf.assoc inf_bot_right)
end
class domain_boolean_monoid = boolean_monoid +
assumes rdv': "(z \<sqinter> x \<cdot> \<top>) \<cdot> y = z \<cdot> y \<sqinter> x \<cdot> \<top>"
begin
sublocale dblo: domain_bdlo_monoid "1" "(\<cdot>)" "(\<sqinter>)" "(\<le>)" "(<)" "(\<squnion>)" "\<bottom>" "\<top>"
by (standard, simp add: rdv')
definition "a x = 1 \<sqinter> -(dblo.d x)"
lemma a_d_iff: "a x = 1 \<sqinter> -(x \<cdot> \<top>)"
by (clarsimp simp: a_def dblo.d_def inf_sup_distrib1)
lemma topr: "-(x \<cdot> \<top>) \<cdot> \<top> = -(x \<cdot> \<top>)"
-proof (rule antisym)
+proof (rule order.antisym)
show "-(x \<cdot> \<top>) \<le> -(x \<cdot> \<top>) \<cdot> \<top>"
by (metis mult_isol_var mult_oner order_refl top_greatest)
have "-(x \<cdot> \<top>) \<sqinter> (x \<cdot> \<top>) = \<bottom>"
by simp
hence "(-(x \<cdot> \<top>) \<sqinter> (x \<cdot> \<top>)) \<cdot> \<top> = \<bottom>"
by simp
hence "-(x \<cdot> \<top>) \<cdot> \<top> \<sqinter> (x \<cdot> \<top>) = \<bottom>"
by (metis rdv')
thus "-(x \<cdot> \<top>) \<cdot> \<top> \<le> -(x \<cdot> \<top>)"
by (simp add: inf_bot_iff_le)
qed
lemma dd_a: "dblo.d x = a (a x)"
by (metis a_d_iff dblo.d_def double_compl inf_top.left_neutral mult_1_left rdv' topr)
lemma ad_a [simp]: "a (dblo.d x) = a x"
by (simp add: a_def)
lemma da_a [simp]: "dblo.d (a x) = a x"
using ad_a dd_a by auto
lemma a1 [simp]: "a x \<cdot> x = \<bottom>"
proof -
have "a x \<cdot> x \<cdot> \<top> = \<bottom>"
by (metis a_d_iff inf_compl_bot mult_1_left rdv' topr)
then show ?thesis
by (metis (no_types) dblo.d_def dblo.ds.domain_very_strict inf_bot_right)
qed
lemma a2 [simp]: "a (x \<cdot> y) \<squnion> a (x \<cdot> a (a y)) = a (x \<cdot> a (a y))"
by (metis a_def dblo.ds.dsr2 dd_a sup.idem)
lemma a3 [simp]: "a (a x) \<squnion> a x = 1"
by (metis a_def da_a inf.commute sup.commute sup_compl_top sup_inf_absorb sup_inf_distrib1)
subclass domain_bdlo_monoid ..
text \<open>The next statement shows that every boolean monoid with domain is an antidomain semiring.
In this setting the domain operation has been defined explicitly.\<close>
sublocale ad: antidomain_semiring a "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" "(\<le>)" "(<)"
rewrites ad_eq: "ad.ads_d x = d x"
proof -
show "class.antidomain_semiring a (\<squnion>) (\<cdot>) 1 \<bottom> (\<le>) (<)"
by (standard, simp_all)
then interpret ad: antidomain_semiring a "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" "(\<le>)" "(<)" .
show "ad.ads_d x = d x"
by (simp add: ad.ads_d_def dd_a)
qed
end
subsection\<open>Boolean Monoids with Range\<close>
class range_boolean_monoid = boolean_monoid +
assumes ldv': "y \<cdot> (z \<sqinter> \<top> \<cdot> x) = y \<cdot> z \<sqinter> \<top> \<cdot> x"
begin
definition "r x = 1 \<sqinter> \<top> \<cdot> x"
definition "ar x = 1 \<sqinter> -(r x)"
lemma ar_r_iff: "ar x = 1 \<sqinter> -(\<top> \<cdot> x)"
by (simp add: ar_def inf_sup_distrib1 r_def)
lemma topl: "\<top>\<cdot>(-(\<top> \<cdot> x)) = -(\<top> \<cdot> x)"
-proof (rule antisym)
+proof (rule order.antisym)
show "\<top> \<cdot> - (\<top> \<cdot> x) \<le> - (\<top> \<cdot> x)"
by (metis bot_annir' compl_inf_bot inf_bot_iff_le ldv')
show "- (\<top> \<cdot> x) \<le> \<top> \<cdot> - (\<top> \<cdot> x)"
by (metis inf_le2 inf_top.right_neutral mult_1_left mult_isor)
qed
lemma r_ar: "r x = ar (ar x)"
by (metis ar_r_iff double_compl inf.commute inf_top.right_neutral ldv' mult_1_right r_def topl)
lemma ar_ar [simp]: "ar (r x) = ar x"
by (simp add: ar_def ldv' r_def)
lemma rar_ar [simp]: "r (ar x) = ar x"
using r_ar ar_ar by force
lemma ar1 [simp]: "x \<cdot> ar x = \<bottom>"
proof -
have "\<top> \<cdot> x \<cdot> ar x = \<bottom>"
by (metis ar_r_iff inf_compl_bot ldv' mult_oner topl)
then show ?thesis
by (metis inf_bot_iff_le inf_le2 inf_top.right_neutral mult_1_left mult_isor mult_oner topl)
qed
lemma ars: "r (r x \<cdot> y) = r (x \<cdot> y)"
by (metis inf.commute inf_top.right_neutral ldv' mult_oner mult_assoc r_def)
lemma ar2 [simp]: "ar (x \<cdot> y) \<squnion> ar (ar (ar x) \<cdot> y) = ar (ar (ar x) \<cdot> y)"
by (metis ar_def ars r_ar sup.idem)
lemma ar3 [simp]: "ar (ar x) \<squnion> ar x = 1 "
by (metis ar_def rar_ar inf.commute sup.commute sup_compl_top sup_inf_absorb sup_inf_distrib1)
sublocale ar: antirange_semiring "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" ar "(\<le>)" "(<)"
rewrites ar_eq: "ar.ars_r x = r x"
proof -
show "class.antirange_semiring (\<squnion>) (\<cdot>) 1 \<bottom> ar (\<le>) (<)"
by (standard, simp_all)
then interpret ar: antirange_semiring "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" ar "(\<le>)" "(<)" .
show "ar.ars_r x = r x"
by (simp add: ar.ars_r_def r_ar)
qed
end
subsection \<open>Quantales\<close>
text \<open>This part will eventually move into an AFP quantale entry.\<close>
class quantale = complete_lattice + monoid_mult +
assumes Sup_distr: "Sup X \<cdot> y = Sup {z. \<exists>x \<in> X. z = x \<cdot> y}"
and Sup_distl: "x \<cdot> Sup Y = Sup {z. \<exists>y \<in> Y. z = x \<cdot> y}"
begin
lemma bot_annil'' [simp]: "\<bottom> \<cdot> x = \<bottom>"
using Sup_distr[where X="{}"] by auto
lemma bot_annirr'' [simp]: "x \<cdot> \<bottom> = \<bottom>"
using Sup_distl[where Y="{}"] by auto
lemma sup_distl: "x \<cdot> (y \<squnion> z) = x \<cdot> y \<squnion> x \<cdot> z"
using Sup_distl[where Y="{y, z}"] by (fastforce intro!: Sup_eqI)
lemma sup_distr: "(x \<squnion> y) \<cdot> z = x \<cdot> z \<squnion> y \<cdot> z"
using Sup_distr[where X="{x, y}"] by (fastforce intro!: Sup_eqI)
sublocale semiring_one_zero "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>"
by (standard, auto simp: sup.assoc sup.commute sup_left_commute sup_distl sup_distr)
sublocale dioid_one_zero "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" "(\<le>)" "(<)"
by (standard, simp add: le_iff_sup, auto)
lemma Sup_sup_pred: "x \<squnion> Sup{y. P y} = Sup{y. y = x \<or> P y}"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: Collect_mono Sup_subset_mono Sup_upper)
using Sup_least Sup_upper le_supI2 by fastforce
definition star :: "'a \<Rightarrow> 'a" where
"star x = (SUP i. x ^ i)"
lemma star_def_var1: "star x = Sup{y. \<exists>i. y = x ^ i}"
by (simp add: star_def full_SetCompr_eq)
lemma star_def_var2: "star x = Sup{x ^ i |i. True}"
by (simp add: star_def full_SetCompr_eq)
lemma star_unfoldl' [simp]: "1 \<squnion> x \<cdot> (star x) = star x"
proof -
have "1 \<squnion> x \<cdot> (star x) = x ^ 0 \<squnion> x \<cdot> Sup{y. \<exists>i. y = x ^ i}"
by (simp add: star_def_var1)
also have "... = x ^ 0 \<squnion> Sup{y. \<exists>i. y = x ^ (i + 1)}"
by (simp add: Sup_distl, metis)
also have "... = Sup{y. y = x ^ 0 \<or> (\<exists>i. y = x ^ (i + 1))}"
using Sup_sup_pred by simp
also have "... = Sup{y. \<exists>i. y = x ^ i}"
by (metis Suc_eq_plus1 power.power.power_Suc power.power_eq_if)
finally show ?thesis
by (simp add: star_def_var1)
qed
lemma star_unfoldr' [simp]: "1 \<squnion> (star x) \<cdot> x = star x"
proof -
have "1 \<squnion> (star x) \<cdot> x = x ^ 0 \<squnion> Sup{y. \<exists>i. y = x ^ i} \<cdot> x"
by (simp add: star_def_var1)
also have "... = x ^ 0 \<squnion> Sup{y. \<exists>i. y = x ^ i \<cdot> x}"
by (simp add: Sup_distr, metis)
also have "... = x ^ 0 \<squnion> Sup{y. \<exists>i. y = x ^ (i + 1)}"
using power_Suc2 by simp
also have "... = Sup{y. y = x ^ 0 \<or> (\<exists>i. y = x ^ (i + 1))}"
using Sup_sup_pred by simp
also have "... = Sup{y. \<exists>i. y = x ^ i}"
by (metis Suc_eq_plus1 power.power.power_Suc power.power_eq_if)
finally show ?thesis
by (simp add: star_def_var1)
qed
lemma (in dioid_one_zero) power_inductl: "z + x \<cdot> y \<le> y \<Longrightarrow> (x ^ n) \<cdot> z \<le> y"
proof (induct n)
case 0 show ?case
using "0.prems" by simp
case Suc thus ?case
by (simp, metis mult.assoc mult_isol order_trans)
qed
lemma (in dioid_one_zero) power_inductr: "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> (x ^ n) \<le> y"
proof (induct n)
case 0 show ?case
using "0.prems" by auto
case Suc
{
fix n
assume "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x ^ n \<le> y"
and "z + y \<cdot> x \<le> y"
hence "z \<cdot> x ^ n \<le> y"
by simp
also have "z \<cdot> x ^ Suc n = z \<cdot> x \<cdot> x ^ n"
by (metis mult.assoc power_Suc)
moreover have "... = (z \<cdot> x ^ n) \<cdot> x"
by (metis mult.assoc power_commutes)
moreover have "... \<le> y \<cdot> x"
by (metis calculation(1) mult_isor)
moreover have "... \<le> y"
using \<open>z + y \<cdot> x \<le> y\<close> by simp
ultimately have "z \<cdot> x ^ Suc n \<le> y" by simp
}
thus ?case
by (metis Suc)
qed
lemma star_inductl': "z \<squnion> x \<cdot> y \<le> y \<Longrightarrow> (star x) \<cdot> z \<le> y"
proof -
assume "z \<squnion> x \<cdot> y \<le> y"
hence "\<forall>i. x ^ i \<cdot> z \<le> y"
by (simp add: power_inductl)
hence "Sup{w. \<exists>i. w = x ^ i \<cdot> z} \<le> y"
by (intro Sup_least, fast)
hence "Sup{w. \<exists>i. w = x ^ i} \<cdot> z \<le> y"
using Sup_distr Sup_le_iff by auto
thus "(star x) \<cdot> z \<le> y"
by (simp add: star_def_var1)
qed
lemma star_inductr': "z \<squnion> y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> (star x) \<le> y"
proof -
assume "z \<squnion> y \<cdot> x \<le> y"
hence "\<forall>i. z \<cdot> x ^ i \<le> y"
by (simp add: power_inductr)
hence "Sup{w. \<exists>i. w = z \<cdot> x ^ i} \<le> y"
by (intro Sup_least, fast)
hence "z \<cdot> Sup{w. \<exists>i. w = x ^ i} \<le> y"
using Sup_distl Sup_le_iff by auto
thus "z \<cdot> (star x) \<le> y"
by (simp add: star_def_var1)
qed
sublocale ka: kleene_algebra "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" "(\<le>)" "(<)" star
by standard (simp_all add: star_inductl' star_inductr')
end
text \<open>Distributive quantales are often assumed to satisfy infinite distributivity laws between
joins and meets, but finite ones suffice for our purposes.\<close>
class distributive_quantale = quantale + distrib_lattice
begin
subclass bd_lattice_ordered_monoid
by (standard, simp_all add: distrib_left)
lemma "(1 \<sqinter> x \<cdot> \<top>) \<cdot> x = x"
(* nitpick [expect=genuine]*)
oops
end
subsection \<open>Domain Quantales\<close>
class domain_quantale = distributive_quantale +
assumes rdv'': "(z \<sqinter> x \<cdot> \<top>) \<cdot> y = z \<cdot> y \<sqinter> x \<cdot> \<top>"
begin
subclass domain_bdlo_monoid
by (standard, simp add: rdv'')
end
class range_quantale = distributive_quantale +
assumes ldv'': "y \<cdot> (z \<sqinter> \<top> \<cdot> x) = y \<cdot> z \<sqinter> \<top> \<cdot> x"
class boolean_quantale = quantale + complete_boolean_algebra
begin
subclass boolean_monoid
by (standard, simp_all add: sup_distl)
lemma "(1 \<sqinter> x \<cdot> \<top>) \<cdot> x = x"
(*nitpick[expect=genuine]*)
oops
lemma "(1 \<sqinter> -(x \<cdot> \<top>)) \<cdot> x = \<bottom>"
(*nitpick[expect=genuine]*)
oops
end
subsection\<open>Boolean Domain Quantales\<close>
class domain_boolean_quantale = domain_quantale + boolean_quantale
begin
subclass domain_boolean_monoid
by (standard, simp add: rdv'')
lemma fbox_eq: "ad.fbox x q = Sup{d p |p. d p \<cdot> x \<le> x \<cdot> d q}"
apply (rule Sup_eqI[symmetric])
apply clarsimp
using ad.fbox_demodalisation3 ad.fbox_simp apply auto[1]
apply clarsimp
by (metis ad.fbox_def ad.fbox_demodalisation3 ad.fbox_simp da_a eq_refl)
lemma fdia_eq: "ad.fdia x p = Inf{d q |q. x \<cdot> d p \<le> d q \<cdot> x}"
apply (rule Inf_eqI[symmetric])
apply clarsimp
using ds.fdemodalisation2 apply auto[1]
apply clarsimp
by (metis ad.fd_eq_fdia ad.fdia_def da_a double_compl ds.fdemodalisation2 inf_bot_iff_le inf_compl_bot)
text \<open>The specification statement can be defined explicitly.\<close>
definition R :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"R p q \<equiv> Sup{x. (d p) \<cdot> x \<le> x \<cdot> d q}"
lemma "x \<le> R p q \<Longrightarrow> d p \<le> ad.fbox x (d q)"
proof (simp add: R_def ad.kat_1_equiv ad.kat_2_equiv)
assume "x \<le> Sup{x. d p \<cdot> x \<cdot> a q = \<bottom>}"
hence "d p \<cdot> x \<cdot> a q \<le> d p \<cdot> Sup{x. d p \<cdot> x \<cdot> a q = \<bottom>} \<cdot> a q "
using mult_double_iso by blast
also have "... = Sup{d p \<cdot> x \<cdot> a q |x. d p \<cdot> x \<cdot> a q = \<bottom>}"
apply (subst Sup_distl)
apply (subst Sup_distr)
apply clarsimp
by metis
also have "... = \<bottom>"
by (auto simp: Sup_eqI)
finally show ?thesis
using ad.fbox_demodalisation3 ad.kat_3 ad.kat_4 le_bot by blast
qed
lemma "d p \<le> ad.fbox x (d q) \<Longrightarrow> x \<le> R p q"
apply (simp add: R_def)
apply (rule Sup_upper)
apply simp
using ad.fbox_demodalisation3 ad.fbox_simp apply auto[1]
done
end
subsection\<open>Relational Model of Boolean Domain Quantales\<close>
interpretation rel_dbq: domain_boolean_quantale Inter Union "(\<inter>)" "(\<subseteq>)" "(\<subset>)" "(\<union>)" "{}" UNIV minus uminus Id "(O)"
apply (standard, simp_all add: O_assoc)
by blast +
subsection\<open>Modal Boolean Quantales\<close>
class range_boolean_quantale = range_quantale + boolean_quantale
begin
subclass range_boolean_monoid
by (standard, simp add: ldv'')
lemma fbox_eq: "ar.bbox x (r q) = Sup{r p |p. x \<cdot> r p \<le> (r q) \<cdot> x}"
apply (rule Sup_eqI[symmetric])
apply clarsimp
using ar.ardual.fbox_demodalisation3 ar.ardual.fbox_simp apply auto[1]
apply clarsimp
by (metis ar.ardual.fbox_def ar.ardual.fbox_demodalisation3 eq_refl rar_ar)
lemma fdia_eq: "ar.bdia x (r p) = Inf{r q |q. (r p) \<cdot> x \<le> x \<cdot> r q}"
apply (rule Inf_eqI[symmetric])
apply clarsimp
using ar.ars_r_def ar.ardual.fdemodalisation22 ar.ardual.kat_3_equiv_opp ar.ardual.kat_4_equiv_opp apply auto[1]
apply clarsimp
using ar.bdia_def ar.ardual.ds.fdemodalisation2 r_ar by fastforce
end
class modal_boolean_quantale = domain_boolean_quantale + range_boolean_quantale +
assumes domrange' [simp]: "d (r x) = r x"
and rangedom' [simp]: "r (d x) = d x"
begin
sublocale mka: modal_kleene_algebra "(\<squnion>)" "(\<cdot>)" 1 \<bottom> "(\<le>)" "(<)" star a ar
by standard (simp_all add: ar_eq ad_eq)
end
no_notation fbox ("( |_] _)" [61,81] 82)
and antidomain_semiringl_class.fbox ("( |_] _)" [61,81] 82)
notation ad.fbox ("( |_] _)" [61,81] 82)
subsection \<open>Recursion Rule\<close>
lemma recursion: "mono (f :: 'a \<Rightarrow> 'a :: domain_boolean_quantale) \<Longrightarrow>
(\<And>x. d p \<le> |x] d q \<Longrightarrow> d p \<le> |f x] d q) \<Longrightarrow> d p \<le> |lfp f] d q"
apply (erule lfp_ordinal_induct [where f=f], simp)
by (auto simp: ad.addual.ardual.fbox_demodalisation3 Sup_distr Sup_distl intro: Sup_mono)
text \<open>We have already tested this rule in the context of test quantales~\cite{ArmstrongGS15}, which is based
on a formalisation of quantales that is currently not in the AFP. The two theories will be merged as
soon as the quantale is available in the AFP.\<close>
end
diff --git a/thys/Card_Multisets/Card_Multisets.thy b/thys/Card_Multisets/Card_Multisets.thy
--- a/thys/Card_Multisets/Card_Multisets.thy
+++ b/thys/Card_Multisets/Card_Multisets.thy
@@ -1,227 +1,227 @@
(* Author: Lukas Bulwahn <lukas.bulwahn-at-gmail.com> *)
section \<open>Cardinality of Multisets\<close>
theory Card_Multisets
imports
"HOL-Library.Multiset"
begin
subsection \<open>Additions to Multiset Theory\<close>
lemma mset_set_set_mset_subseteq:
"mset_set (set_mset M) \<subseteq># M"
proof (induct M)
case empty
show ?case by simp
next
case (add x M)
from this show ?case
proof (cases "x \<in># M")
assume "x \<in># M"
from this have "mset_set (set_mset (M + {#x#})) = mset_set (set_mset M)"
by (simp add: insert_absorb)
from this add.hyps show ?thesis
- using subset_mset.order.trans by fastforce
+ using subset_mset.trans by fastforce
next
assume "\<not> x \<in># M"
from this add.hyps have "{#x#} + mset_set (set_mset M) \<subseteq># M + {#x#}"
by (simp add: insert_subset_eq_iff)
from this \<open>\<not> x \<in># M\<close> show ?thesis by simp
qed
qed
lemma size_mset_set_eq_card:
assumes "finite A"
shows "size (mset_set A) = card A"
using assms by (induct A) auto
lemma card_set_mset_leq:
"card (set_mset M) \<le> size M"
by (induct M) (auto simp add: card_insert_le_m1)
subsection \<open>Lemma to Enumerate Sets of Multisets\<close>
lemma set_of_multisets_eq:
assumes "x \<notin> A"
shows "{M. set_mset M \<subseteq> insert x A \<and> size M = Suc k} =
{M. set_mset M \<subseteq> A \<and> size M = Suc k} \<union>
(\<lambda>M. M + {#x#}) ` {M. set_mset M \<subseteq> insert x A \<and> size M = k}"
proof -
from \<open>x \<notin> A\<close> have "{M. set_mset M \<subseteq> insert x A \<and> size M = Suc k} =
{M. set_mset M \<subseteq> A \<and> size M = Suc k} \<union>
{M. set_mset M \<subseteq> insert x A \<and> size M = Suc k \<and> x \<in># M}"
by auto
moreover have "{M. set_mset M \<subseteq> insert x A \<and> size M = Suc k \<and> x \<in># M} =
(\<lambda>M. M + {#x#}) ` {M. set_mset M \<subseteq> insert x A \<and> size M = k}" (is "?S = ?T")
proof
show "?S \<subseteq> ?T"
proof
fix M
assume "M \<in> ?S"
from this have "M = M - {#x#} + {#x#}" by auto
moreover have "M - {#x#} \<in> {M. set_mset M \<subseteq> insert x A \<and> size M = k}"
proof -
have "set_mset (M - {#x#} + {#x#}) \<subseteq> insert x A"
using \<open>M \<in> ?S\<close> by force
moreover have "size (M - {#x#} + {#x#}) = Suc k \<and> x \<in># M - {#x#} + {#x#}"
using \<open>M \<in> ?S\<close> by force
ultimately show ?thesis by force
qed
ultimately show "M \<in> ?T" by auto
qed
next
show "?T \<subseteq> ?S" by force
qed
ultimately show ?thesis by auto
qed
subsection \<open>Derivation of Suitable Induction Rule\<close>
context
begin
private inductive R :: "'a set \<Rightarrow> nat \<Rightarrow> bool"
where
"finite A \<Longrightarrow> R A 0"
| "R {} k"
| "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> R A (Suc k) \<Longrightarrow> R (insert x A) k \<Longrightarrow> R (insert x A) (Suc k)"
private lemma R_eq_finite:
"R A k \<longleftrightarrow> finite A"
proof
assume "R A k"
from this show "finite A" by cases auto
next
assume "finite A"
from this show "R A k"
proof (induct A)
case empty
from this show ?case by (rule R.intros(2))
next
case insert
from this show ?case
proof (induct k)
case 0
from this show ?case
by (intro R.intros(1) finite.insertI)
next
case Suc
from this show ?case
by (metis R.simps Zero_neq_Suc diff_Suc_1)
qed
qed
qed
lemma finite_set_and_nat_induct[consumes 1, case_names zero empty step]:
assumes "finite A"
assumes "\<And>A. finite A \<Longrightarrow> P A 0"
assumes "\<And>k. P {} k"
assumes "\<And>A k x. finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> P A (Suc k) \<Longrightarrow> P (insert x A) k \<Longrightarrow> P (insert x A) (Suc k)"
shows "P A k"
proof -
from \<open>finite A\<close> have "R A k" by (subst R_eq_finite)
from this assms(2-4) show ?thesis by (induct A k) auto
qed
end
subsection \<open>Finiteness of Sets of Multisets\<close>
lemma finite_multisets:
assumes "finite A"
shows "finite {M. set_mset M \<subseteq> A \<and> size M = k}"
using assms
proof (induct A k rule: finite_set_and_nat_induct)
case zero
from this show ?case by auto
next
case empty
from this show ?case by auto
next
case (step A k x)
from this show ?case
using set_of_multisets_eq[OF \<open>x \<notin> A\<close>] by simp
qed
subsection \<open>Cardinality of Multisets\<close>
lemma card_multisets:
assumes "finite A"
shows "card {M. set_mset M \<subseteq> A \<and> size M = k} = (card A + k - 1) choose k"
using assms
proof (induct A k rule: finite_set_and_nat_induct)
case (zero A)
assume "finite (A :: 'a set)"
have "{M. set_mset M \<subseteq> A \<and> size M = 0} = {{#}}" by auto
from this show "card {M. set_mset M \<subseteq> A \<and> size M = 0} = card A + 0 - 1 choose 0"
by simp
next
case (empty k)
show "card {M. set_mset M \<subseteq> {} \<and> size M = k} = card {} + k - 1 choose k"
by (cases k) (auto simp add: binomial_eq_0)
next
case (step A k x)
let ?S\<^sub>1 = "{M. set_mset M \<subseteq> A \<and> size M = Suc k}"
and ?S\<^sub>2 = "{M. set_mset M \<subseteq> insert x A \<and> size M = k}"
assume hyps1: "card ?S\<^sub>1 = card A + Suc k - 1 choose Suc k"
assume hyps2: "card ?S\<^sub>2 = card (insert x A) + k - 1 choose k"
have finite_sets: "finite ?S\<^sub>1" "finite ((\<lambda>M. M + {#x#}) ` ?S\<^sub>2)"
using \<open>finite A\<close> by (auto simp add: finite_multisets)
have inj: "inj_on (\<lambda>M. M + {#x#}) ?S\<^sub>2" by (rule inj_onI) auto
have "card {M. set_mset M \<subseteq> insert x A \<and> size M = Suc k} =
card (?S\<^sub>1 \<union> (\<lambda>M. M + {#x#}) ` ?S\<^sub>2)"
using set_of_multisets_eq \<open>x \<notin> A\<close> by fastforce
also have "\<dots> = card ?S\<^sub>1 + card ((\<lambda>M. M + {#x#}) ` ?S\<^sub>2)"
using finite_sets \<open>x \<notin> A\<close> by (subst card_Un_disjoint) auto
also have "\<dots> = card ?S\<^sub>1 + card ?S\<^sub>2"
using inj by (auto intro: card_image)
also have "\<dots> = card A + Suc k - 1 choose Suc k + (card (insert x A) + k - 1 choose k)"
using hyps1 hyps2 by simp
also have "\<dots> = card (insert x A) + Suc k - 1 choose Suc k"
using \<open>x \<notin> A\<close> \<open>finite A\<close> by simp
finally show ?case .
qed
lemma card_too_small_multisets_covering_set:
assumes "finite A"
assumes "k < card A"
shows "card {M. set_mset M = A \<and> size M = k} = 0"
proof -
from \<open>k < card A\<close> have eq: "{M. set_mset M = A \<and> size M = k} = {}"
using card_set_mset_leq Collect_empty_eq leD by auto
from this show ?thesis by (metis card.empty)
qed
lemma card_multisets_covering_set:
assumes "finite A"
assumes "card A \<le> k"
shows "card {M. set_mset M = A \<and> size M = k} = (k - 1) choose (k - card A)"
proof -
have "{M. set_mset M = A \<and> size M = k} = (\<lambda>M. M + mset_set A) `
{M. set_mset M \<subseteq> A \<and> size M = k - card A}" (is "?S = ?f ` ?T")
proof
show "?S \<subseteq> ?f ` ?T"
proof
fix M
assume "M \<in> ?S"
from this have "M = M - mset_set A + mset_set A"
by (auto simp add: mset_set_set_mset_subseteq subset_mset.diff_add)
moreover from \<open>M \<in> ?S\<close> have "M - mset_set A \<in> ?T"
by (auto simp add: mset_set_set_mset_subseteq size_Diff_submset size_mset_set_eq_card in_diffD)
ultimately show "M \<in> ?f ` ?T" by auto
qed
next
from \<open>finite A\<close> \<open>card A \<le> k\<close> show "?f ` ?T \<subseteq> ?S"
by (auto simp add: size_mset_set_eq_card)+
qed
moreover have "inj_on ?f ?T" by (rule inj_onI) auto
ultimately have "card ?S = card ?T" by (simp add: card_image)
also have "\<dots> = card A + (k - card A) - 1 choose (k - card A)"
using \<open>finite A\<close> by (simp only: card_multisets)
also have "\<dots> = (k - 1) choose (k - card A)"
using \<open>card A \<le> k\<close> by auto
finally show ?thesis .
qed
end
diff --git a/thys/Coinductive/Examples/CCPO_Topology.thy b/thys/Coinductive/Examples/CCPO_Topology.thy
--- a/thys/Coinductive/Examples/CCPO_Topology.thy
+++ b/thys/Coinductive/Examples/CCPO_Topology.thy
@@ -1,335 +1,335 @@
(* Title: CCPO_Topology.thy
Author: Johannes Hölzl, TU Munich
*)
section \<open>CCPO topologies\<close>
theory CCPO_Topology
imports
"HOL-Analysis.Extended_Real_Limits"
"../Coinductive_Nat"
begin
lemma dropWhile_append:
"dropWhile P (xs @ ys) = (if \<forall>x\<in>set xs. P x then dropWhile P ys else dropWhile P xs @ ys)"
by auto
lemma dropWhile_False: "(\<And>x. x \<in> set xs \<Longrightarrow> P x) \<Longrightarrow> dropWhile P xs = []"
by simp
abbreviation (in order) "chain \<equiv> Complete_Partial_Order.chain (\<le>)"
lemma (in linorder) chain_linorder: "chain C"
by (simp add: chain_def linear)
lemma continuous_add_ereal:
assumes "0 \<le> t"
shows "continuous_on {-\<infinity>::ereal <..} (\<lambda>x. t + x)"
proof (subst continuous_on_open_vimage, (intro open_greaterThan allI impI)+)
fix B :: "ereal set" assume "open B"
show "open ((\<lambda>x. t + x) -` B \<inter> {- \<infinity><..})"
proof (cases t)
case (real t')
then have *: "(\<lambda>x. t + x) -` B \<inter> {- \<infinity><..} = (\<lambda>x. 1 * x + (-t)) ` (B \<inter> {-\<infinity> <..})"
apply (simp add: set_eq_iff image_iff Bex_def)
apply (intro allI iffI)
apply (rule_tac x= "x + ereal t'" in exI)
apply (case_tac x)
apply (auto simp: ac_simps)
done
show ?thesis
unfolding *
apply (rule ereal_open_affinity_pos)
using \<open>open B\<close>
apply (auto simp: real)
done
qed (insert \<open>0 \<le> t\<close>, auto)
qed
lemma tendsto_add_ereal:
"0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> (f \<longlongrightarrow> y) F \<Longrightarrow> ((\<lambda>z. x + f z :: ereal) \<longlongrightarrow> x + y) F"
apply (rule tendsto_compose[where f=f])
using continuous_add_ereal[where t=x]
unfolding continuous_on_def
apply (auto simp add: at_within_open[where S="{- \<infinity> <..}"])
done
lemma tendsto_LimI: "(f \<longlongrightarrow> y) F \<Longrightarrow> (f \<longlongrightarrow> Lim F f) F"
by (metis tendsto_Lim tendsto_bot)
subsection \<open>The filter \<open>at'\<close>\<close>
abbreviation (in ccpo) "compact_element \<equiv> ccpo.compact Sup (\<le>)"
lemma tendsto_unique_eventually:
fixes x x' :: "'a :: t2_space"
shows "F \<noteq> bot \<Longrightarrow> eventually (\<lambda>x. f x = g x) F \<Longrightarrow> (f \<longlongrightarrow> x) F \<Longrightarrow> (g \<longlongrightarrow> x') F \<Longrightarrow> x = x'"
by (metis tendsto_unique filterlim_cong)
lemma (in ccpo) ccpo_Sup_upper2: "chain C \<Longrightarrow> x \<in> C \<Longrightarrow> y \<le> x \<Longrightarrow> y \<le> Sup C"
by (blast intro: ccpo_Sup_upper order_trans)
lemma tendsto_open_vimage: "(\<And>B. open B \<Longrightarrow> open (f -` B)) \<Longrightarrow> f \<midarrow>l\<rightarrow> f l"
using continuous_on_open_vimage[of UNIV f] continuous_on_def[of UNIV f] by simp
lemma open_vimageI: "(\<And>x. f \<midarrow>x\<rightarrow> f x) \<Longrightarrow> open A \<Longrightarrow> open (f -` A)"
using continuous_on_open_vimage[of UNIV f] continuous_on_def[of UNIV f] by simp
lemma principal_bot: "principal x = bot \<longleftrightarrow> x = {}"
by (auto simp: filter_eq_iff eventually_principal)
definition "at' x = (if open {x} then principal {x} else at x)"
lemma at'_bot: "at' x \<noteq> bot"
by (simp add: at'_def at_eq_bot_iff principal_bot)
lemma tendsto_id_at'[simp, intro]: "((\<lambda>x. x) \<longlongrightarrow> x) (at' x)"
by (simp add: at'_def topological_tendstoI eventually_principal tendsto_ident_at)
lemma cont_at': "(f \<longlongrightarrow> f x) (at' x) \<longleftrightarrow> f \<midarrow>x\<rightarrow> f x"
using at_eq_bot_iff[of x] by (auto split: if_split_asm intro!: topological_tendstoI simp: eventually_principal at'_def)
subsection \<open>The type class \<open>ccpo_topology\<close>\<close>
text \<open>Temporarily relax type constraints for @{term "open"}.\<close>
setup \<open>Sign.add_const_constraint
(@{const_name "open"}, SOME @{typ "'a::open set \<Rightarrow> bool"})\<close>
class ccpo_topology = "open" + ccpo +
assumes open_ccpo: "open A \<longleftrightarrow> (\<forall>C. chain C \<longrightarrow> C \<noteq> {} \<longrightarrow> Sup C \<in> A \<longrightarrow> C \<inter> A \<noteq> {})"
begin
lemma open_ccpoD:
assumes "open A" "chain C" "C \<noteq> {}" "Sup C \<in> A"
shows "\<exists>c\<in>C. \<forall>c'\<in>C. c \<le> c' \<longrightarrow> c' \<in> A"
proof (rule ccontr)
assume "\<not> ?thesis"
then have *: "\<And>c. c \<in> C \<Longrightarrow> \<exists>c'\<in>C. c \<le> c' \<and> c' \<notin> A"
by auto
with \<open>chain C\<close> \<open>C \<noteq> {}\<close> have "chain (C - A)" "C - A \<noteq> {}"
by (auto intro: chain_Diff)
moreover have "Sup C = Sup (C - A)"
- proof (safe intro!: antisym ccpo_Sup_least \<open>chain C\<close> chain_Diff)
+ proof (safe intro!: order.antisym ccpo_Sup_least \<open>chain C\<close> chain_Diff)
fix c assume "c \<in> C"
with * obtain c' where "c' \<in> C" "c \<le> c'" "c' \<notin> A"
by auto
with \<open>c\<in>C\<close> show "c \<le> \<Squnion>(C - A)"
by (intro ccpo_Sup_upper2 \<open>chain (C - A)\<close>) auto
qed (auto intro: \<open>chain C\<close> ccpo_Sup_upper)
ultimately show False
using \<open>open A\<close> \<open>Sup C \<in> A\<close> by (auto simp: open_ccpo)
qed
lemma open_ccpo_Iic: "open {.. b}"
by (auto simp: open_ccpo) (metis Int_iff atMost_iff ccpo_Sup_upper empty_iff order_trans)
subclass topological_space
proof
show "open (UNIV::'a set)"
unfolding open_ccpo by auto
next
fix S T :: "'a set" assume "open S" "open T"
show "open (S \<inter> T)"
unfolding open_ccpo
proof (intro allI impI)
fix C assume C: "chain C" "C \<noteq> {}" and "\<Squnion>C \<in> S \<inter> T"
with open_ccpoD[OF \<open>open S\<close> C] open_ccpoD[OF \<open>open T\<close> C]
show "C \<inter> (S \<inter> T) \<noteq> {}"
unfolding chain_def by blast
qed
next
fix K :: "'a set set" assume *: "\<forall>D\<in>K. open D"
show "open (\<Union>K)"
unfolding open_ccpo
proof (intro allI impI)
fix C assume "chain C" "C \<noteq> {}" "\<Squnion>C \<in> (\<Union>K)"
with * obtain D where "D \<in> K" "\<Squnion>C \<in> D" "C \<inter> D \<noteq> {}"
by (auto simp: open_ccpo)
then show "C \<inter> (\<Union>K) \<noteq> {}"
by auto
qed
qed
lemma closed_ccpo: "closed A \<longleftrightarrow> (\<forall>C. chain C \<longrightarrow> C \<noteq> {} \<longrightarrow> C \<subseteq> A \<longrightarrow> Sup C \<in> A)"
unfolding closed_def open_ccpo by auto
lemma closed_admissible: "closed {x. P x} \<longleftrightarrow> ccpo.admissible Sup (\<le>) P"
unfolding closed_ccpo ccpo.admissible_def by auto
lemma open_singletonI_compact: "compact_element x \<Longrightarrow> open {x}"
using admissible_compact_neq[of Sup "(\<le>)" x]
by (simp add: closed_admissible[symmetric] open_closed Collect_neg_eq)
lemma closed_Ici: "closed {.. b}"
by (auto simp: closed_ccpo intro: ccpo_Sup_least)
lemma closed_Iic: "closed {b ..}"
by (auto simp: closed_ccpo intro: ccpo_Sup_upper2)
text \<open>
@{class ccpo_topology}s are also @{class t2_space}s.
This is necessary to have a unique continuous extension.
\<close>
subclass t2_space
proof
fix x y :: 'a assume "x \<noteq> y"
show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
proof cases
{ fix x y assume "x \<noteq> y" "x \<le> y"
then have "open {..x} \<and> open (- {..x}) \<and> x \<in> {..x} \<and> y \<in> - {..x} \<and> {..x} \<inter> - {..x} = {}"
by (auto intro: open_ccpo_Iic closed_Ici) }
moreover assume "x \<le> y \<or> y \<le> x"
ultimately show ?thesis
using \<open>x \<noteq> y\<close> by (metis Int_commute)
next
assume "\<not> (x \<le> y \<or> y \<le> x)"
then have "open ({..x} \<inter> - {..y}) \<and> open ({..y} \<inter> - {..x}) \<and>
x \<in> {..x} \<inter> - {..y} \<and> y \<in> {..y} \<inter> - {..x} \<and> ({..x} \<inter> - {..y}) \<inter> ({..y} \<inter> - {..x}) = {}"
by (auto intro: open_ccpo_Iic closed_Ici)
then show ?thesis by auto
qed
qed
end
lemma tendsto_le_ccpo:
fixes f g :: "'a \<Rightarrow> 'b::ccpo_topology"
assumes F: "\<not> trivial_limit F"
assumes x: "(f \<longlongrightarrow> x) F" and y: "(g \<longlongrightarrow> y) F"
assumes ev: "eventually (\<lambda>x. g x \<le> f x) F"
shows "y \<le> x"
proof (rule ccontr)
assume "\<not> y \<le> x"
show False
proof cases
assume "x \<le> y"
with \<open>\<not> y \<le> x\<close>
have "open {..x}" "open (- {..x})" "x \<in> {..x}" "y \<in> - {..x}" "{..x} \<inter> - {..x} = {}"
by (auto intro: open_ccpo_Iic closed_Ici)
with topological_tendstoD[OF x, of "{..x}"] topological_tendstoD[OF y, of "- {..x}"]
have "eventually (\<lambda>z. f z \<le> x) F" "eventually (\<lambda>z. \<not> g z \<le> x) F"
by auto
with ev have "eventually (\<lambda>x. False) F" by eventually_elim (auto intro: order_trans)
with F show False by (auto simp: eventually_False)
next
assume "\<not> x \<le> y"
with \<open>\<not> y \<le> x\<close> have "open ({..x} \<inter> - {..y})" "open ({..y} \<inter> - {..x})"
"x \<in> {..x} \<inter> - {..y}" "y \<in> {..y} \<inter> - {..x}" "({..x} \<inter> - {..y}) \<inter> ({..y} \<inter> - {..x}) = {}"
by (auto intro: open_ccpo_Iic closed_Ici)
with topological_tendstoD[OF x, of "{..x} \<inter> - {..y}"]
topological_tendstoD[OF y, of "{..y} \<inter> - {..x}"]
have "eventually (\<lambda>z. f z \<le> x \<and> \<not> f z \<le> y) F" "eventually (\<lambda>z. g z \<le> y \<and> \<not> g z \<le> x) F"
by auto
with ev have "eventually (\<lambda>x. False) F" by eventually_elim (auto intro: order_trans)
with F show False by (auto simp: eventually_False)
qed
qed
lemma tendsto_ccpoI:
fixes f :: "'a::ccpo_topology \<Rightarrow> 'b::ccpo_topology"
shows "(\<And>C. chain C \<Longrightarrow> C \<noteq> {} \<Longrightarrow> chain (f ` C) \<and> f (Sup C) = Sup (f`C)) \<Longrightarrow> f \<midarrow>x\<rightarrow> f x"
by (intro tendsto_open_vimage) (auto simp: open_ccpo)
lemma tendsto_mcont:
assumes mcont: "mcont Sup (\<le>) Sup (\<le>) (f :: 'a :: ccpo_topology \<Rightarrow> 'b :: ccpo_topology)"
shows "f \<midarrow>l\<rightarrow> f l"
proof (intro tendsto_ccpoI conjI)
fix C :: "'a set" assume C: "chain C" "C \<noteq> {}"
show "chain (f`C)"
using mcont
by (intro chain_imageI[where le_a="(\<le>)"] C) (simp add: mcont_def monotone_def)
show "f (\<Squnion>C) = \<Squnion>(f ` C)"
using mcont C by (simp add: mcont_def cont_def)
qed
subsection \<open>Instances for @{class ccpo_topology}s and continuity theorems\<close>
instantiation set :: (type) ccpo_topology
begin
definition open_set :: "'a set set \<Rightarrow> bool" where
"open_set A \<longleftrightarrow> (\<forall>C. chain C \<longrightarrow> C \<noteq> {} \<longrightarrow> Sup C \<in> A \<longrightarrow> C \<inter> A \<noteq> {})"
instance
by intro_classes (simp add: open_set_def)
end
instantiation enat :: ccpo_topology
begin
instance
proof
fix A :: "enat set"
show "open A = (\<forall>C. chain C \<longrightarrow> C \<noteq> {} \<longrightarrow> \<Squnion>C \<in> A \<longrightarrow> C \<inter> A \<noteq> {})"
proof (intro iffI allI impI)
fix C x assume "open A" "chain C" "C \<noteq> {}" "\<Squnion>C \<in> A"
show "C \<inter> A \<noteq> {}"
proof cases
assume "\<Squnion>C = \<infinity>"
with \<open>\<Squnion>C \<in> A\<close> \<open>open A\<close> obtain n where "{enat n <..} \<subseteq> A"
unfolding open_enat_iff by auto
with \<open>\<Squnion>C = \<infinity>\<close> Sup_eq_top_iff[of C] show ?thesis
by (auto simp: top_enat_def)
next
assume "\<Squnion>C \<noteq> \<infinity>"
then obtain n where "C \<subseteq> {.. enat n}"
unfolding Sup_eq_top_iff top_enat_def[symmetric] by (auto simp: not_less top_enat_def)
moreover have "finite {.. enat n}"
by (auto intro: finite_enat_bounded)
ultimately have "finite C"
by (auto intro: finite_subset)
from in_chain_finite[OF \<open>chain C\<close> \<open>finite C\<close> \<open>C \<noteq> {}\<close>] \<open>\<Squnion>C \<in> A\<close> show ?thesis
by auto
qed
next
assume C: "\<forall>C. chain C \<longrightarrow> C \<noteq> {} \<longrightarrow> \<Squnion>C \<in> A \<longrightarrow> C \<inter> A \<noteq> {}"
show "open A"
unfolding open_enat_iff
proof safe
assume "\<infinity> \<in> A"
{ fix C :: "enat set" assume "infinite C"
then have "\<Squnion>C = \<infinity>"
by (auto simp: Sup_enat_def)
with \<open>infinite C\<close> C[THEN spec, of C] \<open>\<infinity> \<in> A\<close> have "C \<inter> A \<noteq> {}"
by auto }
note inf_C = this
show "\<exists>x. {enat x<..} \<subseteq> A"
proof (rule ccontr)
assume "\<not> (\<exists>x. {enat x<..} \<subseteq> A)"
with \<open>\<infinity> \<in> A\<close> have "\<And>x. \<exists>y>x. enat y \<notin> A"
by (simp add: subset_eq Bex_def) (metis enat.exhaust enat_ord_simps(2))
then have "infinite {n. enat n \<notin> A}"
unfolding infinite_nat_iff_unbounded by auto
then have "infinite (enat ` {n. enat n \<notin> A})"
by (auto dest!: finite_imageD)
from inf_C[OF this] show False
by auto
qed
qed
qed
qed
end
lemmas tendsto_inf2[THEN tendsto_compose, tendsto_intros] =
tendsto_mcont[OF mcont_inf2]
lemma isCont_inf2[THEN isCont_o2[rotated]]:
"isCont (\<lambda>x. x \<sqinter> y) (z :: _ :: {ccpo_topology, complete_distrib_lattice})"
by(simp add: isCont_def tendsto_inf2 tendsto_ident_at)
lemmas tendsto_sup1[THEN tendsto_compose, tendsto_intros] =
tendsto_mcont[OF mcont_sup1]
lemma isCont_If: "isCont f x \<Longrightarrow> isCont g x \<Longrightarrow> isCont (\<lambda>x. if Q then f x else g x) x"
by (cases Q) auto
lemma isCont_enat_case: "isCont (f (epred n)) x \<Longrightarrow> isCont g x \<Longrightarrow> isCont (\<lambda>x. co.case_enat (g x) (\<lambda>n. f n x) n) x"
by (cases n rule: enat_coexhaust) auto
end
diff --git a/thys/Coinductive/Examples/Hamming_Stream.thy b/thys/Coinductive/Examples/Hamming_Stream.thy
--- a/thys/Coinductive/Examples/Hamming_Stream.thy
+++ b/thys/Coinductive/Examples/Hamming_Stream.thy
@@ -1,340 +1,340 @@
(* Title: Hamming_Stream.thy
Author: Andreas Lochbihler, ETH Zurich
*)
section \<open>The Hamming stream defined as a least fixpoint\<close>
theory Hamming_Stream imports
"../Coinductive_List"
"HOL-Computational_Algebra.Primes"
begin
lemma infinity_inf_enat [simp]:
fixes n :: enat
shows "\<infinity> \<sqinter> n = n" "n \<sqinter> \<infinity> = n"
by(simp_all add: inf_enat_def)
lemma eSuc_inf_eSuc [simp]: "eSuc n \<sqinter> eSuc m = eSuc (n \<sqinter> m)"
by(simp add: inf_enat_def)
lemma if_pull2: "(if b then f x x' else f y y') = f (if b then x else y) (if b then x' else y')"
by simp
context ord begin
primcorec lmerge :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist"
where
"lmerge xs ys =
(case xs of LNil \<Rightarrow> LNil | LCons x xs' \<Rightarrow>
case ys of LNil \<Rightarrow> LNil | LCons y ys' \<Rightarrow>
if lhd xs < lhd ys then LCons x (lmerge xs' ys)
else LCons y (if lhd ys < lhd xs then lmerge xs ys' else lmerge xs' ys'))"
lemma lnull_lmerge [simp]: "lnull (lmerge xs ys) \<longleftrightarrow> (lnull xs \<or> lnull ys)"
by(simp add: lmerge_def)
lemma lmerge_eq_LNil_iff: "lmerge xs ys = LNil \<longleftrightarrow> (xs = LNil \<or> ys = LNil)"
using lnull_lmerge unfolding lnull_def .
lemma lhd_lmerge: "\<lbrakk> \<not> lnull xs; \<not> lnull ys \<rbrakk> \<Longrightarrow> lhd (lmerge xs ys) = (if lhd xs < lhd ys then lhd xs else lhd ys)"
by(auto split: llist.split)
lemma ltl_lmerge:
"\<lbrakk> \<not> lnull xs; \<not> lnull ys \<rbrakk> \<Longrightarrow>
ltl (lmerge xs ys) =
(if lhd xs < lhd ys then lmerge (ltl xs) ys
else if lhd ys < lhd xs then lmerge xs (ltl ys)
else lmerge (ltl xs) (ltl ys))"
by(auto split: llist.split)
declare lmerge.sel [simp del]
lemma lmerge_simps:
"lmerge (LCons x xs) (LCons y ys) =
(if x < y then LCons x (lmerge xs (LCons y ys))
else if y < x then LCons y (lmerge (LCons x xs) ys)
else LCons y (lmerge xs ys))"
by(rule llist.expand)(simp_all add: lhd_lmerge ltl_lmerge)
lemma lmerge_LNil [simp]:
"lmerge LNil ys = LNil"
"lmerge xs LNil = LNil"
by simp_all
lemma lprefix_lmergeI:
"\<lbrakk> lprefix xs xs'; lprefix ys ys' \<rbrakk>
\<Longrightarrow> lprefix (lmerge xs ys) (lmerge xs' ys')"
by(coinduction arbitrary: xs xs' ys ys')(fastforce simp add: lhd_lmerge ltl_lmerge dest: lprefix_lhdD lprefix_lnullD simp add: not_lnull_conv split: if_split_asm)
lemma [partial_function_mono]:
assumes F: "mono_llist F" and G: "mono_llist G"
shows "mono_llist (\<lambda>f. lmerge (F f) (G f))"
by(blast intro: monotoneI lprefix_lmergeI monotoneD[OF F] monotoneD[OF G])
lemma in_lset_lmergeD: "x \<in> lset (lmerge xs ys) \<Longrightarrow> x \<in> lset xs \<or> x \<in> lset ys"
by(induct zs\<equiv>"lmerge xs ys" arbitrary: xs ys rule: llist_set_induct)(auto simp add: lhd_lmerge ltl_lmerge split: if_split_asm dest: in_lset_ltlD)
lemma lset_lmerge: "lset (lmerge xs ys) \<subseteq> lset xs \<union> lset ys"
by(auto dest: in_lset_lmergeD)
lemma lfinite_lmergeD: "lfinite (lmerge xs ys) \<Longrightarrow> lfinite xs \<or> lfinite ys"
by(induct zs\<equiv>"lmerge xs ys" arbitrary: xs ys rule: lfinite_induct)(fastforce simp add: ltl_lmerge if_pull2 split: if_split_asm)+
lemma fixes F
defines "F \<equiv> \<lambda>lmerge (xs, ys). case xs of LNil \<Rightarrow> LNil | LCons x xs' \<Rightarrow> case ys of LNil \<Rightarrow> LNil | LCons y ys' \<Rightarrow> (if x < y then LCons x (curry lmerge xs' ys) else if y < x then LCons y (curry lmerge xs ys') else LCons y (curry lmerge xs' ys'))"
shows lmerge_conv_fixp: "lmerge \<equiv> curry (ccpo.fixp (fun_lub lSup) (fun_ord lprefix) F)" (is "?lhs \<equiv> ?rhs")
and lmerge_mono: "mono_llist (\<lambda>lmerge. F lmerge xs)" (is "?mono xs")
proof(intro eq_reflection ext)
show mono: "\<And>xs. ?mono xs" unfolding F_def by(tactic \<open>Partial_Function.mono_tac @{context} 1\<close>)
fix xs ys
show "lmerge xs ys = ?rhs xs ys"
proof(coinduction arbitrary: xs ys)
case Eq_llist thus ?case
by(subst (1 3 4) llist.mono_body_fixp[OF mono])(auto simp add: F_def lmerge_simps split: llist.split)
qed
qed
lemma monotone_lmerge: "monotone (rel_prod lprefix lprefix) lprefix (case_prod lmerge)"
apply(rule llist.fixp_preserves_mono2[OF lmerge_mono lmerge_conv_fixp])
apply(erule conjE|rule allI conjI cont_intro|simp|erule allE, erule llist.mono2mono)+
done
lemma mono2mono_lmerge1 [THEN llist.mono2mono, cont_intro, simp]:
shows monotone_lmerge1: "monotone lprefix lprefix (\<lambda>xs. lmerge xs ys)"
by(simp add: monotone_lmerge[simplified])
lemma mono2mono_lmerge2 [THEN llist.mono2mono, cont_intro, simp]:
shows monotone_lmerge2: "monotone lprefix lprefix (\<lambda>ys. lmerge xs ys)"
by(simp add: monotone_lmerge[simplified])
lemma mcont_lmerge: "mcont (prod_lub lSup lSup) (rel_prod lprefix lprefix) lSup lprefix (case_prod lmerge)"
apply(rule llist.fixp_preserves_mcont2[OF lmerge_mono lmerge_conv_fixp])
apply(erule conjE|rule allI conjI cont_intro|simp|erule allE, erule llist.mcont2mcont)+
done
lemma mcont2mcont_lmerge1 [THEN llist.mcont2mcont, cont_intro, simp]:
shows mcont_lmerge1: "mcont lSup lprefix lSup lprefix (\<lambda>xs. lmerge xs ys)"
by(simp add: mcont_lmerge[simplified])
lemma mcont2mcont_lmerge2 [THEN llist.mcont2mcont, cont_intro, simp]:
shows mcont_lmerge2: "mcont lSup lprefix lSup lprefix (\<lambda>ys. lmerge xs ys)"
by(simp add: mcont_lmerge[simplified])
lemma lfinite_lmergeI [simp]: "\<lbrakk> lfinite xs; lfinite ys \<rbrakk> \<Longrightarrow> lfinite (lmerge xs ys)"
proof(induction arbitrary: ys rule: lfinite_induct)
case (LCons xs)
from LCons.prems LCons.hyps LCons.IH
show ?case by induct(clarsimp simp add: not_lnull_conv lmerge_simps)+
qed simp
lemma linfinite_lmerge [simp]: "\<lbrakk> \<not> lfinite xs; \<not> lfinite ys \<rbrakk> \<Longrightarrow> \<not> lfinite (lmerge xs ys)"
by(auto dest: lfinite_lmergeD)
lemma llength_lmerge_above: "llength xs \<sqinter> llength ys \<le> llength (lmerge xs ys)"
proof(induction xs arbitrary: ys)
case (LCons x xs)
show ?case
proof(induct ys)
case LCons thus ?case using LCons.IH
by(fastforce simp add: bot_enat_def[symmetric] lmerge_simps le_infI1 le_infI2 intro: order_trans[rotated])
qed(simp_all add: bot_enat_def[symmetric])
qed(simp_all add: bot_enat_def[symmetric])
end
context linorder begin
lemma in_lset_lmergeI1:
"\<lbrakk> x \<in> lset xs; lsorted xs; \<not> lfinite ys; \<exists>y\<in>lset ys. x \<le> y \<rbrakk>
\<Longrightarrow> x \<in> lset (lmerge xs ys)"
proof(induction arbitrary: ys rule: lset_induct)
case (find xs)
then obtain y where "y \<in> lset ys" "x \<le> y" by blast
thus ?case by induct(auto simp add: lmerge_simps not_less)
next
case (step x' xs)
then obtain y where "y \<in> lset ys" "x \<le> y" by blast
moreover from \<open>lsorted (LCons x' xs)\<close> \<open>x \<in> lset xs\<close> \<open>x \<noteq> x'\<close>
have "x' < x" "lsorted xs" by(auto simp add: lsorted_LCons)
ultimately show ?case using \<open>\<not> lfinite ys\<close>
by induct(auto 4 3 simp add: lmerge_simps \<open>x \<noteq> x'\<close> not_less intro: step.IH dest: in_lset_lmergeD)
qed
lemma in_lset_lmergeI2:
"\<lbrakk> x \<in> lset ys; lsorted ys; \<not> lfinite xs; \<exists>y\<in>lset xs. x \<le> y \<rbrakk>
\<Longrightarrow> x \<in> lset (lmerge xs ys)"
proof(induction arbitrary: xs rule: lset_induct)
case (find ys)
then obtain y where "y \<in> lset xs" "x \<le> y" by blast
thus ?case by induct(auto simp add: lmerge_simps not_less)
next
case (step x' ys)
then obtain y where "y \<in> lset xs" "x \<le> y" by blast
moreover from \<open>lsorted (LCons x' ys)\<close> \<open>x \<in> lset ys\<close> \<open>x \<noteq> x'\<close>
have "x' < x" "lsorted ys" by(auto simp add: lsorted_LCons)
ultimately show ?case using \<open>\<not> lfinite xs\<close>
by induct(auto 4 3 simp add: lmerge_simps \<open>x \<noteq> x'\<close> not_less intro: step.IH dest: in_lset_lmergeD)
qed
lemma lsorted_lmerge: "\<lbrakk> lsorted xs; lsorted ys \<rbrakk> \<Longrightarrow> lsorted (lmerge xs ys)"
proof(coinduction arbitrary: xs ys)
case (lsorted xs ys)
have ?lhd using lsorted
by(auto 4 3 simp add: not_lnull_conv lsorted_LCons lhd_lmerge ltl_lmerge dest!: in_lset_lmergeD)
moreover have ?ltl using lsorted by(auto simp add: ltl_lmerge intro: lsorted_ltlI)
ultimately show ?case ..
qed
lemma ldistinct_lmerge:
"\<lbrakk> lsorted xs; lsorted ys; ldistinct xs; ldistinct ys \<rbrakk>
\<Longrightarrow> ldistinct (lmerge xs ys)"
-by(coinduction arbitrary: xs ys)(auto 4 3 simp add: lhd_lmerge ltl_lmerge not_lnull_conv lsorted_LCons not_less dest!: in_lset_lmergeD dest: antisym)
+by(coinduction arbitrary: xs ys)(auto 4 3 simp add: lhd_lmerge ltl_lmerge not_lnull_conv lsorted_LCons not_less dest!: in_lset_lmergeD dest: order.antisym)
end
partial_function (llist) hamming' :: "unit \<Rightarrow> nat llist"
where
"hamming' _ =
LCons 1 (lmerge (lmap ((*) 2) (hamming' ())) (lmerge (lmap ((*) 3) (hamming' ())) (lmap ((*) 5) (hamming' ()))))"
definition hamming :: "nat llist"
where "hamming = hamming' ()"
lemma lnull_hamming [simp]: "\<not> lnull hamming"
unfolding hamming_def by(subst hamming'.simps) simp
lemma hamming_eq_LNil_iff [simp]: "hamming = LNil \<longleftrightarrow> False"
using lnull_hamming unfolding lnull_def by simp
lemma lhd_hamming [simp]: "lhd hamming = 1"
unfolding hamming_def by(subst hamming'.simps) simp
lemma ltl_hamming [simp]:
"ltl hamming = lmerge (lmap ((*) 2) hamming) (lmerge (lmap ((*) 3) hamming) (lmap ((*) 5) hamming))"
unfolding hamming_def by(subst hamming'.simps) simp
lemma hamming_unfold:
"hamming = LCons 1 (lmerge (lmap ((*) 2) hamming) (lmerge (lmap ((*) 3) hamming) (lmap ((*) 5) hamming)))"
by(rule llist.expand) simp_all
definition smooth :: "nat \<Rightarrow> bool"
where "smooth n \<longleftrightarrow> (\<forall>p. prime p \<longrightarrow> p dvd n \<longrightarrow> p \<le> 5)"
lemma smooth_0 [simp]: "\<not> smooth 0"
by(auto simp add: smooth_def intro: exI[where x=7])
lemma smooth_Suc0 [simp]: "smooth (Suc 0)"
by(auto simp add: smooth_def)
lemma smooth_gt0: "smooth n \<Longrightarrow> n > 0"
by(cases n) simp_all
lemma smooth_ge_Suc0: "smooth n \<Longrightarrow> n \<ge> Suc 0"
by(cases n) simp_all
lemma prime_nat_dvdD: "prime p \<Longrightarrow> (n :: nat) dvd p \<Longrightarrow> n = 1 \<or> n = p"
unfolding prime_nat_iff by simp
lemma smooth_times [simp]: "smooth (x * y) \<longleftrightarrow> smooth x \<and> smooth y"
by(auto simp add: smooth_def prime_dvd_mult_iff)
lemma smooth2 [simp]: "smooth 2"
by(auto simp add: smooth_def dest: prime_nat_dvdD[of 2, simplified])
lemma smooth3 [simp]: "smooth 3"
by(auto simp add: smooth_def dest: prime_nat_dvdD[of 3, simplified])
lemma smooth5 [simp]: "smooth 5"
by(auto simp add: smooth_def dest: prime_nat_dvdD[of 5, simplified])
lemma hamming_in_smooth: "lset hamming \<subseteq> {n. smooth n}"
unfolding hamming_def by(induct rule: hamming'.fixp_induct)(auto 6 6 dest: in_lset_lmergeD)
lemma lfinite_hamming [simp]: "\<not> lfinite hamming"
proof
assume "lfinite hamming"
then obtain n where n: "llength hamming = enat n" unfolding lfinite_conv_llength_enat by fastforce
have "llength (lmap ((*) 3) hamming) \<sqinter> llength (lmap ((*) 5) hamming) \<le> llength (lmerge (lmap ((*) 3) hamming) (lmap ((*) 5) hamming))"
by(rule llength_lmerge_above)
hence "llength hamming \<le> \<dots>" by simp
moreover have "llength (lmap ((*) 2) hamming) \<sqinter> \<dots> \<le>
llength (lmerge (lmap ((*) 2) hamming) (lmerge (lmap ((*) 3) hamming) (lmap ((*) 5) hamming)))"
by(rule llength_lmerge_above)
ultimately have "llength hamming \<le> \<dots>" by(simp add: inf.absorb1)
also from n have "\<dots> < enat n"
by(subst (asm) hamming_unfold)(cases n, auto simp add: zero_enat_def[symmetric] eSuc_enat[symmetric])
finally show False unfolding n by simp
qed
lemma lsorted_hamming [simp]: "lsorted hamming"
and ldistinct_hamming [simp]: "ldistinct hamming"
proof -
have "lsorted hamming \<and> ldistinct hamming \<and> lset hamming \<subseteq> {1..}"
unfolding hamming_def
by(induct rule: hamming'.fixp_induct)(auto 6 6 simp add: lsorted_LCons dest: in_lset_lmergeD intro: smooth_ge_Suc0 lsorted_lmerge lsorted_lmap monotoneI ldistinct_lmerge inj_onI)
thus "lsorted hamming" "ldistinct hamming" by simp_all
qed
lemma smooth_hamming:
assumes "smooth n"
shows "n \<in> lset hamming"
using assms
proof(induction n rule: less_induct)
have [simp]:
"monotone (\<le>) (\<le>) ((*) 2 :: nat \<Rightarrow> nat)"
"monotone (\<le>) (\<le>) ((*) 3 :: nat \<Rightarrow> nat)"
"monotone (\<le>) (\<le>) ((*) 5 :: nat \<Rightarrow> nat)"
by(simp_all add: monotone_def)
case (less n)
show ?case
proof(cases "n > 1")
case False
moreover from \<open>smooth n\<close> have "n \<noteq> 0" by(rule contrapos_pn) simp
ultimately have "n = 1" by(simp)
thus ?thesis by(subst hamming_unfold) simp
next
case True
hence "\<exists>p. prime p \<and> p dvd n" by(metis less_numeral_extra(4) prime_factor_nat)
with \<open>smooth n\<close> obtain p where "prime p" "p dvd n" "p \<le> 5" by(auto simp add: smooth_def)
from \<open>p dvd n\<close> obtain n' where n: "n = p * n'" by(auto simp add: dvd_def)
with \<open>smooth n\<close> \<open>prime p\<close> have "smooth n'" by(simp add: smooth_def)
with \<open>prime p\<close> n have "n' < n" by(simp add: smooth_gt0 prime_gt_Suc_0_nat)
hence n': "n' \<in> lset hamming" using \<open>smooth n'\<close> by(rule less.IH)
from \<open>p \<le> 5\<close> have "p = 0 \<or> p = 1 \<or> p = 2 \<or> p = 3 \<or> p = 4 \<or> p = 5"
by presburger
with \<open>prime p\<close> have "p = 2 \<or> p = 3 \<or> p = 5" by auto
thus ?thesis
proof(elim disjE)
assume "p = 2"
with n n' show ?thesis
by(subst hamming_unfold)(simp_all add: in_lset_lmergeI1 not_less lsorted_lmap bexI[where x="n'"] bexI[where x="3*n'"])
next
assume "p = 3"
moreover with n \<open>smooth n'\<close> have "2 * n' < n" by(simp add: smooth_gt0)
hence "2 * n' \<in> lset hamming" by(rule less.IH)(simp add: \<open>smooth n'\<close>)
ultimately show ?thesis using n n'
by(subst hamming_unfold)(auto 4 4 simp add: not_less lsorted_lmap lsorted_lmerge intro: in_lset_lmergeI1 in_lset_lmergeI2 bexI[where x="4*n'"] bexI[where x="5*n'"])
next
assume "p = 5"
moreover with n \<open>smooth n'\<close> have "2 * (2 * n') < n" by(simp add: smooth_gt0)
hence "2 * (2 * n') \<in> lset hamming" by(rule less.IH)(simp only: smooth_times smooth2 \<open>smooth n'\<close> simp_thms)
moreover from \<open>p = 5\<close> n \<open>smooth n'\<close> have "3 * n' < n" by(simp add: smooth_gt0)
hence "3 * n' \<in> lset hamming" by(rule less.IH)(simp add: \<open>smooth n'\<close>)
ultimately show ?thesis using n n'
by(subst hamming_unfold)(auto 4 4 simp add: lsorted_lmap lsorted_lmerge intro: in_lset_lmergeI2 bexI[where x="9*n'"] bexI[where x="8 * n'"])
qed
qed
qed
corollary hamming_smooth: "lset hamming = {n. smooth n}"
using hamming_in_smooth by(blast intro: smooth_hamming)
lemma hamming_THE:
"(THE xs. lsorted xs \<and> ldistinct xs \<and> lset xs = {n. smooth n}) = hamming"
by(rule the_equality)(simp_all add: hamming_smooth lsorted_ldistinct_lset_unique)
end
diff --git a/thys/Coinductive/TLList_CCPO.thy b/thys/Coinductive/TLList_CCPO.thy
--- a/thys/Coinductive/TLList_CCPO.thy
+++ b/thys/Coinductive/TLList_CCPO.thy
@@ -1,506 +1,506 @@
(* Title: TLList_CCPO.thy
Author: Andreas Lochbihler, ETH Zurich
*)
section \<open>Ccpo structure for terminated lazy lists\<close>
theory TLList_CCPO imports TLList begin
lemma Set_is_empty_parametric [transfer_rule]:
includes lifting_syntax
shows "(rel_set A ===> (=)) Set.is_empty Set.is_empty"
by(auto simp add: rel_fun_def Set.is_empty_def dest: rel_setD1 rel_setD2)
lemma monotone_comp: "\<lbrakk> monotone orda ordb g; monotone ordb ordc f \<rbrakk> \<Longrightarrow> monotone orda ordc (f \<circ> g)"
by(rule monotoneI)(simp add: monotoneD)
lemma cont_comp: "\<lbrakk> mcont luba orda lubb ordb g; cont lubb ordb lubc ordc f \<rbrakk> \<Longrightarrow> cont luba orda lubc ordc (f \<circ> g)"
apply(rule contI)
apply(frule (2) mcont_contD)
apply(simp)
apply(drule (1) contD[OF _ chain_imageI])
apply(erule (1) mcont_monoD)
apply(simp_all add: image_image o_def)
done
lemma mcont_comp: "\<lbrakk> mcont luba orda lubb ordb g; mcont lubb ordb lubc ordc f \<rbrakk> \<Longrightarrow> mcont luba orda lubc ordc (f \<circ> g)"
by(auto simp add: mcont_def intro: cont_comp monotone_comp)
context includes lifting_syntax
begin
lemma monotone_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_total A"
shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) monotone monotone"
unfolding monotone_def[abs_def] by transfer_prover
lemma cont_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_total A" "bi_unique B"
shows "((rel_set A ===> A) ===> (A ===> A ===> (=)) ===> (rel_set B ===> B) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) cont cont"
unfolding cont_def[abs_def] Set.is_empty_def[symmetric] by transfer_prover
lemma mcont_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_total A" "bi_unique B"
shows "((rel_set A ===> A) ===> (A ===> A ===> (=)) ===> (rel_set B ===> B) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) mcont mcont"
unfolding mcont_def[abs_def] by transfer_prover
end
lemma (in ccpo) Sup_Un_less:
assumes chain: "Complete_Partial_Order.chain (\<le>) (A \<union> B)"
and AB: "\<forall>x\<in>A. \<exists>y\<in>B. x \<le> y"
shows "Sup (A \<union> B) = Sup B"
-proof(rule antisym)
+proof(rule order.antisym)
from chain have chain': "Complete_Partial_Order.chain (\<le>) B"
by(blast intro: chain_subset)
show "Sup (A \<union> B) \<le> Sup B" using chain
proof(rule ccpo_Sup_least)
fix x
assume "x \<in> A \<union> B"
thus "x \<le> Sup B"
proof
assume "x \<in> A"
then obtain y where "x \<le> y" "y \<in> B" using AB by blast
note \<open>x \<le> y\<close>
also from chain' \<open>y \<in> B\<close> have "y \<le> Sup B" by(rule ccpo_Sup_upper)
finally show ?thesis .
qed(rule ccpo_Sup_upper[OF chain'])
qed
show "Sup B \<le> Sup (A \<union> B)"
using chain chain' by(blast intro: ccpo_Sup_least ccpo_Sup_upper)
qed
subsection \<open>The ccpo structure\<close>
context includes tllist.lifting fixes b :: 'b begin
lift_definition tllist_ord :: "('a, 'b) tllist \<Rightarrow> ('a, 'b) tllist \<Rightarrow> bool"
is "\<lambda>(xs1, b1) (xs2, b2). if lfinite xs1 then b1 = b \<and> lprefix xs1 xs2 \<or> xs1 = xs2 \<and> flat_ord b b1 b2 else xs1 = xs2"
by auto
lift_definition tSup :: "('a, 'b) tllist set \<Rightarrow> ('a, 'b) tllist"
is "\<lambda>A. (lSup (fst ` A), flat_lub b (snd ` (A \<inter> {(xs, _). lfinite xs})))"
proof goal_cases
case (1 A1 A2)
hence "fst ` A1 = fst ` A2" "snd ` (A1 \<inter> {(xs, _). lfinite xs}) = snd ` (A2 \<inter> {(xs, _). lfinite xs})"
by(auto 4 3 simp add: rel_set_def intro: rev_image_eqI)
thus ?case by simp
qed
lemma tllist_ord_simps [simp, code]: (* FIXME: does not work with transfer *)
shows tllist_ord_TNil_TNil: "tllist_ord (TNil b1) (TNil b2) \<longleftrightarrow> flat_ord b b1 b2"
and tllist_ord_TNil_TCons: "tllist_ord (TNil b1) (TCons y ys) \<longleftrightarrow> b1 = b"
and tllist_ord_TCons_TNil: "tllist_ord (TCons x xs) (TNil b2) \<longleftrightarrow> False"
and tllist_ord_TCons_TCons: "tllist_ord (TCons x xs) (TCons y ys) \<longleftrightarrow> x = y \<and> tllist_ord xs ys"
by(auto simp add: tllist_ord.rep_eq flat_ord_def)
lemma tllist_ord_refl [simp]: "tllist_ord xs xs"
by transfer(auto simp add: flat_ord_def)
lemma tllist_ord_antisym: "\<lbrakk> tllist_ord xs ys; tllist_ord ys xs \<rbrakk> \<Longrightarrow> xs = ys"
by transfer(auto simp add: flat_ord_def split: if_split_asm intro: lprefix_antisym)
lemma tllist_ord_trans [trans]: "\<lbrakk> tllist_ord xs ys; tllist_ord ys zs \<rbrakk> \<Longrightarrow> tllist_ord xs zs"
by transfer(auto simp add: flat_ord_def split: if_split_asm intro: lprefix_trans)
lemma chain_tllist_llist_of_tllist:
assumes "Complete_Partial_Order.chain tllist_ord A"
shows "Complete_Partial_Order.chain lprefix (llist_of_tllist ` A)"
by(rule chainI)(auto 4 3 simp add: tllist_ord.rep_eq split: if_split_asm dest: chainD[OF assms])
lemma chain_tllist_terminal:
assumes "Complete_Partial_Order.chain tllist_ord A"
shows "Complete_Partial_Order.chain (flat_ord b) {terminal xs|xs. xs \<in> A \<and> tfinite xs}"
by(rule chainI)(auto simp add: tllist_ord.rep_eq flat_ord_def dest: chainD[OF assms])
lemma flat_ord_chain_finite:
assumes "Complete_Partial_Order.chain (flat_ord b) A"
shows "finite A"
proof -
from assms have "\<exists>z. \<forall>x\<in>A. x = b \<or> x = z"
by(clarsimp simp add: chain_def flat_ord_def) metis
then obtain z where "\<And>x. x \<in> A \<Longrightarrow> x = b \<or> x = z" by blast
hence "A \<subseteq> {b, z}" by auto
thus ?thesis by(rule finite_subset) simp
qed
lemma tSup_empty [simp]: "tSup {} = TNil b"
by(transfer)(simp add: flat_lub_def)
lemma is_TNil_tSup [simp]: "is_TNil (tSup A) \<longleftrightarrow> (\<forall>x\<in>A. is_TNil x)"
by transfer(simp add: split_beta)
lemma chain_tllist_ord_tSup:
assumes chain: "Complete_Partial_Order.chain tllist_ord A"
and A: "xs \<in> A"
shows "tllist_ord xs (tSup A)"
proof(cases "tfinite xs")
case True
show ?thesis
proof(cases "llist_of_tllist xs = llist_of_tllist (tSup A)")
case True
with \<open>tfinite xs\<close> have "lfinite (lSup (llist_of_tllist ` A))"
by(simp add: tSup_def image_image)
hence "terminal (tSup A) = flat_lub b {terminal xs|xs. xs \<in> A \<and> tfinite xs}" (is "_ = flat_lub _ ?A")
by(simp add: tSup_def terminal_tllist_of_llist image_image)(auto intro: rev_image_eqI intro!: arg_cong[where f="flat_lub b"])
moreover have "flat_ord b (terminal xs) (flat_lub b ?A)"
by(rule ccpo.ccpo_Sup_upper[OF Partial_Function.ccpo[OF flat_interpretation]])(blast intro: chain_tllist_terminal[OF chain] A \<open>tfinite xs\<close>)+
ultimately show ?thesis using True by(simp add: tllist_ord.rep_eq)
next
case False
hence "\<exists>ys\<in>A. \<not> tllist_ord ys xs"
by(rule contrapos_np)(auto intro!: lprefix_antisym chain_lSup_lprefix chain_lprefix_lSup simp add: tSup_def image_image A chain_tllist_llist_of_tllist[OF chain] tllist_ord.rep_eq split: if_split_asm)
then obtain ys where "ys \<in> A" "\<not> tllist_ord ys xs" by blast
with A have "tllist_ord xs ys" "xs \<noteq> ys" by(auto dest: chainD[OF chain])
with True have "terminal xs = b" by transfer(auto simp add: flat_ord_def)
with True False show ?thesis
by(simp add: tllist_ord.rep_eq tSup_def image_image chain_lprefix_lSup chain_tllist_llist_of_tllist chain A)
qed
next
case False
thus ?thesis using assms by(simp add: tllist_ord.rep_eq tSup_def image_image chain_lprefix_lSup chain_tllist_llist_of_tllist not_lfinite_lprefix_conv_eq[THEN iffD1] terminal_tllist_of_llist split: if_split_asm)
qed
lemma chain_tSup_tllist_ord:
assumes chain: "Complete_Partial_Order.chain tllist_ord A"
and lub: "\<And>xs'. xs' \<in> A \<Longrightarrow> tllist_ord xs' xs"
shows "tllist_ord (tSup A) xs"
proof -
have "\<And>xs'. xs' \<in> llist_of_tllist ` A \<Longrightarrow> lprefix xs' (llist_of_tllist xs)"
by(auto dest!: lub simp add: tllist_ord.rep_eq split: if_split_asm)
with chain_tllist_llist_of_tllist[OF chain]
have prefix: "lprefix (lSup (llist_of_tllist ` A)) (llist_of_tllist xs)"
by(rule chain_lSup_lprefix)
show ?thesis
proof(cases "tfinite (tSup A)")
case False thus ?thesis using prefix
by(simp add: tllist_ord.rep_eq tSup_def image_image not_lfinite_lprefix_conv_eq[THEN iffD1])
next
case True
from True have fin: "lfinite (lSup (llist_of_tllist ` A))" by(simp add: tSup_def image_image)
have eq: "terminal (tSup A) = flat_lub b {terminal xs|xs. xs \<in> A \<and> tfinite xs}" (is "_ = flat_lub _ ?A")
by(simp add: tSup_def terminal_tllist_of_llist image_image fin)(auto intro: rev_image_eqI intro!: arg_cong[where f="flat_lub b"])
show ?thesis
proof(cases "lprefix (llist_of_tllist xs) (lSup (llist_of_tllist ` A))")
case True
with prefix have "lSup (llist_of_tllist ` A) = llist_of_tllist xs" by(rule lprefix_antisym)
moreover have "flat_ord b (flat_lub b ?A) (terminal xs)"
by(rule ccpo.ccpo_Sup_least[OF Partial_Function.ccpo[OF flat_interpretation]])(auto intro: chain_tllist_terminal[OF chain] dest: lub simp add: tllist_ord.rep_eq flat_ord_def)
ultimately show ?thesis using eq by(simp add: tllist_ord.rep_eq tSup_def image_image)
next
case False
{ fix xs'
assume "xs' \<in> A"
with False have "\<not> lprefix (llist_of_tllist xs) (llist_of_tllist xs')"
by-(erule contrapos_nn, auto 4 4 intro: lprefix_trans chain_lprefix_lSup chain_tllist_llist_of_tllist chain)
with lub[OF \<open>xs' \<in> A\<close>] have "terminal xs' = b"
by(auto simp add: tllist_ord.rep_eq split: if_split_asm) }
with chain_tllist_terminal[OF chain] have "flat_ord b (flat_lub b ?A) b"
by -(rule ccpo.ccpo_Sup_least[OF Partial_Function.ccpo[OF flat_interpretation]], auto simp add: flat_ord_def)
hence "flat_lub b ?A = b" by(simp add: flat_ord_def)
thus ?thesis using True eq prefix
by(simp add: tSup_def terminal_tllist_of_llist tllist_ord.rep_eq image_image)
qed
qed
qed
lemma tllist_ord_ccpo [simp, cont_intro]:
"class.ccpo tSup tllist_ord (mk_less tllist_ord)"
by unfold_locales(auto simp add: mk_less_def intro: tllist_ord_antisym tllist_ord_trans chain_tllist_ord_tSup chain_tSup_tllist_ord)
lemma tllist_ord_partial_function_definitions: "partial_function_definitions tllist_ord tSup"
by unfold_locales(auto simp add: mk_less_def intro: tllist_ord_antisym tllist_ord_trans chain_tllist_ord_tSup chain_tSup_tllist_ord)
interpretation tllist: partial_function_definitions "tllist_ord" "tSup"
by(rule tllist_ord_partial_function_definitions)
lemma admissible_mcont_is_TNil [THEN admissible_subst, cont_intro, simp]:
shows admissible_is_TNil: "ccpo.admissible tSup tllist_ord is_TNil"
by(rule ccpo.admissibleI)(simp)
lemma terminal_tSup:
"\<forall>xs\<in>Y. is_TNil xs \<Longrightarrow> terminal (tSup Y) = flat_lub b (terminal ` Y)"
including tllist.lifting apply transfer
apply (auto simp add: split_def)
apply (rule arg_cong [where f = "flat_lub b"])
apply force
done
lemma thd_tSup:
"\<exists>xs \<in> Y. \<not> is_TNil xs
\<Longrightarrow> thd (tSup Y) = (THE x. x \<in> thd ` (Y \<inter> {xs. \<not> is_TNil xs}))"
apply(simp add: tSup_def image_image)
apply(rule arg_cong[where f=The])
apply(auto intro: rev_image_eqI intro!: ext)
done
lemma ex_TCons_raw_parametric:
includes lifting_syntax
shows "(rel_set (rel_prod (llist_all2 A) B) ===> (=)) (\<lambda>Y. \<exists>(xs, b) \<in> Y. \<not> lnull xs) (\<lambda>Y. \<exists>(xs, b) \<in> Y. \<not> lnull xs)"
by(auto 4 4 simp add: rel_fun_def dest: rel_setD1 rel_setD2 llist_all2_lnullD intro: rev_bexI)
lift_definition ex_TCons :: "('a, 'b) tllist set \<Rightarrow> bool"
is "\<lambda>Y. \<exists>(xs, b) \<in> Y. \<not> lnull xs" parametric ex_TCons_raw_parametric
by (blast dest: rel_setD1 rel_setD2)+
lemma ex_TCons_iff: "ex_TCons Y \<longleftrightarrow> (\<exists>xs \<in> Y. \<not> is_TNil xs)"
by transfer auto
lemma retain_TCons_raw_parametric:
includes lifting_syntax
shows "(rel_set (rel_prod (llist_all2 A) B) ===> rel_set (rel_prod (llist_all2 A) B))
(\<lambda>A. A \<inter> {(xs, b). \<not> lnull xs}) (\<lambda>A. A \<inter> {(xs, b). \<not> lnull xs})"
by(rule rel_funI rel_setI)+(auto 4 4 dest: llist_all2_lnullD rel_setD2 rel_setD1 intro: rev_bexI)
lift_definition retain_TCons :: "('a, 'b) tllist set \<Rightarrow> ('a, 'b) tllist set"
is "\<lambda>A. A \<inter> {(xs, b). \<not> lnull xs}" parametric retain_TCons_raw_parametric
by(rule rel_setI)(fastforce dest: rel_setD1 rel_setD2)+
lemma retain_TCons_conv: "retain_TCons A = A \<inter> {xs. \<not> is_TNil xs}"
by(auto simp add: retain_TCons_def intro: rev_image_eqI)
lemma ttl_tSup:
"\<lbrakk> Complete_Partial_Order.chain tllist_ord Y; \<exists>xs \<in> Y. \<not> is_TNil xs \<rbrakk>
\<Longrightarrow> ttl (tSup Y) = tSup (ttl ` (Y \<inter> {xs. \<not> is_TNil xs}))"
unfolding ex_TCons_iff[symmetric] retain_TCons_conv[symmetric]
proof (transfer, goal_cases)
case prems: (1 Y)
then obtain xs b' where xsb: "(xs, b') \<in> Y" "\<not> lnull xs" by blast
note chain = prems(1)
have "flat_lub b (snd ` (Y \<inter> {(xs, _). lfinite xs})) = flat_lub b (insert b (snd ` (Y \<inter> {(xs, _). lfinite xs})))"
by(auto simp add: flat_lub_def)
also have "insert b (snd ` (Y \<inter> {(xs, _). lfinite xs})) = insert b (snd ` (apfst ltl ` (Y \<inter> {(xs, b). \<not> lnull xs}) \<inter> {(xs, _). lfinite xs}))"
apply(auto intro: rev_image_eqI)
apply(erule contrapos_np)
apply(frule chainD[OF chain \<open>(xs, b') \<in> Y\<close>])
using \<open>\<not> lnull xs\<close> xsb
by(fastforce split: if_split_asm simp add: lprefix_lnull intro!: rev_image_eqI)
also have "flat_lub b \<dots> = flat_lub b (snd ` (apfst ltl ` (Y \<inter> {(xs, b). \<not> lnull xs}) \<inter> {(xs, _). lfinite xs}))"
by(auto simp add: flat_lub_def)
also
have "ltl ` (fst ` Y \<inter> {xs. \<not> lnull xs}) = fst ` apfst ltl ` (Y \<inter> {(xs, b). \<not> lnull xs})"
by(auto simp add: image_image)
ultimately show ?case using prems by clarsimp
qed
lemma tSup_TCons: "A \<noteq> {} \<Longrightarrow> tSup (TCons x ` A) = TCons x (tSup A)"
unfolding Set.is_empty_def[symmetric]
apply transfer
unfolding Set.is_empty_def
apply(clarsimp simp add: image_image lSup_LCons[symmetric])
apply(rule arg_cong[where f="flat_lub b"])
apply(auto 4 3 intro: rev_image_eqI)
done
lemma tllist_ord_terminalD:
"\<lbrakk> tllist_ord xs ys; is_TNil ys \<rbrakk> \<Longrightarrow> flat_ord b (terminal xs) (terminal ys)"
by(cases xs)(auto simp add: is_TNil_def)
lemma tllist_ord_bot [simp]: "tllist_ord (TNil b) xs"
by(cases xs)(simp_all add: flat_ord_def)
lemma tllist_ord_ttlI:
"tllist_ord xs ys \<Longrightarrow> tllist_ord (ttl xs) (ttl ys)"
by(cases xs ys rule: tllist.exhaust[case_product tllist.exhaust])(simp_all)
lemma not_is_TNil_conv: "\<not> is_TNil xs \<longleftrightarrow> (\<exists>x xs'. xs = TCons x xs')"
by(cases xs) simp_all
subsection \<open>Continuity of predefined constants\<close>
lemma mono_tllist_ord_case:
fixes bot
assumes mono: "\<And>x. monotone tllist_ord ord (\<lambda>xs. f x xs (TCons x xs))"
and ord: "class.preorder ord (mk_less ord)"
and bot: "\<And>x. ord (bot b) x"
shows "monotone tllist_ord ord (\<lambda>xs. case xs of TNil b \<Rightarrow> bot b | TCons x xs' \<Rightarrow> f x xs' xs)"
proof -
interpret preorder ord "mk_less ord" by (fact ord)
show ?thesis by(rule monotoneI)(auto split: tllist.split dest: monotoneD[OF mono] simp add: flat_ord_def bot)
qed
lemma mcont_lprefix_case_aux:
fixes f bot ord
defines "g \<equiv> \<lambda>xs. f (thd xs) (ttl xs) (TCons (thd xs) (ttl xs))"
assumes mcont: "\<And>x. mcont tSup tllist_ord lub ord (\<lambda>xs. f x xs (TCons x xs))"
and ccpo: "class.ccpo lub ord (mk_less ord)"
and bot: "\<And>x. ord (bot b) x"
and cont_bot: "cont (flat_lub b) (flat_ord b) lub ord bot"
shows "mcont tSup tllist_ord lub ord (\<lambda>xs. case xs of TNil b \<Rightarrow> bot b | TCons x xs' \<Rightarrow> f x xs' xs)"
(is "mcont _ _ _ _ ?f")
proof(rule mcontI)
interpret b: ccpo lub ord "mk_less ord" by(rule ccpo)
show "cont tSup tllist_ord lub ord ?f"
proof(rule contI)
fix Y :: "('a, 'b) tllist set"
assume chain: "Complete_Partial_Order.chain tllist_ord Y"
and Y: "Y \<noteq> {}"
from chain have chain': "Complete_Partial_Order.chain ord (?f ` Y)"
by(rule chain_imageI)(auto split: tllist.split simp add: flat_ord_def intro!: bot mcont_monoD[OF mcont])
show "?f (tSup Y) = lub (?f ` Y)"
proof(cases "is_TNil (tSup Y)")
case True
hence "?f ` Y = bot ` terminal ` Y"
by(auto 4 3 split: tllist.split intro: rev_image_eqI intro!: imageI)
moreover from True have "tSup Y = TNil (flat_lub b (terminal ` Y))"
by -(rule tllist.expand, simp_all add: terminal_tSup)
moreover have "Complete_Partial_Order.chain (flat_ord b) (terminal ` Y)"
using chain True by(auto intro: chain_imageI tllist_ord_terminalD)
ultimately show ?thesis using Y
by (simp add: contD [OF cont_bot] cong del: b.SUP_cong_simp)
next
case False
hence eq: "tSup Y = TCons (thd (tSup Y)) (ttl (tSup Y))" by simp
have eq':
"?f ` Y =
(\<lambda>x. bot (terminal x)) ` (Y \<inter> {xs. is_TNil xs}) \<union>
(\<lambda>xs. f (thd xs) (ttl xs) xs) ` (Y \<inter> {xs. \<not> is_TNil xs})"
by(auto 4 3 split: tllist.splits intro: rev_image_eqI)
from False obtain xs where xs: "xs \<in> Y" "\<not> is_TNil xs" by auto
{ fix ys
assume "ys \<in> Y" "is_TNil ys"
hence "terminal ys = b" using xs
by(cases xs ys rule: tllist.exhaust[case_product tllist.exhaust])(auto dest: chainD[OF chain]) }
then have lub: "lub (?f ` Y) = lub ((\<lambda>xs. f (thd xs) (ttl xs) xs) ` (Y \<inter> {xs. \<not> is_TNil xs}))"
using xs chain' unfolding eq'
by -(erule ccpo.Sup_Un_less[OF ccpo], auto simp add: intro!: bot)
{ fix xs
assume xs: "xs \<in> Y \<inter> {xs. \<not> is_TNil xs}"
hence "(THE x. x \<in> thd ` (Y \<inter> {xs. \<not> is_TNil xs})) = thd xs"
by(auto dest: chainD[OF chain] simp add: not_is_TNil_conv intro!: the_equality)
hence "f (THE x. x \<in> thd ` (Y \<inter> {xs. \<not> is_TNil xs})) (ttl xs) (TCons (THE x. x \<in> thd ` (Y \<inter> {xs. \<not> is_TNil xs})) (ttl xs)) = f (thd xs) (ttl xs) xs"
using xs by simp }
moreover have "Complete_Partial_Order.chain tllist_ord (ttl ` (Y \<inter> {xs. \<not> is_TNil xs}))"
using chain by(rule chain_imageI[OF chain_subset])(auto simp add: tllist_ord_ttlI)
moreover from False have "Y \<inter> {xs. \<not> is_TNil xs} \<noteq> {}" by auto
ultimately show ?thesis
apply(subst (1 2) eq)
using False
apply(simp add: thd_tSup ttl_tSup[OF chain] mcont_contD[OF mcont] image_image lub)
done
qed
qed
from mcont_mono[OF mcont] _ bot
show "monotone tllist_ord ord ?f"
by(rule mono_tllist_ord_case)(unfold_locales)
qed
lemma cont_TNil [simp, cont_intro]: "cont (flat_lub b) (flat_ord b) tSup tllist_ord TNil"
apply(rule contI)
apply transfer
apply(simp add: image_image image_constant_conv)
apply(rule arg_cong[where f="flat_lub b"])
apply(auto intro: rev_image_eqI)
done
lemma monotone_TCons: "monotone tllist_ord tllist_ord (TCons x)"
by(rule monotoneI) simp
lemmas mono2mono_TCons[cont_intro] = monotone_TCons[THEN tllist.mono2mono]
lemma mcont_TCons: "mcont tSup tllist_ord tSup tllist_ord (TCons x)"
apply(rule mcontI)
apply(rule monotone_TCons)
apply(rule contI)
apply(simp add: tSup_TCons)
done
lemmas mcont2mcont_TCons[cont_intro] = mcont_TCons[THEN tllist.mcont2mcont]
lemmas [transfer_rule del] = tllist_ord.transfer tSup.transfer
lifting_update tllist.lifting
lifting_forget tllist.lifting
lemmas [transfer_rule] = tllist_ord.transfer tSup.transfer
lemma mono2mono_tset[THEN lfp.mono2mono, cont_intro]:
shows smonotone_tset: "monotone tllist_ord (\<subseteq>) tset"
including tllist.lifting
by transfer(rule monotone_comp[OF _ monotone_lset], auto intro: monotoneI)
lemma mcont2mcont_tset [THEN lfp.mcont2mcont, cont_intro]:
shows mcont_tset: "mcont tSup tllist_ord Union (\<subseteq>) tset"
including tllist.lifting
apply transfer
apply(rule mcont_comp[OF _ mcont_lset])
unfolding mcont_def by(auto intro: monotoneI contI )
end
context includes lifting_syntax
begin
lemma rel_fun_lift:
"(\<And>x. A (f x) (g x)) \<Longrightarrow> ((=) ===> A) f g"
by(simp add: rel_fun_def)
lemma tllist_ord_transfer [transfer_rule]:
"((=) ===> pcr_tllist (=) (=) ===> pcr_tllist (=) (=) ===> (=))
(\<lambda>b (xs1, b1) (xs2, b2). if lfinite xs1 then b1 = b \<and> lprefix xs1 xs2 \<or> xs1 = xs2 \<and> flat_ord b b1 b2 else xs1 = xs2)
tllist_ord"
by(rule rel_fun_lift)(rule tllist_ord.transfer)
lemma tSup_transfer [transfer_rule]:
"((=) ===> rel_set (pcr_tllist (=) (=)) ===> pcr_tllist (=) (=))
(\<lambda>b A. (lSup (fst ` A), flat_lub b (snd ` (A \<inter> {(xs, _). lfinite xs}))))
tSup"
by(rule rel_fun_lift)(rule tSup.transfer)
end
lifting_update tllist.lifting
lifting_forget tllist.lifting
interpretation tllist: partial_function_definitions "tllist_ord b" "tSup b" for b
by(rule tllist_ord_partial_function_definitions)
lemma tllist_case_mono [partial_function_mono, cont_intro]:
assumes tnil: "\<And>b. monotone orda ordb (\<lambda>f. tnil f b)"
and tcons: "\<And>x xs. monotone orda ordb (\<lambda>f. tcons f x xs)"
shows "monotone orda ordb (\<lambda>f. case_tllist (tnil f) (tcons f) xs)"
by(rule monotoneI)(auto split: tllist.split dest: monotoneD[OF tnil] monotoneD[OF tcons])
subsection \<open>Definition of recursive functions\<close>
locale tllist_pf = fixes b :: 'b
begin
declaration \<open>Partial_Function.init "tllist" @{term "tllist.fixp_fun b"}
@{term "tllist.mono_body b"} @{thm tllist.fixp_rule_uc[where b=b]} @{thm tllist.fixp_induct_uc[where b=b]} NONE\<close>
abbreviation mono_tllist where "mono_tllist \<equiv> monotone (fun_ord (tllist_ord b)) (tllist_ord b)"
lemma LCons_mono [partial_function_mono, cont_intro]:
"mono_tllist A \<Longrightarrow> mono_tllist (\<lambda>f. TCons x (A f))"
by(rule monotoneI)(auto dest: monotoneD)
end
lemma mono_tllist_lappendt2 [partial_function_mono]:
"tllist_pf.mono_tllist b A \<Longrightarrow> tllist_pf.mono_tllist b (\<lambda>f. lappendt xs (A f))"
apply(rule monotoneI)
apply(drule (1) monotoneD)
apply(simp add: tllist_ord.rep_eq split: if_split_asm)
apply(auto simp add: lappend_inf)
done
lemma mono_tllist_tappend2 [partial_function_mono]:
assumes "\<And>y. tllist_pf.mono_tllist b (C y)"
shows "tllist_pf.mono_tllist b (\<lambda>f. tappend xs (\<lambda>y. C y f))"
apply(cases "tfinite xs")
apply(rule monotoneI)
apply(drule monotoneD[OF assms[where y="terminal xs"]])
including tllist.lifting
apply transfer
apply(fastforce split: if_split_asm)
apply(simp add: tappend_inf)
done
end
diff --git a/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy b/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy
--- a/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy
+++ b/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy
@@ -1,115 +1,115 @@
section \<open>Refinement Lattice \label{S:lattice}\<close>
theory Refinement_Lattice
imports
Main
"HOL-Library.Lattice_Syntax"
begin
text \<open>
The underlying lattice of commands is complete and distributive.
We follow the refinement calculus tradition so that $\nondet$
is non-deterministic choice and $c \refsto d$ means $c$ is refined
(or implemented) by $d$.
\<close>
declare [[show_sorts]]
text \<open>Remove existing notation for quotient as it interferes with the rely quotient\<close>
no_notation Equiv_Relations.quotient (infixl "'/'/" 90)
class refinement_lattice = complete_distrib_lattice
begin
text \<open>The refinement lattice infimum corresponds to non-deterministic choice for commands.\<close>
abbreviation
refine :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
where
"c \<sqsubseteq> d \<equiv> less_eq c d"
abbreviation
refine_strict :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubset>" 50)
where
"c \<sqsubset> d \<equiv> less c d"
text \<open>Non-deterministic choice is monotonic in both arguments\<close>
lemma inf_mono_left: "a \<sqsubseteq> b \<Longrightarrow> a \<sqinter> c \<sqsubseteq> b \<sqinter> c"
using inf_mono by auto
lemma inf_mono_right: "c \<sqsubseteq> d \<Longrightarrow> a \<sqinter> c \<sqsubseteq> a \<sqinter> d"
using inf_mono by auto
text \<open>Binary choice is a special case of choice over a set.\<close>
lemma Inf2_inf: "\<Sqinter>{ f x | x. x \<in> {c, d}} = f c \<sqinter> f d"
proof -
have "{ f x | x. x \<in> {c, d}} = {f c, f d}" by blast
then have "\<Sqinter>{ f x | x. x \<in> {c, d}} = \<Sqinter>{f c, f d}" by simp
also have "... = f c \<sqinter> f d" by simp
finally show ?thesis .
qed
text \<open>Helper lemma for choice over indexed set.\<close>
lemma INF_Inf: "(\<Sqinter>x\<in>X. f x) = (\<Sqinter>{f x |x. x \<in> X})"
by (simp add: Setcompr_eq_image)
lemma (in -) INF_absorb_args: "(\<Sqinter>i j. (f::nat \<Rightarrow> 'c::complete_lattice) (i + j)) = (\<Sqinter>k. f k)"
proof (rule order_class.order.antisym)
show "(\<Sqinter>k. f k) \<le> (\<Sqinter>i j. f (i + j))"
by (simp add: complete_lattice_class.INF_lower complete_lattice_class.le_INF_iff)
next
have "\<And>k. \<exists>i j. f (i + j) \<le> f k"
- by (metis add.left_neutral order_class.eq_iff)
+ by (metis Nat.add_0_right order_refl)
then have "\<And>k. \<exists>i. (\<Sqinter>j. f (i + j)) \<le> f k"
by (meson UNIV_I complete_lattice_class.INF_lower2)
then show "(\<Sqinter>i j. f (i + j)) \<le> (\<Sqinter>k. f k)"
by (simp add: complete_lattice_class.INF_mono)
qed
lemma (in -) nested_Collect: "{f y |y. y \<in> {g x |x. x \<in> X}} = {f (g x) |x. x \<in> X}"
by blast
text \<open>A transition lemma for INF distributivity properties, going from Inf to INF,
qualified version followed by a straightforward one.\<close>
lemma Inf_distrib_INF_qual:
fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
assumes qual: "P {d x |x. x \<in> X}"
assumes f_Inf_distrib: "\<And>c D. P D \<Longrightarrow> f c (\<Sqinter> D) = \<Sqinter> {f c d | d . d \<in> D }"
shows "f c (\<Sqinter>x\<in>X. d x) = (\<Sqinter>x\<in>X. f c (d x))"
proof -
have "f c (\<Sqinter>x\<in>X. d x) = f c (\<Sqinter>{d x |x. x \<in> X})" by (simp add: INF_Inf)
also have "... = (\<Sqinter>{f c dx |dx. dx \<in> {d x | x. x \<in> X}})" by (simp add: qual f_Inf_distrib)
also have "... = (\<Sqinter>{f c (d x) |x. x \<in> X})" by (simp only: nested_Collect)
also have "... = (\<Sqinter>x\<in>X. f c (d x))" by (simp add: INF_Inf)
finally show ?thesis .
qed
lemma Inf_distrib_INF:
fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
assumes f_Inf_distrib: "\<And>c D. f c (\<Sqinter> D) = \<Sqinter> {f c d | d . d \<in> D }"
shows "f c (\<Sqinter>x\<in>X. d x) = (\<Sqinter>x\<in>X. f c (d x))"
by (simp add: Setcompr_eq_image f_Inf_distrib image_comp)
end
lemmas refine_trans = order.trans
text \<open>More transitivity rules to make calculational reasoning smoother\<close>
declare ord_eq_le_trans[trans]
declare ord_le_eq_trans[trans]
declare dual_order.trans[trans]
abbreviation
dist_over_sup :: "('a::refinement_lattice \<Rightarrow> 'a) \<Rightarrow> bool"
where
"dist_over_sup F \<equiv> (\<forall> X . F (\<Squnion> X) = (\<Squnion>x\<in>X. F (x)))"
abbreviation
dist_over_inf :: "('a::refinement_lattice \<Rightarrow> 'a) \<Rightarrow> bool"
where
"dist_over_inf F \<equiv> (\<forall> X . F (\<Sqinter> X) = (\<Sqinter>x\<in>X. F (x)))"
end
diff --git a/thys/Containers/RBT_ext.thy b/thys/Containers/RBT_ext.thy
--- a/thys/Containers/RBT_ext.thy
+++ b/thys/Containers/RBT_ext.thy
@@ -1,358 +1,358 @@
(* Title: Containers/RBT_ext.thy
Author: Andreas Lochbihler, KIT *)
theory RBT_ext
imports
"HOL-Library.RBT_Impl"
Containers_Auxiliary
List_Fusion
begin
section \<open>More on red-black trees\<close>
subsection \<open>More lemmas\<close>
context linorder begin
lemma is_rbt_fold_rbt_insert_impl:
"is_rbt t \<Longrightarrow> is_rbt (RBT_Impl.fold rbt_insert t' t)"
by(simp add: rbt_insert_def is_rbt_fold_rbt_insertwk)
lemma rbt_sorted_fold_insert: "rbt_sorted t \<Longrightarrow> rbt_sorted (RBT_Impl.fold rbt_insert t' t)"
by(induct t' arbitrary: t)(simp_all add: rbt_insert_rbt_sorted)
lemma rbt_lookup_rbt_insert': "rbt_sorted t \<Longrightarrow> rbt_lookup (rbt_insert k v t) = rbt_lookup t(k \<mapsto> v)"
by(simp add: rbt_insert_def rbt_lookup_rbt_insertwk fun_eq_iff split: option.split)
lemma rbt_lookup_fold_rbt_insert_impl:
"rbt_sorted t2 \<Longrightarrow>
rbt_lookup (RBT_Impl.fold rbt_insert t1 t2) = rbt_lookup t2 ++ map_of (rev (RBT_Impl.entries t1))"
proof(induction t1 arbitrary: t2)
case Empty thus ?case by simp
next
case (Branch c l x k r)
show ?case using Branch.prems
by(simp add: map_add_def Branch.IH rbt_insert_rbt_sorted rbt_sorted_fold_insert rbt_lookup_rbt_insert' fun_eq_iff split: option.split)
qed
end
subsection \<open>Build the cross product of two RBTs\<close>
context fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'd \<Rightarrow> 'e" begin
definition alist_product :: "('a \<times> 'b) list \<Rightarrow> ('c \<times> 'd) list \<Rightarrow> (('a \<times> 'c) \<times> 'e) list"
where "alist_product xs ys = concat (map (\<lambda>(a, b). map (\<lambda>(c, d). ((a, c), f a b c d)) ys) xs)"
lemma alist_product_simps [simp]:
"alist_product [] ys = []"
"alist_product xs [] = []"
"alist_product ((a, b) # xs) ys = map (\<lambda>(c, d). ((a, c), f a b c d)) ys @ alist_product xs ys"
by(simp_all add: alist_product_def)
lemma append_alist_product_conv_fold:
"zs @ alist_product xs ys = rev (fold (\<lambda>(a, b). fold (\<lambda>(c, d) rest. ((a, c), f a b c d) # rest) ys) xs (rev zs))"
proof(induction xs arbitrary: zs)
case Nil thus ?case by simp
next
case (Cons x xs)
obtain a b where x: "x = (a, b)" by(cases x)
have "\<And>zs. fold (\<lambda>(c, d). (#) ((a, c), f a b c d)) ys zs =
rev (map (\<lambda>(c, d). ((a, c), f a b c d)) ys) @ zs"
by(induct ys) auto
with Cons.IH[of "zs @ map (\<lambda>(c, d). ((a, c), f a b c d)) ys"] x
show ?case by simp
qed
lemma alist_product_code [code]:
"alist_product xs ys =
rev (fold (\<lambda>(a, b). fold (\<lambda>(c, d) rest. ((a, c), f a b c d) # rest) ys) xs [])"
using append_alist_product_conv_fold[of "[]" xs ys]
by simp
lemma set_alist_product:
"set (alist_product xs ys) =
(\<lambda>((a, b), (c, d)). ((a, c), f a b c d)) ` (set xs \<times> set ys)"
by(auto 4 3 simp add: alist_product_def intro: rev_image_eqI rev_bexI)
lemma distinct_alist_product:
"\<lbrakk> distinct (map fst xs); distinct (map fst ys) \<rbrakk>
\<Longrightarrow> distinct (map fst (alist_product xs ys))"
proof(induct xs)
case Nil thus ?case by simp
next
case (Cons x xs)
obtain a b where x: "x = (a, b)" by(cases x)
have "distinct (map (fst \<circ> (\<lambda>(c, d). ((a, c), f a b c d))) ys)"
using \<open>distinct (map fst ys)\<close> by(induct ys)(auto intro: rev_image_eqI)
thus ?case using Cons x by(auto simp add: set_alist_product intro: rev_image_eqI)
qed
lemma map_of_alist_product:
"map_of (alist_product xs ys) (a, c) =
(case map_of xs a of None \<Rightarrow> None
| Some b \<Rightarrow> map_option (f a b c) (map_of ys c))"
proof(induction xs)
case Nil thus ?case by simp
next
case (Cons x xs)
obtain a b where x: "x = (a, b)" by (cases x)
let ?map = "map (\<lambda>(c, d). ((a, c), f a b c d)) ys"
have "map_of ?map (a, c) = map_option (f a b c) (map_of ys c)"
by(induct ys) auto
moreover {
fix a' assume "a' \<noteq> a"
hence "map_of ?map (a', c) = None"
by(induct ys) auto }
ultimately show ?case using x Cons.IH
by(auto simp add: map_add_def split: option.split)
qed
definition rbt_product :: "('a, 'b) rbt \<Rightarrow> ('c, 'd) rbt \<Rightarrow> ('a \<times> 'c, 'e) rbt"
where
"rbt_product rbt1 rbt2 = rbtreeify (alist_product (RBT_Impl.entries rbt1) (RBT_Impl.entries rbt2))"
lemma rbt_product_code [code]:
"rbt_product rbt1 rbt2 =
rbtreeify (rev (RBT_Impl.fold (\<lambda>a b. RBT_Impl.fold (\<lambda>c d rest. ((a, c), f a b c d) # rest) rbt2) rbt1 []))"
unfolding rbt_product_def alist_product_code RBT_Impl.fold_def ..
end
context
fixes leq_a :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>\<^sub>a" 50)
and less_a :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubset>\<^sub>a" 50)
and leq_b :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix "\<sqsubseteq>\<^sub>b" 50)
and less_b :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix "\<sqsubset>\<^sub>b" 50)
assumes lin_a: "class.linorder leq_a less_a"
and lin_b: "class.linorder leq_b less_b"
begin
abbreviation (input) less_eq_prod' :: "('a \<times> 'b) \<Rightarrow> ('a \<times> 'b) \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
where "less_eq_prod' \<equiv> less_eq_prod leq_a less_a leq_b"
abbreviation (input) less_prod' :: "('a \<times> 'b) \<Rightarrow> ('a \<times> 'b) \<Rightarrow> bool" (infix "\<sqsubset>" 50)
where "less_prod' \<equiv> less_prod leq_a less_a less_b"
lemmas linorder_prod = linorder_prod[OF lin_a lin_b]
lemma sorted_alist_product:
assumes xs: "linorder.sorted leq_a (map fst xs)" "distinct (map fst xs)"
and ys: "linorder.sorted (\<sqsubseteq>\<^sub>b) (map fst ys)"
shows "linorder.sorted (\<sqsubseteq>) (map fst (alist_product f xs ys))"
proof -
interpret a: linorder "(\<sqsubseteq>\<^sub>a)" "(\<sqsubset>\<^sub>a)" by(fact lin_a)
note [simp] =
linorder.sorted.simps(1)[OF linorder_prod] linorder.sorted.simps(2)[OF linorder_prod]
linorder.sorted_append[OF linorder_prod]
linorder.sorted.simps(2)[OF lin_b]
show ?thesis using xs
proof(induction xs)
case Nil show ?case by simp
next
case (Cons x xs)
obtain a b where x: "x = (a, b)" by(cases x)
have "linorder.sorted (\<sqsubseteq>) (map fst (map (\<lambda>(c, d). ((a, c), f a b c d)) ys))"
using ys by(induct ys) auto
thus ?case using x Cons
- by(fastforce simp add: set_alist_product a.not_less dest: bspec a.antisym intro: rev_image_eqI)
+ by(fastforce simp add: set_alist_product a.not_less dest: bspec a.order_antisym intro: rev_image_eqI)
qed
qed
lemma is_rbt_rbt_product:
"\<lbrakk> ord.is_rbt (\<sqsubset>\<^sub>a) rbt1; ord.is_rbt (\<sqsubset>\<^sub>b) rbt2 \<rbrakk>
\<Longrightarrow> ord.is_rbt (\<sqsubset>) (rbt_product f rbt1 rbt2)"
unfolding rbt_product_def
by(blast intro: linorder.is_rbt_rbtreeify[OF linorder_prod] sorted_alist_product linorder.rbt_sorted_entries[OF lin_a] ord.is_rbt_rbt_sorted linorder.distinct_entries[OF lin_a] linorder.rbt_sorted_entries[OF lin_b] distinct_alist_product linorder.distinct_entries[OF lin_b])
lemma rbt_lookup_rbt_product:
"\<lbrakk> ord.is_rbt (\<sqsubset>\<^sub>a) rbt1; ord.is_rbt (\<sqsubset>\<^sub>b) rbt2 \<rbrakk>
\<Longrightarrow> ord.rbt_lookup (\<sqsubset>) (rbt_product f rbt1 rbt2) (a, c) =
(case ord.rbt_lookup (\<sqsubset>\<^sub>a) rbt1 a of None \<Rightarrow> None
| Some b \<Rightarrow> map_option (f a b c) (ord.rbt_lookup (\<sqsubset>\<^sub>b) rbt2 c))"
by(simp add: rbt_product_def linorder.rbt_lookup_rbtreeify[OF linorder_prod] linorder.is_rbt_rbtreeify[OF linorder_prod] sorted_alist_product linorder.rbt_sorted_entries[OF lin_a] ord.is_rbt_rbt_sorted linorder.distinct_entries[OF lin_a] linorder.rbt_sorted_entries[OF lin_b] distinct_alist_product linorder.distinct_entries[OF lin_b] map_of_alist_product linorder.map_of_entries[OF lin_a] linorder.map_of_entries[OF lin_b] cong: option.case_cong)
end
hide_const less_eq_prod' less_prod'
subsection \<open>Build an RBT where keys are paired with themselves\<close>
primrec RBT_Impl_diag :: "('a, 'b) rbt \<Rightarrow> ('a \<times> 'a, 'b) rbt"
where
"RBT_Impl_diag rbt.Empty = rbt.Empty"
| "RBT_Impl_diag (rbt.Branch c l k v r) = rbt.Branch c (RBT_Impl_diag l) (k, k) v (RBT_Impl_diag r)"
lemma entries_RBT_Impl_diag:
"RBT_Impl.entries (RBT_Impl_diag t) = map (\<lambda>(k, v). ((k, k), v)) (RBT_Impl.entries t)"
by(induct t) simp_all
lemma keys_RBT_Impl_diag:
"RBT_Impl.keys (RBT_Impl_diag t) = map (\<lambda>k. (k, k)) (RBT_Impl.keys t)"
by(simp add: RBT_Impl.keys_def entries_RBT_Impl_diag split_beta)
lemma rbt_sorted_RBT_Impl_diag:
"ord.rbt_sorted lt t \<Longrightarrow> ord.rbt_sorted (less_prod leq lt lt) (RBT_Impl_diag t)"
by(induct t)(auto simp add: ord.rbt_sorted.simps ord.rbt_less_prop ord.rbt_greater_prop keys_RBT_Impl_diag)
lemma bheight_RBT_Impl_diag:
"bheight (RBT_Impl_diag t) = bheight t"
by(induct t) simp_all
lemma inv_RBT_Impl_diag:
assumes "inv1 t" "inv2 t"
shows "inv1 (RBT_Impl_diag t)" "inv2 (RBT_Impl_diag t)"
and "color_of t = color.B \<Longrightarrow> color_of (RBT_Impl_diag t) = color.B"
using assms by(induct t)(auto simp add: bheight_RBT_Impl_diag)
lemma is_rbt_RBT_Impl_diag:
"ord.is_rbt lt t \<Longrightarrow> ord.is_rbt (less_prod leq lt lt) (RBT_Impl_diag t)"
by(simp add: ord.is_rbt_def rbt_sorted_RBT_Impl_diag inv_RBT_Impl_diag)
lemma (in linorder) rbt_lookup_RBT_Impl_diag:
"ord.rbt_lookup (less_prod (\<le>) (<) (<)) (RBT_Impl_diag t) =
(\<lambda>(k, k'). if k = k' then ord.rbt_lookup (<) t k else None)"
by(induct t)(auto simp add: ord.rbt_lookup.simps fun_eq_iff)
subsection \<open>Folding and quantifiers over RBTs\<close>
definition RBT_Impl_fold1 :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a, unit) RBT_Impl.rbt \<Rightarrow> 'a"
where "RBT_Impl_fold1 f rbt = fold f (tl (RBT_Impl.keys rbt)) (hd (RBT_Impl.keys rbt))"
lemma RBT_Impl_fold1_simps [simp, code]:
"RBT_Impl_fold1 f rbt.Empty = undefined"
"RBT_Impl_fold1 f (Branch c rbt.Empty k v r) = RBT_Impl.fold (\<lambda>k v. f k) r k"
"RBT_Impl_fold1 f (Branch c (Branch c' l' k' v' r') k v r) =
RBT_Impl.fold (\<lambda>k v. f k) r (f k (RBT_Impl_fold1 f (Branch c' l' k' v' r')))"
by(simp_all add: RBT_Impl_fold1_def RBT_Impl.keys_def fold_map RBT_Impl.fold_def split_def o_def tl_append hd_def split: list.split)
definition RBT_Impl_rbt_all :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"
where [code del]: "RBT_Impl_rbt_all P rbt = (\<forall>(k, v) \<in> set (RBT_Impl.entries rbt). P k v)"
lemma RBT_Impl_rbt_all_simps [simp, code]:
"RBT_Impl_rbt_all P rbt.Empty \<longleftrightarrow> True"
"RBT_Impl_rbt_all P (Branch c l k v r) \<longleftrightarrow> P k v \<and> RBT_Impl_rbt_all P l \<and> RBT_Impl_rbt_all P r"
by(auto simp add: RBT_Impl_rbt_all_def)
definition RBT_Impl_rbt_ex :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool"
where [code del]: "RBT_Impl_rbt_ex P rbt = (\<exists>(k, v) \<in> set (RBT_Impl.entries rbt). P k v)"
lemma RBT_Impl_rbt_ex_simps [simp, code]:
"RBT_Impl_rbt_ex P rbt.Empty \<longleftrightarrow> False"
"RBT_Impl_rbt_ex P (Branch c l k v r) \<longleftrightarrow> P k v \<or> RBT_Impl_rbt_ex P l \<or> RBT_Impl_rbt_ex P r"
by(auto simp add: RBT_Impl_rbt_ex_def)
subsection \<open>List fusion for RBTs\<close>
type_synonym ('a, 'b, 'c) rbt_generator_state = "('c \<times> ('a, 'b) RBT_Impl.rbt) list \<times> ('a, 'b) RBT_Impl.rbt"
fun rbt_has_next :: "('a, 'b, 'c) rbt_generator_state \<Rightarrow> bool"
where
"rbt_has_next ([], rbt.Empty) = False"
| "rbt_has_next _ = True"
fun rbt_keys_next :: "('a, 'b, 'a) rbt_generator_state \<Rightarrow> 'a \<times> ('a, 'b, 'a) rbt_generator_state"
where
"rbt_keys_next ((k, t) # kts, rbt.Empty) = (k, kts, t)"
| "rbt_keys_next (kts, rbt.Branch c l k v r) = rbt_keys_next ((k, r) # kts, l)"
lemma rbt_generator_induct [case_names empty split shuffle]:
assumes "P ([], rbt.Empty)"
and "\<And>k t kts. P (kts, t) \<Longrightarrow> P ((k, t) # kts, rbt.Empty)"
and "\<And>kts c l k v r. P ((f k v, r) # kts, l) \<Longrightarrow> P (kts, Branch c l k v r)"
shows "P ktst"
using assms
apply(induction_schema)
apply pat_completeness
apply(relation "measure (\<lambda>(kts, t). size_list (\<lambda>(k, t). size_rbt (\<lambda>_. 1) (\<lambda>_. 1) t) kts + size_rbt (\<lambda>_. 1) (\<lambda>_. 1) t)")
apply simp_all
done
lemma terminates_rbt_keys_generator:
"terminates (rbt_has_next, rbt_keys_next)"
proof
fix ktst :: "('a \<times> ('a, 'b) rbt) list \<times> ('a, 'b) rbt"
show "ktst \<in> terminates_on (rbt_has_next, rbt_keys_next)"
by(induct ktst taking: "\<lambda>k _. k" rule: rbt_generator_induct)(auto 4 3 intro: terminates_on.intros elim: terminates_on.cases)
qed
lift_definition rbt_keys_generator :: "('a, ('a, 'b, 'a) rbt_generator_state) generator"
is "(rbt_has_next, rbt_keys_next)"
by(rule terminates_rbt_keys_generator)
definition rbt_init :: "('a, 'b) rbt \<Rightarrow> ('a, 'b, 'c) rbt_generator_state"
where "rbt_init = Pair []"
lemma has_next_rbt_keys_generator [simp]:
"list.has_next rbt_keys_generator = rbt_has_next"
by transfer simp
lemma next_rbt_keys_generator [simp]:
"list.next rbt_keys_generator = rbt_keys_next"
by transfer simp
lemma unfoldr_rbt_keys_generator_aux:
"list.unfoldr rbt_keys_generator (kts, t) =
RBT_Impl.keys t @ concat (map (\<lambda>(k, t). k # RBT_Impl.keys t) kts)"
proof(induct "(kts, t)" arbitrary: kts t taking: "\<lambda>k _. k" rule: rbt_generator_induct)
case empty thus ?case
by(simp add: list.unfoldr.simps)
next
case split thus ?case
by(subst list.unfoldr.simps) simp
next
case shuffle thus ?case
by(subst list.unfoldr.simps)(subst (asm) list.unfoldr.simps, simp)
qed
corollary unfoldr_rbt_keys_generator:
"list.unfoldr rbt_keys_generator (rbt_init t) = RBT_Impl.keys t"
by(simp add: unfoldr_rbt_keys_generator_aux rbt_init_def)
fun rbt_entries_next ::
"('a, 'b, 'a \<times> 'b) rbt_generator_state \<Rightarrow> ('a \<times> 'b) \<times> ('a, 'b, 'a \<times> 'b) rbt_generator_state"
where
"rbt_entries_next ((kv, t) # kts, rbt.Empty) = (kv, kts, t)"
| "rbt_entries_next (kts, rbt.Branch c l k v r) = rbt_entries_next (((k, v), r) # kts, l)"
lemma terminates_rbt_entries_generator:
"terminates (rbt_has_next, rbt_entries_next)"
proof(rule terminatesI)
fix ktst :: "('a, 'b, 'a \<times> 'b) rbt_generator_state"
show "ktst \<in> terminates_on (rbt_has_next, rbt_entries_next)"
by(induct ktst taking: Pair rule: rbt_generator_induct)(auto 4 3 intro: terminates_on.intros elim: terminates_on.cases)
qed
lift_definition rbt_entries_generator :: "('a \<times> 'b, ('a, 'b, 'a \<times> 'b) rbt_generator_state) generator"
is "(rbt_has_next, rbt_entries_next)"
by(rule terminates_rbt_entries_generator)
lemma has_next_rbt_entries_generator [simp]:
"list.has_next rbt_entries_generator = rbt_has_next"
by transfer simp
lemma next_rbt_entries_generator [simp]:
"list.next rbt_entries_generator = rbt_entries_next"
by transfer simp
lemma unfoldr_rbt_entries_generator_aux:
"list.unfoldr rbt_entries_generator (kts, t) =
RBT_Impl.entries t @ concat (map (\<lambda>(k, t). k # RBT_Impl.entries t) kts)"
proof(induct "(kts, t)" arbitrary: kts t taking: Pair rule: rbt_generator_induct)
case empty thus ?case
by(simp add: list.unfoldr.simps)
next
case split thus ?case
by(subst list.unfoldr.simps) simp
next
case shuffle thus ?case
by(subst list.unfoldr.simps)(subst (asm) list.unfoldr.simps, simp)
qed
corollary unfoldr_rbt_entries_generator:
"list.unfoldr rbt_entries_generator (rbt_init t) = RBT_Impl.entries t"
by(simp add: unfoldr_rbt_entries_generator_aux rbt_init_def)
end
diff --git a/thys/Containers/Set_Impl.thy b/thys/Containers/Set_Impl.thy
--- a/thys/Containers/Set_Impl.thy
+++ b/thys/Containers/Set_Impl.thy
@@ -1,2069 +1,2069 @@
(* Title: Containers/Set_Impl.thy
Author: Andreas Lochbihler, KIT
René Thiemann, UIBK *)
theory Set_Impl imports
Collection_Enum
DList_Set
RBT_Set2
Closure_Set
Containers_Generator
Complex_Main
begin
section \<open>Different implementations of sets\<close>
subsection \<open>Auxiliary functions\<close>
text \<open>A simple quicksort implementation\<close>
context ord begin
function (sequential) quicksort_acc :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
and quicksort_part :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
"quicksort_acc ac [] = ac"
| "quicksort_acc ac [x] = x # ac"
| "quicksort_acc ac (x # xs) = quicksort_part ac x [] [] [] xs"
| "quicksort_part ac x lts eqs gts [] = quicksort_acc (eqs @ x # quicksort_acc ac gts) lts"
| "quicksort_part ac x lts eqs gts (z # zs) =
(if z > x then quicksort_part ac x lts eqs (z # gts) zs
else if z < x then quicksort_part ac x (z # lts) eqs gts zs
else quicksort_part ac x lts (z # eqs) gts zs)"
by pat_completeness simp_all
lemma length_quicksort_accp:
"quicksort_acc_quicksort_part_dom (Inl (ac, xs)) \<Longrightarrow> length (quicksort_acc ac xs) = length ac + length xs"
and length_quicksort_partp:
"quicksort_acc_quicksort_part_dom (Inr (ac, x, lts, eqs, gts, zs))
\<Longrightarrow> length (quicksort_part ac x lts eqs gts zs) = length ac + 1 + length lts + length eqs + length gts + length zs"
apply(induct rule: quicksort_acc_quicksort_part.pinduct)
apply(simp_all add: quicksort_acc.psimps quicksort_part.psimps)
done
termination
apply(relation "measure (case_sum (\<lambda>(_, xs). 2 * length xs ^ 2) (\<lambda>(_, _, lts, eqs, gts, zs). 2 * (length lts + length eqs + length gts + length zs) ^ 2 + length zs + 1))")
apply(simp_all add: power2_eq_square add_mult_distrib add_mult_distrib2 length_quicksort_accp)
done
definition quicksort :: "'a list \<Rightarrow> 'a list"
where "quicksort = quicksort_acc []"
lemma set_quicksort_acc [simp]: "set (quicksort_acc ac xs) = set ac \<union> set xs"
and set_quicksort_part [simp]:
"set (quicksort_part ac x lts eqs gts zs) =
set ac \<union> {x} \<union> set lts \<union> set eqs \<union> set gts \<union> set zs"
by(induct ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct)(auto split: if_split_asm)
lemma set_quicksort [simp]: "set (quicksort xs) = set xs"
by(simp add: quicksort_def)
lemma distinct_quicksort_acc:
"distinct (quicksort_acc ac xs) = distinct (ac @ xs)"
and distinct_quicksort_part:
"distinct (quicksort_part ac x lts eqs gts zs) = distinct (ac @ [x] @ lts @ eqs @ gts @ zs)"
by(induct ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct) auto
lemma distinct_quicksort [simp]: "distinct (quicksort xs) = distinct xs"
by(simp add: quicksort_def distinct_quicksort_acc)
end
lemmas [code] =
ord.quicksort_acc.simps quicksort_acc.simps
ord.quicksort_part.simps quicksort_part.simps
ord.quicksort_def quicksort_def
context linorder begin
lemma sorted_quicksort_acc:
"\<lbrakk> sorted ac; \<forall>x \<in> set xs. \<forall>a \<in> set ac. x < a \<rbrakk>
\<Longrightarrow> sorted (quicksort_acc ac xs)"
and sorted_quicksort_part:
"\<lbrakk> sorted ac; \<forall>y \<in> set lts \<union> {x} \<union> set eqs \<union> set gts \<union> set zs. \<forall>a \<in> set ac. y < a;
\<forall>y \<in> set lts. y < x; \<forall>y \<in> set eqs. y = x; \<forall>y \<in> set gts. y > x \<rbrakk>
\<Longrightarrow> sorted (quicksort_part ac x lts eqs gts zs)"
proof(induction ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct)
case 1 thus ?case by simp
next
case 2 thus ?case by(auto)
next
case 3 thus ?case by simp
next
case (4 ac x lts eqs gts)
note ac_greater = \<open>\<forall>y\<in>set lts \<union> {x} \<union> set eqs \<union> set gts \<union> set []. \<forall>a\<in>set ac. y < a\<close>
have "sorted eqs" "set eqs \<subseteq> {x}" using \<open>\<forall>y\<in>set eqs. y = x\<close>
by(induct eqs)(simp_all)
moreover have "\<forall>y \<in> set ac \<union> set gts. x \<le> y"
using \<open>\<forall>a\<in>set gts. x < a\<close> ac_greater by auto
moreover have "sorted (quicksort_acc ac gts)"
using \<open>sorted ac\<close> ac_greater by(auto intro: "4.IH")
ultimately have "sorted (eqs @ x # quicksort_acc ac gts)"
by(auto simp add: sorted_append)
moreover have "\<forall>y\<in>set lts. \<forall>a\<in>set (eqs @ x # quicksort_acc ac gts). y < a"
using \<open>\<forall>y\<in>set lts. y < x\<close> ac_greater \<open>\<forall>a\<in>set gts. x < a\<close> \<open>\<forall>y\<in>set eqs. y = x\<close>
by fastforce
ultimately show ?case by(simp add: "4.IH")
next
- case 5 thus ?case by(simp add: not_less eq_iff)
+ case 5 thus ?case by(simp add: not_less order_eq_iff)
qed
lemma sorted_quicksort [simp]: "sorted (quicksort xs)"
by(simp add: quicksort_def sorted_quicksort_acc)
lemma insort_key_append1:
"\<forall>y \<in> set ys. f x < f y \<Longrightarrow> insort_key f x (xs @ ys) = insort_key f x xs @ ys"
proof(induct xs)
case Nil
thus ?case by(cases ys) auto
qed simp
lemma insort_key_append2:
"\<forall>y \<in> set xs. f x > f y \<Longrightarrow> insort_key f x (xs @ ys) = xs @ insort_key f x ys"
by(induct xs) auto
lemma sort_key_append:
"\<forall>x\<in>set xs. \<forall>y\<in>set ys. f x < f y \<Longrightarrow> sort_key f (xs @ ys) = sort_key f xs @ sort_key f ys"
by(induct xs)(simp_all add: insort_key_append1)
definition single_list :: "'a \<Rightarrow> 'a list" where "single_list a = [a]"
lemma to_single_list: "x # xs = single_list x @ xs"
by(simp add: single_list_def)
lemma sort_snoc: "sort (xs @ [x]) = insort x (sort xs)"
by(induct xs)(simp_all add: insort_left_comm)
lemma sort_append_swap: "sort (xs @ ys) = sort (ys @ xs)"
by(induct xs arbitrary: ys rule: rev_induct)(simp_all add: sort_snoc[symmetric])
lemma sort_append_swap2: "sort (xs @ ys @ zs) = sort (ys @ xs @ zs)"
by(induct xs)(simp_all, subst (1 2) sort_append_swap, simp)
lemma sort_Cons_append_swap: "sort (x # xs) = sort (xs @ [x])"
by(subst sort_append_swap) simp
lemma sort_append_Cons_swap: "sort (ys @ x # xs) = sort (ys @ xs @ [x])"
apply(induct ys)
apply(simp only: append.simps sort_Cons_append_swap)
apply simp
done
lemma quicksort_acc_conv_sort:
"quicksort_acc ac xs = sort xs @ ac"
and quicksort_part_conv_sort:
"\<lbrakk> \<forall>y \<in> set lts. y < x; \<forall>y \<in> set eqs. y = x; \<forall>y \<in> set gts. y > x \<rbrakk>
\<Longrightarrow> quicksort_part ac x lts eqs gts zs = sort (lts @ eqs @ gts @ x # zs) @ ac"
proof(induct ac xs and ac x lts eqs gts zs rule: quicksort_acc_quicksort_part.induct)
case 1 thus ?case by simp
next
case 2 thus ?case by simp
next
case 3 thus ?case by simp
next
case (4 ac x lts eqs gts)
note eqs = \<open>\<forall>y\<in>set eqs. y = x\<close>
{ fix eqs
assume "\<forall>y\<in>set eqs. y = x"
hence "insort x eqs = x # eqs" by(induct eqs) simp_all }
note [simp] = this
from eqs have [simp]: "sort eqs = eqs" by(induct eqs) simp_all
from eqs have [simp]: "eqs @ [x] = x # eqs" by(induct eqs) simp_all
show ?case using 4
apply(subst sort_key_append)
apply(auto 4 3 dest: bspec)[1]
apply(simp add: append_assoc[symmetric] sort_snoc del: append_assoc)
apply(subst sort_key_append)
apply(auto 4 3 simp add: insort_key_append1 dest: bspec)
done
next
case (5 ac x lts eqs gts z zs)
have "\<lbrakk> \<not> z < x; \<not> x < z \<rbrakk> \<Longrightarrow> z = x" by simp
thus ?case using 5
apply(simp del: sort_key_simps)
apply(safe, simp_all del: sort_key_simps add: to_single_list)
apply(subst sort_append_swap)
apply(fold append_assoc)
apply(subst (2) sort_append_swap)
apply(subst sort_append_swap2)
apply(unfold append_assoc)
apply(rule refl)
apply(subst (1 5) append_assoc[symmetric])
apply(subst (1 2) sort_append_swap)
apply(unfold append_assoc)
apply(subst sort_append_swap2)
apply(subst (1 2) sort_append_swap)
apply(unfold append_assoc)
apply(subst sort_append_swap2)
apply(rule refl)
apply(subst (2 6) append_assoc[symmetric])
apply(subst (2 5) append_assoc[symmetric])
apply(subst (1 2) sort_append_swap2)
apply(subst (4) append_assoc)
apply(subst (2) sort_append_swap2)
apply simp
done
qed
lemma quicksort_conv_sort: "quicksort xs = sort xs"
by(simp add: quicksort_def quicksort_acc_conv_sort)
lemma sort_remdups: "sort (remdups xs) = remdups (sort xs)"
by(rule sorted_distinct_set_unique) simp_all
end
text \<open>Removing duplicates from a sorted list\<close>
context ord begin
fun remdups_sorted :: "'a list \<Rightarrow> 'a list"
where
"remdups_sorted [] = []"
| "remdups_sorted [x] = [x]"
| "remdups_sorted (x#y#xs) = (if x < y then x # remdups_sorted (y#xs) else remdups_sorted (y#xs))"
end
lemmas [code] = ord.remdups_sorted.simps
context linorder begin
lemma [simp]:
assumes "sorted xs"
shows sorted_remdups_sorted: "sorted (remdups_sorted xs)"
and set_remdups_sorted: "set (remdups_sorted xs) = set xs"
using assms by(induct xs rule: remdups_sorted.induct)(auto)
lemma distinct_remdups_sorted [simp]: "sorted xs \<Longrightarrow> distinct (remdups_sorted xs)"
by(induct xs rule: remdups_sorted.induct)(auto)
lemma remdups_sorted_conv_remdups: "sorted xs \<Longrightarrow> remdups_sorted xs = remdups xs"
by(induct xs rule: remdups_sorted.induct)(auto)
end
text \<open>An specialised operation to convert a finite set into a sorted list\<close>
definition csorted_list_of_set :: "'a :: ccompare set \<Rightarrow> 'a list"
where [code del]:
"csorted_list_of_set A =
(if ID CCOMPARE('a) = None \<or> \<not> finite A then undefined else linorder.sorted_list_of_set cless_eq A)"
lemma csorted_list_of_set_set [simp]:
"\<lbrakk> ID CCOMPARE('a :: ccompare) = Some c; linorder.sorted (le_of_comp c) xs; distinct xs \<rbrakk>
\<Longrightarrow> linorder.sorted_list_of_set (le_of_comp c) (set xs) = xs"
by(simp add: distinct_remdups_id linorder.sorted_list_of_set_sort_remdups[OF ID_ccompare] linorder.sorted_sort_id[OF ID_ccompare])
lemma csorted_list_of_set_split:
fixes A :: "'a :: ccompare set" shows
"P (csorted_list_of_set A) \<longleftrightarrow>
(\<forall>xs. ID CCOMPARE('a) \<noteq> None \<longrightarrow> finite A \<longrightarrow> A = set xs \<longrightarrow> distinct xs \<longrightarrow> linorder.sorted cless_eq xs \<longrightarrow> P xs) \<and>
(ID CCOMPARE('a) = None \<or> \<not> finite A \<longrightarrow> P undefined)"
by(auto simp add: csorted_list_of_set_def linorder.sorted_list_of_set[OF ID_ccompare])
code_identifier code_module Set \<rightharpoonup> (SML) Set_Impl
| code_module Set_Impl \<rightharpoonup> (SML) Set_Impl
subsection \<open>Delete code equation with set as constructor\<close>
lemma is_empty_unfold [code_unfold]:
"set_eq A {} = Set.is_empty A"
"set_eq {} A = Set.is_empty A"
by(auto simp add: Set.is_empty_def set_eq_def)
definition is_UNIV :: "'a set \<Rightarrow> bool"
where [code del]: "is_UNIV A \<longleftrightarrow> A = UNIV"
lemma is_UNIV_unfold [code_unfold]:
"A = UNIV \<longleftrightarrow> is_UNIV A"
"UNIV = A \<longleftrightarrow> is_UNIV A"
"set_eq A UNIV \<longleftrightarrow> is_UNIV A"
"set_eq UNIV A \<longleftrightarrow> is_UNIV A"
by(auto simp add: is_UNIV_def set_eq_def)
lemma [code_unfold del, symmetric, code_post del]:
"x \<in> set xs \<equiv> List.member xs x"
by(simp add: List.member_def)
lemma [code_unfold del, symmetric, code_post del]:
"finite \<equiv> Cardinality.finite'" by(simp)
lemma [code_unfold del, symmetric, code_post del]:
"card \<equiv> Cardinality.card'" by simp
declare [[code drop:
Set.empty
Set.is_empty
uminus_set_inst.uminus_set
Set.member
Set.insert
Set.remove
UNIV
Set.filter
image
Set.subset_eq
Ball
Bex
Set.union
minus_set_inst.minus_set
Set.inter
card
Set.bind
the_elem
Pow
sum
Gcd
Lcm
Product_Type.product
Id_on
Image
trancl
relcomp
wf
Min
Inf_fin
Max
Sup_fin
"Inf :: 'a set set \<Rightarrow> 'a set"
"Sup :: 'a set set \<Rightarrow> 'a set"
sorted_list_of_set
List.map_project
Sup_pred_inst.Sup_pred
finite
Cardinality.finite'
card
Cardinality.card'
Inf_pred_inst.Inf_pred
pred_of_set
Cardinality.subset'
Cardinality.eq_set
Wellfounded.acc
Bleast
can_select
"set_eq :: 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
irrefl
bacc
set_of_pred
set_of_seq
]]
declare
Cardinality.finite'_def[code]
Cardinality.card'_def[code]
subsection \<open>Set implementations\<close>
definition Collect_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set"
where [simp]: "Collect_set = Collect"
definition DList_set :: "'a :: ceq set_dlist \<Rightarrow> 'a set"
where "DList_set = Collect o DList_Set.member"
definition RBT_set :: "'a :: ccompare set_rbt \<Rightarrow> 'a set"
where "RBT_set = Collect o RBT_Set2.member"
definition Complement :: "'a set \<Rightarrow> 'a set"
where [simp]: "Complement A = - A"
definition Set_Monad :: "'a list \<Rightarrow> 'a set"
where [simp]: "Set_Monad = set"
code_datatype Collect_set DList_set RBT_set Set_Monad Complement
lemma DList_set_empty [simp]: "DList_set DList_Set.empty = {}"
by(simp add: DList_set_def)
lemma RBT_set_empty [simp]: "RBT_set RBT_Set2.empty = {}"
by(simp add: RBT_set_def)
lemma RBT_set_conv_keys:
"ID CCOMPARE('a :: ccompare) \<noteq> None
\<Longrightarrow> RBT_set (t :: 'a set_rbt) = set (RBT_Set2.keys t)"
by(clarsimp simp add: RBT_set_def member_conv_keys)
subsection \<open>Set operations\<close>
text \<open>
A collection of all the theorems about @{const Complement}.
\<close>
ML \<open>
structure Set_Complement_Eqs = Named_Thms
(
val name = @{binding set_complement_code}
val description = "Code equations involving set complement"
)
\<close>
setup \<open>Set_Complement_Eqs.setup\<close>
text \<open>Various fold operations over sets\<close>
typedef ('a, 'b) comp_fun_commute = "{f :: 'a \<Rightarrow> 'b \<Rightarrow> 'b. comp_fun_commute f}"
morphisms comp_fun_commute_apply Abs_comp_fun_commute
by(rule exI[where x="\<lambda>_. id"])(simp, unfold_locales, auto)
setup_lifting type_definition_comp_fun_commute
lemma comp_fun_commute_apply' [simp]:
"comp_fun_commute (comp_fun_commute_apply f)"
using comp_fun_commute_apply[of f] by simp
lift_definition set_fold_cfc :: "('a, 'b) comp_fun_commute \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b" is "Finite_Set.fold" .
declare [[code drop: set_fold_cfc]]
lemma set_fold_cfc_code [code]:
fixes xs :: "'a :: ceq list"
and dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt"
shows set_fold_cfc_Complement[set_complement_code]:
"set_fold_cfc f''' b (Complement A) = Code.abort (STR ''set_fold_cfc not supported on Complement'') (\<lambda>_. set_fold_cfc f''' b (Complement A))"
and
"set_fold_cfc f''' b (Collect_set P) = Code.abort (STR ''set_fold_cfc not supported on Collect_set'') (\<lambda>_. set_fold_cfc f''' b (Collect_set P))"
"set_fold_cfc f b (Set_Monad xs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''set_fold_cfc Set_Monad: ceq = None'') (\<lambda>_. set_fold_cfc f b (Set_Monad xs))
| Some eq \<Rightarrow> List.fold (comp_fun_commute_apply f) (equal_base.list_remdups eq xs) b)"
(is ?Set_Monad)
"set_fold_cfc f' b (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''set_fold_cfc DList_set: ceq = None'') (\<lambda>_. set_fold_cfc f' b (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.fold (comp_fun_commute_apply f') dxs b)"
(is ?DList_set)
"set_fold_cfc f'' b (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''set_fold_cfc RBT_set: ccompare = None'') (\<lambda>_. set_fold_cfc f'' b (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.fold (comp_fun_commute_apply f'') rbt b)"
(is ?RBT_set)
proof -
show ?Set_Monad
by(auto split: option.split dest!: Collection_Eq.ID_ceq simp add: set_fold_cfc_def comp_fun_commute.fold_set_fold_remdups)
show ?DList_set
apply(auto split: option.split simp add: DList_set_def)
apply transfer
apply(auto dest: Collection_Eq.ID_ceq simp add: List.member_def[abs_def] comp_fun_commute.fold_set_fold_remdups distinct_remdups_id)
done
show ?RBT_set
apply(auto split: option.split simp add: RBT_set_conv_keys fold_conv_fold_keys)
apply transfer
apply(simp add: comp_fun_commute.fold_set_fold_remdups distinct_remdups_id linorder.distinct_keys[OF ID_ccompare] ord.is_rbt_rbt_sorted)
done
qed simp_all
typedef ('a, 'b) comp_fun_idem = "{f :: 'a \<Rightarrow> 'b \<Rightarrow> 'b. comp_fun_idem f}"
morphisms comp_fun_idem_apply Abs_comp_fun_idem
by(rule exI[where x="\<lambda>_. id"])(simp, unfold_locales, auto)
setup_lifting type_definition_comp_fun_idem
lemma comp_fun_idem_apply' [simp]:
"comp_fun_idem (comp_fun_idem_apply f)"
using comp_fun_idem_apply[of f] by simp
lift_definition set_fold_cfi :: "('a, 'b) comp_fun_idem \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b" is "Finite_Set.fold" .
declare [[code drop: set_fold_cfi]]
lemma set_fold_cfi_code [code]:
fixes xs :: "'a list"
and dxs :: "'b :: ceq set_dlist"
and rbt :: "'c :: ccompare set_rbt" shows
"set_fold_cfi f b (Complement A) = Code.abort (STR ''set_fold_cfi not supported on Complement'') (\<lambda>_. set_fold_cfi f b (Complement A))"
"set_fold_cfi f b (Collect_set P) = Code.abort (STR ''set_fold_cfi not supported on Collect_set'') (\<lambda>_. set_fold_cfi f b (Collect_set P))"
"set_fold_cfi f b (Set_Monad xs) = List.fold (comp_fun_idem_apply f) xs b"
(is ?Set_Monad)
"set_fold_cfi f' b (DList_set dxs) =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''set_fold_cfi DList_set: ceq = None'') (\<lambda>_. set_fold_cfi f' b (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.fold (comp_fun_idem_apply f') dxs b)"
(is ?DList_set)
"set_fold_cfi f'' b (RBT_set rbt) =
(case ID CCOMPARE('c) of None \<Rightarrow> Code.abort (STR ''set_fold_cfi RBT_set: ccompare = None'') (\<lambda>_. set_fold_cfi f'' b (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.fold (comp_fun_idem_apply f'') rbt b)"
(is ?RBT_set)
proof -
show ?Set_Monad
by(auto split: option.split dest!: Collection_Eq.ID_ceq simp add: set_fold_cfi_def comp_fun_idem.fold_set_fold)
show ?DList_set
apply(auto split: option.split simp add: DList_set_def)
apply transfer
apply(auto dest: Collection_Eq.ID_ceq simp add: List.member_def[abs_def] comp_fun_idem.fold_set_fold)
done
show ?RBT_set
apply(auto split: option.split simp add: RBT_set_conv_keys fold_conv_fold_keys)
apply transfer
apply(simp add: comp_fun_idem.fold_set_fold)
done
qed simp_all
typedef 'a semilattice_set = "{f :: 'a \<Rightarrow> 'a \<Rightarrow> 'a. semilattice_set f}"
morphisms semilattice_set_apply Abs_semilattice_set
proof
show "(\<lambda>x y. if x = y then x else undefined) \<in> ?semilattice_set"
unfolding mem_Collect_eq by(unfold_locales) simp_all
qed
setup_lifting type_definition_semilattice_set
lemma semilattice_set_apply' [simp]:
"semilattice_set (semilattice_set_apply f)"
using semilattice_set_apply[of f] by simp
lemma comp_fun_idem_semilattice_set_apply [simp]:
"comp_fun_idem (semilattice_set_apply f)"
proof -
interpret semilattice_set "semilattice_set_apply f" by simp
show ?thesis by(unfold_locales)(simp_all add: fun_eq_iff left_commute)
qed
lift_definition set_fold1 :: "'a semilattice_set \<Rightarrow> 'a set \<Rightarrow> 'a" is "semilattice_set.F" .
lemma (in semilattice_set) F_set_conv_fold:
"xs \<noteq> [] \<Longrightarrow> F (set xs) = Finite_Set.fold f (hd xs) (set (tl xs))"
by(clarsimp simp add: neq_Nil_conv eq_fold)
lemma set_fold1_code [code]:
fixes rbt :: "'a :: {ccompare, lattice} set_rbt"
and dxs :: "'b :: {ceq, lattice} set_dlist" shows
set_fold1_Complement[set_complement_code]:
"set_fold1 f (Complement A) = Code.abort (STR ''set_fold1: Complement'') (\<lambda>_. set_fold1 f (Complement A))"
and "set_fold1 f (Collect_set P) = Code.abort (STR ''set_fold1: Collect_set'') (\<lambda>_. set_fold1 f (Collect_set P))"
and "set_fold1 f (Set_Monad (x # xs)) = fold (semilattice_set_apply f) xs x" (is "?Set_Monad")
and
"set_fold1 f' (DList_set dxs) =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''set_fold1 DList_set: ceq = None'') (\<lambda>_. set_fold1 f' (DList_set dxs))
| Some _ \<Rightarrow> if DList_Set.null dxs then Code.abort (STR ''set_fold1 DList_set: empty set'') (\<lambda>_. set_fold1 f' (DList_set dxs))
else DList_Set.fold (semilattice_set_apply f') (DList_Set.tl dxs) (DList_Set.hd dxs))"
(is "?DList_set")
and
"set_fold1 f'' (RBT_set rbt) =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''set_fold1 RBT_set: ccompare = None'') (\<lambda>_. set_fold1 f'' (RBT_set rbt))
| Some _ \<Rightarrow> if RBT_Set2.is_empty rbt then Code.abort (STR ''set_fold1 RBT_set: empty set'') (\<lambda>_. set_fold1 f'' (RBT_set rbt))
else RBT_Set2.fold1 (semilattice_set_apply f'') rbt)"
(is "?RBT_set")
proof -
show ?Set_Monad
by(simp add: set_fold1_def semilattice_set.eq_fold comp_fun_idem.fold_set_fold)
show ?DList_set
by(simp add: set_fold1_def semilattice_set.F_set_conv_fold comp_fun_idem.fold_set_fold DList_set_def DList_Set.Collect_member split: option.split)(transfer, simp)
show ?RBT_set
by(simp add: set_fold1_def semilattice_set.F_set_conv_fold comp_fun_idem.fold_set_fold RBT_set_def RBT_Set2.member_conv_keys RBT_Set2.fold1_conv_fold split: option.split)
qed simp_all
text \<open>Implementation of set operations\<close>
lemma Collect_code [code]:
fixes P :: "'a :: cenum \<Rightarrow> bool" shows
"Collect P =
(case ID CENUM('a) of None \<Rightarrow> Collect_set P
| Some (enum, _) \<Rightarrow> Set_Monad (filter P enum))"
by(auto split: option.split dest: in_cenum)
lemma finite_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt"
and A :: "'c :: finite_UNIV set" and P :: "'c \<Rightarrow> bool" shows
"finite (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''finite DList_set: ceq = None'') (\<lambda>_. finite (DList_set dxs))
| Some _ \<Rightarrow> True)"
"finite (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''finite RBT_set: ccompare = None'') (\<lambda>_. finite (RBT_set rbt))
| Some _ \<Rightarrow> True)"
and finite_Complement [set_complement_code]:
"finite (Complement A) \<longleftrightarrow>
(if of_phantom (finite_UNIV :: 'c finite_UNIV) then True
else if finite A then False
else Code.abort (STR ''finite Complement: infinite set'') (\<lambda>_. finite (Complement A)))"
and
"finite (Set_Monad xs) = True"
"finite (Collect_set P) \<longleftrightarrow>
of_phantom (finite_UNIV :: 'c finite_UNIV) \<or> Code.abort (STR ''finite Collect_set'') (\<lambda>_. finite (Collect_set P))"
by(auto simp add: DList_set_def RBT_set_def member_conv_keys card_gt_0_iff finite_UNIV split: option.split elim: finite_subset[rotated 1])
lemma card_code [code]:
fixes dxs :: "'a :: ceq set_dlist" and xs :: "'a list"
and rbt :: "'b :: ccompare set_rbt"
and A :: "'c :: card_UNIV set" shows
"card (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''card DList_set: ceq = None'') (\<lambda>_. card (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.length dxs)"
"card (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''card RBT_set: ccompare = None'') (\<lambda>_. card (RBT_set rbt))
| Some _ \<Rightarrow> length (RBT_Set2.keys rbt))"
"card (Set_Monad xs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''card Set_Monad: ceq = None'') (\<lambda>_. card (Set_Monad xs))
| Some eq \<Rightarrow> length (equal_base.list_remdups eq xs))"
and card_Complement [set_complement_code]:
"card (Complement A) =
(let a = card A; s = CARD('c)
in if s > 0 then s - a
else if finite A then 0
else Code.abort (STR ''card Complement: infinite'') (\<lambda>_. card (Complement A)))"
by(auto simp add: RBT_set_def member_conv_keys distinct_card DList_set_def Let_def card_UNIV Compl_eq_Diff_UNIV card_Diff_subset_Int card_gt_0_iff finite_subset[of A UNIV] List.card_set dest: Collection_Eq.ID_ceq split: option.split)
lemma is_UNIV_code [code]:
fixes rbt :: "'a :: {cproper_interval, finite_UNIV} set_rbt"
and A :: "'b :: card_UNIV set" shows
"is_UNIV A \<longleftrightarrow>
(let a = CARD('b);
b = card A
in if a > 0 then a = b
else if b > 0 then False
else Code.abort (STR ''is_UNIV called on infinite type and set'') (\<lambda>_. is_UNIV A))"
(is ?generic)
"is_UNIV (RBT_set rbt) =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''is_UNIV RBT_set: ccompare = None'') (\<lambda>_. is_UNIV (RBT_set rbt))
| Some _ \<Rightarrow> of_phantom (finite_UNIV :: 'a finite_UNIV) \<and> proper_intrvl.exhaustive_fusion cproper_interval rbt_keys_generator (RBT_Set2.init rbt))"
(is ?rbt)
proof -
{
fix c
assume linorder: "ID CCOMPARE('a) = Some c"
have "is_UNIV (RBT_set rbt) =
(finite (UNIV :: 'a set) \<and> proper_intrvl.exhaustive cproper_interval (RBT_Set2.keys rbt))"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
have "finite (UNIV :: 'a set)"
unfolding \<open>?lhs\<close>[unfolded is_UNIV_def, symmetric]
using linorder
by(simp add: finite_code)
moreover
hence "proper_intrvl.exhaustive cproper_interval (RBT_Set2.keys rbt)"
using linorder \<open>?lhs\<close>
by(simp add: linorder_proper_interval.exhaustive_correct[OF ID_ccompare_interval[OF linorder]] sorted_RBT_Set_keys is_UNIV_def RBT_set_conv_keys)
ultimately show ?rhs ..
next
assume ?rhs
thus ?lhs using linorder
by(auto simp add: linorder_proper_interval.exhaustive_correct[OF ID_ccompare_interval[OF linorder]] sorted_RBT_Set_keys is_UNIV_def RBT_set_conv_keys)
qed }
thus ?rbt
by(auto simp add: finite_UNIV proper_intrvl.exhaustive_fusion_def unfoldr_rbt_keys_generator is_UNIV_def split: option.split)
show ?generic
by(auto simp add: Let_def is_UNIV_def dest: card_seteq[of UNIV A] dest!: card_ge_0_finite)
qed
lemma is_empty_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt"
and A :: "'c set" shows
"Set.is_empty (Set_Monad xs) \<longleftrightarrow> xs = []"
"Set.is_empty (DList_set dxs) \<longleftrightarrow>
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''is_empty DList_set: ceq = None'') (\<lambda>_. Set.is_empty (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.null dxs)" (is ?DList_set)
"Set.is_empty (RBT_set rbt) \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''is_empty RBT_set: ccompare = None'') (\<lambda>_. Set.is_empty (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.is_empty rbt)" (is ?RBT_set)
and is_empty_Complement [set_complement_code]:
"Set.is_empty (Complement A) \<longleftrightarrow> is_UNIV A" (is ?Complement)
proof -
show ?DList_set
by(clarsimp simp add: DList_set_def Set.is_empty_def DList_Set.member_empty_empty split: option.split)
show ?RBT_set
by(clarsimp simp add: RBT_set_def Set.is_empty_def RBT_Set2.member_empty_empty[symmetric] fun_eq_iff simp del: RBT_Set2.member_empty_empty split: option.split)
show ?Complement
by(auto simp add: is_UNIV_def Set.is_empty_def)
qed(simp_all add: Set.is_empty_def List.null_def)
lemma Set_insert_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"\<And>x. Set.insert x (Collect_set A) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''insert Collect_set: ceq = None'') (\<lambda>_. Set.insert x (Collect_set A))
| Some eq \<Rightarrow> Collect_set (equal_base.fun_upd eq A x True))"
"\<And>x. Set.insert x (Set_Monad xs) = Set_Monad (x # xs)"
"\<And>x. Set.insert x (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''insert DList_set: ceq = None'') (\<lambda>_. Set.insert x (DList_set dxs))
| Some _ \<Rightarrow> DList_set (DList_Set.insert x dxs))"
"\<And>x. Set.insert x (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''insert RBT_set: ccompare = None'') (\<lambda>_. Set.insert x (RBT_set rbt))
| Some _ \<Rightarrow> RBT_set (RBT_Set2.insert x rbt))"
and insert_Complement [set_complement_code]:
"\<And>x. Set.insert x (Complement X) = Complement (Set.remove x X)"
by(auto split: option.split dest: equal.equal_eq[OF ID_ceq] simp add: DList_set_def DList_Set.member_insert RBT_set_def)
lemma Set_member_code [code]:
fixes xs :: "'a :: ceq list" shows
"\<And>x. x \<in> Collect_set A \<longleftrightarrow> A x"
"\<And>x. x \<in> DList_set dxs \<longleftrightarrow> DList_Set.member dxs x"
"\<And>x. x \<in> RBT_set rbt \<longleftrightarrow> RBT_Set2.member rbt x"
and mem_Complement [set_complement_code]:
"\<And>x. x \<in> Complement X \<longleftrightarrow> x \<notin> X"
and
"\<And>x. x \<in> Set_Monad xs \<longleftrightarrow>
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''member Set_Monad: ceq = None'') (\<lambda>_. x \<in> Set_Monad xs)
| Some eq \<Rightarrow> equal_base.list_member eq xs x)"
by(auto simp add: DList_set_def RBT_set_def List.member_def split: option.split dest!: Collection_Eq.ID_ceq)
lemma Set_remove_code [code]:
fixes rbt :: "'a :: ccompare set_rbt"
and dxs :: "'b :: ceq set_dlist" shows
"\<And>x. Set.remove x (Collect_set A) =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''remove Collect: ceq = None'') (\<lambda>_. Set.remove x (Collect_set A))
| Some eq \<Rightarrow> Collect_set (equal_base.fun_upd eq A x False))"
"\<And>x. Set.remove x (DList_set dxs) =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''remove DList_set: ceq = None'') (\<lambda>_. Set.remove x (DList_set dxs))
| Some _ \<Rightarrow> DList_set (DList_Set.remove x dxs))"
"\<And>x. Set.remove x (RBT_set rbt) =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''remove RBT_set: ccompare = None'') (\<lambda>_. Set.remove x (RBT_set rbt))
| Some _ \<Rightarrow> RBT_set (RBT_Set2.remove x rbt))"
and remove_Complement [set_complement_code]:
"\<And>x A. Set.remove x (Complement A) = Complement (Set.insert x A)"
by(auto split: option.split if_split_asm dest: equal.equal_eq[OF ID_ceq] simp add: DList_set_def DList_Set.member_remove RBT_set_def)
lemma Set_uminus_code [code, set_complement_code]:
"- A = Complement A"
"- (Collect_set P) = Collect_set (\<lambda>x. \<not> P x)"
"- (Complement B) = B"
by auto
text \<open>
These equations represent complements as true complements.
If you want that the complement operations returns an explicit enumeration of the elements, use the following set of equations which use @{class cenum}.
\<close>
lemma Set_uminus_cenum:
fixes A :: "'a :: cenum set" shows
"- A =
(case ID CENUM('a) of None \<Rightarrow> Complement A
| Some (enum, _) \<Rightarrow> Set_Monad (filter (\<lambda>x. x \<notin> A) enum))"
and "- (Complement B) = B"
by(auto split: option.split dest: ID_cEnum)
lemma Set_minus_code [code]:
fixes rbt1 rbt2 :: "'a :: ccompare set_rbt"
shows "A - B = A \<inter> (- B)"
"RBT_set rbt1 - RBT_set rbt2 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''minus RBT_set RBT_set: ccompare = None'') (\<lambda>_. RBT_set rbt1 - RBT_set rbt2)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.minus rbt1 rbt2))"
by (auto simp: Set_member_code(3) split: option.splits)
lemma Set_union_code [code]:
fixes rbt1 rbt2 :: "'a :: ccompare set_rbt"
and rbt :: "'b :: {ccompare, ceq} set_rbt"
and dxs :: "'b set_dlist"
and dxs1 dxs2 :: "'c :: ceq set_dlist" shows
"RBT_set rbt1 \<union> RBT_set rbt2 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''union RBT_set RBT_set: ccompare = None'') (\<lambda>_. RBT_set rbt1 \<union> RBT_set rbt2)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.union rbt1 rbt2))" (is ?RBT_set_RBT_set)
"RBT_set rbt \<union> DList_set dxs =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''union RBT_set DList_set: ccompare = None'') (\<lambda>_. RBT_set rbt \<union> DList_set dxs)
| Some _ \<Rightarrow>
case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''union RBT_set DList_set: ceq = None'') (\<lambda>_. RBT_set rbt \<union> DList_set dxs)
| Some _ \<Rightarrow> RBT_set (DList_Set.fold RBT_Set2.insert dxs rbt))" (is ?RBT_set_DList_set)
"DList_set dxs \<union> RBT_set rbt =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''union DList_set RBT_set: ccompare = None'') (\<lambda>_. RBT_set rbt \<union> DList_set dxs)
| Some _ \<Rightarrow>
case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''union DList_set RBT_set: ceq = None'') (\<lambda>_. RBT_set rbt \<union> DList_set dxs)
| Some _ \<Rightarrow> RBT_set (DList_Set.fold RBT_Set2.insert dxs rbt))" (is ?DList_set_RBT_set)
"DList_set dxs1 \<union> DList_set dxs2 =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''union DList_set DList_set: ceq = None'') (\<lambda>_. DList_set dxs1 \<union> DList_set dxs2)
| Some _ \<Rightarrow> DList_set (DList_Set.union dxs1 dxs2))" (is ?DList_set_DList_set)
"Set_Monad zs \<union> RBT_set rbt2 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''union Set_Monad RBT_set: ccompare = None'') (\<lambda>_. Set_Monad zs \<union> RBT_set rbt2)
| Some _ \<Rightarrow> RBT_set (fold RBT_Set2.insert zs rbt2))" (is ?Set_Monad_RBT_set)
"RBT_set rbt1 \<union> Set_Monad zs =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''union RBT_set Set_Monad: ccompare = None'') (\<lambda>_. RBT_set rbt1 \<union> Set_Monad zs)
| Some _ \<Rightarrow> RBT_set (fold RBT_Set2.insert zs rbt1))" (is ?RBT_set_Set_Monad)
"Set_Monad ws \<union> DList_set dxs2 =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''union Set_Monad DList_set: ceq = None'') (\<lambda>_. Set_Monad ws \<union> DList_set dxs2)
| Some _ \<Rightarrow> DList_set (fold DList_Set.insert ws dxs2))" (is ?Set_Monad_DList_set)
"DList_set dxs1 \<union> Set_Monad ws =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''union DList_set Set_Monad: ceq = None'') (\<lambda>_. DList_set dxs1 \<union> Set_Monad ws)
| Some _ \<Rightarrow> DList_set (fold DList_Set.insert ws dxs1))" (is ?DList_set_Set_Monad)
"Set_Monad xs \<union> Set_Monad ys = Set_Monad (xs @ ys)"
"Collect_set A \<union> B = Collect_set (\<lambda>x. A x \<or> x \<in> B)"
"B \<union> Collect_set A = Collect_set (\<lambda>x. A x \<or> x \<in> B)"
and Set_union_Complement [set_complement_code]:
"Complement B \<union> B' = Complement (B \<inter> - B')"
"B' \<union> Complement B = Complement (- B' \<inter> B)"
proof -
show ?RBT_set_RBT_set ?Set_Monad_RBT_set ?RBT_set_Set_Monad
by(auto split: option.split simp add: RBT_set_def)
show ?RBT_set_DList_set ?DList_set_RBT_set
by(auto split: option.split simp add: RBT_set_def DList_set_def DList_Set.fold_def DList_Set.member_def List.member_def dest: equal.equal_eq[OF ID_ceq])
show ?DList_set_Set_Monad ?Set_Monad_DList_set
by(auto split: option.split simp add: DList_set_def DList_Set.member_fold_insert)
show ?DList_set_DList_set
by(auto split: option.split simp add: DList_set_def DList_Set.member_union)
qed(auto)
lemma Set_inter_code [code]:
fixes rbt1 rbt2 :: "'a :: ccompare set_rbt"
and rbt :: "'b :: {ccompare, ceq} set_rbt"
and dxs :: "'b set_dlist"
and dxs1 dxs2 :: "'c :: ceq set_dlist"
and xs1 xs2 :: "'c list"
shows
"Collect_set A'' \<inter> J = Collect_set (\<lambda>x. A'' x \<and> x \<in> J)" (is ?collect1)
"J \<inter> Collect_set A'' = Collect_set (\<lambda>x. A'' x \<and> x \<in> J)" (is ?collect2)
"Set_Monad xs'' \<inter> I = Set_Monad (filter (\<lambda>x. x \<in> I) xs'')" (is ?monad1)
"I \<inter> Set_Monad xs'' = Set_Monad (filter (\<lambda>x. x \<in> I) xs'')" (is ?monad2)
"DList_set dxs1 \<inter> H =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''inter DList_set1: ceq = None'') (\<lambda>_. DList_set dxs1 \<inter> H)
| Some eq \<Rightarrow> DList_set (DList_Set.filter (\<lambda>x. x \<in> H) dxs1))" (is ?dlist1)
"H \<inter> DList_set dxs2 =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''inter DList_set2: ceq = None'') (\<lambda>_. H \<inter> DList_set dxs2)
| Some eq \<Rightarrow> DList_set (DList_Set.filter (\<lambda>x. x \<in> H) dxs2))" (is ?dlist2)
"RBT_set rbt1 \<inter> G =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''inter RBT_set1: ccompare = None'') (\<lambda>_. RBT_set rbt1 \<inter> G)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.filter (\<lambda>x. x \<in> G) rbt1))" (is ?rbt1)
"G \<inter> RBT_set rbt2 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''inter RBT_set2: ccompare = None'') (\<lambda>_. G \<inter> RBT_set rbt2)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.filter (\<lambda>x. x \<in> G) rbt2))" (is ?rbt2)
and Set_inter_Complement [set_complement_code]:
"Complement B'' \<inter> Complement B''' = Complement (B'' \<union> B''')" (is ?complement)
and
"Set_Monad xs \<inter> RBT_set rbt1 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''inter Set_Monad RBT_set: ccompare = None'') (\<lambda>_. RBT_set rbt1 \<inter> Set_Monad xs)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.inter_list rbt1 xs))" (is ?monad_rbt)
"Set_Monad xs' \<inter> DList_set dxs2 =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''inter Set_Monad DList_set: ceq = None'') (\<lambda>_. Set_Monad xs' \<inter> DList_set dxs2)
| Some eq \<Rightarrow> DList_set (DList_Set.filter (equal_base.list_member eq xs') dxs2))" (is ?monad_dlist)
"Set_Monad xs1 \<inter> Set_Monad xs2 =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''inter Set_Monad Set_Monad: ceq = None'') (\<lambda>_. Set_Monad xs1 \<inter> Set_Monad xs2)
| Some eq \<Rightarrow> Set_Monad (filter (equal_base.list_member eq xs2) xs1))" (is ?monad)
"DList_set dxs \<inter> RBT_set rbt =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''inter DList_set RBT_set: ccompare = None'') (\<lambda>_. DList_set dxs \<inter> RBT_set rbt)
| Some _ \<Rightarrow>
case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''inter DList_set RBT_set: ceq = None'') (\<lambda>_. DList_set dxs \<inter> RBT_set rbt)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.inter_list rbt (list_of_dlist dxs)))" (is ?dlist_rbt)
"DList_set dxs1 \<inter> DList_set dxs2 =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''inter DList_set DList_set: ceq = None'') (\<lambda>_. DList_set dxs1 \<inter> DList_set dxs2)
| Some _ \<Rightarrow> DList_set (DList_Set.filter (DList_Set.member dxs2) dxs1))" (is ?dlist)
"DList_set dxs1 \<inter> Set_Monad xs' =
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''inter DList_set Set_Monad: ceq = None'') (\<lambda>_. DList_set dxs1 \<inter> Set_Monad xs')
| Some eq \<Rightarrow> DList_set (DList_Set.filter (equal_base.list_member eq xs') dxs1))" (is ?dlist_monad)
"RBT_set rbt1 \<inter> RBT_set rbt2 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''inter RBT_set RBT_set: ccompare = None'') (\<lambda>_. RBT_set rbt1 \<inter> RBT_set rbt2)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.inter rbt1 rbt2))" (is ?rbt_rbt)
"RBT_set rbt \<inter> DList_set dxs =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''inter RBT_set DList_set: ccompare = None'') (\<lambda>_. RBT_set rbt \<inter> DList_set dxs)
| Some _ \<Rightarrow>
case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''inter RBT_set DList_set: ceq = None'') (\<lambda>_. RBT_set rbt \<inter> DList_set dxs)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.inter_list rbt (list_of_dlist dxs)))" (is ?rbt_dlist)
"RBT_set rbt1 \<inter> Set_Monad xs =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''inter RBT_set Set_Monad: ccompare = None'') (\<lambda>_. RBT_set rbt1 \<inter> Set_Monad xs)
| Some _ \<Rightarrow> RBT_set (RBT_Set2.inter_list rbt1 xs))" (is ?rbt_monad)
proof -
show ?rbt_rbt ?rbt1 ?rbt2 ?rbt_dlist ?rbt_monad ?dlist_rbt ?monad_rbt
by(auto simp add: RBT_set_def DList_set_def DList_Set.member_def List.member_def dest: equal.equal_eq[OF ID_ceq] split: option.split)
show ?dlist ?dlist1 ?dlist2 ?dlist_monad ?monad_dlist ?monad ?monad1 ?monad2 ?collect1 ?collect2 ?complement
by(auto simp add: DList_set_def List.member_def dest!: Collection_Eq.ID_ceq split: option.splits)
qed
lemma Set_bind_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"Set.bind (Set_Monad xs) f = fold ((\<union>) \<circ> f) xs (Set_Monad [])" (is ?Set_Monad)
"Set.bind (DList_set dxs) f' =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''bind DList_set: ceq = None'') (\<lambda>_. Set.bind (DList_set dxs) f')
| Some _ \<Rightarrow> DList_Set.fold (union \<circ> f') dxs {})" (is ?DList)
"Set.bind (RBT_set rbt) f'' =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''bind RBT_set: ccompare = None'') (\<lambda>_. Set.bind (RBT_set rbt) f'')
| Some _ \<Rightarrow> RBT_Set2.fold (union \<circ> f'') rbt {})" (is ?RBT)
proof -
show ?Set_Monad by(simp add: set_bind_conv_fold)
show ?DList by(auto simp add: DList_set_def DList_Set.member_def List.member_def List.member_def[abs_def] set_bind_conv_fold DList_Set.fold_def split: option.split dest: equal.equal_eq[OF ID_ceq] ID_ceq)
show ?RBT by(clarsimp split: option.split simp add: RBT_set_def RBT_Set2.fold_conv_fold_keys RBT_Set2.member_conv_keys set_bind_conv_fold)
qed
lemma UNIV_code [code]: "UNIV = - {}"
by(simp)
lift_definition inf_sls :: "'a :: lattice semilattice_set" is "inf" by unfold_locales
lemma Inf_fin_code [code]: "Inf_fin A = set_fold1 inf_sls A"
by transfer(simp add: Inf_fin_def)
lift_definition sup_sls :: "'a :: lattice semilattice_set" is "sup" by unfold_locales
lemma Sup_fin_code [code]: "Sup_fin A = set_fold1 sup_sls A"
by transfer(simp add: Sup_fin_def)
lift_definition inf_cfi :: "('a :: lattice, 'a) comp_fun_idem" is "inf"
by(rule comp_fun_idem_inf)
lemma Inf_code:
fixes A :: "'a :: complete_lattice set" shows
"Inf A = (if finite A then set_fold_cfi inf_cfi top A else Code.abort (STR ''Inf: infinite'') (\<lambda>_. Inf A))"
by transfer(simp add: Inf_fold_inf)
lift_definition sup_cfi :: "('a :: lattice, 'a) comp_fun_idem" is "sup"
by(rule comp_fun_idem_sup)
lemma Sup_code:
fixes A :: "'a :: complete_lattice set" shows
"Sup A = (if finite A then set_fold_cfi sup_cfi bot A else Code.abort (STR ''Sup: infinite'') (\<lambda>_. Sup A))"
by transfer(simp add: Sup_fold_sup)
lemmas Inter_code [code] = Inf_code[where ?'a = "_ :: type set"]
lemmas Union_code [code] = Sup_code[where ?'a = "_ :: type set"]
lemmas Predicate_Inf_code [code] = Inf_code[where ?'a = "_ :: type Predicate.pred"]
lemmas Predicate_Sup_code [code] = Sup_code[where ?'a = "_ :: type Predicate.pred"]
lemmas Inf_fun_code [code] = Inf_code[where ?'a = "_ :: type \<Rightarrow> _ :: complete_lattice"]
lemmas Sup_fun_code [code] = Sup_code[where ?'a = "_ :: type \<Rightarrow> _ :: complete_lattice"]
lift_definition min_sls :: "'a :: linorder semilattice_set" is min by unfold_locales
lemma Min_code [code]: "Min A = set_fold1 min_sls A"
by transfer(simp add: Min_def)
lift_definition max_sls :: "'a :: linorder semilattice_set" is max by unfold_locales
lemma Max_code [code]: "Max A = set_fold1 max_sls A"
by transfer(simp add: Max_def)
text \<open>
We do not implement @{term Ball}, @{term Bex}, and @{term sorted_list_of_set} for @{term Collect_set} using @{term cEnum},
because it should already have been converted to an explicit list of elements if that is possible.
\<close>
lemma Ball_code [code]:
fixes rbt :: "'a :: ccompare set_rbt"
and dxs :: "'b :: ceq set_dlist" shows
"Ball (Set_Monad xs) P = list_all P xs"
"Ball (DList_set dxs) P' =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''Ball DList_set: ceq = None'') (\<lambda>_. Ball (DList_set dxs) P')
| Some _ \<Rightarrow> DList_Set.dlist_all P' dxs)"
"Ball (RBT_set rbt) P'' =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''Ball RBT_set: ccompare = None'') (\<lambda>_. Ball (RBT_set rbt) P'')
| Some _ \<Rightarrow> RBT_Set2.all P'' rbt)"
by(simp_all add: DList_set_def RBT_set_def list_all_iff dlist_all_conv_member RBT_Set2.all_conv_all_member split: option.splits)
lemma Bex_code [code]:
fixes rbt :: "'a :: ccompare set_rbt"
and dxs :: "'b :: ceq set_dlist" shows
"Bex (Set_Monad xs) P = list_ex P xs"
"Bex (DList_set dxs) P' =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''Bex DList_set: ceq = None'') (\<lambda>_. Bex (DList_set dxs) P')
| Some _ \<Rightarrow> DList_Set.dlist_ex P' dxs)"
"Bex (RBT_set rbt) P'' =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''Bex RBT_set: ccompare = None'') (\<lambda>_. Bex (RBT_set rbt) P'')
| Some _ \<Rightarrow> RBT_Set2.ex P'' rbt)"
by(simp_all add: DList_set_def RBT_set_def list_ex_iff dlist_ex_conv_member RBT_Set2.ex_conv_ex_member split: option.splits)
lemma csorted_list_of_set_code [code]:
fixes rbt :: "'a :: ccompare set_rbt"
and dxs :: "'b :: {ccompare, ceq} set_dlist"
and xs :: "'a :: ccompare list" shows
"csorted_list_of_set (RBT_set rbt) =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''csorted_list_of_set RBT_set: ccompare = None'') (\<lambda>_. csorted_list_of_set (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.keys rbt)"
"csorted_list_of_set (DList_set dxs) =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''csorted_list_of_set DList_set: ceq = None'') (\<lambda>_. csorted_list_of_set (DList_set dxs))
| Some _ \<Rightarrow>
case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''csorted_list_of_set DList_set: ccompare = None'') (\<lambda>_. csorted_list_of_set (DList_set dxs))
| Some c \<Rightarrow> ord.quicksort (lt_of_comp c) (list_of_dlist dxs))"
"csorted_list_of_set (Set_Monad xs) =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''csorted_list_of_set Set_Monad: ccompare = None'') (\<lambda>_. csorted_list_of_set (Set_Monad xs))
| Some c \<Rightarrow> ord.remdups_sorted (lt_of_comp c) (ord.quicksort (lt_of_comp c) xs))"
by(auto split: option.split simp add: RBT_set_def DList_set_def DList_Set.Collect_member member_conv_keys sorted_RBT_Set_keys linorder.sorted_list_of_set_sort_remdups[OF ID_ccompare] linorder.quicksort_conv_sort[OF ID_ccompare] distinct_remdups_id distinct_list_of_dlist linorder.remdups_sorted_conv_remdups[OF ID_ccompare] linorder.sorted_sort[OF ID_ccompare] linorder.sort_remdups[OF ID_ccompare] csorted_list_of_set_def)
lemma cless_set_code [code]:
fixes rbt rbt' :: "'a :: ccompare set_rbt"
and rbt1 rbt2 :: "'b :: cproper_interval set_rbt"
and A B :: "'a set"
and A' B' :: "'b set" shows
"cless_set A B \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cless_set: ccompare = None'') (\<lambda>_. cless_set A B)
| Some c \<Rightarrow>
if finite A \<and> finite B then ord.lexordp (\<lambda>x y. lt_of_comp c y x) (csorted_list_of_set A) (csorted_list_of_set B)
else Code.abort (STR ''cless_set: infinite set'') (\<lambda>_. cless_set A B))"
(is "?fin_fin")
and cless_set_Complement2 [set_complement_code]:
"cless_set A' (Complement B') \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_set Complement2: ccompare = None'') (\<lambda>_. cless_set A' (Complement B'))
| Some c \<Rightarrow>
if finite A' \<and> finite B' then
finite (UNIV :: 'b set) \<longrightarrow>
proper_intrvl.set_less_aux_Compl (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B')
else Code.abort (STR ''cless_set Complement2: infinite set'') (\<lambda>_. cless_set A' (Complement B')))"
(is "?fin_Compl_fin")
and cless_set_Complement1 [set_complement_code]:
"cless_set (Complement A') B' \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_set Complement1: ccompare = None'') (\<lambda>_. cless_set (Complement A') B')
| Some c \<Rightarrow>
if finite A' \<and> finite B' then
finite (UNIV :: 'b set) \<and>
proper_intrvl.Compl_set_less_aux (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B')
else Code.abort (STR ''cless_set Complement1: infinite set'') (\<lambda>_. cless_set (Complement A') B'))"
(is "?Compl_fin_fin")
and cless_set_Complement12 [set_complement_code]:
"cless_set (Complement A) (Complement B) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cless_set Complement Complement: ccompare = None'') (\<lambda>_. cless_set (Complement A) (Complement B))
| Some _ \<Rightarrow> cless B A)" (is ?Compl_Compl)
and
"cless_set (RBT_set rbt) (RBT_set rbt') \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cless_set RBT_set RBT_set: ccompare = None'') (\<lambda>_. cless_set (RBT_set rbt) (RBT_set rbt'))
| Some c \<Rightarrow> ord.lexord_fusion (\<lambda>x y. lt_of_comp c y x) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt) (RBT_Set2.init rbt'))"
(is ?rbt_rbt)
and cless_set_rbt_Complement2 [set_complement_code]:
"cless_set (RBT_set rbt1) (Complement (RBT_set rbt2)) \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_set RBT_set (Complement RBT_set): ccompare = None'') (\<lambda>_. cless_set (RBT_set rbt1) (Complement (RBT_set rbt2)))
| Some c \<Rightarrow>
finite (UNIV :: 'b set) \<longrightarrow>
proper_intrvl.set_less_aux_Compl_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))"
(is ?rbt_Compl)
and cless_set_rbt_Complement1 [set_complement_code]:
"cless_set (Complement (RBT_set rbt1)) (RBT_set rbt2) \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_set (Complement RBT_set) RBT_set: ccompare = None'') (\<lambda>_. cless_set (Complement (RBT_set rbt1)) (RBT_set rbt2))
| Some c \<Rightarrow>
finite (UNIV :: 'b set) \<and>
proper_intrvl.Compl_set_less_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))"
(is ?Compl_rbt)
proof -
note [split] = option.split csorted_list_of_set_split
and [simp] =
le_of_comp_of_ords_linorder[OF ID_ccompare]
lt_of_comp_of_ords
finite_subset[OF subset_UNIV] ccompare_set_def ID_Some
ord.lexord_fusion_def
proper_intrvl.Compl_set_less_aux_fusion_def
proper_intrvl.set_less_aux_Compl_fusion_def
unfoldr_rbt_keys_generator
RBT_set_def sorted_RBT_Set_keys member_conv_keys
linorder.set_less_finite_iff[OF ID_ccompare]
linorder.set_less_aux_code[OF ID_ccompare, symmetric]
linorder.Compl_set_less_Compl[OF ID_ccompare]
linorder.infinite_set_less_Complement[OF ID_ccompare]
linorder.infinite_Complement_set_less[OF ID_ccompare]
linorder_proper_interval.set_less_aux_Compl2_conv_set_less_aux_Compl[OF ID_ccompare_interval, symmetric]
linorder_proper_interval.Compl1_set_less_aux_conv_Compl_set_less_aux[OF ID_ccompare_interval, symmetric]
show ?Compl_Compl by simp
show ?rbt_rbt
by auto
show ?rbt_Compl by(cases "finite (UNIV :: 'b set)") auto
show ?Compl_rbt by(cases "finite (UNIV :: 'b set)") auto
show ?fin_fin by auto
show ?fin_Compl_fin by(cases "finite (UNIV :: 'b set)", auto)
show ?Compl_fin_fin by(cases "finite (UNIV :: 'b set)") auto
qed
lemma le_of_comp_set_less_eq:
"le_of_comp (comp_of_ords (ord.set_less_eq le) (ord.set_less le)) = ord.set_less_eq le"
by (rule le_of_comp_of_ords_gen, simp add: ord.set_less_def)
lemma cless_eq_set_code [code]:
fixes rbt rbt' :: "'a :: ccompare set_rbt"
and rbt1 rbt2 :: "'b :: cproper_interval set_rbt"
and A B :: "'a set"
and A' B' :: "'b set" shows
"cless_eq_set A B \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cless_eq_set: ccompare = None'') (\<lambda>_. cless_eq_set A B)
| Some c \<Rightarrow>
if finite A \<and> finite B then
ord.lexordp_eq (\<lambda>x y. lt_of_comp c y x) (csorted_list_of_set A) (csorted_list_of_set B)
else Code.abort (STR ''cless_eq_set: infinite set'') (\<lambda>_. cless_eq_set A B))"
(is "?fin_fin")
and cless_eq_set_Complement2 [set_complement_code]:
"cless_eq_set A' (Complement B') \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_eq_set Complement2: ccompare = None'') (\<lambda>_. cless_eq_set A' (Complement B'))
| Some c \<Rightarrow>
if finite A' \<and> finite B' then
finite (UNIV :: 'b set) \<longrightarrow>
proper_intrvl.set_less_eq_aux_Compl (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B')
else Code.abort (STR ''cless_eq_set Complement2: infinite set'') (\<lambda>_. cless_eq_set A' (Complement B')))"
(is "?fin_Compl_fin")
and cless_eq_set_Complement1 [set_complement_code]:
"cless_eq_set (Complement A') B' \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_eq_set Complement1: ccompare = None'') (\<lambda>_. cless_eq_set (Complement A') B')
| Some c \<Rightarrow>
if finite A' \<and> finite B' then
finite (UNIV :: 'b set) \<and>
proper_intrvl.Compl_set_less_eq_aux (lt_of_comp c) cproper_interval None (csorted_list_of_set A') (csorted_list_of_set B')
else Code.abort (STR ''cless_eq_set Complement1: infinite set'') (\<lambda>_. cless_eq_set (Complement A') B'))"
(is "?Compl_fin_fin")
and cless_eq_set_Complement12 [set_complement_code]:
"cless_eq_set (Complement A) (Complement B) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cless_eq_set Complement Complement: ccompare = None'') (\<lambda>_. cless_eq (Complement A) (Complement B))
| Some c \<Rightarrow> cless_eq_set B A)"
(is ?Compl_Compl)
"cless_eq_set (RBT_set rbt) (RBT_set rbt') \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cless_eq_set RBT_set RBT_set: ccompare = None'') (\<lambda>_. cless_eq_set (RBT_set rbt) (RBT_set rbt'))
| Some c \<Rightarrow> ord.lexord_eq_fusion (\<lambda>x y. lt_of_comp c y x) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt) (RBT_Set2.init rbt'))"
(is ?rbt_rbt)
and cless_eq_set_rbt_Complement2 [set_complement_code]:
"cless_eq_set (RBT_set rbt1) (Complement (RBT_set rbt2)) \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_eq_set RBT_set (Complement RBT_set): ccompare = None'') (\<lambda>_. cless_eq_set (RBT_set rbt1) (Complement (RBT_set rbt2)))
| Some c \<Rightarrow>
finite (UNIV :: 'b set) \<longrightarrow>
proper_intrvl.set_less_eq_aux_Compl_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))"
(is ?rbt_Compl)
and cless_eq_set_rbt_Complement1 [set_complement_code]:
"cless_eq_set (Complement (RBT_set rbt1)) (RBT_set rbt2) \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''cless_eq_set (Complement RBT_set) RBT_set: ccompare = None'') (\<lambda>_. cless_eq_set (Complement (RBT_set rbt1)) (RBT_set rbt2))
| Some c \<Rightarrow>
finite (UNIV :: 'b set) \<and>
proper_intrvl.Compl_set_less_eq_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))"
(is ?Compl_rbt)
proof -
note [split] = option.split csorted_list_of_set_split
and [simp] =
le_of_comp_set_less_eq
finite_subset[OF subset_UNIV] ccompare_set_def ID_Some
ord.lexord_eq_fusion_def proper_intrvl.Compl_set_less_eq_aux_fusion_def
proper_intrvl.set_less_eq_aux_Compl_fusion_def
unfoldr_rbt_keys_generator
RBT_set_def sorted_RBT_Set_keys member_conv_keys
linorder.set_less_eq_finite_iff[OF ID_ccompare]
linorder.set_less_eq_aux_code[OF ID_ccompare, symmetric]
linorder.Compl_set_less_eq_Compl[OF ID_ccompare]
linorder.infinite_set_less_eq_Complement[OF ID_ccompare]
linorder.infinite_Complement_set_less_eq[OF ID_ccompare]
linorder_proper_interval.set_less_eq_aux_Compl2_conv_set_less_eq_aux_Compl[OF ID_ccompare_interval, symmetric]
linorder_proper_interval.Compl1_set_less_eq_aux_conv_Compl_set_less_eq_aux[OF ID_ccompare_interval, symmetric]
show ?Compl_Compl by simp
show ?rbt_rbt
by auto
show ?rbt_Compl by(cases "finite (UNIV :: 'b set)") auto
show ?Compl_rbt by(cases "finite (UNIV :: 'b set)") auto
show ?fin_fin by auto
show ?fin_Compl_fin by (cases "finite (UNIV :: 'b set)", auto)
show ?Compl_fin_fin by(cases "finite (UNIV :: 'b set)") auto
qed
lemma cproper_interval_set_Some_Some_code [code]:
fixes rbt1 rbt2 :: "'a :: cproper_interval set_rbt"
and A B :: "'a set" shows
"cproper_interval (Some A) (Some B) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cproper_interval: ccompare = None'') (\<lambda>_. cproper_interval (Some A) (Some B))
| Some c \<Rightarrow>
finite (UNIV :: 'a set) \<and> proper_intrvl.proper_interval_set_aux (lt_of_comp c) cproper_interval (csorted_list_of_set A) (csorted_list_of_set B))"
(is ?fin_fin)
and cproper_interval_set_Some_Some_Complement [set_complement_code]:
"cproper_interval (Some A) (Some (Complement B)) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cproper_interval Complement2: ccompare = None'') (\<lambda>_. cproper_interval (Some A) (Some (Complement B)))
| Some c \<Rightarrow>
finite (UNIV :: 'a set) \<and> proper_intrvl.proper_interval_set_Compl_aux (lt_of_comp c) cproper_interval None 0 (csorted_list_of_set A) (csorted_list_of_set B))"
(is ?fin_Compl_fin)
and cproper_interval_set_Some_Complement_Some [set_complement_code]:
"cproper_interval (Some (Complement A)) (Some B) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cproper_interval Complement1: ccompare = None'') (\<lambda>_. cproper_interval (Some (Complement A)) (Some B))
| Some c \<Rightarrow>
finite (UNIV :: 'a set) \<and> proper_intrvl.proper_interval_Compl_set_aux (lt_of_comp c) cproper_interval None (csorted_list_of_set A) (csorted_list_of_set B))"
(is ?Compl_fin_fin)
and cproper_interval_set_Some_Complement_Some_Complement [set_complement_code]:
"cproper_interval (Some (Complement A)) (Some (Complement B)) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cproper_interval Complement Complement: ccompare = None'') (\<lambda>_. cproper_interval (Some (Complement A)) (Some (Complement B)))
| Some _ \<Rightarrow> cproper_interval (Some B) (Some A))"
(is ?Compl_Compl)
"cproper_interval (Some (RBT_set rbt1)) (Some (RBT_set rbt2)) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cproper_interval RBT_set RBT_set: ccompare = None'') (\<lambda>_. cproper_interval (Some (RBT_set rbt1)) (Some (RBT_set rbt2)))
| Some c \<Rightarrow>
finite (UNIV :: 'a set) \<and> proper_intrvl.proper_interval_set_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))"
(is ?rbt_rbt)
and cproper_interval_set_Some_rbt_Some_Complement [set_complement_code]:
"cproper_interval (Some (RBT_set rbt1)) (Some (Complement (RBT_set rbt2))) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cproper_interval RBT_set (Complement RBT_set): ccompare = None'') (\<lambda>_. cproper_interval (Some (RBT_set rbt1)) (Some (Complement (RBT_set rbt2))))
| Some c \<Rightarrow>
finite (UNIV :: 'a set) \<and> proper_intrvl.proper_interval_set_Compl_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None 0 (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))"
(is ?rbt_Compl_rbt)
and cproper_interval_set_Some_Complement_Some_rbt [set_complement_code]:
"cproper_interval (Some (Complement (RBT_set rbt1))) (Some (RBT_set rbt2)) \<longleftrightarrow>
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''cproper_interval (Complement RBT_set) RBT_set: ccompare = None'') (\<lambda>_. cproper_interval (Some (Complement (RBT_set rbt1))) (Some (RBT_set rbt2)))
| Some c \<Rightarrow>
finite (UNIV :: 'a set) \<and> proper_intrvl.proper_interval_Compl_set_aux_fusion (lt_of_comp c) cproper_interval rbt_keys_generator rbt_keys_generator None (RBT_Set2.init rbt1) (RBT_Set2.init rbt2))"
(is ?Compl_rbt_rbt)
proof -
note [split] = option.split csorted_list_of_set_split
and [simp] =
lt_of_comp_of_ords
finite_subset[OF subset_UNIV] ccompare_set_def ID_Some
linorder.set_less_finite_iff[OF ID_ccompare]
RBT_set_def sorted_RBT_Set_keys member_conv_keys
linorder.distinct_entries[OF ID_ccompare]
unfoldr_rbt_keys_generator
proper_intrvl.proper_interval_set_aux_fusion_def
proper_intrvl.proper_interval_set_Compl_aux_fusion_def
proper_intrvl.proper_interval_Compl_set_aux_fusion_def
linorder_proper_interval.proper_interval_set_aux[OF ID_ccompare_interval]
linorder_proper_interval.proper_interval_set_Compl_aux[OF ID_ccompare_interval]
linorder_proper_interval.proper_interval_Compl_set_aux[OF ID_ccompare_interval]
and [cong] = conj_cong
show ?Compl_Compl
by(clarsimp simp add: Complement_cproper_interval_set_Complement simp del: cproper_interval_set_Some_Some)
show ?rbt_rbt ?rbt_Compl_rbt ?Compl_rbt_rbt by auto
show ?fin_fin ?fin_Compl_fin ?Compl_fin_fin by auto
qed
context ord begin
fun sorted_list_subset :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where
"sorted_list_subset eq [] ys = True"
| "sorted_list_subset eq (x # xs) [] = False"
| "sorted_list_subset eq (x # xs) (y # ys) \<longleftrightarrow>
(if eq x y then sorted_list_subset eq xs ys
else x > y \<and> sorted_list_subset eq (x # xs) ys)"
end
context linorder begin
lemma sorted_list_subset_correct:
"\<lbrakk> sorted xs; distinct xs; sorted ys; distinct ys \<rbrakk>
\<Longrightarrow> sorted_list_subset (=) xs ys \<longleftrightarrow> set xs \<subseteq> set ys"
-apply(induct "(=) :: 'a \<Rightarrow> 'a \<Rightarrow> bool" xs ys rule: sorted_list_subset.induct)
-apply(auto 6 2)
-apply auto
-by (metis eq_iff insert_iff subsetD)
+ apply(induct "(=) :: 'a \<Rightarrow> 'a \<Rightarrow> bool" xs ys rule: sorted_list_subset.induct)
+ apply(auto 6 2)
+ using order_antisym apply auto
+ done
end
context ord begin
definition sorted_list_subset_fusion :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where "sorted_list_subset_fusion eq g1 g2 s1 s2 = sorted_list_subset eq (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
lemma sorted_list_subset_fusion_code:
"sorted_list_subset_fusion eq g1 g2 s1 s2 =
(if list.has_next g1 s1 then
let (x, s1') = list.next g1 s1
in list.has_next g2 s2 \<and> (
let (y, s2') = list.next g2 s2
in if eq x y then sorted_list_subset_fusion eq g1 g2 s1' s2'
else y < x \<and> sorted_list_subset_fusion eq g1 g2 s1 s2')
else True)"
unfolding sorted_list_subset_fusion_def
by(subst (1 2 5) list.unfoldr.simps)(simp add: split_beta Let_def)
end
lemmas [code] = ord.sorted_list_subset_fusion_code
text \<open>
Define a new constant for the subset operation
because @{theory "HOL-Library.Cardinality"} introduces @{const "Cardinality.subset'"}
and rewrites @{const "subset"} to @{const "Cardinality.subset'"}
based on the sort of the element type.
\<close>
definition subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
where [simp, code del]: "subset_eq = (\<subseteq>)"
lemma subseteq_code [code]: "(\<subseteq>) = subset_eq"
by simp
lemma subset'_code [code]: "Cardinality.subset' = subset_eq"
by simp
lemma subset_eq_code [folded subset_eq_def, code]:
fixes A1 A2 :: "'a set"
and rbt :: "'b :: ccompare set_rbt"
and rbt1 rbt2 :: "'d :: {ccompare, ceq} set_rbt"
and dxs :: "'c :: ceq set_dlist"
and xs :: "'c list" shows
"RBT_set rbt \<subseteq> B \<longleftrightarrow>
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''subset RBT_set1: ccompare = None'') (\<lambda>_. RBT_set rbt \<subseteq> B)
| Some _ \<Rightarrow> list_all_fusion rbt_keys_generator (\<lambda>x. x \<in> B) (RBT_Set2.init rbt))" (is ?rbt)
"DList_set dxs \<subseteq> C \<longleftrightarrow>
(case ID CEQ('c) of None \<Rightarrow> Code.abort (STR ''subset DList_set1: ceq = None'') (\<lambda>_. DList_set dxs \<subseteq> C)
| Some _ \<Rightarrow> DList_Set.dlist_all (\<lambda>x. x \<in> C) dxs)" (is ?dlist)
"Set_Monad xs \<subseteq> C \<longleftrightarrow> list_all (\<lambda>x. x \<in> C) xs" (is ?Set_Monad)
and Collect_subset_eq_Complement [folded subset_eq_def, set_complement_code]:
"Collect_set P \<subseteq> Complement A \<longleftrightarrow> A \<subseteq> {x. \<not> P x}" (is ?Collect_set_Compl)
and Complement_subset_eq_Complement [folded subset_eq_def, set_complement_code]:
"Complement A1 \<subseteq> Complement A2 \<longleftrightarrow> A2 \<subseteq> A1" (is ?Compl)
and
"RBT_set rbt1 \<subseteq> RBT_set rbt2 \<longleftrightarrow>
(case ID CCOMPARE('d) of None \<Rightarrow> Code.abort (STR ''subset RBT_set RBT_set: ccompare = None'') (\<lambda>_. RBT_set rbt1 \<subseteq> RBT_set rbt2)
| Some c \<Rightarrow>
(case ID CEQ('d) of None \<Rightarrow> ord.sorted_list_subset_fusion (lt_of_comp c) (\<lambda> x y. c x y = Eq) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2)
| Some eq \<Rightarrow> ord.sorted_list_subset_fusion (lt_of_comp c) eq rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2)))"
(is ?rbt_rbt)
proof -
show ?rbt_rbt
by (auto simp add: comparator.eq[OF ID_ccompare'] RBT_set_def member_conv_keys unfoldr_rbt_keys_generator ord.sorted_list_subset_fusion_def linorder.sorted_list_subset_correct[OF ID_ccompare] sorted_RBT_Set_keys split: option.split dest!: ID_ceq[THEN equal.equal_eq] del: iffI)
show ?rbt
by(auto simp add: RBT_set_def member_conv_keys list_all_fusion_def unfoldr_rbt_keys_generator keys.rep_eq list_all_iff split: option.split)
show ?dlist by(auto simp add: DList_set_def dlist_all_conv_member split: option.split)
show ?Set_Monad by(auto simp add: list_all_iff split: option.split)
show ?Collect_set_Compl ?Compl by auto
qed
hide_const (open) subset_eq
hide_fact (open) subset_eq_def
lemma eq_set_code [code]: "Cardinality.eq_set = set_eq"
by(simp add: set_eq_def)
lemma set_eq_code [code]:
fixes rbt1 rbt2 :: "'b :: {ccompare, ceq} set_rbt" shows
"set_eq A B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
and set_eq_Complement_Complement [set_complement_code]:
"set_eq (Complement A) (Complement B) = set_eq A B"
and
"set_eq (RBT_set rbt1) (RBT_set rbt2) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''set_eq RBT_set RBT_set: ccompare = None'') (\<lambda>_. set_eq (RBT_set rbt1) (RBT_set rbt2))
| Some c \<Rightarrow>
(case ID CEQ('b) of None \<Rightarrow> list_all2_fusion (\<lambda> x y. c x y = Eq) rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2)
| Some eq \<Rightarrow> list_all2_fusion eq rbt_keys_generator rbt_keys_generator (RBT_Set2.init rbt1) (RBT_Set2.init rbt2)))"
(is ?rbt_rbt)
proof -
show ?rbt_rbt
by (auto 4 3 split: option.split simp add: comparator.eq[OF ID_ccompare'] sorted_RBT_Set_keys list_all2_fusion_def unfoldr_rbt_keys_generator RBT_set_conv_keys set_eq_def list.rel_eq dest!: ID_ceq[THEN equal.equal_eq] intro: linorder.sorted_distinct_set_unique[OF ID_ccompare])
qed(auto simp add: set_eq_def)
lemma Set_project_code [code]:
"Set.filter P A = A \<inter> Collect_set P"
by(auto simp add: Set.filter_def)
lemma Set_image_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"image f (Set_Monad xs) = Set_Monad (map f xs)"
"image f (Collect_set A) = Code.abort (STR ''image Collect_set'') (\<lambda>_. image f (Collect_set A))"
and image_Complement_Complement [set_complement_code]:
"image f (Complement (Complement B)) = image f B"
and
"image g (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''image DList_set: ceq = None'') (\<lambda>_. image g (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.fold (insert \<circ> g) dxs {})"
(is ?dlist)
"image h (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''image RBT_set: ccompare = None'') (\<lambda>_. image h (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.fold (insert \<circ> h) rbt {})"
(is ?rbt)
proof -
{ fix xs have "fold (insert \<circ> g) xs {} = g ` set xs"
by(induct xs rule: rev_induct) simp_all }
thus ?dlist
by(simp add: DList_set_def DList_Set.fold_def DList_Set.Collect_member split: option.split)
{ fix xs have "fold (insert \<circ> h) xs {} = h ` set xs"
by(induct xs rule: rev_induct) simp_all }
thus ?rbt by(auto simp add: RBT_set_def fold_conv_fold_keys member_conv_keys split: option.split)
qed simp_all
lemma the_elem_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"the_elem (Set_Monad [x]) = x"
"the_elem (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''the_elem DList_set: ceq = None'') (\<lambda>_. the_elem (DList_set dxs))
| Some _ \<Rightarrow>
case list_of_dlist dxs of [x] \<Rightarrow> x
| _ \<Rightarrow> Code.abort (STR ''the_elem DList_set: not unique'') (\<lambda>_. the_elem (DList_set dxs)))"
"the_elem (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''the_elem RBT_set: ccompare = None'') (\<lambda>_. the_elem (RBT_set rbt))
| Some _ \<Rightarrow>
case RBT_Mapping2.impl_of rbt of RBT_Impl.Branch _ RBT_Impl.Empty x _ RBT_Impl.Empty \<Rightarrow> x
| _ \<Rightarrow> Code.abort (STR ''the_elem RBT_set: not unique'') (\<lambda>_. the_elem (RBT_set rbt)))"
by(auto simp add: RBT_set_def DList_set_def DList_Set.Collect_member the_elem_def member_conv_keys split: option.split list.split rbt.split)(simp add: RBT_Set2.keys_def)
lemma Pow_set_conv_fold:
"Pow (set xs \<union> A) = fold (\<lambda>x A. A \<union> insert x ` A) xs (Pow A)"
by(induct xs rule: rev_induct)(auto simp add: Pow_insert)
lemma Pow_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"Pow A = Collect_set (\<lambda>B. B \<subseteq> A)"
"Pow (Set_Monad xs) = fold (\<lambda>x A. A \<union> insert x ` A) xs {{}}"
"Pow (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''Pow DList_set: ceq = None'') (\<lambda>_. Pow (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.fold (\<lambda>x A. A \<union> insert x ` A) dxs {{}})"
"Pow (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''Pow RBT_set: ccompare = None'') (\<lambda>_. Pow (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.fold (\<lambda>x A. A \<union> insert x ` A) rbt {{}})"
by(auto simp add: DList_set_def DList_Set.Collect_member DList_Set.fold_def RBT_set_def fold_conv_fold_keys member_conv_keys Pow_set_conv_fold[where A="{}", simplified] split: option.split)
lemma fold_singleton: "Finite_Set.fold f x {y} = f y x"
by(fastforce simp add: Finite_Set.fold_def intro: fold_graph.intros elim: fold_graph.cases)
lift_definition sum_cfc :: "('a \<Rightarrow> 'b :: comm_monoid_add) \<Rightarrow> ('a, 'b) comp_fun_commute"
is "\<lambda>f :: 'a \<Rightarrow> 'b. plus \<circ> f"
by(unfold_locales)(simp add: fun_eq_iff add.left_commute)
lemma sum_code [code]:
"sum f A = (if finite A then set_fold_cfc (sum_cfc f) 0 A else 0)"
by transfer(simp add: sum.eq_fold)
lemma product_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and dys :: "'b :: ceq set_dlist"
and rbt1 :: "'c :: ccompare set_rbt"
and rbt2 :: "'d :: ccompare set_rbt" shows
"Product_Type.product A B = Collect_set (\<lambda>(x, y). x \<in> A \<and> y \<in> B)"
"Product_Type.product (Set_Monad xs) (Set_Monad ys) =
Set_Monad (fold (\<lambda>x. fold (\<lambda>y rest. (x, y) # rest) ys) xs [])"
(is ?Set_Monad)
"Product_Type.product (DList_set dxs) B1 =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''product DList_set1: ceq = None'') (\<lambda>_. Product_Type.product (DList_set dxs) B1)
| Some _ \<Rightarrow> DList_Set.fold (\<lambda>x rest. Pair x ` B1 \<union> rest) dxs {})"
(is "?dlist1")
"Product_Type.product A1 (DList_set dys) =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''product DList_set2: ceq = None'') (\<lambda>_. Product_Type.product A1 (DList_set dys))
| Some _ \<Rightarrow> DList_Set.fold (\<lambda>y rest. (\<lambda>x. (x, y)) ` A1 \<union> rest) dys {})"
(is "?dlist2")
"Product_Type.product (DList_set dxs) (DList_set dys) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''product DList_set DList_set: ceq1 = None'') (\<lambda>_. Product_Type.product (DList_set dxs) (DList_set dys))
| Some _ \<Rightarrow>
case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''product DList_set DList_set: ceq2 = None'') (\<lambda>_. Product_Type.product (DList_set dxs) (DList_set dys))
| Some _ \<Rightarrow> DList_set (DList_Set.product dxs dys))"
"Product_Type.product (RBT_set rbt1) B2 =
(case ID CCOMPARE('c) of None \<Rightarrow> Code.abort (STR ''product RBT_set: ccompare1 = None'') (\<lambda>_. Product_Type.product (RBT_set rbt1) B2)
| Some _ \<Rightarrow> RBT_Set2.fold (\<lambda>x rest. Pair x ` B2 \<union> rest) rbt1 {})"
(is "?rbt1")
"Product_Type.product A2 (RBT_set rbt2) =
(case ID CCOMPARE('d) of None \<Rightarrow> Code.abort (STR ''product RBT_set: ccompare2 = None'') (\<lambda>_. Product_Type.product A2 (RBT_set rbt2))
| Some _ \<Rightarrow> RBT_Set2.fold (\<lambda>y rest. (\<lambda>x. (x, y)) ` A2 \<union> rest) rbt2 {})"
(is "?rbt2")
"Product_Type.product (RBT_set rbt1) (RBT_set rbt2) =
(case ID CCOMPARE('c) of None \<Rightarrow> Code.abort (STR ''product RBT_set RBT_set: ccompare1 = None'') (\<lambda>_. Product_Type.product (RBT_set rbt1) (RBT_set rbt2))
| Some _ \<Rightarrow>
case ID CCOMPARE('d) of None \<Rightarrow> Code.abort (STR ''product RBT_set RBT_set: ccompare2 = None'') (\<lambda>_. Product_Type.product (RBT_set rbt1) (RBT_set rbt2))
| Some _ \<Rightarrow> RBT_set (RBT_Set2.product rbt1 rbt2))"
proof -
have [simp]: "\<And>a zs. fold (\<lambda>y. (#) (a, y)) ys zs = rev (map (Pair a) ys) @ zs"
by(induct ys) simp_all
have [simp]: "\<And>zs. fold (\<lambda>x. fold (\<lambda>y rest. (x, y) # rest) ys) xs zs = rev (concat (map (\<lambda>x. map (Pair x) ys) xs)) @ zs"
by(induct xs) simp_all
show ?Set_Monad by(auto simp add: Product_Type.product_def)
{ fix xs :: "'a list"
have "fold (\<lambda>x. (\<union>) (Pair x ` B1)) xs {} = set xs \<times> B1"
by(induct xs rule: rev_induct) auto }
thus ?dlist1
by(simp add: Product_Type.product_def DList_set_def DList_Set.fold.rep_eq DList_Set.Collect_member split: option.split)
{ fix ys :: "'b list"
have "fold (\<lambda>y. (\<union>) ((\<lambda>x. (x, y)) ` A1)) ys {} = A1 \<times> set ys"
by(induct ys rule: rev_induct) auto }
thus ?dlist2
by(simp add: Product_Type.product_def DList_set_def DList_Set.fold.rep_eq DList_Set.Collect_member split: option.split)
{ fix xs :: "'c list"
have "fold (\<lambda>x. (\<union>) (Pair x ` B2)) xs {} = set xs \<times> B2"
by(induct xs rule: rev_induct) auto }
thus ?rbt1
by(simp add: Product_Type.product_def RBT_set_def RBT_Set2.member_product RBT_Set2.member_conv_keys fold_conv_fold_keys split: option.split)
{ fix ys :: "'d list"
have "fold (\<lambda>y. (\<union>) ((\<lambda>x. (x, y)) ` A2)) ys {} = A2 \<times> set ys"
by(induct ys rule: rev_induct) auto }
thus ?rbt2
by(simp add: Product_Type.product_def RBT_set_def RBT_Set2.member_product RBT_Set2.member_conv_keys fold_conv_fold_keys split: option.split)
qed(auto simp add: RBT_set_def DList_set_def Product_Type.product_def DList_Set.product_member RBT_Set2.member_product split: option.split)
lemma Id_on_code [code]:
fixes A :: "'a :: ceq set"
and dxs :: "'a set_dlist"
and P :: "'a \<Rightarrow> bool"
and rbt :: "'b :: ccompare set_rbt" shows
"Id_on B = (\<lambda>x. (x, x)) ` B"
and Id_on_Complement [set_complement_code]:
"Id_on (Complement A) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''Id_on Complement: ceq = None'') (\<lambda>_. Id_on (Complement A))
| Some eq \<Rightarrow> Collect_set (\<lambda>(x, y). eq x y \<and> x \<notin> A))"
and
"Id_on (Collect_set P) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''Id_on Collect_set: ceq = None'') (\<lambda>_. Id_on (Collect_set P))
| Some eq \<Rightarrow> Collect_set (\<lambda>(x, y). eq x y \<and> P x))"
"Id_on (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''Id_on DList_set: ceq = None'') (\<lambda>_. Id_on (DList_set dxs))
| Some _ \<Rightarrow> DList_set (DList_Set.Id_on dxs))"
"Id_on (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''Id_on RBT_set: ccompare = None'') (\<lambda>_. Id_on (RBT_set rbt))
| Some _ \<Rightarrow> RBT_set (RBT_Set2.Id_on rbt))"
by(auto simp add: DList_set_def RBT_set_def DList_Set.member_Id_on RBT_Set2.member_Id_on dest: equal.equal_eq[OF ID_ceq] split: option.split)
lemma Image_code [code]:
fixes dxs :: "('a :: ceq \<times> 'b :: ceq) set_dlist"
and rbt :: "('c :: ccompare \<times> 'd :: ccompare) set_rbt" shows
"X `` Y = snd ` Set.filter (\<lambda>(x, y). x \<in> Y) X"
(is ?generic)
"Set_Monad rxs `` A = Set_Monad (fold (\<lambda>(x, y) rest. if x \<in> A then y # rest else rest) rxs [])"
(is ?Set_Monad)
"DList_set dxs `` B =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''Image DList_set: ceq1 = None'') (\<lambda>_. DList_set dxs `` B)
| Some _ \<Rightarrow>
case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''Image DList_set: ceq2 = None'') (\<lambda>_. DList_set dxs `` B)
| Some _ \<Rightarrow>
DList_Set.fold (\<lambda>(x, y) acc. if x \<in> B then insert y acc else acc) dxs {})"
(is ?DList_set)
"RBT_set rbt `` C =
(case ID CCOMPARE('c) of None \<Rightarrow> Code.abort (STR ''Image RBT_set: ccompare1 = None'') (\<lambda>_. RBT_set rbt `` C)
| Some _ \<Rightarrow>
case ID CCOMPARE('d) of None \<Rightarrow> Code.abort (STR ''Image RBT_set: ccompare2 = None'') (\<lambda>_. RBT_set rbt `` C)
| Some _ \<Rightarrow>
RBT_Set2.fold (\<lambda>(x, y) acc. if x \<in> C then insert y acc else acc) rbt {})"
(is ?RBT_set)
proof -
show ?generic by(auto intro: rev_image_eqI)
have "set (fold (\<lambda>(x, y) rest. if x \<in> A then y # rest else rest) rxs []) = set rxs `` A"
by(induct rxs rule: rev_induct)(auto split: if_split_asm)
thus ?Set_Monad by(auto)
{ fix dxs :: "('a \<times> 'b) list"
have "fold (\<lambda>(x, y) acc. if x \<in> B then insert y acc else acc) dxs {} = set dxs `` B"
by(induct dxs rule: rev_induct)(auto split: if_split_asm) }
thus ?DList_set
by(clarsimp simp add: DList_set_def Collect_member ceq_prod_def ID_Some DList_Set.fold.rep_eq split: option.split)
{ fix rbt :: "(('c \<times> 'd) \<times> unit) list"
have "fold (\<lambda>(a, _). case a of (x, y) \<Rightarrow> \<lambda>acc. if x \<in> C then insert y acc else acc) rbt {} = (fst ` set rbt) `` C"
by(induct rbt rule: rev_induct)(auto simp add: split_beta split: if_split_asm) }
thus ?RBT_set
by(clarsimp simp add: RBT_set_def ccompare_prod_def ID_Some RBT_Set2.fold.rep_eq member_conv_keys RBT_Set2.keys.rep_eq RBT_Impl.fold_def RBT_Impl.keys_def split: option.split)
qed
lemma insert_relcomp: "insert (a, b) A O B = A O B \<union> {a} \<times> {c. (b, c) \<in> B}"
by auto
lemma trancl_code [code]:
"trancl A =
(if finite A then ntrancl (card A - 1) A else Code.abort (STR ''trancl: infinite set'') (\<lambda>_. trancl A))"
by (simp add: finite_trancl_ntranl)
lemma set_relcomp_set:
"set xs O set ys = fold (\<lambda>(x, y). fold (\<lambda>(y', z) A. if y = y' then insert (x, z) A else A) ys) xs {}"
proof(induct xs rule: rev_induct)
case Nil show ?case by simp
next
case (snoc x xs)
note [[show_types]]
{ fix a :: 'a and b :: 'c and X :: "('a \<times> 'b) set"
have "fold (\<lambda>(y', z) A. if b = y' then insert (a, z) A else A) ys X = X \<union> {a} \<times> {c. (b, c) \<in> set ys}"
by(induct ys arbitrary: X rule: rev_induct)(auto split: if_split_asm) }
thus ?case using snoc by(cases x)(simp add: insert_relcomp)
qed
lemma If_not: "(if \<not> a then b else c) = (if a then c else b)"
by auto
lemma relcomp_code [code]:
fixes rbt1 :: "('a :: ccompare \<times> 'b :: ccompare) set_rbt"
and rbt2 :: "('b \<times> 'c :: ccompare) set_rbt"
and rbt3 :: "('a \<times> 'd :: {ccompare, ceq}) set_rbt"
and rbt4 :: "('d \<times> 'a) set_rbt"
and rbt5 :: "('b \<times> 'a) set_rbt"
and dxs1 :: "('d \<times> 'e :: ceq) set_dlist"
and dxs2 :: "('e \<times> 'd) set_dlist"
and dxs3 :: "('e \<times> 'f :: ceq) set_dlist"
and dxs4 :: "('f \<times> 'g :: ceq) set_dlist"
and xs1 :: "('h \<times> 'i :: ceq) list"
and xs2 :: "('i \<times> 'j) list"
and xs3 :: "('b \<times> 'h) list"
and xs4 :: "('h \<times> 'b) list"
and xs5 :: "('f \<times> 'h) list"
and xs6 :: "('h \<times> 'f) list"
shows
"RBT_set rbt1 O RBT_set rbt2 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set RBT_set: ccompare1 = None'') (\<lambda>_. RBT_set rbt1 O RBT_set rbt2)
| Some _ \<Rightarrow>
case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set RBT_set: ccompare2 = None'') (\<lambda>_. RBT_set rbt1 O RBT_set rbt2)
| Some c_b \<Rightarrow>
case ID CCOMPARE('c) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set RBT_set: ccompare3 = None'') (\<lambda>_. RBT_set rbt1 O RBT_set rbt2)
| Some _ \<Rightarrow> RBT_Set2.fold (\<lambda>(x, y). RBT_Set2.fold (\<lambda>(y', z) A. if c_b y y' \<noteq> Eq then A else insert (x, z) A) rbt2) rbt1 {})"
(is ?rbt_rbt)
"RBT_set rbt3 O DList_set dxs1 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set DList_set: ccompare1 = None'') (\<lambda>_. RBT_set rbt3 O DList_set dxs1)
| Some _ \<Rightarrow>
case ID CCOMPARE('d) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set DList_set: ccompare2 = None'') (\<lambda>_. RBT_set rbt3 O DList_set dxs1)
| Some _ \<Rightarrow>
case ID CEQ('d) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set DList_set: ceq2 = None'') (\<lambda>_. RBT_set rbt3 O DList_set dxs1)
| Some eq \<Rightarrow>
case ID CEQ('e) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set DList_set: ceq3 = None'') (\<lambda>_. RBT_set rbt3 O DList_set dxs1)
| Some _ \<Rightarrow> RBT_Set2.fold (\<lambda>(x, y). DList_Set.fold (\<lambda>(y', z) A. if eq y y' then insert (x, z) A else A) dxs1) rbt3 {})"
(is ?rbt_dlist)
"DList_set dxs2 O RBT_set rbt4 =
(case ID CEQ('e) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set RBT_set: ceq1 = None'') (\<lambda>_. DList_set dxs2 O RBT_set rbt4)
| Some _ \<Rightarrow>
case ID CCOMPARE('d) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set RBT_set: ceq2 = None'') (\<lambda>_. DList_set dxs2 O RBT_set rbt4)
| Some _ \<Rightarrow>
case ID CEQ('d) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set RBT_set: ccompare2 = None'') (\<lambda>_. DList_set dxs2 O RBT_set rbt4)
| Some eq \<Rightarrow>
case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set RBT_set: ccompare3 = None'') (\<lambda>_. DList_set dxs2 O RBT_set rbt4)
| Some _ \<Rightarrow> DList_Set.fold (\<lambda>(x, y). RBT_Set2.fold (\<lambda>(y', z) A. if eq y y' then insert (x, z) A else A) rbt4) dxs2 {})"
(is ?dlist_rbt)
"DList_set dxs3 O DList_set dxs4 =
(case ID CEQ('e) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set DList_set: ceq1 = None'') (\<lambda>_. DList_set dxs3 O DList_set dxs4)
| Some _ \<Rightarrow>
case ID CEQ('f) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set DList_set: ceq2 = None'') (\<lambda>_. DList_set dxs3 O DList_set dxs4)
| Some eq \<Rightarrow>
case ID CEQ('g) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set DList_set: ceq3 = None'') (\<lambda>_. DList_set dxs3 O DList_set dxs4)
| Some _ \<Rightarrow> DList_Set.fold (\<lambda>(x, y). DList_Set.fold (\<lambda>(y', z) A. if eq y y' then insert (x, z) A else A) dxs4) dxs3 {})"
(is ?dlist_dlist)
"Set_Monad xs1 O Set_Monad xs2 =
(case ID CEQ('i) of None \<Rightarrow> Code.abort (STR ''relcomp Set_Monad Set_Monad: ceq = None'') (\<lambda>_. Set_Monad xs1 O Set_Monad xs2)
| Some eq \<Rightarrow> fold (\<lambda>(x, y). fold (\<lambda>(y', z) A. if eq y y' then insert (x, z) A else A) xs2) xs1 {})"
(is ?monad_monad)
"RBT_set rbt1 O Set_Monad xs3 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set Set_Monad: ccompare1 = None'') (\<lambda>_. RBT_set rbt1 O Set_Monad xs3)
| Some _ \<Rightarrow>
case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''relcomp RBT_set Set_Monad: ccompare2 = None'') (\<lambda>_. RBT_set rbt1 O Set_Monad xs3)
| Some c_b \<Rightarrow> RBT_Set2.fold (\<lambda>(x, y). fold (\<lambda>(y', z) A. if c_b y y' \<noteq> Eq then A else insert (x, z) A) xs3) rbt1 {})"
(is ?rbt_monad)
"Set_Monad xs4 O RBT_set rbt5 =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''relcomp Set_Monad RBT_set: ccompare1 = None'') (\<lambda>_. Set_Monad xs4 O RBT_set rbt5)
| Some _ \<Rightarrow>
case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''relcomp Set_Monad RBT_set: ccompare2 = None'') (\<lambda>_. Set_Monad xs4 O RBT_set rbt5)
| Some c_b \<Rightarrow> fold (\<lambda>(x, y). RBT_Set2.fold (\<lambda>(y', z) A. if c_b y y' \<noteq> Eq then A else insert (x, z) A) rbt5) xs4 {})"
(is ?monad_rbt)
"DList_set dxs3 O Set_Monad xs5 =
(case ID CEQ('e) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set Set_Monad: ceq1 = None'') (\<lambda>_. DList_set dxs3 O Set_Monad xs5)
| Some _ \<Rightarrow>
case ID CEQ('f) of None \<Rightarrow> Code.abort (STR ''relcomp DList_set Set_Monad: ceq2 = None'') (\<lambda>_. DList_set dxs3 O Set_Monad xs5)
| Some eq \<Rightarrow> DList_Set.fold (\<lambda>(x, y). fold (\<lambda>(y', z) A. if eq y y' then insert (x, z) A else A) xs5) dxs3 {})"
(is ?dlist_monad)
"Set_Monad xs6 O DList_set dxs4 =
(case ID CEQ('f) of None \<Rightarrow> Code.abort (STR ''relcomp Set_Monad DList_set: ceq1 = None'') (\<lambda>_. Set_Monad xs6 O DList_set dxs4)
| Some eq \<Rightarrow>
case ID CEQ('g) of None \<Rightarrow> Code.abort (STR ''relcomp Set_Monad DList_set: ceq2 = None'') (\<lambda>_. Set_Monad xs6 O DList_set dxs4)
| Some _ \<Rightarrow> fold (\<lambda>(x, y). DList_Set.fold (\<lambda>(y', z) A. if eq y y' then insert (x, z) A else A) dxs4) xs6 {})"
(is ?monad_dlist)
proof -
show ?rbt_rbt ?rbt_monad ?monad_rbt
by(auto simp add: comparator.eq[OF ID_ccompare'] RBT_set_def ccompare_prod_def member_conv_keys ID_Some RBT_Set2.fold_conv_fold_keys' RBT_Set2.keys.rep_eq If_not set_relcomp_set split: option.split del: equalityI)
show ?rbt_dlist ?dlist_rbt ?dlist_dlist ?monad_monad ?dlist_monad ?monad_dlist
by(auto simp add: RBT_set_def DList_set_def member_conv_keys ID_Some ccompare_prod_def ceq_prod_def Collect_member RBT_Set2.fold_conv_fold_keys' RBT_Set2.keys.rep_eq DList_Set.fold.rep_eq set_relcomp_set dest: equal.equal_eq[OF ID_ceq] split: option.split del: equalityI)
qed
lemma irrefl_code [code]:
fixes r :: "('a :: {ceq, ccompare} \<times> 'a) set" shows
"irrefl r \<longleftrightarrow>
(case ID CEQ('a) of Some eq \<Rightarrow> (\<forall>(x, y) \<in> r. \<not> eq x y) | None \<Rightarrow>
case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''irrefl: ceq = None & ccompare = None'') (\<lambda>_. irrefl r)
| Some c \<Rightarrow> (\<forall>(x, y) \<in> r. c x y \<noteq> Eq))"
apply(auto simp add: irrefl_distinct comparator.eq[OF ID_ccompare'] split: option.split dest!: ID_ceq[THEN equal.equal_eq])
done
lemma wf_code [code]:
fixes rbt :: "('a :: ccompare \<times> 'a) set_rbt"
and dxs :: "('b :: ceq \<times> 'b) set_dlist" shows
"wf (Set_Monad xs) = acyclic (Set_Monad xs)"
"wf (RBT_set rbt) =
(case ID CCOMPARE('a) of None \<Rightarrow> Code.abort (STR ''wf RBT_set: ccompare = None'') (\<lambda>_. wf (RBT_set rbt))
| Some _ \<Rightarrow> acyclic (RBT_set rbt))"
"wf (DList_set dxs) =
(case ID CEQ('b) of None \<Rightarrow> Code.abort (STR ''wf DList_set: ceq = None'') (\<lambda>_. wf (DList_set dxs))
| Some _ \<Rightarrow> acyclic (DList_set dxs))"
by(auto simp add: wf_iff_acyclic_if_finite split: option.split del: iffI)(simp_all add: wf_iff_acyclic_if_finite finite_code ccompare_prod_def ceq_prod_def ID_Some)
lemma bacc_code [code]:
"bacc R 0 = - snd ` R"
"bacc R (Suc n) = (let rec = bacc R n in rec \<union> - snd ` (Set.filter (\<lambda>(y, x). y \<notin> rec) R))"
by(auto intro: rev_image_eqI simp add: Let_def)
(* TODO: acc could also be computed for infinite universes if r is finite *)
lemma acc_code [code]:
fixes A :: "('a :: {finite, card_UNIV} \<times> 'a) set" shows
"Wellfounded.acc A = bacc A (of_phantom (card_UNIV :: 'a card_UNIV))"
by(simp add: card_UNIV acc_bacc_eq)
lemma sorted_list_of_set_code [code]:
fixes dxs :: "'a :: {linorder, ceq} set_dlist"
and rbt :: "'b :: {linorder, ccompare} set_rbt"
shows
"sorted_list_of_set (Set_Monad xs) = sort (remdups xs)"
"sorted_list_of_set (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''sorted_list_of_set DList_set: ceq = None'') (\<lambda>_. sorted_list_of_set (DList_set dxs))
| Some _ \<Rightarrow> sort (list_of_dlist dxs))"
"sorted_list_of_set (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''sorted_list_of_set RBT_set: ccompare = None'') (\<lambda>_. sorted_list_of_set (RBT_set rbt))
| Some _ \<Rightarrow> sort (RBT_Set2.keys rbt))"
\<comment> \<open>We must sort the keys because @{term ccompare}'s ordering need not coincide with @{term linorder}'s.\<close>
by(auto simp add: DList_set_def RBT_set_def sorted_list_of_set_sort_remdups Collect_member distinct_remdups_id distinct_list_of_dlist member_conv_keys split: option.split)
lemma map_project_set: "List.map_project f (set xs) = set (List.map_filter f xs)"
by(auto simp add: List.map_project_def List.map_filter_def intro: rev_image_eqI)
lemma map_project_simps:
shows map_project_empty: "List.map_project f {} = {}"
and map_project_insert:
"List.map_project f (insert x A) =
(case f x of None \<Rightarrow> List.map_project f A
| Some y \<Rightarrow> insert y (List.map_project f A))"
by(auto simp add: List.map_project_def split: option.split)
lemma map_project_conv_fold:
"List.map_project f (set xs) =
fold (\<lambda>x A. case f x of None \<Rightarrow> A | Some y \<Rightarrow> insert y A) xs {}"
by(induct xs rule: rev_induct)(simp_all add: map_project_simps cong: option.case_cong)
lemma map_project_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"List.map_project f (Set_Monad xs) = Set_Monad (List.map_filter f xs)"
"List.map_project g (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''map_project DList_set: ceq = None'') (\<lambda>_. List.map_project g (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.fold (\<lambda>x A. case g x of None \<Rightarrow> A | Some y \<Rightarrow> insert y A) dxs {})"
(is ?dlist)
"List.map_project h (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''map_project RBT_set: ccompare = None'') (\<lambda>_. List.map_project h (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.fold (\<lambda>x A. case h x of None \<Rightarrow> A | Some y \<Rightarrow> insert y A) rbt {})"
(is ?rbt)
proof -
show ?dlist ?rbt
by(auto split: option.split simp add: RBT_set_def DList_set_def DList_Set.fold.rep_eq Collect_member map_project_conv_fold RBT_Set2.fold_conv_fold_keys member_conv_keys del: equalityI)
qed(auto simp add: List.map_project_def List.map_filter_def intro: rev_image_eqI)
lemma Bleast_code [code]:
"Bleast A P =
(if finite A then case filter P (sorted_list_of_set A) of [] \<Rightarrow> abort_Bleast A P | x # xs \<Rightarrow> x
else abort_Bleast A P)"
proof(cases "finite A")
case True
hence *: "A = set (sorted_list_of_set A)" by(simp add: sorted_list_of_set)
show ?thesis using True
by(subst (1 3) *)(unfold Bleast_code, simp add: sorted_sort_id)
qed(simp add: abort_Bleast_def Bleast_def)
lemma can_select_code [code]:
fixes xs :: "'a :: ceq list"
and dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"can_select P (Set_Monad xs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''can_select Set_Monad: ceq = None'') (\<lambda>_. can_select P (Set_Monad xs))
| Some eq \<Rightarrow> case filter P xs of Nil \<Rightarrow> False | x # xs \<Rightarrow> list_all (eq x) xs)"
(is ?Set_Monad)
"can_select Q (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''can_select DList_set: ceq = None'') (\<lambda>_. can_select Q (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.length (DList_Set.filter Q dxs) = 1)"
(is ?dlist)
"can_select R (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''can_select RBT_set: ccompare = None'') (\<lambda>_. can_select R (RBT_set rbt))
| Some _ \<Rightarrow> singleton_list_fusion (filter_generator R rbt_keys_generator) (RBT_Set2.init rbt))"
(is ?rbt)
proof -
show ?Set_Monad
apply(auto split: option.split list.split dest!: ID_ceq[THEN equal.equal_eq] dest: filter_eq_ConsD simp add: can_select_def filter_empty_conv list_all_iff)
apply(drule filter_eq_ConsD, fastforce)
apply(drule filter_eq_ConsD, clarsimp, blast)
done
show ?dlist
by(clarsimp simp add: can_select_def card_eq_length[symmetric] Set_member_code card_eq_Suc_0_ex1 simp del: card_eq_length split: option.split)
note [simp del] = distinct_keys
show ?rbt
using distinct_keys[of rbt]
apply(auto simp add: can_select_def singleton_list_fusion_def unfoldr_filter_generator unfoldr_rbt_keys_generator Set_member_code member_conv_keys filter_empty_conv empty_filter_conv split: option.split list.split dest: filter_eq_ConsD)
apply(drule filter_eq_ConsD, fastforce)
apply(drule filter_eq_ConsD, fastforce simp add: empty_filter_conv)
apply(drule filter_eq_ConsD)
apply clarsimp
apply(drule Cons_eq_filterD)
apply clarify
apply(simp (no_asm_use))
apply blast
done
qed
lemma pred_of_set_code [code]:
fixes dxs :: "'a :: ceq set_dlist"
and rbt :: "'b :: ccompare set_rbt" shows
"pred_of_set (Set_Monad xs) = fold (sup \<circ> Predicate.single) xs bot"
"pred_of_set (DList_set dxs) =
(case ID CEQ('a) of None \<Rightarrow> Code.abort (STR ''pred_of_set DList_set: ceq = None'') (\<lambda>_. pred_of_set (DList_set dxs))
| Some _ \<Rightarrow> DList_Set.fold (sup \<circ> Predicate.single) dxs bot)"
"pred_of_set (RBT_set rbt) =
(case ID CCOMPARE('b) of None \<Rightarrow> Code.abort (STR ''pred_of_set RBT_set: ccompare = None'') (\<lambda>_. pred_of_set (RBT_set rbt))
| Some _ \<Rightarrow> RBT_Set2.fold (sup \<circ> Predicate.single) rbt bot)"
by(auto simp add: pred_of_set_set_fold_sup fold_map DList_set_def RBT_set_def Collect_member member_conv_keys DList_Set.fold.rep_eq fold_conv_fold_keys split: option.split)
text \<open>
@{typ "'a Predicate.pred"} is implemented as a monad,
so we keep the monad when converting to @{typ "'a set"}.
For this case, @{term insert_monad} and @{term union_monad}
avoid the unnecessary dictionary construction.
\<close>
definition insert_monad :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set"
where [simp]: "insert_monad = insert"
definition union_monad :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set"
where [simp]: "union_monad = (\<union>)"
lemma insert_monad_code [code]:
"insert_monad x (Set_Monad xs) = Set_Monad (x # xs)"
by simp
lemma union_monad_code [code]:
"union_monad (Set_Monad xs) (Set_Monad ys) = Set_Monad (xs @ ys)"
by(simp)
lemma set_of_pred_code [code]:
"set_of_pred (Predicate.Seq f) =
(case f () of seq.Empty \<Rightarrow> Set_Monad []
| seq.Insert x P \<Rightarrow> insert_monad x (set_of_pred P)
| seq.Join P xq \<Rightarrow> union_monad (set_of_pred P) (set_of_seq xq))"
by(simp add: of_pred_code cong: seq.case_cong)
lemma set_of_seq_code [code]:
"set_of_seq seq.Empty = Set_Monad []"
"set_of_seq (seq.Insert x P) = insert_monad x (set_of_pred P)"
"set_of_seq (seq.Join P xq) = union_monad (set_of_pred P) (set_of_seq xq)"
by(simp_all add: of_seq_code)
hide_const (open) insert_monad union_monad
subsection \<open>Type class instantiations\<close>
datatype set_impl = Set_IMPL
declare
set_impl.eq.simps [code del]
set_impl.size [code del]
set_impl.rec [code del]
set_impl.case [code del]
lemma [code]:
fixes x :: set_impl
shows "size x = 0"
and "size_set_impl x = 0"
by(case_tac [!] x) simp_all
definition set_Choose :: set_impl where [simp]: "set_Choose = Set_IMPL"
definition set_Collect :: set_impl where [simp]: "set_Collect = Set_IMPL"
definition set_DList :: set_impl where [simp]: "set_DList = Set_IMPL"
definition set_RBT :: set_impl where [simp]: "set_RBT = Set_IMPL"
definition set_Monad :: set_impl where [simp]: "set_Monad = Set_IMPL"
code_datatype set_Choose set_Collect set_DList set_RBT set_Monad
definition set_empty_choose :: "'a set" where [simp]: "set_empty_choose = {}"
lemma set_empty_choose_code [code]:
"(set_empty_choose :: 'a :: {ceq, ccompare} set) =
(case CCOMPARE('a) of Some _ \<Rightarrow> RBT_set RBT_Set2.empty
| None \<Rightarrow> case CEQ('a) of None \<Rightarrow> Set_Monad [] | Some _ \<Rightarrow> DList_set (DList_Set.empty))"
by(simp split: option.split)
definition set_impl_choose2 :: "set_impl \<Rightarrow> set_impl \<Rightarrow> set_impl"
where [simp]: "set_impl_choose2 = (\<lambda>_ _. Set_IMPL)"
lemma set_impl_choose2_code [code]:
"set_impl_choose2 x y = set_Choose"
"set_impl_choose2 set_Collect set_Collect = set_Collect"
"set_impl_choose2 set_DList set_DList = set_DList"
"set_impl_choose2 set_RBT set_RBT = set_RBT"
"set_impl_choose2 set_Monad set_Monad = set_Monad"
by(simp_all)
definition set_empty :: "set_impl \<Rightarrow> 'a set"
where [simp]: "set_empty = (\<lambda>_. {})"
lemma set_empty_code [code]:
"set_empty set_Collect = Collect_set (\<lambda>_. False)"
"set_empty set_DList = DList_set DList_Set.empty"
"set_empty set_RBT = RBT_set RBT_Set2.empty"
"set_empty set_Monad = Set_Monad []"
"set_empty set_Choose = set_empty_choose"
by(simp_all)
class set_impl =
fixes set_impl :: "('a, set_impl) phantom"
syntax (input)
"_SET_IMPL" :: "type => logic" ("(1SET'_IMPL/(1'(_')))")
parse_translation \<open>
let
fun set_impl_tr [ty] =
(Syntax.const @{syntax_const "_constrain"} $ Syntax.const @{const_syntax "set_impl"} $
(Syntax.const @{type_syntax phantom} $ ty $ Syntax.const @{type_syntax set_impl}))
| set_impl_tr ts = raise TERM ("set_impl_tr", ts);
in [(@{syntax_const "_SET_IMPL"}, K set_impl_tr)] end
\<close>
declare [[code drop: "{}"]]
lemma empty_code [code, code_unfold]:
"({} :: 'a :: set_impl set) = set_empty (of_phantom SET_IMPL('a))"
by simp
subsection \<open>Generator for the @{class set_impl}-class\<close>
text \<open>
This generator registers itself at the derive-manager for the classes @{class set_impl}.
Here, one can choose
the desired implementation via the parameter.
\begin{itemize}
\item \texttt{instantiation type :: (type,\ldots,type) (rbt,dlist,collect,monad,choose, or arbitrary constant name) set-impl}
\end{itemize}
\<close>
text \<open>
This generator can be used for arbitrary types, not just datatypes.
\<close>
ML_file \<open>set_impl_generator.ML\<close>
derive (dlist) set_impl unit bool
derive (rbt) set_impl nat
derive (set_RBT) set_impl int (* shows usage of constant names *)
derive (dlist) set_impl Enum.finite_1 Enum.finite_2 Enum.finite_3
derive (rbt) set_impl integer natural
derive (rbt) set_impl char
instantiation sum :: (set_impl, set_impl) set_impl begin
definition "SET_IMPL('a + 'b) = Phantom('a + 'b)
(set_impl_choose2 (of_phantom SET_IMPL('a)) (of_phantom SET_IMPL('b)))"
instance ..
end
instantiation prod :: (set_impl, set_impl) set_impl begin
definition "SET_IMPL('a * 'b) = Phantom('a * 'b)
(set_impl_choose2 (of_phantom SET_IMPL('a)) (of_phantom SET_IMPL('b)))"
instance ..
end
derive (choose) set_impl list
derive (rbt) set_impl String.literal
instantiation option :: (set_impl) set_impl begin
definition "SET_IMPL('a option) = Phantom('a option) (of_phantom SET_IMPL('a))"
instance ..
end
derive (monad) set_impl "fun"
derive (choose) set_impl set
instantiation phantom :: (type, set_impl) set_impl begin
definition "SET_IMPL(('a, 'b) phantom) = Phantom (('a, 'b) phantom) (of_phantom SET_IMPL('b))"
instance ..
end
text \<open>
We enable automatic implementation selection for sets constructed by @{const set},
although they could be directly converted using @{const Set_Monad} in constant time.
However, then it is more likely that the parameters of binary operators have
different implementations, which can lead to less efficient execution.
However, we test whether automatic selection picks @{const Set_Monad} anyway and
take a short-cut.
\<close>
definition set_aux :: "set_impl \<Rightarrow> 'a list \<Rightarrow> 'a set"
where [simp, code del]: "set_aux _ = set"
lemma set_aux_code [code]:
defines "conv \<equiv> foldl (\<lambda>s (x :: 'a). insert x s)"
shows
"set_aux impl = conv (set_empty impl)" (is "?thesis1")
"set_aux set_Choose =
(case CCOMPARE('a :: {ccompare, ceq}) of Some _ \<Rightarrow> conv (RBT_set RBT_Set2.empty)
| None \<Rightarrow> case CEQ('a) of None \<Rightarrow> Set_Monad
| Some _ \<Rightarrow> conv (DList_set DList_Set.empty))" (is "?thesis2")
"set_aux set_Monad = Set_Monad"
proof -
have "conv {} = set"
by(rule ext)(induct_tac x rule: rev_induct, simp_all add: conv_def)
thus ?thesis1 ?thesis2
by(simp_all split: option.split)
qed simp
lemma set_code [code]:
fixes xs :: "'a :: set_impl list"
shows "set xs = set_aux (of_phantom (ID SET_IMPL('a))) xs"
by(simp)
subsection \<open>Pretty printing for sets\<close>
text \<open>
@{term code_post} marks contexts (as hypothesis) in which we use code\_post as a
decision procedure rather than a pretty-printing engine.
The intended use is to enable more rules when proving assumptions of rewrite rules.
\<close>
definition code_post :: bool where "code_post = True"
lemma conj_code_post [code_post]:
assumes code_post
shows "True & x \<longleftrightarrow> x" "False & x \<longleftrightarrow> False"
by simp_all
text \<open>
A flag to switch post-processing of sets on and off.
Use \verb$declare pretty_sets[code_post del]$ to disable pretty printing of sets in value.
\<close>
definition code_post_set :: bool
where pretty_sets [code_post, simp]: "code_post_set = True"
definition collapse_RBT_set :: "'a set_rbt \<Rightarrow> 'a :: ccompare set \<Rightarrow> 'a set"
where "collapse_RBT_set r M = set (RBT_Set2.keys r) \<union> M"
lemma RBT_set_collapse_RBT_set [code_post]:
fixes r :: "'a :: ccompare set_rbt"
assumes "code_post \<Longrightarrow> is_ccompare TYPE('a)" and code_post_set
shows "RBT_set r = collapse_RBT_set r {}"
using assms
by(clarsimp simp add: code_post_def is_ccompare_def RBT_set_def member_conv_keys collapse_RBT_set_def)
lemma collapse_RBT_set_Branch [code_post]:
"collapse_RBT_set (Mapping_RBT (Branch c l x v r)) M =
collapse_RBT_set (Mapping_RBT l) (insert x (collapse_RBT_set (Mapping_RBT r) M))"
unfolding collapse_RBT_set_def
by(auto simp add: is_ccompare_def set_keys_Mapping_RBT)
lemma collapse_RBT_set_Empty [code_post]:
"collapse_RBT_set (Mapping_RBT rbt.Empty) M = M"
by(auto simp add: collapse_RBT_set_def set_keys_Mapping_RBT)
definition collapse_DList_set :: "'a :: ceq set_dlist \<Rightarrow> 'a set"
where "collapse_DList_set dxs = set (DList_Set.list_of_dlist dxs)"
lemma DList_set_collapse_DList_set [code_post]:
fixes dxs :: "'a :: ceq set_dlist"
assumes "code_post \<Longrightarrow> is_ceq TYPE('a)" and code_post_set
shows "DList_set dxs = collapse_DList_set dxs"
using assms
by(clarsimp simp add: code_post_def DList_set_def is_ceq_def collapse_DList_set_def Collect_member)
lemma collapse_DList_set_empty [code_post]: "collapse_DList_set (Abs_dlist []) = {}"
by(simp add: collapse_DList_set_def Abs_dlist_inverse)
lemma collapse_DList_set_Cons [code_post]:
"collapse_DList_set (Abs_dlist (x # xs)) = insert x (collapse_DList_set (Abs_dlist xs))"
by(simp add: collapse_DList_set_def set_list_of_dlist_Abs_dlist)
lemma Set_Monad_code_post [code_post]:
assumes code_post_set
shows "Set_Monad [] = {}"
and "Set_Monad (x#xs) = insert x (Set_Monad xs)"
by simp_all
end
diff --git a/thys/Containers/Set_Linorder.thy b/thys/Containers/Set_Linorder.thy
--- a/thys/Containers/Set_Linorder.thy
+++ b/thys/Containers/Set_Linorder.thy
@@ -1,2481 +1,2481 @@
(* Title: Containers/Set_Linorder.thy
Author: Andreas Lochbihler, KIT *)
theory Set_Linorder
imports
Containers_Auxiliary
Lexicographic_Order
Extend_Partial_Order
"HOL-Library.Cardinality"
begin
section \<open>An executable linear order on sets\<close>
subsection \<open>Definition of the linear order\<close>
subsubsection \<open>Extending finite and cofinite sets\<close>
text \<open>
Partition sets into finite and cofinite sets and distribute the rest arbitrarily such that
complement switches between the two.
\<close>
consts infinite_complement_partition :: "'a set set"
specification (infinite_complement_partition)
finite_complement_partition: "finite (A :: 'a set) \<Longrightarrow> A \<in> infinite_complement_partition"
complement_partition: "\<not> finite (UNIV :: 'a set)
\<Longrightarrow> (A :: 'a set) \<in> infinite_complement_partition \<longleftrightarrow> - A \<notin> infinite_complement_partition"
proof(cases "finite (UNIV :: 'a set)")
case False
define Field_r where "Field_r = {\<P> :: 'a set set. (\<forall>C \<in> \<P>. - C \<notin> \<P>) \<and> (\<forall>A. finite A \<longrightarrow> A \<in> \<P>)}"
define r where "r = {(\<P>1, \<P>2). \<P>1 \<subseteq> \<P>2 \<and> \<P>1 \<in> Field_r \<and> \<P>2 \<in> Field_r}"
have in_Field_r [simp]: "\<And>\<P>. \<P> \<in> Field_r \<longleftrightarrow> (\<forall>C \<in> \<P>. - C \<notin> \<P>) \<and> (\<forall>A. finite A \<longrightarrow> A \<in> \<P>)"
unfolding Field_r_def by simp
have in_r [simp]: "\<And>\<P>1 \<P>2. (\<P>1, \<P>2) \<in> r \<longleftrightarrow> \<P>1 \<subseteq> \<P>2 \<and> \<P>1 \<in> Field_r \<and> \<P>2 \<in> Field_r"
unfolding r_def by simp
have Field_r [simp]: "Field r = Field_r" by(auto simp add: Field_def Field_r_def)
have "Partial_order r"
by(auto simp add: Field_def r_def partial_order_on_def preorder_on_def intro!: refl_onI transI antisymI)
moreover have "\<exists>\<B> \<in> Field r. \<forall>\<A> \<in> \<CC>. (\<A>, \<B>) \<in> r" if \<CC>: "\<CC> \<in> Chains r" for \<CC>
proof -
let ?\<B> = "\<Union>\<CC> \<union> {A. finite A}"
have *: "?\<B> \<in> Field r" using False \<CC>
by clarsimp(safe, drule (2) ChainsD, auto 4 4 dest: Chains_Field)
hence "\<And>\<A>. \<A> \<in> \<CC> \<Longrightarrow> (\<A>, ?\<B>) \<in> r"
using \<CC> by(auto simp del: in_Field_r dest: Chains_Field)
with * show "\<exists>\<B> \<in> Field r. \<forall>\<A> \<in> \<CC>. (\<A>, \<B>) \<in> r" by blast
qed
ultimately have "\<exists>\<P> \<in> Field r. \<forall>\<A> \<in> Field r. (\<P>, \<A>) \<in> r \<longrightarrow> \<A> = \<P>"
by(rule Zorns_po_lemma)
then obtain \<P> where \<P>: "\<P> \<in> Field r"
and max: "\<And>\<A>. \<lbrakk> \<A> \<in> Field r; (\<P>, \<A>) \<in> r \<rbrakk> \<Longrightarrow> \<A> = \<P>" by blast
have "\<forall>A. finite A \<longrightarrow> A \<in> \<P>" using \<P> by simp
moreover {
fix C
have "C \<in> \<P> \<or> - C \<in> \<P>"
proof(rule ccontr)
assume "\<not> ?thesis"
hence C: "C \<notin> \<P>" "- C \<notin> \<P>" by simp_all
let ?\<A> = "insert C \<P>"
have *: "?\<A> \<in> Field r" using C \<P> by auto
hence "(\<P>, ?\<A>) \<in> r" using \<P> by auto
with * have "?\<A> = \<P>" by(rule max)
thus False using C by auto
qed
hence "C \<in> \<P> \<longleftrightarrow> - C \<notin> \<P>" using \<P> by auto }
ultimately have "\<exists>\<P> :: 'a set set. (\<forall>A. finite A \<longrightarrow> A \<in> \<P>) \<and> (\<forall>C. C \<in> \<P> \<longleftrightarrow> - C \<notin> \<P>)"
by blast
thus ?thesis by metis
qed auto
lemma not_in_complement_partition:
"\<not> finite (UNIV :: 'a set)
\<Longrightarrow> (A :: 'a set) \<notin> infinite_complement_partition \<longleftrightarrow> - A \<in> infinite_complement_partition"
by(metis complement_partition)
lemma not_in_complement_partition_False:
"\<lbrakk> (A :: 'a set) \<in> infinite_complement_partition; \<not> finite (UNIV :: 'a set) \<rbrakk>
\<Longrightarrow> - A \<in> infinite_complement_partition = False"
by(simp add: not_in_complement_partition)
lemma infinite_complement_partition_finite [simp]:
"finite (UNIV :: 'a set) \<Longrightarrow> infinite_complement_partition = (UNIV :: 'a set set)"
by(auto intro: finite_subset finite_complement_partition)
lemma Compl_eq_empty_iff: "- A = {} \<longleftrightarrow> A = UNIV"
by auto
subsubsection \<open>A lexicographic-style order on finite subsets\<close>
context ord begin
definition set_less_aux :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "\<sqsubset>''" 50)
where "A \<sqsubset>' B \<longleftrightarrow> finite A \<and> finite B \<and> (\<exists>y \<in> B - A. \<forall>z \<in> (A - B) \<union> (B - A). y \<le> z \<and> (z \<le> y \<longrightarrow> y = z))"
definition set_less_eq_aux :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "\<sqsubseteq>''" 50)
where "A \<sqsubseteq>' B \<longleftrightarrow> A \<in> infinite_complement_partition \<and> A = B \<or> A \<sqsubset>' B"
lemma set_less_aux_irrefl [iff]: "\<not> A \<sqsubset>' A"
by(auto simp add: set_less_aux_def)
lemma set_less_eq_aux_refl [iff]: "A \<sqsubseteq>' A \<longleftrightarrow> A \<in> infinite_complement_partition"
by(simp add: set_less_eq_aux_def)
lemma set_less_aux_empty [simp]: "\<not> A \<sqsubset>' {}"
by(auto simp add: set_less_aux_def intro: finite_subset finite_complement_partition)
lemma set_less_eq_aux_empty [simp]: "A \<sqsubseteq>' {} \<longleftrightarrow> A = {}"
by(auto simp add: set_less_eq_aux_def finite_complement_partition)
lemma set_less_aux_antisym: "\<lbrakk> A \<sqsubset>' B; B \<sqsubset>' A \<rbrakk> \<Longrightarrow> False"
by(auto simp add: set_less_aux_def split: if_split_asm)
lemma set_less_aux_conv_set_less_eq_aux:
"A \<sqsubset>' B \<longleftrightarrow> A \<sqsubseteq>' B \<and> \<not> B \<sqsubseteq>' A"
by(auto simp add: set_less_eq_aux_def dest: set_less_aux_antisym)
lemma set_less_eq_aux_antisym: "\<lbrakk> A \<sqsubseteq>' B; B \<sqsubseteq>' A \<rbrakk> \<Longrightarrow> A = B"
by(auto simp add: set_less_eq_aux_def dest: set_less_aux_antisym)
lemma set_less_aux_finiteD: "A \<sqsubset>' B \<Longrightarrow> finite A \<and> B \<in> infinite_complement_partition"
by(auto simp add: set_less_aux_def finite_complement_partition)
lemma set_less_eq_aux_infinite_complement_partitionD:
"A \<sqsubseteq>' B \<Longrightarrow> A \<in> infinite_complement_partition \<and> B \<in> infinite_complement_partition"
by(auto simp add: set_less_eq_aux_def dest: set_less_aux_finiteD intro: finite_complement_partition)
lemma Compl_set_less_aux_Compl:
"finite (UNIV :: 'a set) \<Longrightarrow> - A \<sqsubset>' - B \<longleftrightarrow> B \<sqsubset>' A"
by(auto simp add: set_less_aux_def finite_subset[OF subset_UNIV])
lemma Compl_set_less_eq_aux_Compl:
"finite (UNIV :: 'a set) \<Longrightarrow> - A \<sqsubseteq>' - B \<longleftrightarrow> B \<sqsubseteq>' A"
by(auto simp add: set_less_eq_aux_def Compl_set_less_aux_Compl)
lemma set_less_aux_insert_same:
"x \<in> A \<longleftrightarrow> x \<in> B \<Longrightarrow> insert x A \<sqsubset>' insert x B \<longleftrightarrow> A \<sqsubset>' B"
by(auto simp add: set_less_aux_def)
lemma set_less_eq_aux_insert_same:
"\<lbrakk> A \<in> infinite_complement_partition; insert x B \<in> infinite_complement_partition;
x \<in> A \<longleftrightarrow> x \<in> B \<rbrakk>
\<Longrightarrow> insert x A \<sqsubseteq>' insert x B \<longleftrightarrow> A \<sqsubseteq>' B"
by(auto simp add: set_less_eq_aux_def set_less_aux_insert_same)
end
context order begin
lemma set_less_aux_singleton_iff: "A \<sqsubset>' {x} \<longleftrightarrow> finite A \<and> (\<forall>a\<in>A. x < a)"
by(auto simp add: set_less_aux_def less_le Bex_def)
end
context linorder begin
lemma wlog_le [case_names sym le]:
assumes "\<And>a b. P a b \<Longrightarrow> P b a"
and "\<And>a b. a \<le> b \<Longrightarrow> P a b"
shows "P b a"
by (metis assms linear)
lemma empty_set_less_aux [simp]: "{} \<sqsubset>' A \<longleftrightarrow> A \<noteq> {} \<and> finite A"
by(auto 4 3 simp add: set_less_aux_def intro!: Min_eqI intro: bexI[where x="Min A"] order_trans[where y="Min A"] Min_in)
lemma empty_set_less_eq_aux [simp]: "{} \<sqsubseteq>' A \<longleftrightarrow> finite A"
by(auto simp add: set_less_eq_aux_def finite_complement_partition)
lemma set_less_aux_trans:
assumes AB: "A \<sqsubset>' B" and BC: "B \<sqsubset>' C"
shows "A \<sqsubset>' C"
proof -
from AB BC have A: "finite A" and B: "finite B" and C: "finite C"
by(auto simp add: set_less_aux_def)
from AB A B obtain ab where ab: "ab \<in> B - A"
and minAB: "\<And>x. \<lbrakk> x \<in> A; x \<notin> B \<rbrakk> \<Longrightarrow> ab \<le> x \<and> (x \<le> ab \<longrightarrow> ab = x)"
and minBA: "\<And>x. \<lbrakk> x \<in> B; x \<notin> A \<rbrakk> \<Longrightarrow> ab \<le> x \<and> (x \<le> ab \<longrightarrow> ab = x)"
unfolding set_less_aux_def by auto
from BC B C obtain bc where bc: "bc \<in> C - B"
and minBC: "\<And>x. \<lbrakk> x \<in> B; x \<notin> C \<rbrakk> \<Longrightarrow> bc \<le> x \<and> (x \<le> bc \<longrightarrow> bc = x)"
and minCB: "\<And>x. \<lbrakk> x \<in> C; x \<notin> B \<rbrakk> \<Longrightarrow> bc \<le> x \<and> (x \<le> bc \<longrightarrow> bc = x)"
unfolding set_less_aux_def by auto
show ?thesis
proof(cases "ab \<le> bc")
case True
hence "ab \<in> C - A" "ab \<notin> A - C"
using ab bc by(auto dest: minBC antisym)
moreover {
fix x
assume x: "x \<in> (C - A) \<union> (A - C)"
hence "ab \<le> x"
by(cases "x \<in> B")(auto dest: minAB minBA minBC minCB intro: order_trans[OF True])
moreover hence "ab \<noteq> x \<longrightarrow> \<not> x \<le> ab" using ab bc x
by(cases "x \<in> B")(auto dest: antisym)
moreover note calculation }
ultimately show ?thesis using A C unfolding set_less_aux_def by auto
next
case False
with linear[of ab bc] have "bc \<le> ab" by simp
with ab bc have "bc \<in> C - A" "bc \<notin> A - C" by(auto dest: minAB antisym)
moreover {
fix x
assume x: "x \<in> (C - A) \<union> (A - C)"
hence "bc \<le> x"
by(cases "x \<in> B")(auto dest: minAB minBA minBC minCB intro: order_trans[OF \<open>bc \<le> ab\<close>])
moreover hence "bc \<noteq> x \<longrightarrow> \<not> x \<le> bc" using ab bc x
by(cases "x \<in> B")(auto dest: antisym)
moreover note calculation }
ultimately show ?thesis using A C unfolding set_less_aux_def by auto
qed
qed
lemma set_less_eq_aux_trans [trans]:
"\<lbrakk> A \<sqsubseteq>' B; B \<sqsubseteq>' C \<rbrakk> \<Longrightarrow> A \<sqsubseteq>' C"
by(auto simp add: set_less_eq_aux_def dest: set_less_aux_trans)
lemma set_less_trans_set_less_eq [trans]:
"\<lbrakk> A \<sqsubset>' B; B \<sqsubseteq>' C \<rbrakk> \<Longrightarrow> A \<sqsubset>' C"
by(auto simp add: set_less_eq_aux_def dest: set_less_aux_trans)
lemma set_less_eq_aux_porder: "partial_order_on infinite_complement_partition {(A, B). A \<sqsubseteq>' B}"
by(auto simp add: preorder_on_def partial_order_on_def intro!: refl_onI transI antisymI dest: set_less_eq_aux_infinite_complement_partitionD intro: set_less_eq_aux_antisym set_less_eq_aux_trans del: equalityI)
lemma psubset_finite_imp_set_less_aux:
assumes AsB: "A \<subset> B" and B: "finite B"
shows "A \<sqsubset>' B"
proof -
from AsB B have A: "finite A" by(auto intro: finite_subset)
moreover from AsB B have "Min (B - A) \<in> B - A" by - (rule Min_in, auto)
ultimately show ?thesis using B AsB
by(auto simp add: set_less_aux_def intro!: rev_bexI[where x="Min (B - A)"] Min_eqI dest: Min_ge_iff[THEN iffD1, rotated 2])
qed
lemma subset_finite_imp_set_less_eq_aux:
"\<lbrakk> A \<subseteq> B; finite B \<rbrakk> \<Longrightarrow> A \<sqsubseteq>' B"
by(cases "A = B")(auto simp add: set_less_eq_aux_def finite_complement_partition intro: psubset_finite_imp_set_less_aux)
lemma empty_set_less_aux_finite_iff:
"finite A \<Longrightarrow> {} \<sqsubset>' A \<longleftrightarrow> A \<noteq> {}"
by(auto intro: psubset_finite_imp_set_less_aux)
lemma set_less_aux_finite_total:
assumes A: "finite A" and B: "finite B"
shows "A \<sqsubset>' B \<or> A = B \<or> B \<sqsubset>' A"
proof(cases "A \<subseteq> B \<or> B \<subseteq> A")
case True thus ?thesis
using A B psubset_finite_imp_set_less_aux[of A B] psubset_finite_imp_set_less_aux[of B A]
by auto
next
case False
hence A': "\<not> A \<subseteq> B" and B': "\<not> B \<subseteq> A" and AnB: "A \<noteq> B" by auto
thus ?thesis using A B
proof(induct "Min (B - A)" "Min (A - B)" arbitrary: A B rule: wlog_le)
case (sym m n)
from sym.hyps[OF refl refl] sym.prems show ?case by blast
next
case (le A B)
note A = \<open>finite A\<close> and B = \<open>finite B\<close>
and A' = \<open>\<not> A \<subseteq> B\<close> and B' = \<open>\<not> B \<subseteq> A\<close>
{ fix z
assume z: "z \<in> (A - B) \<union> (B - A)"
hence "Min (B - A) \<le> z \<and> (z \<le> Min (B - A) \<longrightarrow> Min (B - A) = z)"
proof
assume "z \<in> B - A"
hence "Min (B - A) \<le> z" by(simp add: B)
thus ?thesis by auto
next
assume "z \<in> A - B"
hence "Min (A - B) \<le> z" by(simp add: A)
with le.hyps show ?thesis by(auto)
qed }
moreover have "Min (B - A) \<in> B - A" by(rule Min_in)(simp_all add: B B')
ultimately have "A \<sqsubset>' B" using A B by(auto simp add: set_less_aux_def)
thus ?case ..
qed
qed
lemma set_less_eq_aux_finite_total:
"\<lbrakk> finite A; finite B \<rbrakk> \<Longrightarrow> A \<sqsubseteq>' B \<or> A = B \<or> B \<sqsubseteq>' A"
by(drule (1) set_less_aux_finite_total)(auto simp add: set_less_eq_aux_def)
lemma set_less_eq_aux_finite_total2:
"\<lbrakk> finite A; finite B \<rbrakk> \<Longrightarrow> A \<sqsubseteq>' B \<or> B \<sqsubseteq>' A"
by(drule (1) set_less_eq_aux_finite_total)(auto simp add: finite_complement_partition)
lemma set_less_aux_rec:
assumes A: "finite A" and B: "finite B"
and A': "A \<noteq> {}" and B': "B \<noteq> {}"
shows "A \<sqsubset>' B \<longleftrightarrow> Min B < Min A \<or> Min A = Min B \<and> A - {Min A} \<sqsubset>' B - {Min A}"
proof(cases "Min A = Min B")
case True
from A A' B B' have "insert (Min A) A = A" "insert (Min B) B = B"
by(auto simp add: ex_in_conv[symmetric] exI)
with True show ?thesis
by(subst (2) set_less_aux_insert_same[symmetric, where x="Min A"]) simp_all
next
case False
have "A \<sqsubset>' B \<longleftrightarrow> Min B < Min A"
proof
assume AB: "A \<sqsubset>' B"
with B A obtain ab where ab: "ab \<in> B - A"
and AB: "\<And>x. \<lbrakk> x \<in> A; x \<notin> B \<rbrakk> \<Longrightarrow> ab \<le> x"
by(auto simp add: set_less_aux_def)
{ fix a assume "a \<in> A"
hence "Min B \<le> a" using A A' B B' ab
by(cases "a \<in> B")(auto intro: order_trans[where y=ab] dest: AB) }
hence "Min B \<le> Min A" using A A' by simp
with False show "Min B < Min A" using False by auto
next
assume "Min B < Min A"
hence "\<forall>z\<in>A - B \<union> (B - A). Min B \<le> z \<and> (z \<le> Min B \<longrightarrow> Min B = z)"
using A B A' B' by(auto 4 4 intro: Min_in Min_eqI dest: bspec bspec[where x="Min B"])
moreover have "Min B \<notin> A" using \<open>Min B < Min A\<close> by (metis A Min_le not_less)
ultimately show "A \<sqsubset>' B" using A B A' B' by(simp add: set_less_aux_def bexI[where x="Min B"])
qed
thus ?thesis using False by simp
qed
lemma set_less_eq_aux_rec:
assumes "finite A" "finite B" "A \<noteq> {}" "B \<noteq> {}"
shows "A \<sqsubseteq>' B \<longleftrightarrow> Min B < Min A \<or> Min A = Min B \<and> A - {Min A} \<sqsubseteq>' B - {Min A}"
proof(cases "A = B")
case True thus ?thesis using assms by(simp add: finite_complement_partition)
next
case False
moreover
hence "Min A = Min B \<Longrightarrow> A - {Min A} \<noteq> B - {Min B}"
by (metis (lifting) assms Min_in insert_Diff)
ultimately show ?thesis using set_less_aux_rec[OF assms]
by(simp add: set_less_eq_aux_def cong: conj_cong)
qed
lemma set_less_aux_Min_antimono:
"\<lbrakk> Min A < Min B; finite A; finite B; A \<noteq> {} \<rbrakk> \<Longrightarrow> B \<sqsubset>' A"
using set_less_aux_rec[of B A]
by(cases "B = {}")(simp_all add: empty_set_less_aux_finite_iff)
lemma sorted_Cons_Min: "sorted (x # xs) \<Longrightarrow> Min (insert x (set xs)) = x"
by(auto simp add: intro: Min_eqI)
lemma set_less_aux_code:
"\<lbrakk> sorted xs; distinct xs; sorted ys; distinct ys \<rbrakk>
\<Longrightarrow> set xs \<sqsubset>' set ys \<longleftrightarrow> ord.lexordp (>) xs ys"
apply(induct xs ys rule: list_induct2')
apply(simp_all add: empty_set_less_aux_finite_iff sorted_Cons_Min set_less_aux_rec neq_Nil_conv)
apply(auto cong: conj_cong)
done
lemma set_less_eq_aux_code:
assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys"
shows "set xs \<sqsubseteq>' set ys \<longleftrightarrow> ord.lexordp_eq (>) xs ys"
proof -
have dual: "class.linorder (\<ge>) (>)"
by(rule linorder.dual_linorder) unfold_locales
from assms show ?thesis
by(auto simp add: set_less_eq_aux_def finite_complement_partition linorder.lexordp_eq_conv_lexord[OF dual] set_less_aux_code intro: sorted_distinct_set_unique)
qed
end
subsubsection \<open>Extending @{term set_less_eq_aux} to have @{term "{}"} as least element\<close>
context ord begin
definition set_less_eq_aux' :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "\<sqsubseteq>''''" 50)
where "A \<sqsubseteq>'' B \<longleftrightarrow> A \<sqsubseteq>' B \<or> A = {} \<and> B \<in> infinite_complement_partition"
lemma set_less_eq_aux'_refl:
"A \<sqsubseteq>'' A \<longleftrightarrow> A \<in> infinite_complement_partition"
by(auto simp add: set_less_eq_aux'_def)
lemma set_less_eq_aux'_antisym: "\<lbrakk> A \<sqsubseteq>'' B; B \<sqsubseteq>'' A \<rbrakk> \<Longrightarrow> A = B"
by(auto simp add: set_less_eq_aux'_def intro: set_less_eq_aux_antisym del: equalityI)
lemma set_less_eq_aux'_infinite_complement_partitionD:
"A \<sqsubseteq>'' B \<Longrightarrow> A \<in> infinite_complement_partition \<and> B \<in> infinite_complement_partition"
by(auto simp add: set_less_eq_aux'_def intro: finite_complement_partition dest: set_less_eq_aux_infinite_complement_partitionD)
lemma empty_set_less_eq_def [simp]: "{} \<sqsubseteq>'' B \<longleftrightarrow> B \<in> infinite_complement_partition"
by(auto simp add: set_less_eq_aux'_def dest: set_less_eq_aux_infinite_complement_partitionD)
end
context linorder begin
lemma set_less_eq_aux'_trans: "\<lbrakk> A \<sqsubseteq>'' B; B \<sqsubseteq>'' C \<rbrakk> \<Longrightarrow> A \<sqsubseteq>'' C"
by(auto simp add: set_less_eq_aux'_def del: equalityI intro: set_less_eq_aux_trans dest: set_less_eq_aux_infinite_complement_partitionD)
lemma set_less_eq_aux'_porder: "partial_order_on infinite_complement_partition {(A, B). A \<sqsubseteq>'' B}"
by(auto simp add: partial_order_on_def preorder_on_def intro!: refl_onI transI antisymI dest: set_less_eq_aux'_antisym set_less_eq_aux'_infinite_complement_partitionD simp add: set_less_eq_aux'_refl intro: set_less_eq_aux'_trans)
end
subsubsection \<open>Extend @{term set_less_eq_aux'} to a total order on @{term infinite_complement_partition}\<close>
context ord begin
definition set_less_eq_aux'' :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "\<sqsubseteq>''''''" 50)
where "set_less_eq_aux'' =
(SOME sleq.
(linear_order_on UNIV {(a, b). a \<le> b} \<longrightarrow> linear_order_on infinite_complement_partition {(A, B). sleq A B}) \<and> order_consistent {(A, B). A \<sqsubseteq>'' B} {(A, B). sleq A B})"
lemma set_less_eq_aux''_spec:
shows "linear_order {(a, b). a \<le> b} \<Longrightarrow> linear_order_on infinite_complement_partition {(A, B). A \<sqsubseteq>''' B}"
(is "PROP ?thesis1")
and "order_consistent {(A, B). A \<sqsubseteq>'' B} {(A, B). A \<sqsubseteq>''' B}" (is ?thesis2)
proof -
let ?P = "\<lambda>sleq. (linear_order {(a, b). a \<le> b} \<longrightarrow> linear_order_on infinite_complement_partition {(A, B). sleq A B}) \<and>
order_consistent {(A, B). A \<sqsubseteq>'' B} {(A, B). sleq A B}"
have "Ex ?P"
proof(cases "linear_order {(a, b). a \<le> b}")
case False
have "antisym {(a, b). a \<sqsubseteq>'' b}"
by (rule antisymI) (simp add: set_less_eq_aux'_antisym)
then show ?thesis using False
by (auto intro: antisym_order_consistent_self)
next
case True
hence "partial_order_on infinite_complement_partition {(A, B). A \<sqsubseteq>'' B}"
by(rule linorder.set_less_eq_aux'_porder[OF linear_order_imp_linorder])
then obtain s where "linear_order_on infinite_complement_partition s"
and "order_consistent {(A, B). A \<sqsubseteq>'' B} s" by(rule porder_extend_to_linorder)
thus ?thesis by(auto intro: exI[where x="\<lambda>A B. (A, B) \<in> s"])
qed
hence "?P (Eps ?P)" by(rule someI_ex)
thus "PROP ?thesis1" ?thesis2 by(simp_all add: set_less_eq_aux''_def)
qed
end
context linorder begin
lemma set_less_eq_aux''_linear_order:
"linear_order_on infinite_complement_partition {(A, B). A \<sqsubseteq>''' B}"
by(rule set_less_eq_aux''_spec)(rule linear_order)
lemma set_less_eq_aux''_refl [iff]: "A \<sqsubseteq>''' A \<longleftrightarrow> A \<in> infinite_complement_partition"
using set_less_eq_aux''_linear_order
by(auto simp add: linear_order_on_def partial_order_on_def preorder_on_def dest: refl_onD refl_onD1)
lemma set_less_eq_aux'_into_set_less_eq_aux'':
assumes "A \<sqsubseteq>'' B"
shows "A \<sqsubseteq>''' B"
proof(rule ccontr)
assume nleq: "\<not> ?thesis"
moreover from assms have A: "A \<in> infinite_complement_partition"
and B: "B \<in> infinite_complement_partition"
by(auto dest: set_less_eq_aux'_infinite_complement_partitionD)
with set_less_eq_aux''_linear_order have "A \<sqsubseteq>''' B \<or> A = B \<or> B \<sqsubseteq>''' A"
by(auto simp add: linear_order_on_def dest: total_onD)
ultimately have "B \<sqsubseteq>''' A" using B by auto
with assms have "A = B" using set_less_eq_aux''_spec(2)
by(simp add: order_consistent_def)
with A nleq show False by simp
qed
lemma finite_set_less_eq_aux''_finite:
assumes "finite A" and "finite B"
shows "A \<sqsubseteq>''' B \<longleftrightarrow> A \<sqsubseteq>'' B"
proof
assume "A \<sqsubseteq>''' B"
from assms have "A \<sqsubseteq>' B \<or> B \<sqsubseteq>' A" by(rule set_less_eq_aux_finite_total2)
hence "A \<sqsubseteq>'' B \<or> B \<sqsubseteq>'' A" by(auto simp add: set_less_eq_aux'_def)
thus "A \<sqsubseteq>'' B"
proof
assume "B \<sqsubseteq>'' A"
hence "B \<sqsubseteq>''' A" by(rule set_less_eq_aux'_into_set_less_eq_aux'')
with \<open>A \<sqsubseteq>''' B\<close> set_less_eq_aux''_linear_order have "A = B"
by(auto simp add: linear_order_on_def partial_order_on_def dest: antisymD)
thus ?thesis using assms by(simp add: finite_complement_partition set_less_eq_aux'_def)
qed
qed(rule set_less_eq_aux'_into_set_less_eq_aux'')
lemma set_less_eq_aux''_finite:
"finite (UNIV :: 'a set) \<Longrightarrow> set_less_eq_aux'' = set_less_eq_aux"
by(auto simp add: fun_eq_iff finite_set_less_eq_aux''_finite set_less_eq_aux'_def finite_subset[OF subset_UNIV])
lemma set_less_eq_aux''_antisym:
"\<lbrakk> A \<sqsubseteq>''' B; B \<sqsubseteq>''' A;
A \<in> infinite_complement_partition; B \<in> infinite_complement_partition \<rbrakk>
\<Longrightarrow> A = B"
using set_less_eq_aux''_linear_order
by(auto simp add: linear_order_on_def partial_order_on_def dest: antisymD del: equalityI)
lemma set_less_eq_aux''_trans: "\<lbrakk> A \<sqsubseteq>''' B; B \<sqsubseteq>''' C \<rbrakk> \<Longrightarrow> A \<sqsubseteq>''' C"
using set_less_eq_aux''_linear_order
by(auto simp add: linear_order_on_def partial_order_on_def preorder_on_def dest: transD)
lemma set_less_eq_aux''_total:
"\<lbrakk> A \<in> infinite_complement_partition; B \<in> infinite_complement_partition \<rbrakk>
\<Longrightarrow> A \<sqsubseteq>''' B \<or> B \<sqsubseteq>''' A"
using set_less_eq_aux''_linear_order
by(auto simp add: linear_order_on_def dest: total_onD)
end
subsubsection \<open>Extend @{term set_less_eq_aux''} to cofinite sets\<close>
context ord begin
definition set_less_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
where
"A \<sqsubseteq> B \<longleftrightarrow>
(if A \<in> infinite_complement_partition then A \<sqsubseteq>''' B \<or> B \<notin> infinite_complement_partition
else B \<notin> infinite_complement_partition \<and> - B \<sqsubseteq>''' - A)"
definition set_less :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "\<sqsubset>" 50)
where "A \<sqsubset> B \<longleftrightarrow> A \<sqsubseteq> B \<and> \<not> B \<sqsubseteq> A"
lemma set_less_eq_def2:
"A \<sqsubseteq> B \<longleftrightarrow>
(if finite (UNIV :: 'a set) then A \<sqsubseteq>''' B
else if A \<in> infinite_complement_partition then A \<sqsubseteq>''' B \<or> B \<notin> infinite_complement_partition
else B \<notin> infinite_complement_partition \<and> - B \<sqsubseteq>''' - A)"
by(simp add: set_less_eq_def)
end
context linorder begin
lemma set_less_eq_refl [iff]: "A \<sqsubseteq> A"
by(auto simp add: set_less_eq_def2 not_in_complement_partition)
lemma set_less_eq_antisym: "\<lbrakk> A \<sqsubseteq> B; B \<sqsubseteq> A \<rbrakk> \<Longrightarrow> A = B"
by(auto simp add: set_less_eq_def2 set_less_eq_aux''_finite not_in_complement_partition not_in_complement_partition_False del: equalityI split: if_split_asm dest: set_less_eq_aux_antisym set_less_eq_aux''_antisym)
lemma set_less_eq_trans: "\<lbrakk> A \<sqsubseteq> B; B \<sqsubseteq> C \<rbrakk> \<Longrightarrow> A \<sqsubseteq> C"
by(auto simp add: set_less_eq_def split: if_split_asm intro: set_less_eq_aux''_trans)
lemma set_less_eq_total: "A \<sqsubseteq> B \<or> B \<sqsubseteq> A"
by(auto simp add: set_less_eq_def2 set_less_eq_aux''_finite not_in_complement_partition not_in_complement_partition_False intro: set_less_eq_aux_finite_total2 finite_subset[OF subset_UNIV] del: disjCI dest: set_less_eq_aux''_total)
lemma set_less_eq_linorder: "class.linorder (\<sqsubseteq>) (\<sqsubset>)"
by(unfold_locales)(auto simp add: set_less_def set_less_eq_antisym set_less_eq_total intro: set_less_eq_trans)
lemma set_less_eq_conv_set_less: "set_less_eq A B \<longleftrightarrow> A = B \<or> set_less A B"
by(auto simp add: set_less_def del: equalityI dest: set_less_eq_antisym)
lemma Compl_set_less_eq_Compl: "- A \<sqsubseteq> - B \<longleftrightarrow> B \<sqsubseteq> A"
by(auto simp add: set_less_eq_def2 not_in_complement_partition_False not_in_complement_partition set_less_eq_aux''_finite Compl_set_less_eq_aux_Compl)
lemma Compl_set_less_Compl: "- A \<sqsubset> - B \<longleftrightarrow> B \<sqsubset> A"
by(simp add: set_less_def Compl_set_less_eq_Compl)
lemma set_less_eq_finite_iff: "\<lbrakk> finite A; finite B \<rbrakk> \<Longrightarrow> A \<sqsubseteq> B \<longleftrightarrow> A \<sqsubseteq>' B"
by(auto simp add: set_less_eq_def finite_complement_partition set_less_eq_aux'_def finite_set_less_eq_aux''_finite)
lemma set_less_finite_iff: "\<lbrakk> finite A; finite B \<rbrakk> \<Longrightarrow> A \<sqsubset> B \<longleftrightarrow> A \<sqsubset>' B"
by(simp add: set_less_def set_less_aux_conv_set_less_eq_aux set_less_eq_finite_iff)
lemma infinite_set_less_eq_Complement:
"\<lbrakk> finite A; finite B; \<not> finite (UNIV :: 'a set) \<rbrakk> \<Longrightarrow> A \<sqsubseteq> - B"
by(simp add: set_less_eq_def finite_complement_partition not_in_complement_partition)
lemma infinite_set_less_Complement:
"\<lbrakk> finite A; finite B; \<not> finite (UNIV :: 'a set) \<rbrakk> \<Longrightarrow> A \<sqsubset> - B"
by(auto simp add: set_less_def dest: set_less_eq_antisym intro: infinite_set_less_eq_Complement)
lemma infinite_Complement_set_less_eq:
"\<lbrakk> finite A; finite B; \<not> finite (UNIV :: 'a set) \<rbrakk> \<Longrightarrow> \<not> - A \<sqsubseteq> B"
using infinite_set_less_eq_Complement[of A B] Compl_set_less_eq_Compl[of A "- B"]
by(auto dest: set_less_eq_antisym)
lemma infinite_Complement_set_less:
"\<lbrakk> finite A; finite B; \<not> finite (UNIV :: 'a set) \<rbrakk> \<Longrightarrow> \<not> - A \<sqsubset> B"
using infinite_Complement_set_less_eq[of A B]
by(simp add: set_less_def)
lemma empty_set_less_eq [iff]: "{} \<sqsubseteq> A"
by(auto simp add: set_less_eq_def finite_complement_partition intro: set_less_eq_aux'_into_set_less_eq_aux'')
lemma set_less_eq_empty [iff]: "A \<sqsubseteq> {} \<longleftrightarrow> A = {}"
by (metis empty_set_less_eq set_less_eq_antisym)
lemma empty_set_less_iff [iff]: "{} \<sqsubset> A \<longleftrightarrow> A \<noteq> {}"
by(simp add: set_less_def)
lemma not_set_less_empty [simp]: "\<not> A \<sqsubset> {}"
by(simp add: set_less_def)
lemma set_less_eq_UNIV [iff]: "A \<sqsubseteq> UNIV"
using Compl_set_less_eq_Compl[of "- A" "{}"] by simp
lemma UNIV_set_less_eq [iff]: "UNIV \<sqsubseteq> A \<longleftrightarrow> A = UNIV"
using Compl_set_less_eq_Compl[of "{}" "- A"]
by(simp add: Compl_eq_empty_iff)
lemma set_less_UNIV_iff [iff]: "A \<sqsubset> UNIV \<longleftrightarrow> A \<noteq> UNIV"
by(simp add: set_less_def)
lemma not_UNIV_set_less [simp]: "\<not> UNIV \<sqsubset> A"
by(simp add: set_less_def)
end
subsection \<open>Implementation based on sorted lists\<close>
type_synonym 'a proper_interval = "'a option \<Rightarrow> 'a option \<Rightarrow> bool"
class proper_intrvl = ord +
fixes proper_interval :: "'a proper_interval"
class proper_interval = proper_intrvl +
assumes proper_interval_simps:
"proper_interval None None = True"
"proper_interval None (Some y) = (\<exists>z. z < y)"
"proper_interval (Some x) None = (\<exists>z. x < z)"
"proper_interval (Some x) (Some y) = (\<exists>z. x < z \<and> z < y)"
context proper_intrvl begin
function set_less_eq_aux_Compl :: "'a option \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where
"set_less_eq_aux_Compl ao [] ys \<longleftrightarrow> True"
| "set_less_eq_aux_Compl ao xs [] \<longleftrightarrow> True"
| "set_less_eq_aux_Compl ao (x # xs) (y # ys) \<longleftrightarrow>
(if x < y then proper_interval ao (Some x) \<or> set_less_eq_aux_Compl (Some x) xs (y # ys)
else if y < x then proper_interval ao (Some y) \<or> set_less_eq_aux_Compl (Some y) (x # xs) ys
else proper_interval ao (Some y))"
by(pat_completeness) simp_all
termination by(lexicographic_order)
fun Compl_set_less_eq_aux :: "'a option \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where
"Compl_set_less_eq_aux ao [] [] \<longleftrightarrow> \<not> proper_interval ao None"
| "Compl_set_less_eq_aux ao [] (y # ys) \<longleftrightarrow> \<not> proper_interval ao (Some y) \<and> Compl_set_less_eq_aux (Some y) [] ys"
| "Compl_set_less_eq_aux ao (x # xs) [] \<longleftrightarrow> \<not> proper_interval ao (Some x) \<and> Compl_set_less_eq_aux (Some x) xs []"
| "Compl_set_less_eq_aux ao (x # xs) (y # ys) \<longleftrightarrow>
(if x < y then \<not> proper_interval ao (Some x) \<and> Compl_set_less_eq_aux (Some x) xs (y # ys)
else if y < x then \<not> proper_interval ao (Some y) \<and> Compl_set_less_eq_aux (Some y) (x # xs) ys
else \<not> proper_interval ao (Some y))"
fun set_less_aux_Compl :: "'a option \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"set_less_aux_Compl ao [] [] \<longleftrightarrow> proper_interval ao None"
| "set_less_aux_Compl ao [] (y # ys) \<longleftrightarrow> proper_interval ao (Some y) \<or> set_less_aux_Compl (Some y) [] ys"
| "set_less_aux_Compl ao (x # xs) [] \<longleftrightarrow> proper_interval ao (Some x) \<or> set_less_aux_Compl (Some x) xs []"
| "set_less_aux_Compl ao (x # xs) (y # ys) \<longleftrightarrow>
(if x < y then proper_interval ao (Some x) \<or> set_less_aux_Compl (Some x) xs (y # ys)
else if y < x then proper_interval ao (Some y) \<or> set_less_aux_Compl (Some y) (x # xs) ys
else proper_interval ao (Some y))"
function Compl_set_less_aux :: "'a option \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"Compl_set_less_aux ao [] ys \<longleftrightarrow> False"
| "Compl_set_less_aux ao xs [] \<longleftrightarrow> False"
| "Compl_set_less_aux ao (x # xs) (y # ys) \<longleftrightarrow>
(if x < y then \<not> proper_interval ao (Some x) \<and> Compl_set_less_aux (Some x) xs (y # ys)
else if y < x then \<not> proper_interval ao (Some y) \<and> Compl_set_less_aux (Some y) (x # xs) ys
else \<not> proper_interval ao (Some y))"
by pat_completeness simp_all
termination by lexicographic_order
end
lemmas [code] =
proper_intrvl.set_less_eq_aux_Compl.simps
proper_intrvl.set_less_aux_Compl.simps
proper_intrvl.Compl_set_less_eq_aux.simps
proper_intrvl.Compl_set_less_aux.simps
class linorder_proper_interval = linorder + proper_interval
begin
theorem assumes fin: "finite (UNIV :: 'a set)"
and xs: "sorted xs" "distinct xs"
and ys: "sorted ys" "distinct ys"
shows set_less_eq_aux_Compl2_conv_set_less_eq_aux_Compl:
"set xs \<sqsubseteq>' - set ys \<longleftrightarrow> set_less_eq_aux_Compl None xs ys" (is ?Compl2)
and Compl1_set_less_eq_aux_conv_Compl_set_less_eq_aux:
"- set xs \<sqsubseteq>' set ys \<longleftrightarrow> Compl_set_less_eq_aux None xs ys" (is ?Compl1)
proof -
note fin' [simp] = finite_subset[OF subset_UNIV fin]
define above where "above = case_option UNIV (Collect \<circ> less)"
have above_simps [simp]: "above None = UNIV" "\<And>x. above (Some x) = {y. x < y}"
and above_upclosed: "\<And>x y ao. \<lbrakk> x \<in> above ao; x < y \<rbrakk> \<Longrightarrow> y \<in> above ao"
and proper_interval_Some2: "\<And>x ao. proper_interval ao (Some x) \<longleftrightarrow> (\<exists>z\<in>above ao. z < x)"
and proper_interval_None2: "\<And>ao. proper_interval ao None \<longleftrightarrow> above ao \<noteq> {}"
by(simp_all add: proper_interval_simps above_def split: option.splits)
{ fix ao x A B
assume ex: "proper_interval ao (Some x)"
and A: "\<forall>a \<in> A. x \<le> a"
and B: "\<forall>b \<in> B. x \<le> b"
from ex obtain z where z_ao: "z \<in> above ao" and z: "z < x"
by(auto simp add: proper_interval_Some2)
with A have A_eq: "A \<inter> above ao = A"
by(auto simp add: above_upclosed)
from z_ao z B have B_eq: "B \<inter> above ao = B"
by(auto simp add: above_upclosed)
define w where "w = Min (above ao)"
with z_ao have "w \<le> z" "\<forall>z \<in> above ao. w \<le> z" "w \<in> above ao"
by(auto simp add: Min_le_iff intro: Min_in)
hence "A \<inter> above ao \<sqsubset>' (- B) \<inter> above ao" (is "?lhs \<sqsubset>' ?rhs")
using A B z by(auto simp add: set_less_aux_def intro!: bexI[where x="w"])
hence "A \<sqsubseteq>' ?rhs" unfolding A_eq by(simp add: set_less_eq_aux_def)
moreover
from z_ao z A B have "z \<in> - A \<inter> above ao" "z \<notin> B" by auto
hence neq: "- A \<inter> above ao \<noteq> B \<inter> above ao" by auto
have "\<not> - A \<inter> above ao \<sqsubset>' B \<inter> above ao" (is "\<not> ?lhs' \<sqsubset>' ?rhs'")
using A B z z_ao by(force simp add: set_less_aux_def not_less dest: bspec[where x=z])
with neq have "\<not> ?lhs' \<sqsubseteq>' B" unfolding B_eq by(auto simp add: set_less_eq_aux_def)
moreover note calculation }
note proper_interval_set_less_eqI = this(1)
and proper_interval_not_set_less_eq_auxI = this(2)
{ fix ao
assume "set xs \<union> set ys \<subseteq> above ao"
with xs ys
have "set_less_eq_aux_Compl ao xs ys \<longleftrightarrow> set xs \<sqsubseteq>' (- set ys) \<inter> above ao"
proof(induction ao xs ys rule: set_less_eq_aux_Compl.induct)
case 1 thus ?case by simp
next
case 2 thus ?case by(auto intro: subset_finite_imp_set_less_eq_aux)
next
case (3 ao x xs y ys)
note ao = \<open>set (x # xs) \<union> set (y # ys) \<subseteq> above ao\<close>
hence x_ao: "x \<in> above ao" and y_ao: "y \<in> above ao" by simp_all
note yys = \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close>
hence ys: "sorted ys" "distinct ys" and y_Min: "\<forall>y' \<in> set ys. y < y'"
by(auto simp add: less_le)
note xxs = \<open>sorted (x # xs)\<close> \<open>distinct (x # xs)\<close>
hence xs: "sorted xs" "distinct xs" and x_Min: "\<forall>x' \<in> set xs. x < x'"
by(auto simp add: less_le)
let ?lhs = "set (x # xs)" and ?rhs = "- set (y # ys) \<inter> above ao"
show ?case
proof(cases "x < y")
case True
show ?thesis
proof(cases "proper_interval ao (Some x)")
case True
hence "?lhs \<sqsubseteq>' ?rhs" using x_Min y_Min \<open>x < y\<close>
by(auto intro!: proper_interval_set_less_eqI)
with True show ?thesis using \<open>x < y\<close> by simp
next
case False
have "set xs \<union> set (y # ys) \<subseteq> above (Some x)" using True x_Min y_Min by auto
with True xs yys
have IH: "set_less_eq_aux_Compl (Some x) xs (y # ys) =
(set xs \<sqsubseteq>' - set (y # ys) \<inter> above (Some x))"
by(rule "3.IH")
from True y_Min x_ao have "x \<in> - set (y # ys) \<inter> above ao" by auto
hence "?rhs \<noteq> {}" by blast
moreover have "Min ?lhs = x" using x_Min x_ao by(auto intro!: Min_eqI)
moreover have "Min ?rhs = x" using \<open>x < y\<close> y_Min x_ao False
by(auto intro!: Min_eqI simp add: proper_interval_Some2)
moreover have "set xs = set xs - {x}"
using ao x_Min by auto
moreover have "- set (y # ys) \<inter> above (Some x) = - set (y # ys) \<inter> above ao - {x}"
using False x_ao by(auto simp add: proper_interval_Some2 intro: above_upclosed)
ultimately show ?thesis using True False IH
by(simp del: set_simps)(subst (2) set_less_eq_aux_rec, simp_all add: x_ao)
qed
next
case False
show ?thesis
proof(cases "y < x")
case True
show ?thesis
proof(cases "proper_interval ao (Some y)")
case True
hence "?lhs \<sqsubseteq>' ?rhs" using x_Min y_Min \<open>\<not> x < y\<close>
by(auto intro!: proper_interval_set_less_eqI)
with True show ?thesis using \<open>\<not> x < y\<close> by simp
next
case False
have "set (x # xs) \<union> set ys \<subseteq> above (Some y)"
using \<open>y < x\<close> x_Min y_Min by auto
with \<open>\<not> x < y\<close> \<open>y < x\<close> xxs ys
have IH: "set_less_eq_aux_Compl (Some y) (x # xs) ys =
(set (x # xs) \<sqsubseteq>' - set ys \<inter> above (Some y))"
by(rule "3.IH")
moreover have "- set ys \<inter> above (Some y) = ?rhs"
using y_ao False by(auto intro: above_upclosed simp add: proper_interval_Some2)
ultimately show ?thesis using \<open>\<not> x < y\<close> True False by simp
qed
next
case False with \<open>\<not> x < y\<close> have "x = y" by auto
{ assume "proper_interval ao (Some y)"
hence "?lhs \<sqsubseteq>' ?rhs" using x_Min y_Min \<open>x = y\<close>
by(auto intro!: proper_interval_set_less_eqI) }
moreover
{ assume "?lhs \<sqsubseteq>' ?rhs"
moreover have "?lhs \<noteq> ?rhs"
proof
assume eq: "?lhs = ?rhs"
have "x \<in> ?lhs" using x_ao by simp
also note eq also note \<open>x = y\<close>
finally show False by simp
qed
ultimately obtain z where "z \<in> above ao" "z < y" using \<open>x = y\<close> y_ao
by(fastforce simp add: set_less_eq_aux_def set_less_aux_def not_le dest!: bspec[where x=y])
hence "proper_interval ao (Some y)" by(auto simp add: proper_interval_Some2) }
ultimately show ?thesis using \<open>x = y\<close> \<open>\<not> x < y\<close> \<open>\<not> y < x\<close> by auto
qed
qed
qed }
from this[of None] show ?Compl2 by simp
{ fix ao
assume "set xs \<union> set ys \<subseteq> above ao"
with xs ys
have "Compl_set_less_eq_aux ao xs ys \<longleftrightarrow> (- set xs) \<inter> above ao \<sqsubseteq>' set ys"
proof(induction ao xs ys rule: Compl_set_less_eq_aux.induct)
case 1 thus ?case by(simp add: proper_interval_None2)
next
case (2 ao y ys)
from \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close>
have ys: "sorted ys" "distinct ys" and y_Min: "\<forall>y' \<in> set ys. y < y'"
by(auto simp add: less_le)
show ?case
proof(cases "proper_interval ao (Some y)")
case True
hence "\<not> - set [] \<inter> above ao \<sqsubseteq>' set (y # ys)" using y_Min
by -(erule proper_interval_not_set_less_eq_auxI, auto)
thus ?thesis using True by simp
next
case False
note ao = \<open>set [] \<union> set (y # ys) \<subseteq> above ao\<close>
hence y_ao: "y \<in> above ao" by simp
from ao y_Min have "set [] \<union> set ys \<subseteq> above (Some y)" by auto
with \<open>sorted []\<close> \<open>distinct []\<close> ys
have "Compl_set_less_eq_aux (Some y) [] ys \<longleftrightarrow> - set [] \<inter> above (Some y) \<sqsubseteq>' set ys"
by(rule "2.IH")
moreover have "above ao \<noteq> {}" using y_ao by auto
moreover have "Min (above ao) = y"
and "Min (set (y # ys)) = y"
using y_ao False ao by(auto intro!: Min_eqI simp add: proper_interval_Some2 not_less)
moreover have "above ao - {y} = above (Some y)" using False y_ao
by(auto simp add: proper_interval_Some2 intro: above_upclosed)
moreover have "set ys - {y} = set ys"
using y_Min y_ao by(auto)
ultimately show ?thesis using False y_ao
by(simp)(subst (2) set_less_eq_aux_rec, simp_all)
qed
next
case (3 ao x xs)
from \<open>sorted (x # xs)\<close> \<open>distinct (x # xs)\<close>
have xs: "sorted xs" "distinct xs" and x_Min: "\<forall>x'\<in>set xs. x < x'"
by(auto simp add: less_le)
show ?case
proof(cases "proper_interval ao (Some x)")
case True
then obtain z where "z \<in> above ao" "z < x" by(auto simp add: proper_interval_Some2)
hence "z \<in> - set (x # xs) \<inter> above ao" using x_Min by auto
thus ?thesis using True by auto
next
case False
note ao = \<open>set (x # xs) \<union> set [] \<subseteq> above ao\<close>
hence x_ao: "x \<in> above ao" by simp
from ao have "set xs \<union> set [] \<subseteq> above (Some x)" using x_Min by auto
with xs \<open>sorted []\<close> \<open>distinct []\<close>
have "Compl_set_less_eq_aux (Some x) xs [] \<longleftrightarrow>
- set xs \<inter> above (Some x) \<sqsubseteq>' set []"
by(rule "3.IH")
moreover have "- set (x # xs) \<inter> above ao = - set xs \<inter> above (Some x)"
using False x_ao by(auto simp add: proper_interval_Some2 intro: above_upclosed)
ultimately show ?thesis using False by simp
qed
next
case (4 ao x xs y ys)
note ao = \<open>set (x # xs) \<union> set (y # ys) \<subseteq> above ao\<close>
hence x_ao: "x \<in> above ao" and y_ao: "y \<in> above ao" by simp_all
note xxs = \<open>sorted (x # xs)\<close> \<open>distinct (x # xs)\<close>
hence xs: "sorted xs" "distinct xs" and x_Min: "\<forall>x'\<in>set xs. x < x'"
by(auto simp add: less_le)
note yys = \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close>
hence ys: "sorted ys" "distinct ys" and y_Min: "\<forall>y'\<in>set ys. y < y'"
by(auto simp add: less_le)
let ?lhs = "- set (x # xs) \<inter> above ao" and ?rhs = "set (y # ys)"
show ?case
proof(cases "x < y")
case True
show ?thesis
proof(cases "proper_interval ao (Some x)")
case True
hence "\<not> ?lhs \<sqsubseteq>' ?rhs" using x_Min y_Min \<open>x < y\<close>
by -(erule proper_interval_not_set_less_eq_auxI, auto)
thus ?thesis using True \<open>x < y\<close> by simp
next
case False
have "set xs \<union> set (y # ys) \<subseteq> above (Some x)"
using ao x_Min y_Min True by auto
with True xs yys
have "Compl_set_less_eq_aux (Some x) xs (y # ys) \<longleftrightarrow>
- set xs \<inter> above (Some x) \<sqsubseteq>' set (y # ys)"
by(rule "4.IH")
moreover have "- set xs \<inter> above (Some x) = ?lhs"
using x_ao False by(auto intro: above_upclosed simp add: proper_interval_Some2)
ultimately show ?thesis using False True by simp
qed
next
case False
show ?thesis
proof(cases "y < x")
case True
show ?thesis
proof(cases "proper_interval ao (Some y)")
case True
hence "\<not> ?lhs \<sqsubseteq>' ?rhs" using x_Min y_Min \<open>y < x\<close>
by -(erule proper_interval_not_set_less_eq_auxI, auto)
thus ?thesis using True \<open>y < x\<close> \<open>\<not> x < y\<close> by simp
next
case False
from ao True x_Min y_Min
have "set (x # xs) \<union> set ys \<subseteq> above (Some y)" by auto
with \<open>\<not> x < y\<close> True xxs ys
have "Compl_set_less_eq_aux (Some y) (x # xs) ys \<longleftrightarrow>
- set (x # xs) \<inter> above (Some y) \<sqsubseteq>' set ys"
by(rule "4.IH")
moreover have "y \<in> ?lhs" using True x_Min y_ao by auto
hence "?lhs \<noteq> {}" by auto
moreover have "Min ?lhs = y" using True False x_Min y_ao
by(auto intro!: Min_eqI simp add: not_le not_less proper_interval_Some2)
moreover have "Min ?rhs = y" using y_Min y_ao by(auto intro!: Min_eqI)
moreover have "- set (x # xs) \<inter> above (Some y) = ?lhs - {y}"
using y_ao False by(auto intro: above_upclosed simp add: proper_interval_Some2)
moreover have "set ys = set ys - {y}"
using y_ao y_Min by(auto intro: above_upclosed)
ultimately show ?thesis using True False \<open>\<not> x < y\<close> y_ao
by(simp)(subst (2) set_less_eq_aux_rec, simp_all)
qed
next
case False
with \<open>\<not> x < y\<close> have "x = y" by auto
{ assume "proper_interval ao (Some y)"
hence "\<not> ?lhs \<sqsubseteq>' ?rhs" using x_Min y_Min \<open>x = y\<close>
by -(erule proper_interval_not_set_less_eq_auxI, auto) }
moreover
{ assume "\<not> ?lhs \<sqsubseteq>' ?rhs"
also have "?rhs = set (y # ys) \<inter> above ao"
using ao by auto
finally obtain z where "z \<in> above ao" "z < y"
using \<open>x = y\<close> x_ao x_Min[unfolded Ball_def]
by(fastforce simp add: set_less_eq_aux_def set_less_aux_def simp add: less_le not_le dest!: bspec[where x=y])
hence "proper_interval ao (Some y)"
by(auto simp add: proper_interval_Some2) }
ultimately show ?thesis using \<open>x = y\<close> by auto
qed
qed
qed }
from this[of None] show ?Compl1 by simp
qed
lemma set_less_aux_Compl_iff:
"set_less_aux_Compl ao xs ys \<longleftrightarrow> set_less_eq_aux_Compl ao xs ys \<and> \<not> Compl_set_less_eq_aux ao ys xs"
by(induct ao xs ys rule: set_less_aux_Compl.induct)(auto simp add: not_less_iff_gr_or_eq)
lemma Compl_set_less_aux_Compl_iff:
"Compl_set_less_aux ao xs ys \<longleftrightarrow> Compl_set_less_eq_aux ao xs ys \<and> \<not> set_less_eq_aux_Compl ao ys xs"
by(induct ao xs ys rule: Compl_set_less_aux.induct)(auto simp add: not_less_iff_gr_or_eq)
theorem assumes fin: "finite (UNIV :: 'a set)"
and xs: "sorted xs" "distinct xs"
and ys: "sorted ys" "distinct ys"
shows set_less_aux_Compl2_conv_set_less_aux_Compl:
"set xs \<sqsubset>' - set ys \<longleftrightarrow> set_less_aux_Compl None xs ys" (is ?Compl2)
and Compl1_set_less_aux_conv_Compl_set_less_aux:
"- set xs \<sqsubset>' set ys \<longleftrightarrow> Compl_set_less_aux None xs ys" (is ?Compl1)
using assms
by(simp_all only: set_less_aux_conv_set_less_eq_aux set_less_aux_Compl_iff Compl_set_less_aux_Compl_iff set_less_eq_aux_Compl2_conv_set_less_eq_aux_Compl Compl1_set_less_eq_aux_conv_Compl_set_less_eq_aux)
end
subsection \<open>Implementation of proper intervals for sets\<close>
definition length_last :: "'a list \<Rightarrow> nat \<times> 'a"
where "length_last xs = (length xs, last xs)"
lemma length_last_Nil [code]: "length_last [] = (0, undefined)"
by(simp add: length_last_def last_def)
lemma length_last_Cons_code [symmetric, code]:
"fold (\<lambda>x (n, _) . (n + 1, x)) xs (1, x) = length_last (x # xs)"
by(induct xs rule: rev_induct)(simp_all add: length_last_def)
context proper_intrvl begin
fun exhaustive_above :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" where
"exhaustive_above x [] \<longleftrightarrow> \<not> proper_interval (Some x) None"
| "exhaustive_above x (y # ys) \<longleftrightarrow> \<not> proper_interval (Some x) (Some y) \<and> exhaustive_above y ys"
fun exhaustive :: "'a list \<Rightarrow> bool" where
"exhaustive [] = False"
| "exhaustive (x # xs) \<longleftrightarrow> \<not> proper_interval None (Some x) \<and> exhaustive_above x xs"
fun proper_interval_set_aux :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where
"proper_interval_set_aux xs [] \<longleftrightarrow> False"
| "proper_interval_set_aux [] (y # ys) \<longleftrightarrow> ys \<noteq> [] \<or> proper_interval (Some y) None"
| "proper_interval_set_aux (x # xs) (y # ys) \<longleftrightarrow>
(if x < y then False
else if y < x then proper_interval (Some y) (Some x) \<or> ys \<noteq> [] \<or> \<not> exhaustive_above x xs
else proper_interval_set_aux xs ys)"
fun proper_interval_set_Compl_aux :: "'a option \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where
"proper_interval_set_Compl_aux ao n [] [] \<longleftrightarrow>
CARD('a) > n + 1"
| "proper_interval_set_Compl_aux ao n [] (y # ys) \<longleftrightarrow>
(let m = CARD('a) - n; (len_y, y') = length_last (y # ys)
in m \<noteq> len_y \<and> (m = len_y + 1 \<longrightarrow> \<not> proper_interval (Some y') None))"
| "proper_interval_set_Compl_aux ao n (x # xs) [] \<longleftrightarrow>
(let m = CARD('a) - n; (len_x, x') = length_last (x # xs)
in m \<noteq> len_x \<and> (m = len_x + 1 \<longrightarrow> \<not> proper_interval (Some x') None))"
| "proper_interval_set_Compl_aux ao n (x # xs) (y # ys) \<longleftrightarrow>
(if x < y then
proper_interval ao (Some x) \<or>
proper_interval_set_Compl_aux (Some x) (n + 1) xs (y # ys)
else if y < x then
proper_interval ao (Some y) \<or>
proper_interval_set_Compl_aux (Some y) (n + 1) (x # xs) ys
else proper_interval ao (Some x) \<and>
(let m = card (UNIV :: 'a set) - n in m - length ys \<noteq> 2 \<or> m - length xs \<noteq> 2))"
fun proper_interval_Compl_set_aux :: "'a option \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where
"proper_interval_Compl_set_aux ao (x # xs) (y # ys) \<longleftrightarrow>
(if x < y then
\<not> proper_interval ao (Some x) \<and>
proper_interval_Compl_set_aux (Some x) xs (y # ys)
else if y < x then
\<not> proper_interval ao (Some y) \<and>
proper_interval_Compl_set_aux (Some y) (x # xs) ys
else \<not> proper_interval ao (Some x) \<and> (ys = [] \<longrightarrow> xs \<noteq> []))"
| "proper_interval_Compl_set_aux ao _ _ \<longleftrightarrow> False"
end
lemmas [code] =
proper_intrvl.exhaustive_above.simps
proper_intrvl.exhaustive.simps
proper_intrvl.proper_interval_set_aux.simps
proper_intrvl.proper_interval_set_Compl_aux.simps
proper_intrvl.proper_interval_Compl_set_aux.simps
context linorder_proper_interval begin
lemma exhaustive_above_iff:
"\<lbrakk> sorted xs; distinct xs; \<forall>x'\<in>set xs. x < x' \<rbrakk> \<Longrightarrow> exhaustive_above x xs \<longleftrightarrow> set xs = {z. z > x}"
proof(induction x xs rule: exhaustive_above.induct)
case 1 thus ?case by(simp add: proper_interval_simps)
next
case (2 x y ys)
from \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close>
have ys: "sorted ys" "distinct ys" and y: "\<forall>y'\<in>set ys. y < y'"
by(auto simp add: less_le)
hence "exhaustive_above y ys = (set ys = {z. y < z})" by(rule "2.IH")
moreover from \<open>\<forall>y'\<in>set (y # ys). x < y'\<close> have "x < y" by simp
ultimately show ?case using y
by(fastforce simp add: proper_interval_simps)
qed
lemma exhaustive_correct:
assumes "sorted xs" "distinct xs"
shows "exhaustive xs \<longleftrightarrow> set xs = UNIV"
proof(cases xs)
case Nil thus ?thesis by auto
next
case Cons
show ?thesis using assms unfolding Cons exhaustive.simps
apply(subst exhaustive_above_iff)
- apply(auto simp add: less_le proper_interval_simps not_less)
- by (metis List.set_simps(2) UNIV_I eq_iff set_ConsD)
+ apply(auto simp add: less_le proper_interval_simps not_less intro: order_antisym)
+ done
qed
theorem proper_interval_set_aux:
assumes fin: "finite (UNIV :: 'a set)"
and xs: "sorted xs" "distinct xs"
and ys: "sorted ys" "distinct ys"
shows "proper_interval_set_aux xs ys \<longleftrightarrow> (\<exists>A. set xs \<sqsubset>' A \<and> A \<sqsubset>' set ys)"
proof -
note [simp] = finite_subset[OF subset_UNIV fin]
show ?thesis using xs ys
proof(induction xs ys rule: proper_interval_set_aux.induct)
case 1 thus ?case by simp
next
case (2 y ys)
hence "\<forall>y'\<in>set ys. y < y'" by(auto simp add: less_le)
thus ?case
by(cases ys)(auto simp add: proper_interval_simps set_less_aux_singleton_iff intro: psubset_finite_imp_set_less_aux)
next
case (3 x xs y ys)
from \<open>sorted (x # xs)\<close> \<open>distinct (x # xs)\<close>
have xs: "sorted xs" "distinct xs" and x: "\<forall>x'\<in>set xs. x < x'"
by(auto simp add: less_le)
from \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close>
have ys: "sorted ys" "distinct ys" and y: "\<forall>y'\<in>set ys. y < y'"
by(auto simp add: less_le)
have Minxxs: "Min (set (x # xs)) = x" and xnxs: "x \<notin> set xs"
using x by(auto intro!: Min_eqI)
have Minyys: "Min (set (y # ys)) = y" and ynys: "y \<notin> set ys"
using y by(auto intro!: Min_eqI)
show ?case
proof(cases "x < y")
case True
hence "set (y # ys) \<sqsubset>' set (x # xs)" using Minxxs Minyys
by -(rule set_less_aux_Min_antimono, simp_all)
thus ?thesis using True by(auto dest: set_less_aux_trans set_less_aux_antisym)
next
case False
show ?thesis
proof(cases "y < x")
case True
{ assume "proper_interval (Some y) (Some x)"
then obtain z where z: "y < z" "z < x" by(auto simp add: proper_interval_simps)
hence "set (x # xs) \<sqsubset>' {z}" using x
by -(rule set_less_aux_Min_antimono, auto)
moreover have "{z} \<sqsubset>' set (y # ys)" using z y Minyys
by -(rule set_less_aux_Min_antimono, auto)
ultimately have "\<exists>A. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' set (y # ys)" by blast }
moreover
{ assume "ys \<noteq> []"
hence "{y} \<sqsubset>' set (y # ys)" using y
by -(rule psubset_finite_imp_set_less_aux, auto simp add: neq_Nil_conv)
moreover have "set (x # xs) \<sqsubset>' {y}" using False True x
by -(rule set_less_aux_Min_antimono, auto)
ultimately have "\<exists>A. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' set (y # ys)" by blast }
moreover
{ assume "\<not> exhaustive_above x xs"
then obtain z where z: "z > x" "z \<notin> set xs" using x
by(auto simp add: exhaustive_above_iff[OF xs x])
let ?A = "insert z (set (x # xs))"
have "set (x # xs) \<sqsubset>' ?A" using z
by -(rule psubset_finite_imp_set_less_aux, auto)
moreover have "?A \<sqsubset>' set (y # ys)" using Minyys False True z x
by -(rule set_less_aux_Min_antimono, auto)
ultimately have "\<exists>A. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' set (y # ys)" by blast }
moreover
{ fix A
assume A: "set (x # xs) \<sqsubset>' A" and A': "A \<sqsubset>' {y}"
and pi: "\<not> proper_interval (Some y) (Some x)"
from A have nempty: "A \<noteq> {}" by auto
have "y \<notin> A"
proof
assume "y \<in> A"
moreover with A' have "A \<noteq> {y}" by auto
ultimately have "{y} \<sqsubset>' A" by -(rule psubset_finite_imp_set_less_aux, auto)
with A' show False by(rule set_less_aux_antisym)
qed
have "y < Min A" unfolding not_le[symmetric]
proof
assume "Min A \<le> y"
moreover have "Min A \<noteq> y" using \<open>y \<notin> A\<close> nempty by clarsimp
ultimately have "Min A < Min {y}" by simp
hence "{y} \<sqsubset>' A" by(rule set_less_aux_Min_antimono)(simp_all add: nempty)
with A' show False by(rule set_less_aux_antisym)
qed
with pi nempty have "x \<le> Min A" by(auto simp add: proper_interval_simps)
moreover
from A obtain z where z: "z \<in> A" "z \<notin> set (x # xs)"
by(auto simp add: set_less_aux_def)
with \<open>x \<le> Min A\<close> nempty have "x < z" by auto
with z have "\<not> exhaustive_above x xs"
by(auto simp add: exhaustive_above_iff[OF xs x]) }
ultimately show ?thesis using True False by fastforce
next
case False
with \<open>\<not> x < y\<close> have "x = y" by auto
from \<open>\<not> x < y\<close> False
have "proper_interval_set_aux xs ys = (\<exists>A. set xs \<sqsubset>' A \<and> A \<sqsubset>' set ys)"
using xs ys by(rule "3.IH")
also have "\<dots> = (\<exists>A. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' set (y # ys))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain A where A: "set xs \<sqsubset>' A"
and A': "A \<sqsubset>' set ys" by blast
hence nempty: "A \<noteq> {}" "ys \<noteq> []" by auto
let ?A = "insert y A"
{ assume "Min A \<le> y"
also from y nempty have "y < Min (set ys)" by auto
finally have "set ys \<sqsubset>' A" by(rule set_less_aux_Min_antimono)(simp_all add: nempty)
with A' have False by(rule set_less_aux_antisym) }
hence MinA: "y < Min A" by(metis not_le)
with nempty have "y \<notin> A" by auto
moreover
with MinA nempty have MinyA: "Min ?A = y" by -(rule Min_eqI, auto)
ultimately have A1: "set (x # xs) \<sqsubset>' ?A" using \<open>x = y\<close> A Minxxs xnxs
by(subst set_less_aux_rec) simp_all
moreover
have "?A \<sqsubset>' set (y # ys)" using \<open>x = y\<close> MinyA \<open>y \<notin> A\<close> A' Minyys ynys
by(subst set_less_aux_rec) simp_all
ultimately show ?rhs by blast
next
assume "?rhs"
then obtain A where A: "set (x # xs) \<sqsubset>' A"
and A': "A \<sqsubset>' set (y # ys)" by blast
let ?A = "A - {x}"
from A have nempty: "A \<noteq> {}" by auto
{ assume "x < Min A"
hence "Min (set (x # xs)) < Min A" unfolding Minxxs .
hence "A \<sqsubset>' set (x # xs)"
by(rule set_less_aux_Min_antimono) simp_all
with A have False by(rule set_less_aux_antisym) }
moreover
{ assume "Min A < x"
hence "Min A < Min (set (y # ys))" unfolding \<open>x = y\<close> Minyys .
hence "set (y # ys) \<sqsubset>' A" by(rule set_less_aux_Min_antimono)(simp_all add: nempty)
with A' have False by(rule set_less_aux_antisym) }
ultimately have MinA: "Min A = x" by(metis less_linear)
hence "x \<in> A" using nempty by(metis Min_in \<open>finite A\<close>)
from A nempty Minxxs xnxs have "set xs \<sqsubset>' ?A"
by(subst (asm) set_less_aux_rec)(auto simp add: MinA)
moreover from A' \<open>x = y\<close> nempty Minyys MinA ynys have "?A \<sqsubset>' set ys"
by(subst (asm) set_less_aux_rec) simp_all
ultimately show ?lhs by blast
qed
finally show ?thesis using \<open>x = y\<close> by simp
qed
qed
qed
qed
lemma proper_interval_set_Compl_aux:
assumes fin: "finite (UNIV :: 'a set)"
and xs: "sorted xs" "distinct xs"
and ys: "sorted ys" "distinct ys"
shows "proper_interval_set_Compl_aux None 0 xs ys \<longleftrightarrow> (\<exists>A. set xs \<sqsubset>' A \<and> A \<sqsubset>' - set ys)"
proof -
note [simp] = finite_subset[OF subset_UNIV fin]
define above where "above = case_option UNIV (Collect \<circ> less)"
have above_simps [simp]: "above None = UNIV" "\<And>x. above (Some x) = {y. x < y}"
and above_upclosed: "\<And>x y ao. \<lbrakk> x \<in> above ao; x < y \<rbrakk> \<Longrightarrow> y \<in> above ao"
and proper_interval_Some2: "\<And>x ao. proper_interval ao (Some x) \<longleftrightarrow> (\<exists>z\<in>above ao. z < x)"
by(simp_all add: proper_interval_simps above_def split: option.splits)
{ fix ao n
assume "set xs \<subseteq> above ao" "set ys \<subseteq> above ao"
from xs \<open>set xs \<subseteq> above ao\<close> ys \<open>set ys \<subseteq> above ao\<close>
have "proper_interval_set_Compl_aux ao (card (UNIV - above ao)) xs ys \<longleftrightarrow>
(\<exists>A \<subseteq> above ao. set xs \<sqsubset>' A \<and> A \<sqsubset>' - set ys \<inter> above ao)"
proof(induct ao n\<equiv>"card (UNIV - above ao)" xs ys rule: proper_interval_set_Compl_aux.induct)
case (1 ao)
have "card (UNIV - above ao) + 1 < CARD('a) \<longleftrightarrow> (\<exists>A \<subseteq> above ao. A \<noteq> {} \<and> A \<sqsubset>' above ao)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
hence "card (UNIV - (UNIV - above ao)) > 1" by(simp add: card_Diff_subset)
from card_gt_1D[OF this]
obtain x y where above: "x \<in> above ao" "y \<in> above ao"
and neq: "x \<noteq> y" by blast
hence "{x} \<sqsubset>' {x, y} \<inter> above ao"
by(simp_all add: psubsetI psubset_finite_imp_set_less_aux)
also have "\<dots> \<sqsubseteq>' above ao"
by(auto intro: subset_finite_imp_set_less_eq_aux)
finally show ?rhs using above by blast
next
assume ?rhs
then obtain A where nempty: "A \<inter> above ao \<noteq> {}"
and subset: "A \<subseteq> above ao"
and less: "A \<sqsubset>' above ao" by blast
from nempty obtain x where x: "x \<in> A" "x \<in> above ao" by blast
show ?lhs
proof(rule ccontr)
assume "\<not> ?lhs"
moreover have "CARD('a) \<ge> card (UNIV - above ao)" by(rule card_mono) simp_all
moreover from card_Un_disjoint[of "UNIV - above ao" "above ao"]
have "CARD('a) = card (UNIV - above ao) + card (above ao)" by auto
ultimately have "card (above ao) = 1" using x
by(cases "card (above ao)")(auto simp add: not_less_eq less_Suc_eq_le)
with x have "above ao = {x}" unfolding card_eq_1_iff by auto
moreover with x subset have A: "A = {x}" by auto
ultimately show False using less by simp
qed
qed
thus ?case by simp
next
case (2 ao y ys)
note ys = \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close> \<open>set (y # ys) \<subseteq> above ao\<close>
have len_ys: "length ys = card (set ys)"
using ys by(auto simp add: List.card_set intro: sym)
define m where "m = CARD('a) - card (UNIV - above ao)"
have "CARD('a) = card (above ao) + card (UNIV - above ao)"
using card_Un_disjoint[of "above ao" "UNIV - above ao"] by auto
hence m_eq: "m = card (above ao)" unfolding m_def by simp
have "m \<noteq> length ys + 1 \<and> (m = length ys + 2 \<longrightarrow> \<not> proper_interval (Some (last (y # ys))) None) \<longleftrightarrow>
(\<exists>A \<subseteq> above ao. A \<noteq> {} \<and> A \<sqsubset>' - set (y # ys) \<inter> above ao)" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
hence m: "m \<noteq> length ys + 1"
and pi: "m = length ys + 2 \<Longrightarrow> \<not> proper_interval (Some (last (y # ys))) None"
by simp_all
have "length ys + 1 = card (set (y # ys))" using ys len_ys by simp
also have "\<dots> \<le> m" unfolding m_eq by(rule card_mono)(simp, rule ys)
finally have "length ys + 2 \<le> m" using m by simp
show ?rhs
proof(cases "m = length ys + 2")
case True
hence "card (UNIV - (UNIV - above ao) - set (y # ys)) = 1"
using ys len_ys
by(subst card_Diff_subset)(auto simp add: m_def card_Diff_subset)
then obtain z where z: "z \<in> above ao" "z \<noteq> y" "z \<notin> set ys"
unfolding card_eq_1_iff by auto
from True have "\<not> proper_interval (Some (last (y # ys))) None" by(rule pi)
hence "z \<le> last (y # ys)" by(simp add: proper_interval_simps not_less del: last.simps)
moreover have ly: "last (y # ys) \<in> set (y # ys)" by(rule last_in_set) simp
with z have "z \<noteq> last (y # ys)" by auto
ultimately have "z < last (y # ys)" by simp
hence "{last (y # ys)} \<sqsubset>' {z}"
using z ly ys by(auto 4 3 simp add: set_less_aux_def)
also have "\<dots> \<sqsubseteq>' - set (y # ys) \<inter> above ao"
using z by(auto intro: subset_finite_imp_set_less_eq_aux)
also have "{last (y # ys)} \<noteq> {}" using ly ys by blast
moreover have "{last (y # ys)} \<subseteq> above ao" using ys by auto
ultimately show ?thesis by blast
next
case False
with \<open>length ys + 2 \<le> m\<close> ys len_ys
have "card (UNIV - (UNIV - above ao) - set (y # ys)) > 1"
by(subst card_Diff_subset)(auto simp add: card_Diff_subset m_def)
from card_gt_1D[OF this]
obtain x x' where x: "x \<in> above ao" "x \<noteq> y" "x \<notin> set ys"
and x': "x' \<in> above ao" "x' \<noteq> y" "x' \<notin> set ys"
and neq: "x \<noteq> x'" by auto
hence "{x} \<sqsubset>' {x, x'} \<inter> above ao"
by(simp_all add: psubsetI psubset_finite_imp_set_less_aux)
also have "\<dots> \<sqsubseteq>' - set (y # ys) \<inter> above ao" using x x' ys
by(auto intro: subset_finite_imp_set_less_eq_aux)
also have "{x} \<inter> above ao \<noteq> {}" using x by simp
ultimately show ?rhs by blast
qed
next
assume ?rhs
then obtain A where nempty: "A \<noteq> {}"
and less: "A \<sqsubset>' - set (y # ys) \<inter> above ao"
and subset: "A \<subseteq> above ao" by blast
have "card (set (y # ys)) \<le> card (above ao)" using ys(3) by(simp add: card_mono)
hence "length ys + 1 \<le> m" unfolding m_eq using ys by(simp add: len_ys)
have "m \<noteq> length ys + 1"
proof
assume "m = length ys + 1"
hence "card (above ao) \<le> card (set (y # ys))"
unfolding m_eq using ys len_ys by auto
from card_seteq[OF _ _ this] ys have "set (y # ys) = above ao" by simp
with nempty less show False by(auto simp add: set_less_aux_def)
qed
moreover
{ assume "m = length ys + 2"
hence "card (above ao - set (y # ys)) = 1"
using ys len_ys m_eq by(auto simp add: card_Diff_subset)
then obtain z where z: "above ao - set (y # ys) = {z}"
unfolding card_eq_1_iff ..
hence eq_z: "- set (y # ys) \<inter> above ao = {z}" by auto
with less have "A \<sqsubset>' {z}" by simp
have "\<not> proper_interval (Some (last (y # ys))) None"
proof
assume "proper_interval (Some (last (y # ys))) None"
then obtain z' where z': "last (y # ys) < z'"
by(clarsimp simp add: proper_interval_simps)
have "last (y # ys) \<in> set (y # ys)" by(rule last_in_set) simp
with ys z' have "z' \<in> above ao" "z' \<notin> set (y # ys)"
by(auto simp del: last.simps sorted.simps(2) intro: above_upclosed dest: sorted_last)
with eq_z have "z = z'" by fastforce
from z' have "\<And>x. x \<in> set (y # ys) \<Longrightarrow> x < z'" using ys
by(auto dest: sorted_last simp del: sorted.simps(2))
with eq_z \<open>z = z'\<close>
have "\<And>x. x \<in> above ao \<Longrightarrow> x \<le> z'" by(fastforce)
with \<open>A \<sqsubset>' {z}\<close> nempty \<open>z = z'\<close> subset
show False by(auto simp add: set_less_aux_def)
qed }
ultimately show ?lhs by simp
qed
thus ?case by(simp add: length_last_def m_def Let_def)
next
case (3 ao x xs)
note xs = \<open>sorted (x # xs)\<close> \<open>distinct (x # xs)\<close> \<open>set (x # xs) \<subseteq> above ao\<close>
have len_xs: "length xs = card (set xs)"
using xs by(auto simp add: List.card_set intro: sym)
define m where "m = CARD('a) - card (UNIV - above ao)"
have "CARD('a) = card (above ao) + card (UNIV - above ao)"
using card_Un_disjoint[of "above ao" "UNIV - above ao"] by auto
hence m_eq: "m = card (above ao)" unfolding m_def by simp
have "m \<noteq> length xs + 1 \<and> (m = length xs + 2 \<longrightarrow> \<not> proper_interval (Some (last (x # xs))) None) \<longleftrightarrow>
(\<exists>A \<subseteq> above ao. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' above ao)" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
hence m: "m \<noteq> length xs + 1"
and pi: "m = length xs + 2 \<Longrightarrow> \<not> proper_interval (Some (last (x # xs))) None"
by simp_all
have "length xs + 1 = card (set (x # xs))" using xs len_xs by simp
also have "\<dots> \<le> m" unfolding m_eq by(rule card_mono)(simp, rule xs)
finally have "length xs + 2 \<le> m" using m by simp
show ?rhs
proof(cases "m = length xs + 2")
case True
hence "card (UNIV - (UNIV - above ao) - set (x # xs)) = 1"
using xs len_xs
by(subst card_Diff_subset)(auto simp add: m_def card_Diff_subset)
then obtain z where z: "z \<in> above ao" "z \<noteq> x" "z \<notin> set xs"
unfolding card_eq_1_iff by auto
define A where "A = insert z {y. y \<in> set (x # xs) \<and> y < z}"
from True have "\<not> proper_interval (Some (last (x # xs))) None" by(rule pi)
hence "z \<le> last (x # xs)" by(simp add: proper_interval_simps not_less del: last.simps)
moreover have lx: "last (x # xs) \<in> set (x # xs)" by(rule last_in_set) simp
with z have "z \<noteq> last (x # xs)" by auto
ultimately have "z < last (x # xs)" by simp
hence "set (x # xs) \<sqsubset>' A"
using z xs by(auto simp add: A_def set_less_aux_def intro: rev_bexI[where x=z])
moreover have "last (x # xs) \<notin> A" using xs \<open>z < last (x # xs)\<close>
by(auto simp add: A_def simp del: last.simps)
hence "A \<subset> insert (last (x # xs)) A" by blast
hence less': "A \<sqsubset>' insert (last (x # xs)) A"
by(rule psubset_finite_imp_set_less_aux) simp
have "\<dots> \<subseteq> above ao" using xs lx z
by(auto simp del: last.simps simp add: A_def)
hence "insert (last (x # xs)) A \<sqsubseteq>' above ao"
by(auto intro: subset_finite_imp_set_less_eq_aux)
with less' have "A \<sqsubset>' above ao"
by(rule set_less_trans_set_less_eq)
moreover have "A \<subseteq> above ao" using xs z by(auto simp add: A_def)
ultimately show ?thesis by blast
next
case False
with \<open>length xs + 2 \<le> m\<close> xs len_xs
have "card (UNIV - (UNIV - above ao) - set (x # xs)) > 1"
by(subst card_Diff_subset)(auto simp add: card_Diff_subset m_def)
from card_gt_1D[OF this]
obtain y y' where y: "y \<in> above ao" "y \<noteq> x" "y \<notin> set xs"
and y': "y' \<in> above ao" "y' \<noteq> x" "y' \<notin> set xs"
and neq: "y \<noteq> y'" by auto
define A where "A = insert y (set (x # xs) \<inter> above ao)"
hence "set (x # xs) \<subset> A" using y xs by auto
hence "set (x # xs) \<sqsubset>' \<dots>"
by(fastforce intro: psubset_finite_imp_set_less_aux)
moreover have *: "\<dots> \<subset> above ao"
using y y' neq by(auto simp add: A_def)
moreover from * have "A \<sqsubset>' above ao"
by(auto intro: psubset_finite_imp_set_less_aux)
ultimately show ?thesis by blast
qed
next
assume ?rhs
then obtain A where lessA: "set (x # xs) \<sqsubset>' A"
and Aless: "A \<sqsubset>' above ao" and subset: "A \<subseteq> above ao" by blast
have "card (set (x # xs)) \<le> card (above ao)" using xs(3) by(simp add: card_mono)
hence "length xs + 1 \<le> m" unfolding m_eq using xs by(simp add: len_xs)
have "m \<noteq> length xs + 1"
proof
assume "m = length xs + 1"
hence "card (above ao) \<le> card (set (x # xs))"
unfolding m_eq using xs len_xs by auto
from card_seteq[OF _ _ this] xs have "set (x # xs) = above ao" by simp
with lessA Aless show False by(auto dest: set_less_aux_antisym)
qed
moreover
{ assume "m = length xs + 2"
hence "card (above ao - set (x # xs)) = 1"
using xs len_xs m_eq by(auto simp add: card_Diff_subset)
then obtain z where z: "above ao - set (x # xs) = {z}"
unfolding card_eq_1_iff ..
have "\<not> proper_interval (Some (last (x # xs))) None"
proof
assume "proper_interval (Some (last (x # xs))) None"
then obtain z' where z': "last (x # xs) < z'"
by(clarsimp simp add: proper_interval_simps)
have "last (x # xs) \<in> set (x # xs)" by(rule last_in_set) simp
with xs z' have "z' \<in> above ao" "z' \<notin> set (x # xs)"
by(auto simp del: last.simps sorted.simps(2) intro: above_upclosed dest: sorted_last)
with z have "z = z'" by fastforce
from z' have y_less: "\<And>y. y \<in> set (x # xs) \<Longrightarrow> y < z'" using xs
by(auto simp del: sorted.simps(2) dest: sorted_last)
with z \<open>z = z'\<close> have "\<And>y. y \<in> above ao \<Longrightarrow> y \<le> z'" by(fastforce)
from lessA subset obtain y where y: "y \<in> A" "y \<in> above ao" "y \<notin> set (x # xs)"
and min: "\<And>y'. \<lbrakk> y' \<in> set (x # xs); y' \<in> above ao; y' \<notin> A \<rbrakk> \<Longrightarrow> y \<le> y'"
by(auto simp add: set_less_aux_def)
with z \<open>z = z'\<close> have "y = z'" by auto
have "set (x # xs) \<subseteq> A"
proof
fix y'
assume y': "y' \<in> set (x # xs)"
show "y' \<in> A"
proof(rule ccontr)
assume "y' \<notin> A"
from y' xs have "y' \<in> above ao" by auto
with y' have "y \<le> y'" using \<open>y' \<notin> A\<close> by(rule min)
moreover from y' have "y' < z'" by(rule y_less)
ultimately show False using \<open>y = z'\<close> by simp
qed
qed
moreover from z xs have "above ao = insert z (set (x # xs))" by auto
ultimately have "A = above ao" using y \<open>y = z'\<close> \<open>z = z'\<close> subset by auto
with Aless show False by simp
qed }
ultimately show ?lhs by simp
qed
thus ?case by(simp add: length_last_def m_def Let_def del: last.simps)
next
case (4 ao x xs y ys)
note xxs = \<open>sorted (x # xs)\<close> \<open>distinct (x # xs)\<close>
and yys = \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close>
and xxs_above = \<open>set (x # xs) \<subseteq> above ao\<close>
and yys_above = \<open>set (y # ys) \<subseteq> above ao\<close>
from xxs have xs: "sorted xs" "distinct xs" and x_Min: "\<forall>x'\<in>set xs. x < x'"
by(auto simp add: less_le)
from yys have ys: "sorted ys" "distinct ys" and y_Min: "\<forall>y'\<in>set ys. y < y'"
by(auto simp add: less_le)
have len_xs: "length xs = card (set xs)"
using xs by(auto simp add: List.card_set intro: sym)
have len_ys: "length ys = card (set ys)"
using ys by(auto simp add: List.card_set intro: sym)
show ?case
proof(cases "x < y")
case True
have "proper_interval ao (Some x) \<or>
proper_interval_set_Compl_aux (Some x) (card (UNIV - above ao) + 1) xs (y # ys) \<longleftrightarrow>
(\<exists>A \<subseteq> above ao. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' - set (y # ys) \<inter> above ao)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof(cases "proper_interval ao (Some x)")
case True
then obtain z where z: "z \<in> above ao" "z < x"
by(clarsimp simp add: proper_interval_Some2)
moreover with xxs have "\<forall>x'\<in>set xs. z < x'" by(auto)
ultimately have "set (x # xs) \<sqsubset>' {z}"
by(auto simp add: set_less_aux_def intro!: bexI[where x=z])
moreover {
from z yys \<open>x < y\<close> have "z < y" "\<forall>y'\<in>set ys. z < y'"
by(auto)
hence subset: "{z} \<subseteq> - set (y # ys) \<inter> above ao"
using ys \<open>x < y\<close> z by auto
moreover have "x \<in> \<dots>" using yys xxs \<open>x < y\<close> xxs_above by(auto)
ultimately have "{z} \<subset> \<dots>" using \<open>z < x\<close> by fastforce
hence "{z} \<sqsubset>' \<dots>"
by(fastforce intro: psubset_finite_imp_set_less_aux) }
moreover have "{z} \<subseteq> above ao" using z by simp
ultimately have ?rhs by blast
thus ?thesis using True by simp
next
case False
hence above_eq: "above ao = insert x (above (Some x))" using xxs_above
by(auto simp add: proper_interval_Some2 intro: above_upclosed)
moreover have "card (above (Some x)) < CARD('a)"
by(rule psubset_card_mono)(auto)
ultimately have card_eq: "card (UNIV - above ao) + 1 = card (UNIV - above (Some x))"
by(simp add: card_Diff_subset)
from xxs_above x_Min have xs_above: "set xs \<subseteq> above (Some x)" by(auto)
from \<open>x < y\<close> y_Min have "set (y # ys) \<subseteq> above (Some x)" by(auto)
with \<open>x < y\<close> card_eq xs xs_above yys
have "proper_interval_set_Compl_aux (Some x) (card (UNIV - above ao) + 1) xs (y # ys) \<longleftrightarrow>
(\<exists>A \<subseteq> above (Some x). set xs \<sqsubset>' A \<and> A \<sqsubset>' - set (y # ys) \<inter> above (Some x))"
by(subst card_eq)(rule 4)
also have "\<dots> \<longleftrightarrow> ?rhs" (is "?lhs' \<longleftrightarrow> _")
proof
assume ?lhs'
then obtain A where less_A: "set xs \<sqsubset>' A"
and A_less: "A \<sqsubset>' - set (y # ys) \<inter> above (Some x)"
and subset: "A \<subseteq> above (Some x)" by blast
let ?A = "insert x A"
have Min_A': "Min ?A = x" using xxs_above False subset
by(auto intro!: Min_eqI simp add: proper_interval_Some2)
moreover have "Min (set (x # xs)) = x"
using x_Min by(auto intro!: Min_eqI)
moreover have Amx: "A - {x} = A"
using False subset
by(auto simp add: proper_interval_Some2 intro: above_upclosed)
moreover have "set xs - {x} = set xs" using x_Min by auto
ultimately have less_A': "set (x # xs) \<sqsubset>' ?A"
using less_A xxs_above x_Min by(subst set_less_aux_rec) simp_all
have "x \<in> - insert y (set ys) \<inter> above ao"
using \<open>x < y\<close> xxs_above y_Min by auto
hence "- insert y (set ys) \<inter> above ao \<noteq> {}" by auto
moreover have "Min (- insert y (set ys) \<inter> above ao) = x"
using yys y_Min xxs_above \<open>x < y\<close> False
by(auto intro!: Min_eqI simp add: proper_interval_Some2)
moreover have "- set (y # ys) \<inter> above ao - {x} = - set (y # ys) \<inter> above (Some x)"
using yys_above False xxs_above
by(auto simp add: proper_interval_Some2 intro: above_upclosed)
ultimately have A'_less: "?A \<sqsubset>' - set (y # ys) \<inter> above ao"
using Min_A' A_less Amx xxs_above by(subst set_less_aux_rec) simp_all
moreover have "?A \<subseteq> above ao" using subset xxs_above by(auto intro: above_upclosed)
ultimately show ?rhs using less_A' by blast
next
assume ?rhs
then obtain A where less_A: "set (x # xs) \<sqsubset>' A"
and A_less: "A \<sqsubset>' - set (y # ys) \<inter> above ao"
and subset: "A \<subseteq> above ao" by blast
let ?A = "A - {x}"
from less_A subset xxs_above have "set (x # xs) \<inter> above ao \<sqsubset>' A \<inter> above ao"
by(simp add: Int_absorb2)
with False xxs_above subset have "x \<in> A"
by(auto simp add: set_less_aux_def proper_interval_Some2)
hence "\<dots> \<noteq> {}" by auto
moreover from \<open>x \<in> A\<close> False subset
have Min_A: "Min A = x"
by(auto intro!: Min_eqI simp add: proper_interval_Some2 not_less)
moreover have "Min (set (x # xs)) = x"
using x_Min by(auto intro!: Min_eqI)
moreover have eq_A: "?A \<subseteq> above (Some x)"
using xxs_above False subset
by(fastforce simp add: proper_interval_Some2 not_less intro: above_upclosed)
moreover have "set xs - {x} = set xs"
using x_Min by(auto)
ultimately have less_A': "set xs \<sqsubset>' ?A"
using xxs_above less_A by(subst (asm) set_less_aux_rec)(simp_all cong: conj_cong)
have "x \<in> - insert y (set ys) \<inter> above ao"
using \<open>x < y\<close> xxs_above y_Min by auto
hence "- insert y (set ys) \<inter> above ao \<noteq> {}" by auto
moreover have "Min (- set (y # ys) \<inter> above ao) = x"
using yys y_Min xxs_above \<open>x < y\<close> False
by(auto intro!: Min_eqI simp add: proper_interval_Some2)
moreover have "- set (y # ys) \<inter> above (Some x) = - set (y # ys) \<inter> above ao - {x}"
by(auto simp add: above_eq)
ultimately have "?A \<sqsubset>' - set (y # ys) \<inter> above (Some x)"
using A_less \<open>A \<noteq> {}\<close> eq_A Min_A
by(subst (asm) set_less_aux_rec) simp_all
with less_A' eq_A show ?lhs' by blast
qed
finally show ?thesis using False by simp
qed
thus ?thesis using True by simp
next
case False
show ?thesis
proof(cases "y < x")
case True
have "proper_interval ao (Some y) \<or>
proper_interval_set_Compl_aux (Some y) (card (UNIV - above ao) + 1) (x # xs) ys \<longleftrightarrow>
(\<exists>A \<subseteq> above ao. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' - set (y # ys) \<inter> above ao)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof(cases "proper_interval ao (Some y)")
case True
then obtain z where z: "z \<in> above ao" "z < y"
by(clarsimp simp add: proper_interval_Some2)
from xxs \<open>y < x\<close> have "\<forall>x'\<in>set (x # xs). y < x'" by(auto)
hence less_A: "set (x # xs) \<sqsubset>' {y}"
by(auto simp add: set_less_aux_def intro!: bexI[where x=y])
have "{y} \<sqsubset>' {z}"
using z y_Min by(auto simp add: set_less_aux_def intro: bexI[where x=z])
also have "\<dots> \<subseteq> - set (y # ys) \<inter> above ao" using z y_Min by auto
hence "{z} \<sqsubseteq>' \<dots>" by(auto intro: subset_finite_imp_set_less_eq_aux)
finally have "{y} \<sqsubset>' \<dots>" .
moreover have "{y} \<subseteq> above ao" using yys_above by auto
ultimately have ?rhs using less_A by blast
thus ?thesis using True by simp
next
case False
hence above_eq: "above ao = insert y (above (Some y))" using yys_above
by(auto simp add: proper_interval_Some2 intro: above_upclosed)
moreover have "card (above (Some y)) < CARD('a)"
by(rule psubset_card_mono)(auto)
ultimately have card_eq: "card (UNIV - above ao) + 1 = card (UNIV - above (Some y))"
by(simp add: card_Diff_subset)
from yys_above y_Min have ys_above: "set ys \<subseteq> above (Some y)" by(auto)
have eq_ys: "- set ys \<inter> above (Some y) = - set (y # ys) \<inter> above ao"
by(auto simp add: above_eq)
from \<open>y < x\<close> x_Min have "set (x # xs) \<subseteq> above (Some y)" by(auto)
with \<open>\<not> x < y\<close> \<open>y < x\<close> card_eq xxs ys ys_above
have "proper_interval_set_Compl_aux (Some y) (card (UNIV - above ao) + 1) (x # xs) ys \<longleftrightarrow>
(\<exists>A \<subseteq> above (Some y). set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' - set ys \<inter> above (Some y))"
by(subst card_eq)(rule 4)
also have "\<dots> \<longleftrightarrow> ?rhs" (is "?lhs' \<longleftrightarrow> _")
proof
assume ?lhs'
then obtain A where "set (x # xs) \<sqsubset>' A" and subset: "A \<subseteq> above (Some y)"
and "A \<sqsubset>' - set ys \<inter> above (Some y)" by blast
moreover from subset have "A - {y} = A" by auto
ultimately have "set (x # xs) \<sqsubset>' A - {y}"
and "A - {y} \<sqsubset>' - set (y # ys) \<inter> above ao"
using eq_ys by simp_all
moreover from subset have "A - {y} \<subseteq> above ao"
using yys_above by(auto intro: above_upclosed)
ultimately show ?rhs by blast
next
assume ?rhs
then obtain A where "set (x # xs) \<sqsubset>' A"
and A_less: "A \<sqsubset>' - set (y # ys) \<inter> above ao"
and subset: "A \<subseteq> above ao" by blast
moreover
from A_less False yys_above have "y \<notin> A"
by(auto simp add: set_less_aux_def proper_interval_Some2 not_less)
ultimately have "set (x # xs) \<sqsubset>' A"
and "A \<sqsubset>' - set ys \<inter> above (Some y)"
using eq_ys by simp_all
moreover from \<open>y \<notin> A\<close> subset above_eq have "A \<subseteq> above (Some y)" by auto
ultimately show ?lhs' by blast
qed
finally show ?thesis using False by simp
qed
with False True show ?thesis by simp
next
case False
with \<open>\<not> x < y\<close> have "x = y" by simp
have "proper_interval ao (Some x) \<and>
(CARD('a) - (card (UNIV - above ao) + length ys) \<noteq> 2 \<or>
CARD('a) - (card (UNIV - above ao) + length xs) \<noteq> 2) \<longleftrightarrow>
(\<exists>A \<subseteq> above ao. set (x # xs) \<sqsubset>' A \<and> A \<sqsubset>' - set (y # ys) \<inter> above ao)"
(is "?below \<and> ?card \<longleftrightarrow> ?rhs")
proof(cases "?below")
case False
hence "- set (y # ys) \<inter> above ao \<sqsubset>' set (x # xs)"
using \<open>x = y\<close> yys_above xxs_above y_Min
by(auto simp add: not_less set_less_aux_def proper_interval_Some2 intro!: bexI[where x=y])
with False show ?thesis by(auto dest: set_less_aux_trans)
next
case True
then obtain z where z: "z \<in> above ao" "z < x"
by(clarsimp simp add: proper_interval_Some2)
have "?card \<longleftrightarrow> ?rhs"
proof
assume ?rhs
then obtain A where less_A: "set (x # xs) \<sqsubset>' A"
and A_less: "A \<sqsubset>' - set (y # ys) \<inter> above ao"
and subset: "A \<subseteq> above ao" by blast
{
assume c_ys: "CARD('a) - (card (UNIV - above ao) + length ys) = 2"
and c_xs: "CARD('a) - (card (UNIV - above ao) + length xs) = 2"
from c_ys yys_above len_ys y_Min have "card (UNIV - (UNIV - above ao) - set (y # ys)) = 1"
by(subst card_Diff_subset)(auto simp add: card_Diff_subset)
then obtain z' where eq_y: "- set (y # ys) \<inter> above ao = {z'}"
unfolding card_eq_1_iff by auto
moreover from z have "z \<notin> set (y # ys)" using \<open>x = y\<close> y_Min by auto
ultimately have "z' = z" using z by fastforce
from c_xs xxs_above len_xs x_Min have "card (UNIV - (UNIV - above ao) - set (x # xs)) = 1"
by(subst card_Diff_subset)(auto simp add: card_Diff_subset)
then obtain z'' where eq_x: "- set (x # xs) \<inter> above ao = {z''}"
unfolding card_eq_1_iff by auto
moreover from z have "z \<notin> set (x # xs)" using x_Min by auto
ultimately have "z'' = z" using z by fastforce
from less_A subset obtain q where "q \<in> A" "q \<in> above ao" "q \<notin> set (x # xs)"
by(auto simp add: set_less_aux_def)
hence "q \<in> {z''}" unfolding eq_x[symmetric] by simp
hence "q = z''" by simp
with \<open>q \<in> A\<close> \<open>z' = z\<close> \<open>z'' = z\<close> z
have "- set (y # ys) \<inter> above ao \<subseteq> A"
unfolding eq_y by simp
hence "- set (y # ys) \<inter> above ao \<sqsubseteq>' A"
by(auto intro: subset_finite_imp_set_less_eq_aux)
with A_less have False by(auto dest: set_less_trans_set_less_eq) }
thus ?card by auto
next
assume ?card (is "?card_ys \<or> ?card_xs")
thus ?rhs
proof
assume ?card_ys
let ?YS = "UNIV - (UNIV - above ao) - set (y # ys)"
from \<open>?card_ys\<close> yys_above len_ys y_Min have "card ?YS \<noteq> 1"
by(subst card_Diff_subset)(auto simp add: card_Diff_subset)
moreover have "?YS \<noteq> {}" using True y_Min yys_above \<open>x = y\<close>
by(fastforce simp add: proper_interval_Some2)
hence "card ?YS \<noteq> 0" by simp
ultimately have "card ?YS > 1" by(cases "card ?YS") simp_all
from card_gt_1D[OF this] obtain x' y'
where x': "x' \<in> above ao" "x' \<notin> set (y # ys)"
and y': "y' \<in> above ao" "y' \<notin> set (y # ys)"
and neq: "x' \<noteq> y'" by auto
let ?A = "{z}"
have "set (x # xs) \<sqsubset>' ?A" using z x_Min
by(auto simp add: set_less_aux_def intro!: rev_bexI)
moreover
{ have "?A \<subseteq> - set (y # ys) \<inter> above ao"
using z \<open>x = y\<close> y_Min by(auto)
moreover have "x' \<notin> ?A \<or> y' \<notin> ?A" using neq by auto
with x' y' have "?A \<noteq> - set (y # ys) \<inter> above ao" by auto
ultimately have "?A \<subset> - set (y # ys) \<inter> above ao" by(rule psubsetI)
hence "?A \<sqsubset>' \<dots>" by(fastforce intro: psubset_finite_imp_set_less_aux) }
ultimately show ?thesis using z by blast
next
assume ?card_xs
let ?XS = "UNIV - (UNIV - above ao) - set (x # xs)"
from \<open>?card_xs\<close> xxs_above len_xs x_Min have "card ?XS \<noteq> 1"
by(subst card_Diff_subset)(auto simp add: card_Diff_subset)
moreover have "?XS \<noteq> {}" using True x_Min xxs_above
by(fastforce simp add: proper_interval_Some2)
hence "card ?XS \<noteq> 0" by simp
ultimately have "card ?XS > 1" by(cases "card ?XS") simp_all
from card_gt_1D[OF this] obtain x' y'
where x': "x' \<in> above ao" "x' \<notin> set (x # xs)"
and y': "y' \<in> above ao" "y' \<notin> set (x # xs)"
and neq: "x' \<noteq> y'" by auto
define A
where "A = (if x' = Min (above ao) then insert y' (set (x # xs)) else insert x' (set (x # xs)))"
hence "set (x # xs) \<subseteq> A" by auto
moreover have "set (x # xs) \<noteq> \<dots>"
using neq x' y' by(auto simp add: A_def)
ultimately have "set (x # xs) \<subset> A" ..
hence "set (x # xs) \<sqsubset>' \<dots>"
by(fastforce intro: psubset_finite_imp_set_less_aux)
moreover {
have nempty: "above ao \<noteq> {}" using z by auto
have "A \<sqsubset>' {Min (above ao)}"
using z x' y' neq \<open>x = y\<close> x_Min xxs_above
by(auto 6 4 simp add: set_less_aux_def A_def nempty intro!: rev_bexI Min_eqI)
also have "Min (above ao) \<le> z" using z by(simp)
hence "Min (above ao) < x" using \<open>z < x\<close> by(rule le_less_trans)
with \<open>x = y\<close> y_Min have "Min (above ao) \<notin> set (y # ys)" by auto
hence "{Min (above ao)} \<subseteq> - set (y # ys) \<inter> above ao"
by(auto simp add: nempty)
hence "{Min (above ao)} \<sqsubseteq>' \<dots>" by(auto intro: subset_finite_imp_set_less_eq_aux)
finally have "A \<sqsubset>' \<dots>" . }
moreover have "A \<subseteq> above ao" using xxs_above yys_above x' y'
by(auto simp add: A_def)
ultimately show ?rhs by blast
qed
qed
thus ?thesis using True by simp
qed
thus ?thesis using \<open>x = y\<close> by simp
qed
qed
qed }
from this[of None]
show ?thesis by(simp)
qed
lemma proper_interval_Compl_set_aux:
assumes fin: "finite (UNIV :: 'a set)"
and xs: "sorted xs" "distinct xs"
and ys: "sorted ys" "distinct ys"
shows "proper_interval_Compl_set_aux None xs ys \<longleftrightarrow> (\<exists>A. - set xs \<sqsubset>' A \<and> A \<sqsubset>' set ys)"
proof -
note [simp] = finite_subset[OF subset_UNIV fin]
define above where "above = case_option UNIV (Collect \<circ> less)"
have above_simps [simp]: "above None = UNIV" "\<And>x. above (Some x) = {y. x < y}"
and above_upclosed: "\<And>x y ao. \<lbrakk> x \<in> above ao; x < y \<rbrakk> \<Longrightarrow> y \<in> above ao"
and proper_interval_Some2: "\<And>x ao. proper_interval ao (Some x) \<longleftrightarrow> (\<exists>z\<in>above ao. z < x)"
by(simp_all add: proper_interval_simps above_def split: option.splits)
{ fix ao n
assume "set xs \<subseteq> above ao" "set ys \<subseteq> above ao"
from xs \<open>set xs \<subseteq> above ao\<close> ys \<open>set ys \<subseteq> above ao\<close>
have "proper_interval_Compl_set_aux ao xs ys \<longleftrightarrow>
(\<exists>A. - set xs \<inter> above ao \<sqsubset>' A \<inter> above ao \<and> A \<inter> above ao \<sqsubset>' set ys \<inter> above ao)"
proof(induction ao xs ys rule: proper_interval_Compl_set_aux.induct)
case ("2_1" ao ys)
{ fix A
assume "above ao \<sqsubset>' A \<inter> above ao"
also have "\<dots> \<subseteq> above ao" by simp
hence "A \<inter> above ao \<sqsubseteq>' above ao"
by(auto intro: subset_finite_imp_set_less_eq_aux)
finally have False by simp }
thus ?case by auto
next
case ("2_2" ao xs) thus ?case by simp
next
case (1 ao x xs y ys)
note xxs = \<open>sorted (x # xs)\<close> \<open>distinct (x # xs)\<close>
hence xs: "sorted xs" "distinct xs" and x_Min: "\<forall>x' \<in> set xs. x < x'"
by(auto simp add: less_le)
note yys = \<open>sorted (y # ys)\<close> \<open>distinct (y # ys)\<close>
hence ys: "sorted ys" "distinct ys" and y_Min: "\<forall>y'\<in>set ys. y < y'"
by(auto simp add: less_le)
note xxs_above = \<open>set (x # xs) \<subseteq> above ao\<close>
note yys_above = \<open>set (y # ys) \<subseteq> above ao\<close>
show ?case
proof(cases "x < y")
case True
have "\<not> proper_interval ao (Some x) \<and> proper_interval_Compl_set_aux (Some x) xs (y # ys) \<longleftrightarrow>
(\<exists>A. - set (x # xs) \<inter> above ao \<sqsubset>' A \<inter> above ao \<and> A \<inter> above ao \<sqsubset>' set (y # ys) \<inter> above ao)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof(cases "proper_interval ao (Some x)")
case True
then obtain z where z: "z < x" "z \<in> above ao"
by(auto simp add: proper_interval_Some2)
hence nempty: "above ao \<noteq> {}" by auto
with z have "Min (above ao) \<le> z" by auto
hence "Min (above ao) < x" using \<open>z < x\<close> by(rule le_less_trans)
hence "set (y # ys) \<inter> above ao \<sqsubset>' - set (x # xs) \<inter> above ao"
using y_Min x_Min z \<open>x < y\<close>
by(fastforce simp add: set_less_aux_def nempty intro!: Min_eqI bexI[where x="Min (above ao)"])
thus ?thesis using True by(auto dest: set_less_aux_trans set_less_aux_antisym)
next
case False
hence above_eq: "above ao = insert x (above (Some x))"
using xxs_above by(auto simp add: proper_interval_Some2 intro: above_upclosed)
from x_Min have xs_above: "set xs \<subseteq> above (Some x)" by auto
from \<open>x < y\<close> y_Min have ys_above: "set (y # ys) \<subseteq> above (Some x)" by auto
have eq_xs: "- set xs \<inter> above (Some x) = - set (x # xs) \<inter> above ao"
using above_eq by auto
have eq_ys: "set (y # ys) \<inter> above (Some x) = set (y # ys) \<inter> above ao"
using y_Min \<open>x < y\<close> xxs_above by(auto intro: above_upclosed)
from \<open>x < y\<close> xs xs_above yys ys_above
have "proper_interval_Compl_set_aux (Some x) xs (y # ys) \<longleftrightarrow>
(\<exists>A. - set xs \<inter> above (Some x) \<sqsubset>' A \<inter> above (Some x) \<and>
A \<inter> above (Some x) \<sqsubset>' set (y # ys) \<inter> above (Some x))"
by(rule "1.IH")
also have "\<dots> \<longleftrightarrow> ?rhs" (is "?lhs \<longleftrightarrow> _")
proof
assume ?lhs
then obtain A where "- set xs \<inter> above (Some x) \<sqsubset>' A \<inter> above (Some x)"
and "A \<inter> above (Some x) \<sqsubset>' set (y # ys) \<inter> above (Some x)" by blast
moreover have "A \<inter> above (Some x) = (A - {x}) \<inter> above ao"
using above_eq by auto
ultimately have "- set (x # xs) \<inter> above ao \<sqsubset>' (A - {x}) \<inter> above ao"
and "(A - {x}) \<inter> above ao \<sqsubset>' set (y # ys) \<inter> above ao"
using eq_xs eq_ys by simp_all
thus ?rhs by blast
next
assume ?rhs
then obtain A where "- set (x # xs) \<inter> above ao \<sqsubset>' A \<inter> above ao"
and A_less: "A \<inter> above ao \<sqsubset>' set (y # ys) \<inter> above ao" by blast
moreover have "x \<notin> A"
proof
assume "x \<in> A"
hence "set (y # ys) \<inter> above ao \<sqsubset>' A \<inter> above ao"
using y_Min \<open>x < y\<close> by(auto simp add: above_eq set_less_aux_def intro!: bexI[where x=x])
with A_less show False by(auto dest: set_less_aux_antisym)
qed
hence "A \<inter> above ao = A \<inter> above (Some x)" using above_eq by auto
ultimately show ?lhs using eq_xs eq_ys by auto
qed
finally show ?thesis using False by simp
qed
thus ?thesis using True by simp
next
case False
show ?thesis
proof(cases "y < x")
case True
show ?thesis (is "?lhs \<longleftrightarrow> ?rhs")
proof(cases "proper_interval ao (Some y)")
case True
then obtain z where z: "z < y" "z \<in> above ao"
by(auto simp add: proper_interval_Some2)
hence nempty: "above ao \<noteq> {}" by auto
with z have "Min (above ao) \<le> z" by auto
hence "Min (above ao) < y" using \<open>z < y\<close> by(rule le_less_trans)
hence "set (y # ys) \<inter> above ao \<sqsubset>' - set (x # xs) \<inter> above ao"
using y_Min x_Min z \<open>y < x\<close>
by(fastforce simp add: set_less_aux_def nempty intro!: Min_eqI bexI[where x="Min (above ao)"])
thus ?thesis using True \<open>y < x\<close> by(auto dest: set_less_aux_trans set_less_aux_antisym)
next
case False
hence above_eq: "above ao = insert y (above (Some y))"
using yys_above by(auto simp add: proper_interval_Some2 intro: above_upclosed)
from y_Min have ys_above: "set ys \<subseteq> above (Some y)" by auto
from \<open>y < x\<close> x_Min have xs_above: "set (x # xs) \<subseteq> above (Some y)" by auto
have "y \<in> - set (x # xs) \<inter> above ao" using \<open>y < x\<close> x_Min yys_above by auto
hence nempty: "- set (x # xs) \<inter> above ao \<noteq> {}" by auto
have Min_x: "Min (- set (x # xs) \<inter> above ao) = y"
using above_eq \<open>y < x\<close> x_Min by(auto intro!: Min_eqI)
have Min_y: "Min (set (y # ys) \<inter> above ao) = y"
using y_Min above_eq by(auto intro!: Min_eqI)
have eq_xs: "- set (x # xs) \<inter> above ao - {y} = - set (x # xs) \<inter> above (Some y)"
by(auto simp add: above_eq)
have eq_ys: "set ys \<inter> above ao - {y} = set ys \<inter> above (Some y)"
using y_Min above_eq by auto
from \<open>\<not> x < y\<close> \<open>y < x\<close> xxs xs_above ys ys_above
have "proper_interval_Compl_set_aux (Some y) (x # xs) ys \<longleftrightarrow>
(\<exists>A. - set (x # xs) \<inter> above (Some y) \<sqsubset>' A \<inter> above (Some y) \<and>
A \<inter> above (Some y) \<sqsubset>' set ys \<inter> above (Some y))"
by(rule "1.IH")
also have "\<dots> \<longleftrightarrow> ?rhs" (is "?lhs' \<longleftrightarrow> _")
proof
assume ?lhs'
then obtain A where less_A: "- set (x # xs) \<inter> above (Some y) \<sqsubset>' A \<inter> above (Some y)"
and A_less: "A \<inter> above (Some y) \<sqsubset>' set ys \<inter> above (Some y)" by blast
let ?A = "insert y A"
have Min_A: "Min (?A \<inter> above ao) = y"
using above_eq by(auto intro!: Min_eqI)
moreover have A_eq: "A \<inter> above ao - {y} = A \<inter> above (Some y)"
using above_eq by auto
ultimately have less_A': "- set (x # xs) \<inter> above ao \<sqsubset>' ?A \<inter> above ao"
using nempty yys_above less_A Min_x eq_xs by(subst set_less_aux_rec) simp_all
have A'_less: "?A \<inter> above ao \<sqsubset>' set (y # ys) \<inter> above ao"
using yys_above nempty Min_A A_eq A_less Min_y eq_ys
by(subst set_less_aux_rec) simp_all
with less_A' show ?rhs by blast
next
assume ?rhs
then obtain A where less_A: "- set (x # xs) \<inter> above ao \<sqsubset>' A \<inter> above ao"
and A_less: "A \<inter> above ao \<sqsubset>' set (y # ys) \<inter> above ao" by blast
from less_A have nempty': "A \<inter> above ao \<noteq> {}" by auto
moreover have A_eq: "A \<inter> above ao - {y} = A \<inter> above (Some y)"
using above_eq by auto
moreover have y_in_xxs: "y \<in> - set (x # xs) \<inter> above ao"
using \<open>y < x\<close> x_Min yys_above by auto
moreover have "y \<in> A"
proof(rule ccontr)
assume "y \<notin> A"
hence "A \<inter> above ao \<sqsubset>' - set (x # xs) \<inter> above ao"
using \<open>y < x\<close> x_Min y_in_xxs
by(auto simp add: set_less_aux_def above_eq intro: bexI[where x=y])
with less_A show False by(rule set_less_aux_antisym)
qed
hence Min_A: "Min (A \<inter> above ao) = y" using above_eq y_Min by(auto intro!: Min_eqI)
ultimately have less_A': "- set (x # xs) \<inter> above (Some y) \<sqsubset>' A \<inter> above (Some y)"
using nempty less_A Min_x eq_xs
by(subst (asm) set_less_aux_rec)(auto dest: bspec[where x=y])
have A'_less: "A \<inter> above (Some y) \<sqsubset>' set ys \<inter> above (Some y)"
using A_less nempty' yys_above Min_A Min_y A_eq eq_ys
by(subst (asm) set_less_aux_rec) simp_all
with less_A' show ?lhs' by blast
qed
finally show ?thesis using \<open>\<not> x < y\<close> \<open>y < x\<close> False by simp
qed
next
case False
with \<open>\<not> x < y\<close> have "x = y" by auto
thus ?thesis (is "?lhs \<longleftrightarrow> ?rhs")
proof(cases "proper_interval ao (Some x)")
case True
then obtain z where z: "z < x" "z \<in> above ao"
by(auto simp add: proper_interval_Some2)
hence nempty: "above ao \<noteq> {}" by auto
with z have "Min (above ao) \<le> z" by auto
hence "Min (above ao) < x" using \<open>z < x\<close> by(rule le_less_trans)
hence "set (y # ys) \<inter> above ao \<sqsubset>' - set (x # xs) \<inter> above ao"
using y_Min x_Min z \<open>x = y\<close>
by(fastforce simp add: set_less_aux_def nempty intro!: Min_eqI bexI[where x="Min (above ao)"])
thus ?thesis using True \<open>x = y\<close> by(auto dest: set_less_aux_trans set_less_aux_antisym)
next
case False
hence above_eq: "above ao = insert x (above (Some x))"
using xxs_above by(auto simp add: proper_interval_Some2 intro: above_upclosed)
have "(ys = [] \<longrightarrow> xs \<noteq> []) \<longleftrightarrow> ?rhs" (is "?lhs' \<longleftrightarrow> _")
proof(intro iffI strip notI)
assume ?lhs'
show ?rhs
proof(cases ys)
case Nil
with \<open>?lhs'\<close> obtain x' xs' where xs_eq: "xs = x' # xs'"
by(auto simp add: neq_Nil_conv)
with xs have x'_Min: "\<forall>x'' \<in> set xs'. x' < x''"
by(auto simp add: less_le)
let ?A = "- set (x # xs')"
have "- set (x # xs) \<inter> above ao \<subseteq> ?A \<inter> above ao"
using xs_eq by auto
moreover have "x' \<notin> - set (x # xs) \<inter> above ao" "x' \<in> ?A \<inter> above ao"
using xs_eq xxs_above x'_Min x_Min by auto
ultimately have "- set (x # xs) \<inter> above ao \<subset> ?A \<inter> above ao"
by blast
hence "- set (x # xs) \<inter> above ao \<sqsubset>' \<dots> "
by(fastforce intro: psubset_finite_imp_set_less_aux)
moreover have "\<dots> \<sqsubset>' set (y # ys) \<inter> above ao"
using Nil \<open>x = y\<close> by(auto simp add: set_less_aux_def above_eq)
ultimately show ?thesis by blast
next
case (Cons y' ys')
let ?A = "{y}"
have "- set (x # xs) \<inter> above ao \<sqsubset>' ?A \<inter> above ao"
using \<open>x = y\<close> x_Min by(auto simp add: set_less_aux_def above_eq)
moreover have "\<dots> \<subset> set (y # ys) \<inter> above ao"
using yys_above yys Cons by auto
hence "?A \<inter> above ao \<sqsubset>' \<dots>" by(fastforce intro: psubset_finite_imp_set_less_aux)
ultimately show ?thesis by blast
qed
next
assume Nil: "ys = []" "xs = []" and ?rhs
then obtain A where less_A: "- {x} \<inter> above ao \<sqsubset>' A \<inter> above ao"
and A_less: "A \<inter> above ao \<sqsubset>' {x}" using \<open>x = y\<close> above_eq by auto
have "x \<notin> A" using A_less by(auto simp add: set_less_aux_def above_eq)
hence "A \<inter> above ao \<subseteq> - {x} \<inter> above ao" by auto
hence "A \<inter> above ao \<sqsubseteq>' \<dots>" by(auto intro: subset_finite_imp_set_less_eq_aux)
with less_A have "\<dots> \<sqsubset>' \<dots>" by(rule set_less_trans_set_less_eq)
thus False by simp
qed
with \<open>x = y\<close> False show ?thesis by simp
qed
qed
qed
qed }
from this[of None] show ?thesis by simp
qed
end
subsection \<open>Proper intervals for HOL types\<close>
instantiation unit :: proper_interval begin
fun proper_interval_unit :: "unit proper_interval" where
"proper_interval_unit None None = True"
| "proper_interval_unit _ _ = False"
instance by intro_classes auto
end
instantiation bool :: proper_interval begin
fun proper_interval_bool :: "bool proper_interval" where
"proper_interval_bool (Some x) (Some y) \<longleftrightarrow> False"
| "proper_interval_bool (Some x) None \<longleftrightarrow> \<not> x"
| "proper_interval_bool None (Some y) \<longleftrightarrow> y"
| "proper_interval_bool None None = True"
instance by intro_classes auto
end
instantiation nat :: proper_interval begin
fun proper_interval_nat :: "nat proper_interval" where
"proper_interval_nat no None = True"
| "proper_interval_nat None (Some x) \<longleftrightarrow> x > 0"
| "proper_interval_nat (Some x) (Some y) \<longleftrightarrow> y - x > 1"
instance by intro_classes auto
end
instantiation int :: proper_interval begin
fun proper_interval_int :: "int proper_interval" where
"proper_interval_int (Some x) (Some y) \<longleftrightarrow> y - x > 1"
| "proper_interval_int _ _ = True"
instance by intro_classes (auto intro: less_add_one, metis less_add_one minus_less_iff)
end
instantiation integer :: proper_interval begin
context includes integer.lifting begin
lift_definition proper_interval_integer :: "integer proper_interval" is "proper_interval" .
instance by(intro_classes)(transfer, simp only: proper_interval_simps)+
end
end
lemma proper_interval_integer_simps [code]:
includes integer.lifting fixes x y :: integer and xo yo :: "integer option" shows
"proper_interval (Some x) (Some y) = (1 < y - x)"
"proper_interval None yo = True"
"proper_interval xo None = True"
by(transfer, simp)+
instantiation natural :: proper_interval begin
context includes natural.lifting begin
lift_definition proper_interval_natural :: "natural proper_interval" is "proper_interval" .
instance by(intro_classes)(transfer, simp only: proper_interval_simps)+
end
end
lemma proper_interval_natural_simps [code]:
includes natural.lifting fixes x y :: natural and xo :: "natural option" shows
"proper_interval xo None = True"
"proper_interval None (Some y) \<longleftrightarrow> y > 0"
"proper_interval (Some x) (Some y) \<longleftrightarrow> y - x > 1"
by(transfer, simp)+
lemma char_less_iff_nat_of_char: "x < y \<longleftrightarrow> of_char x < (of_char y :: nat)"
by (fact less_char_def)
lemma nat_of_char_inject [simp]: "of_char x = (of_char y :: nat) \<longleftrightarrow> x = y"
by (fact of_char_eq_iff)
lemma char_le_iff_nat_of_char: "x \<le> y \<longleftrightarrow> of_char x \<le> (of_char y :: nat)"
by (fact less_eq_char_def)
instantiation char :: proper_interval
begin
fun proper_interval_char :: "char proper_interval" where
"proper_interval_char None None \<longleftrightarrow> True"
| "proper_interval_char None (Some x) \<longleftrightarrow> x \<noteq> CHR 0x00"
| "proper_interval_char (Some x) None \<longleftrightarrow> x \<noteq> CHR 0xFF"
| "proper_interval_char (Some x) (Some y) \<longleftrightarrow> of_char y - of_char x > (1 :: nat)"
instance proof
fix y :: char
{ assume "y \<noteq> CHR 0x00"
have "CHR 0x00 < y"
proof (rule ccontr)
assume "\<not> CHR 0x00 < y"
then have "of_char y = (of_char CHR 0x00 :: nat)"
by (simp add: not_less char_le_iff_nat_of_char)
then have "y = CHR 0x00"
using nat_of_char_inject [of y "CHR 0x00"] by simp
with \<open>y \<noteq> CHR 0x00\<close> show False
by simp
qed }
moreover
{ fix z :: char
assume "z < CHR 0x00"
hence False
by (simp add: char_less_iff_nat_of_char of_char_eq_iff [symmetric]) }
ultimately show "proper_interval None (Some y) = (\<exists>z. z < y)"
by auto
fix x :: char
{ assume "x \<noteq> CHR 0xFF"
then have "x < CHR 0xFF"
by (auto simp add: neq_iff char_less_iff_nat_of_char)
(insert nat_of_char_less_256 [of x], simp)
hence "\<exists>z. x < z" .. }
moreover
{ fix z :: char
assume "CHR 0xFF < z"
hence "False"
by (simp add: char_less_iff_nat_of_char)
(insert nat_of_char_less_256 [of z], simp) }
ultimately show "proper_interval (Some x) None = (\<exists>z. x < z)" by auto
{ assume gt: "of_char y - of_char x > (1 :: nat)"
let ?z = "char_of (of_char x + (1 :: nat))"
from gt nat_of_char_less_256 [of y]
have 255: "of_char x < (255 :: nat)" by arith
with gt have "x < ?z" "?z < y"
by (simp_all add: char_less_iff_nat_of_char)
hence "\<exists>z. x < z \<and> z < y" by blast }
moreover
{ fix z
assume "x < z" "z < y"
hence "(1 :: nat) < of_char y - of_char x"
by (simp add: char_less_iff_nat_of_char) }
ultimately show "proper_interval (Some x) (Some y) = (\<exists>z>x. z < y)"
by auto
qed simp
end
instantiation Enum.finite_1 :: proper_interval begin
definition proper_interval_finite_1 :: "Enum.finite_1 proper_interval"
where "proper_interval_finite_1 x y \<longleftrightarrow> x = None \<and> y = None"
instance by intro_classes (simp_all add: proper_interval_finite_1_def less_finite_1_def)
end
instantiation Enum.finite_2 :: proper_interval begin
fun proper_interval_finite_2 :: "Enum.finite_2 proper_interval" where
"proper_interval_finite_2 None None \<longleftrightarrow> True"
| "proper_interval_finite_2 None (Some x) \<longleftrightarrow> x = finite_2.a\<^sub>2"
| "proper_interval_finite_2 (Some x) None \<longleftrightarrow> x = finite_2.a\<^sub>1"
| "proper_interval_finite_2 (Some x) (Some y) \<longleftrightarrow> False"
instance by intro_classes (auto simp add: less_finite_2_def)
end
instantiation Enum.finite_3 :: proper_interval begin
fun proper_interval_finite_3 :: "Enum.finite_3 proper_interval" where
"proper_interval_finite_3 None None \<longleftrightarrow> True"
| "proper_interval_finite_3 None (Some x) \<longleftrightarrow> x \<noteq> finite_3.a\<^sub>1"
| "proper_interval_finite_3 (Some x) None \<longleftrightarrow> x \<noteq> finite_3.a\<^sub>3"
| "proper_interval_finite_3 (Some x) (Some y) \<longleftrightarrow> x = finite_3.a\<^sub>1 \<and> y = finite_3.a\<^sub>3"
instance
proof
fix x y :: Enum.finite_3
show "proper_interval None (Some y) = (\<exists>z. z < y)"
by(cases y)(auto simp add: less_finite_3_def split: finite_3.split)
show "proper_interval (Some x) None = (\<exists>z. x < z)"
by(cases x)(auto simp add: less_finite_3_def)
show "proper_interval (Some x) (Some y) = (\<exists>z>x. z < y)"
by(auto simp add: less_finite_3_def split: finite_3.split_asm)
qed simp
end
subsection \<open>List fusion for the order and proper intervals on @{typ "'a set"}\<close>
definition length_last_fusion :: "('a, 's) generator \<Rightarrow> 's \<Rightarrow> nat \<times> 'a"
where "length_last_fusion g s = length_last (list.unfoldr g s)"
lemma length_last_fusion_code [code]:
"length_last_fusion g s =
(if list.has_next g s then
let (x, s') = list.next g s
in fold_fusion g (\<lambda>x (n, _). (n + 1, x)) s' (1, x)
else (0, undefined))"
unfolding length_last_fusion_def
by(subst list.unfoldr.simps)(simp add: length_last_Nil length_last_Cons_code fold_fusion_def split_beta)
declare length_last_fusion_def [symmetric, code_unfold]
context proper_intrvl begin
definition set_less_eq_aux_Compl_fusion :: "('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 'a option \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where
"set_less_eq_aux_Compl_fusion g1 g2 ao s1 s2 =
set_less_eq_aux_Compl ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
definition Compl_set_less_eq_aux_fusion :: "('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 'a option \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where
"Compl_set_less_eq_aux_fusion g1 g2 ao s1 s2 =
Compl_set_less_eq_aux ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
definition set_less_aux_Compl_fusion :: "('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 'a option \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where
"set_less_aux_Compl_fusion g1 g2 ao s1 s2 =
set_less_aux_Compl ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
definition Compl_set_less_aux_fusion :: "('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 'a option \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where
"Compl_set_less_aux_fusion g1 g2 ao s1 s2 =
Compl_set_less_aux ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
definition exhaustive_above_fusion :: "('a, 's) generator \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
where "exhaustive_above_fusion g a s = exhaustive_above a (list.unfoldr g s)"
definition exhaustive_fusion :: "('a, 's) generator \<Rightarrow> 's \<Rightarrow> bool"
where "exhaustive_fusion g s = exhaustive (list.unfoldr g s)"
definition proper_interval_set_aux_fusion :: "('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where
"proper_interval_set_aux_fusion g1 g2 s1 s2 =
proper_interval_set_aux (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
definition proper_interval_set_Compl_aux_fusion ::
"('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 'a option \<Rightarrow> nat \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where
"proper_interval_set_Compl_aux_fusion g1 g2 ao n s1 s2 =
proper_interval_set_Compl_aux ao n (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
definition proper_interval_Compl_set_aux_fusion ::
"('a, 's1) generator \<Rightarrow> ('a, 's2) generator \<Rightarrow> 'a option \<Rightarrow> 's1 \<Rightarrow> 's2 \<Rightarrow> bool"
where
"proper_interval_Compl_set_aux_fusion g1 g2 ao s1 s2 =
proper_interval_Compl_set_aux ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)"
lemma set_less_eq_aux_Compl_fusion_code:
"set_less_eq_aux_Compl_fusion g1 g2 ao s1 s2 \<longleftrightarrow>
(list.has_next g1 s1 \<longrightarrow> list.has_next g2 s2 \<longrightarrow>
(let (x, s1') = list.next g1 s1;
(y, s2') = list.next g2 s2
in if x < y then proper_interval ao (Some x) \<or> set_less_eq_aux_Compl_fusion g1 g2 (Some x) s1' s2
else if y < x then proper_interval ao (Some y) \<or> set_less_eq_aux_Compl_fusion g1 g2 (Some y) s1 s2'
else proper_interval ao (Some y)))"
unfolding set_less_eq_aux_Compl_fusion_def
by(subst (1 2 4 5) list.unfoldr.simps)(simp add: split_beta)
lemma Compl_set_less_eq_aux_fusion_code:
"Compl_set_less_eq_aux_fusion g1 g2 ao s1 s2 \<longleftrightarrow>
(if list.has_next g1 s1 then
let (x, s1') = list.next g1 s1
in if list.has_next g2 s2 then
let (y, s2') = list.next g2 s2
in if x < y then \<not> proper_interval ao (Some x) \<and> Compl_set_less_eq_aux_fusion g1 g2 (Some x) s1' s2
else if y < x then \<not> proper_interval ao (Some y) \<and> Compl_set_less_eq_aux_fusion g1 g2 (Some y) s1 s2'
else \<not> proper_interval ao (Some y)
else \<not> proper_interval ao (Some x) \<and> Compl_set_less_eq_aux_fusion g1 g2 (Some x) s1' s2
else if list.has_next g2 s2 then
let (y, s2') = list.next g2 s2
in \<not> proper_interval ao (Some y) \<and> Compl_set_less_eq_aux_fusion g1 g2 (Some y) s1 s2'
else \<not> proper_interval ao None)"
unfolding Compl_set_less_eq_aux_fusion_def
by(subst (1 2 4 5 8 9) list.unfoldr.simps)(simp add: split_beta)
lemma set_less_aux_Compl_fusion_code:
"set_less_aux_Compl_fusion g1 g2 ao s1 s2 \<longleftrightarrow>
(if list.has_next g1 s1 then
let (x, s1') = list.next g1 s1
in if list.has_next g2 s2 then
let (y, s2') = list.next g2 s2
in if x < y then proper_interval ao (Some x) \<or> set_less_aux_Compl_fusion g1 g2 (Some x) s1' s2
else if y < x then proper_interval ao (Some y) \<or> set_less_aux_Compl_fusion g1 g2 (Some y) s1 s2'
else proper_interval ao (Some y)
else proper_interval ao (Some x) \<or> set_less_aux_Compl_fusion g1 g2 (Some x) s1' s2
else if list.has_next g2 s2 then
let (y, s2') = list.next g2 s2
in proper_interval ao (Some y) \<or> set_less_aux_Compl_fusion g1 g2 (Some y) s1 s2'
else proper_interval ao None)"
unfolding set_less_aux_Compl_fusion_def
by(subst (1 2 4 5 8 9) list.unfoldr.simps)(simp add: split_beta)
lemma Compl_set_less_aux_fusion_code:
"Compl_set_less_aux_fusion g1 g2 ao s1 s2 \<longleftrightarrow>
list.has_next g1 s1 \<and> list.has_next g2 s2 \<and>
(let (x, s1') = list.next g1 s1;
(y, s2') = list.next g2 s2
in if x < y then \<not> proper_interval ao (Some x) \<and> Compl_set_less_aux_fusion g1 g2 (Some x) s1' s2
else if y < x then \<not> proper_interval ao (Some y) \<and> Compl_set_less_aux_fusion g1 g2 (Some y) s1 s2'
else \<not> proper_interval ao (Some y))"
unfolding Compl_set_less_aux_fusion_def
by(subst (1 2 4 5) list.unfoldr.simps)(simp add: split_beta)
lemma exhaustive_above_fusion_code:
"exhaustive_above_fusion g y s \<longleftrightarrow>
(if list.has_next g s then
let (x, s') = list.next g s
in \<not> proper_interval (Some y) (Some x) \<and> exhaustive_above_fusion g x s'
else \<not> proper_interval (Some y) None)"
unfolding exhaustive_above_fusion_def
by(subst list.unfoldr.simps)(simp add: split_beta)
lemma exhaustive_fusion_code:
"exhaustive_fusion g s =
(list.has_next g s \<and>
(let (x, s') = list.next g s
in \<not> proper_interval None (Some x) \<and> exhaustive_above_fusion g x s'))"
unfolding exhaustive_fusion_def exhaustive_above_fusion_def
by(subst (1) list.unfoldr.simps)(simp add: split_beta)
lemma proper_interval_set_aux_fusion_code:
"proper_interval_set_aux_fusion g1 g2 s1 s2 \<longleftrightarrow>
list.has_next g2 s2 \<and>
(let (y, s2') = list.next g2 s2
in if list.has_next g1 s1 then
let (x, s1') = list.next g1 s1
in if x < y then False
else if y < x then proper_interval (Some y) (Some x) \<or> list.has_next g2 s2' \<or> \<not> exhaustive_above_fusion g1 x s1'
else proper_interval_set_aux_fusion g1 g2 s1' s2'
else list.has_next g2 s2' \<or> proper_interval (Some y) None)"
unfolding proper_interval_set_aux_fusion_def exhaustive_above_fusion_def
by(subst (1 2) list.unfoldr.simps)(simp add: split_beta)
lemma proper_interval_set_Compl_aux_fusion_code:
"proper_interval_set_Compl_aux_fusion g1 g2 ao n s1 s2 \<longleftrightarrow>
(if list.has_next g1 s1 then
let (x, s1') = list.next g1 s1
in if list.has_next g2 s2 then
let (y, s2') = list.next g2 s2
in if x < y then
proper_interval ao (Some x) \<or>
proper_interval_set_Compl_aux_fusion g1 g2 (Some x) (n + 1) s1' s2
else if y < x then
proper_interval ao (Some y) \<or>
proper_interval_set_Compl_aux_fusion g1 g2 (Some y) (n + 1) s1 s2'
else
proper_interval ao (Some x) \<and>
(let m = CARD('a) - n
in m - length_fusion g2 s2' \<noteq> 2 \<or> m - length_fusion g1 s1' \<noteq> 2)
else
let m = CARD('a) - n; (len_x, x') = length_last_fusion g1 s1
in m \<noteq> len_x \<and> (m = len_x + 1 \<longrightarrow> \<not> proper_interval (Some x') None)
else if list.has_next g2 s2 then
let (y, s2') = list.next g2 s2;
m = CARD('a) - n;
(len_y, y') = length_last_fusion g2 s2
in m \<noteq> len_y \<and> (m = len_y + 1 \<longrightarrow> \<not> proper_interval (Some y') None)
else CARD('a) > n + 1)"
unfolding proper_interval_set_Compl_aux_fusion_def length_last_fusion_def length_fusion_def
by(subst (1 2 4 5 9 10) list.unfoldr.simps)(simp add: split_beta)
lemma proper_interval_Compl_set_aux_fusion_code:
"proper_interval_Compl_set_aux_fusion g1 g2 ao s1 s2 \<longleftrightarrow>
list.has_next g1 s1 \<and> list.has_next g2 s2 \<and>
(let (x, s1') = list.next g1 s1;
(y, s2') = list.next g2 s2
in if x < y then
\<not> proper_interval ao (Some x) \<and> proper_interval_Compl_set_aux_fusion g1 g2 (Some x) s1' s2
else if y < x then
\<not> proper_interval ao (Some y) \<and> proper_interval_Compl_set_aux_fusion g1 g2 (Some y) s1 s2'
else \<not> proper_interval ao (Some x) \<and> (list.has_next g2 s2' \<or> list.has_next g1 s1'))"
unfolding proper_interval_Compl_set_aux_fusion_def
by(subst (1 2 4 5) list.unfoldr.simps)(auto simp add: split_beta)
end
lemmas [code] =
set_less_eq_aux_Compl_fusion_code proper_intrvl.set_less_eq_aux_Compl_fusion_code
Compl_set_less_eq_aux_fusion_code proper_intrvl.Compl_set_less_eq_aux_fusion_code
set_less_aux_Compl_fusion_code proper_intrvl.set_less_aux_Compl_fusion_code
Compl_set_less_aux_fusion_code proper_intrvl.Compl_set_less_aux_fusion_code
exhaustive_above_fusion_code proper_intrvl.exhaustive_above_fusion_code
exhaustive_fusion_code proper_intrvl.exhaustive_fusion_code
proper_interval_set_aux_fusion_code proper_intrvl.proper_interval_set_aux_fusion_code
proper_interval_set_Compl_aux_fusion_code proper_intrvl.proper_interval_set_Compl_aux_fusion_code
proper_interval_Compl_set_aux_fusion_code proper_intrvl.proper_interval_Compl_set_aux_fusion_code
lemmas [symmetric, code_unfold] =
set_less_eq_aux_Compl_fusion_def proper_intrvl.set_less_eq_aux_Compl_fusion_def
Compl_set_less_eq_aux_fusion_def proper_intrvl.Compl_set_less_eq_aux_fusion_def
set_less_aux_Compl_fusion_def proper_intrvl.set_less_aux_Compl_fusion_def
Compl_set_less_aux_fusion_def proper_intrvl.Compl_set_less_aux_fusion_def
exhaustive_above_fusion_def proper_intrvl.exhaustive_above_fusion_def
exhaustive_fusion_def proper_intrvl.exhaustive_fusion_def
proper_interval_set_aux_fusion_def proper_intrvl.proper_interval_set_aux_fusion_def
proper_interval_set_Compl_aux_fusion_def proper_intrvl.proper_interval_set_Compl_aux_fusion_def
proper_interval_Compl_set_aux_fusion_def proper_intrvl.proper_interval_Compl_set_aux_fusion_def
subsection \<open>Drop notation\<close>
context ord begin
no_notation set_less_aux (infix "\<sqsubset>''" 50)
and set_less_eq_aux (infix "\<sqsubseteq>''" 50)
and set_less_eq_aux' (infix "\<sqsubseteq>''''" 50)
and set_less_eq_aux'' (infix "\<sqsubseteq>''''''" 50)
and set_less_eq (infix "\<sqsubseteq>" 50)
and set_less (infix "\<sqsubset>" 50)
end
end
diff --git a/thys/CryptHOL/Misc_CryptHOL.thy b/thys/CryptHOL/Misc_CryptHOL.thy
--- a/thys/CryptHOL/Misc_CryptHOL.thy
+++ b/thys/CryptHOL/Misc_CryptHOL.thy
@@ -1,2521 +1,2521 @@
(* Title: Misc_CryptHOL.thy
Author: Andreas Lochbihler, ETH Zurich *)
section \<open>Miscellaneous library additions\<close>
theory Misc_CryptHOL imports
Probabilistic_While.While_SPMF
"HOL-Library.Rewrite"
"HOL-Library.Simps_Case_Conv"
"HOL-Library.Type_Length"
"HOL-Eisbach.Eisbach"
Coinductive.TLList
Monad_Normalisation.Monad_Normalisation
Monomorphic_Monad.Monomorphic_Monad
Applicative_Lifting.Applicative
begin
hide_const (open) Henstock_Kurzweil_Integration.negligible
declare eq_on_def [simp del]
subsection \<open>HOL\<close>
lemma asm_rl_conv: "(PROP P \<Longrightarrow> PROP P) \<equiv> Trueprop True"
by(rule equal_intr_rule) iprover+
named_theorems if_distribs "Distributivity theorems for If"
lemma if_mono_cong: "\<lbrakk>b \<Longrightarrow> x \<le> x'; \<not> b \<Longrightarrow> y \<le> y' \<rbrakk> \<Longrightarrow> If b x y \<le> If b x' y'"
by simp
lemma if_cong_then: "\<lbrakk> b = b'; b' \<Longrightarrow> t = t'; e = e' \<rbrakk> \<Longrightarrow> If b t e = If b' t' e'"
by simp
lemma if_False_eq: "\<lbrakk> b \<Longrightarrow> False; e = e' \<rbrakk> \<Longrightarrow> If b t e = e'"
by auto
lemma imp_OO_imp [simp]: "(\<longrightarrow>) OO (\<longrightarrow>) = (\<longrightarrow>)"
by auto
lemma inj_on_fun_updD: "\<lbrakk> inj_on (f(x := y)) A; x \<notin> A \<rbrakk> \<Longrightarrow> inj_on f A"
by(auto simp add: inj_on_def split: if_split_asm)
lemma disjoint_notin1: "\<lbrakk> A \<inter> B = {}; x \<in> B \<rbrakk> \<Longrightarrow> x \<notin> A" by auto
lemma Least_le_Least:
fixes x :: "'a :: wellorder"
assumes "Q x"
and Q: "\<And>x. Q x \<Longrightarrow> \<exists>y\<le>x. P y"
shows "Least P \<le> Least Q"
proof -
obtain f :: "'a \<Rightarrow> 'a" where "\<forall>a. \<not> Q a \<or> f a \<le> a \<and> P (f a)" using Q by moura
moreover have "Q (Least Q)" using \<open>Q x\<close> by(rule LeastI)
ultimately show ?thesis by (metis (full_types) le_cases le_less less_le_trans not_less_Least)
qed
lemma is_empty_image [simp]: "Set.is_empty (f ` A) = Set.is_empty A"
by(auto simp add: Set.is_empty_def)
subsection \<open>Relations\<close>
inductive Imagep :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool"
for R P
where ImagepI: "\<lbrakk> P x; R x y \<rbrakk> \<Longrightarrow> Imagep R P y"
lemma r_r_into_tranclp: "\<lbrakk> r x y; r y z \<rbrakk> \<Longrightarrow> r^++ x z"
by(rule tranclp.trancl_into_trancl)(rule tranclp.r_into_trancl)
lemma transp_tranclp_id:
assumes "transp R"
shows "tranclp R = R"
proof(intro ext iffI)
fix x y
assume "R^++ x y"
thus "R x y" by induction(blast dest: transpD[OF assms])+
qed simp
lemma transp_inv_image: "transp r \<Longrightarrow> transp (\<lambda>x y. r (f x) (f y))"
using trans_inv_image[where r="{(x, y). r x y}" and f = f]
by(simp add: transp_trans inv_image_def)
lemma Domainp_conversep: "Domainp R\<inverse>\<inverse> = Rangep R"
by(auto)
lemma bi_unique_rel_set_bij_betw:
assumes unique: "bi_unique R"
and rel: "rel_set R A B"
shows "\<exists>f. bij_betw f A B \<and> (\<forall>x\<in>A. R x (f x))"
proof -
from assms obtain f where f: "\<And>x. x \<in> A \<Longrightarrow> R x (f x)" and B: "\<And>x. x \<in> A \<Longrightarrow> f x \<in> B"
apply(atomize_elim)
apply(fold all_conj_distrib)
apply(subst choice_iff[symmetric])
apply(auto dest: rel_setD1)
done
have "inj_on f A" by(rule inj_onI)(auto dest!: f dest: bi_uniqueDl[OF unique])
moreover have "f ` A = B" using rel
by(auto 4 3 intro: B dest: rel_setD2 f bi_uniqueDr[OF unique])
ultimately have "bij_betw f A B" unfolding bij_betw_def ..
thus ?thesis using f by blast
qed
definition restrict_relp :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
("_ \<upharpoonleft> (_ \<otimes> _)" [53, 54, 54] 53)
where "restrict_relp R P Q = (\<lambda>x y. R x y \<and> P x \<and> Q y)"
lemma restrict_relp_apply [simp]: "(R \<upharpoonleft> P \<otimes> Q) x y \<longleftrightarrow> R x y \<and> P x \<and> Q y"
by(simp add: restrict_relp_def)
lemma restrict_relpI [intro?]: "\<lbrakk> R x y; P x; Q y \<rbrakk> \<Longrightarrow> (R \<upharpoonleft> P \<otimes> Q) x y"
by(simp add: restrict_relp_def)
lemma restrict_relpE [elim?, cases pred]:
assumes "(R \<upharpoonleft> P \<otimes> Q) x y"
obtains (restrict_relp) "R x y" "P x" "Q y"
using assms by(simp add: restrict_relp_def)
lemma conversep_restrict_relp [simp]: "(R \<upharpoonleft> P \<otimes> Q)\<inverse>\<inverse> = R\<inverse>\<inverse> \<upharpoonleft> Q \<otimes> P"
by(auto simp add: fun_eq_iff)
lemma restrict_relp_restrict_relp [simp]: "R \<upharpoonleft> P \<otimes> Q \<upharpoonleft> P' \<otimes> Q' = R \<upharpoonleft> inf P P' \<otimes> inf Q Q'"
by(auto simp add: fun_eq_iff)
lemma restrict_relp_cong:
"\<lbrakk> P = P'; Q = Q'; \<And>x y. \<lbrakk> P x; Q y \<rbrakk> \<Longrightarrow> R x y = R' x y \<rbrakk> \<Longrightarrow> R \<upharpoonleft> P \<otimes> Q = R' \<upharpoonleft> P' \<otimes> Q'"
by(auto simp add: fun_eq_iff)
lemma restrict_relp_cong_simp:
"\<lbrakk> P = P'; Q = Q'; \<And>x y. P x =simp=> Q y =simp=> R x y = R' x y \<rbrakk> \<Longrightarrow> R \<upharpoonleft> P \<otimes> Q = R' \<upharpoonleft> P' \<otimes> Q'"
by(rule restrict_relp_cong; simp add: simp_implies_def)
lemma restrict_relp_parametric [transfer_rule]:
includes lifting_syntax shows
"((A ===> B ===> (=)) ===> (A ===> (=)) ===> (B ===> (=)) ===> A ===> B ===> (=)) restrict_relp restrict_relp"
unfolding restrict_relp_def[abs_def] by transfer_prover
lemma restrict_relp_mono: "\<lbrakk> R \<le> R'; P \<le> P'; Q \<le> Q' \<rbrakk> \<Longrightarrow> R \<upharpoonleft> P \<otimes> Q \<le> R' \<upharpoonleft> P' \<otimes> Q'"
by(simp add: le_fun_def)
lemma restrict_relp_mono':
"\<lbrakk> (R \<upharpoonleft> P \<otimes> Q) x y; \<lbrakk> R x y; P x; Q y \<rbrakk> \<Longrightarrow> R' x y &&& P' x &&& Q' y \<rbrakk>
\<Longrightarrow> (R' \<upharpoonleft> P' \<otimes> Q') x y"
by(auto dest: conjunctionD1 conjunctionD2)
lemma restrict_relp_DomainpD: "Domainp (R \<upharpoonleft> P \<otimes> Q) x \<Longrightarrow> Domainp R x \<and> P x"
by(auto simp add: Domainp.simps)
lemma restrict_relp_True: "R \<upharpoonleft> (\<lambda>_. True) \<otimes> (\<lambda>_. True) = R"
by(simp add: fun_eq_iff)
lemma restrict_relp_False1: "R \<upharpoonleft> (\<lambda>_. False) \<otimes> Q = bot"
by(simp add: fun_eq_iff)
lemma restrict_relp_False2: "R \<upharpoonleft> P \<otimes> (\<lambda>_. False) = bot"
by(simp add: fun_eq_iff)
definition rel_prod2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> ('c \<times> 'b) \<Rightarrow> bool"
where "rel_prod2 R a = (\<lambda>(c, b). R a b)"
lemma rel_prod2_simps [simp]: "rel_prod2 R a (c, b) \<longleftrightarrow> R a b"
by(simp add: rel_prod2_def)
lemma restrict_rel_prod:
"rel_prod (R \<upharpoonleft> I1 \<otimes> I2) (S \<upharpoonleft> I1' \<otimes> I2') = rel_prod R S \<upharpoonleft> pred_prod I1 I1' \<otimes> pred_prod I2 I2'"
by(auto simp add: fun_eq_iff)
lemma restrict_rel_prod1:
"rel_prod (R \<upharpoonleft> I1 \<otimes> I2) S = rel_prod R S \<upharpoonleft> pred_prod I1 (\<lambda>_. True) \<otimes> pred_prod I2 (\<lambda>_. True)"
by(simp add: restrict_rel_prod[symmetric] restrict_relp_True)
lemma restrict_rel_prod2:
"rel_prod R (S \<upharpoonleft> I1 \<otimes> I2) = rel_prod R S \<upharpoonleft> pred_prod (\<lambda>_. True) I1 \<otimes> pred_prod (\<lambda>_. True) I2"
by(simp add: restrict_rel_prod[symmetric] restrict_relp_True)
consts relcompp_witness :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'c \<Rightarrow> 'b"
specification (relcompp_witness)
relcompp_witness1: "(A OO B) (fst xy) (snd xy) \<Longrightarrow> A (fst xy) (relcompp_witness A B xy)"
relcompp_witness2: "(A OO B) (fst xy) (snd xy) \<Longrightarrow> B (relcompp_witness A B xy) (snd xy)"
apply(fold all_conj_distrib)
apply(rule choice allI)+
by(auto intro: choice allI)
lemmas relcompp_witness[of _ _ "(x, y)" for x y, simplified] = relcompp_witness1 relcompp_witness2
hide_fact (open) relcompp_witness1 relcompp_witness2
lemma relcompp_witness_eq [simp]: "relcompp_witness (=) (=) (x, x) = x"
using relcompp_witness(1)[of "(=)" "(=)" x x] by(simp add: eq_OO)
subsection \<open>Pairs\<close>
lemma split_apfst [simp]: "case_prod h (apfst f xy) = case_prod (h \<circ> f) xy"
by(cases xy) simp
definition corec_prod :: "('s \<Rightarrow> 'a) \<Rightarrow> ('s \<Rightarrow> 'b) \<Rightarrow> 's \<Rightarrow> 'a \<times> 'b"
where "corec_prod f g = (\<lambda>s. (f s, g s))"
lemma corec_prod_apply: "corec_prod f g s = (f s, g s)"
by(simp add: corec_prod_def)
lemma corec_prod_sel [simp]:
shows fst_corec_prod: "fst (corec_prod f g s) = f s"
and snd_corec_prod: "snd (corec_prod f g s) = g s"
by(simp_all add: corec_prod_apply)
lemma apfst_corec_prod [simp]: "apfst h (corec_prod f g s) = corec_prod (h \<circ> f) g s"
by(simp add: corec_prod_apply)
lemma apsnd_corec_prod [simp]: "apsnd h (corec_prod f g s) = corec_prod f (h \<circ> g) s"
by(simp add: corec_prod_apply)
lemma map_corec_prod [simp]: "map_prod f g (corec_prod h k s) = corec_prod (f \<circ> h) (g \<circ> k) s"
by(simp add: corec_prod_apply)
lemma split_corec_prod [simp]: "case_prod h (corec_prod f g s) = h (f s) (g s)"
by(simp add: corec_prod_apply)
lemma Pair_fst_Unity: "(fst x, ()) = x"
by(cases x) simp
definition rprodl :: "('a \<times> 'b) \<times> 'c \<Rightarrow> 'a \<times> ('b \<times> 'c)" where "rprodl = (\<lambda>((a, b), c). (a, (b, c)))"
lemma rprodl_simps [simp]: "rprodl ((a, b), c) = (a, (b, c))"
by(simp add: rprodl_def)
lemma rprodl_parametric [transfer_rule]: includes lifting_syntax shows
"(rel_prod (rel_prod A B) C ===> rel_prod A (rel_prod B C)) rprodl rprodl"
unfolding rprodl_def by transfer_prover
definition lprodr :: "'a \<times> ('b \<times> 'c) \<Rightarrow> ('a \<times> 'b) \<times> 'c" where "lprodr = (\<lambda>(a, b, c). ((a, b), c))"
lemma lprodr_simps [simp]: "lprodr (a, b, c) = ((a, b), c)"
by(simp add: lprodr_def)
lemma lprodr_parametric [transfer_rule]: includes lifting_syntax shows
"(rel_prod A (rel_prod B C) ===> rel_prod (rel_prod A B) C) lprodr lprodr"
unfolding lprodr_def by transfer_prover
lemma lprodr_inverse [simp]: "rprodl (lprodr x) = x"
by(cases x) auto
lemma rprodl_inverse [simp]: "lprodr (rprodl x) = x"
by(cases x) auto
lemma pred_prod_mono' [mono]:
"pred_prod A B xy \<longrightarrow> pred_prod A' B' xy"
if "\<And>x. A x \<longrightarrow> A' x" "\<And>y. B y \<longrightarrow> B' y"
using that by(cases xy) auto
fun rel_witness_prod :: "('a \<times> 'b) \<times> ('c \<times> 'd) \<Rightarrow> (('a \<times> 'c) \<times> ('b \<times> 'd))" where
"rel_witness_prod ((a, b), (c, d)) = ((a, c), (b, d))"
subsection \<open>Sums\<close>
lemma islE:
assumes "isl x"
obtains l where "x = Inl l"
using assms by(cases x) auto
lemma Inl_in_Plus [simp]: "Inl x \<in> A <+> B \<longleftrightarrow> x \<in> A"
by auto
lemma Inr_in_Plus [simp]: "Inr x \<in> A <+> B \<longleftrightarrow> x \<in> B"
by auto
lemma Inl_eq_map_sum_iff: "Inl x = map_sum f g y \<longleftrightarrow> (\<exists>z. y = Inl z \<and> x = f z)"
by(cases y) auto
lemma Inr_eq_map_sum_iff: "Inr x = map_sum f g y \<longleftrightarrow> (\<exists>z. y = Inr z \<and> x = g z)"
by(cases y) auto
lemma inj_on_map_sum [simp]:
"\<lbrakk> inj_on f A; inj_on g B \<rbrakk> \<Longrightarrow> inj_on (map_sum f g) (A <+> B)"
proof(rule inj_onI, goal_cases)
case (1 x y)
then show ?case by(cases x; cases y; auto simp add: inj_on_def)
qed
lemma inv_into_map_sum:
"inv_into (A <+> B) (map_sum f g) x = map_sum (inv_into A f) (inv_into B g) x"
if "x \<in> f ` A <+> g ` B" "inj_on f A" "inj_on g B"
using that by(cases rule: PlusE[consumes 1])(auto simp add: inv_into_f_eq f_inv_into_f)
fun rsuml :: "('a + 'b) + 'c \<Rightarrow> 'a + ('b + 'c)" where
"rsuml (Inl (Inl a)) = Inl a"
| "rsuml (Inl (Inr b)) = Inr (Inl b)"
| "rsuml (Inr c) = Inr (Inr c)"
fun lsumr :: "'a + ('b + 'c) \<Rightarrow> ('a + 'b) + 'c" where
"lsumr (Inl a) = Inl (Inl a)"
| "lsumr (Inr (Inl b)) = Inl (Inr b)"
| "lsumr (Inr (Inr c)) = Inr c"
lemma rsuml_lsumr [simp]: "rsuml (lsumr x) = x"
by(cases x rule: lsumr.cases) simp_all
lemma lsumr_rsuml [simp]: "lsumr (rsuml x) = x"
by(cases x rule: rsuml.cases) simp_all
subsection \<open>Option\<close>
declare is_none_bind [simp]
lemma case_option_collapse: "case_option x (\<lambda>_. x) y = x"
by(simp split: option.split)
lemma indicator_single_Some: "indicator {Some x} (Some y) = indicator {x} y"
by(simp split: split_indicator)
subsubsection \<open>Predicator and relator\<close>
lemma option_pred_mono_strong:
"\<lbrakk> pred_option P x; \<And>a. \<lbrakk> a \<in> set_option x; P a \<rbrakk> \<Longrightarrow> P' a \<rbrakk> \<Longrightarrow> pred_option P' x"
by(fact option.pred_mono_strong)
lemma option_pred_map [simp]: "pred_option P (map_option f x) = pred_option (P \<circ> f) x"
by(fact option.pred_map)
lemma option_pred_o_map [simp]: "pred_option P \<circ> map_option f = pred_option (P \<circ> f)"
by(simp add: fun_eq_iff)
lemma option_pred_bind [simp]: "pred_option P (Option.bind x f) = pred_option (pred_option P \<circ> f) x"
by(simp add: pred_option_def)
lemma pred_option_conj [simp]:
"pred_option (\<lambda>x. P x \<and> Q x) = (\<lambda>x. pred_option P x \<and> pred_option Q x)"
by(auto simp add: pred_option_def)
lemma pred_option_top [simp]:
"pred_option (\<lambda>_. True) = (\<lambda>_. True)"
by(fact option.pred_True)
lemma rel_option_restrict_relpI [intro?]:
"\<lbrakk> rel_option R x y; pred_option P x; pred_option Q y \<rbrakk> \<Longrightarrow> rel_option (R \<upharpoonleft> P \<otimes> Q) x y"
by(erule option.rel_mono_strong) simp
lemma rel_option_restrict_relpE [elim?]:
assumes "rel_option (R \<upharpoonleft> P \<otimes> Q) x y"
obtains "rel_option R x y" "pred_option P x" "pred_option Q y"
proof
show "rel_option R x y" using assms by(auto elim!: option.rel_mono_strong)
have "pred_option (Domainp (R \<upharpoonleft> P \<otimes> Q)) x" using assms by(fold option.Domainp_rel) blast
then show "pred_option P x" by(rule option_pred_mono_strong)(blast dest!: restrict_relp_DomainpD)
have "pred_option (Domainp (R \<upharpoonleft> P \<otimes> Q)\<inverse>\<inverse>) y" using assms
by(fold option.Domainp_rel)(auto simp only: option.rel_conversep Domainp_conversep)
then show "pred_option Q y" by(rule option_pred_mono_strong)(auto dest!: restrict_relp_DomainpD)
qed
lemma rel_option_restrict_relp_iff:
"rel_option (R \<upharpoonleft> P \<otimes> Q) x y \<longleftrightarrow> rel_option R x y \<and> pred_option P x \<and> pred_option Q y"
by(blast intro: rel_option_restrict_relpI elim: rel_option_restrict_relpE)
lemma option_rel_map_restrict_relp:
shows option_rel_map_restrict_relp1:
"rel_option (R \<upharpoonleft> P \<otimes> Q) (map_option f x) = rel_option (R \<circ> f \<upharpoonleft> P \<circ> f \<otimes> Q) x"
and option_rel_map_restrict_relp2:
"rel_option (R \<upharpoonleft> P \<otimes> Q) x (map_option g y) = rel_option ((\<lambda>x. R x \<circ> g) \<upharpoonleft> P \<otimes> Q \<circ> g) x y"
by(simp_all add: option.rel_map restrict_relp_def fun_eq_iff)
fun rel_witness_option :: "'a option \<times> 'b option \<Rightarrow> ('a \<times> 'b) option" where
"rel_witness_option (Some x, Some y) = Some (x, y)"
| "rel_witness_option (None, None) = None"
| "rel_witness_option _ = None" \<comment> \<open>Just to make the definition complete\<close>
lemma rel_witness_option:
shows set_rel_witness_option: "\<lbrakk> rel_option A x y; (a, b) \<in> set_option (rel_witness_option (x, y)) \<rbrakk> \<Longrightarrow> A a b"
and map1_rel_witness_option: "rel_option A x y \<Longrightarrow> map_option fst (rel_witness_option (x, y)) = x"
and map2_rel_witness_option: "rel_option A x y \<Longrightarrow> map_option snd (rel_witness_option (x, y)) = y"
by(cases "(x, y)" rule: rel_witness_option.cases; simp; fail)+
lemma rel_witness_option1:
assumes "rel_option A x y"
shows "rel_option (\<lambda>a (a', b). a = a' \<and> A a' b) x (rel_witness_option (x, y))"
using map1_rel_witness_option[OF assms, symmetric]
unfolding option.rel_eq[symmetric] option.rel_map
by(rule option.rel_mono_strong)(auto intro: set_rel_witness_option[OF assms])
lemma rel_witness_option2:
assumes "rel_option A x y"
shows "rel_option (\<lambda>(a, b') b. b = b' \<and> A a b') (rel_witness_option (x, y)) y"
using map2_rel_witness_option[OF assms]
unfolding option.rel_eq[symmetric] option.rel_map
by(rule option.rel_mono_strong)(auto intro: set_rel_witness_option[OF assms])
subsubsection \<open>Orders on option\<close>
abbreviation le_option :: "'a option \<Rightarrow> 'a option \<Rightarrow> bool"
where "le_option \<equiv> ord_option (=)"
lemma le_option_bind_mono:
"\<lbrakk> le_option x y; \<And>a. a \<in> set_option x \<Longrightarrow> le_option (f a) (g a) \<rbrakk>
\<Longrightarrow> le_option (Option.bind x f) (Option.bind y g)"
by(cases x) simp_all
lemma le_option_refl [simp]: "le_option x x"
by(cases x) simp_all
lemma le_option_conv_option_ord: "le_option = option_ord"
by(auto simp add: fun_eq_iff flat_ord_def elim: ord_option.cases)
definition pcr_Some :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b option \<Rightarrow> bool"
where "pcr_Some R x y \<longleftrightarrow> (\<exists>z. y = Some z \<and> R x z)"
lemma pcr_Some_simps [simp]: "pcr_Some R x (Some y) \<longleftrightarrow> R x y"
by(simp add: pcr_Some_def)
lemma pcr_SomeE [cases pred]:
assumes "pcr_Some R x y"
obtains (pcr_Some) z where "y = Some z" "R x z"
using assms by(auto simp add: pcr_Some_def)
subsubsection \<open>Filter for option\<close>
fun filter_option :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'a option"
where
"filter_option P None = None"
| "filter_option P (Some x) = (if P x then Some x else None)"
lemma set_filter_option [simp]: "set_option (filter_option P x) = {y \<in> set_option x. P y}"
by(cases x) auto
lemma filter_map_option: "filter_option P (map_option f x) = map_option f (filter_option (P \<circ> f) x)"
by(cases x) simp_all
lemma is_none_filter_option [simp]: "Option.is_none (filter_option P x) \<longleftrightarrow> Option.is_none x \<or> \<not> P (the x)"
by(cases x) simp_all
lemma filter_option_eq_Some_iff [simp]: "filter_option P x = Some y \<longleftrightarrow> x = Some y \<and> P y"
by(cases x) auto
lemma Some_eq_filter_option_iff [simp]: "Some y = filter_option P x \<longleftrightarrow> x = Some y \<and> P y"
by(cases x) auto
lemma filter_conv_bind_option: "filter_option P x = Option.bind x (\<lambda>y. if P y then Some y else None)"
by(cases x) simp_all
subsubsection \<open>Assert for option\<close>
primrec assert_option :: "bool \<Rightarrow> unit option" where
"assert_option True = Some ()"
| "assert_option False = None"
lemma set_assert_option_conv: "set_option (assert_option b) = (if b then {()} else {})"
by(simp)
lemma in_set_assert_option [simp]: "x \<in> set_option (assert_option b) \<longleftrightarrow> b"
by(cases b) simp_all
subsubsection \<open>Join on options\<close>
definition join_option :: "'a option option \<Rightarrow> 'a option"
where "join_option x = (case x of Some y \<Rightarrow> y | None \<Rightarrow> None)"
simps_of_case join_simps [simp, code]: join_option_def
lemma set_join_option [simp]: "set_option (join_option x) = \<Union>(set_option ` set_option x)"
by(cases x)(simp_all)
lemma in_set_join_option: "x \<in> set_option (join_option (Some (Some x)))"
by simp
lemma map_join_option: "map_option f (join_option x) = join_option (map_option (map_option f) x)"
by(cases x) simp_all
lemma bind_conv_join_option: "Option.bind x f = join_option (map_option f x)"
by(cases x) simp_all
lemma join_conv_bind_option: "join_option x = Option.bind x id"
by(cases x) simp_all
lemma join_option_parametric [transfer_rule]:
includes lifting_syntax shows
"(rel_option (rel_option R) ===> rel_option R) join_option join_option"
unfolding join_conv_bind_option[abs_def] by transfer_prover
lemma join_option_eq_Some [simp]: "join_option x = Some y \<longleftrightarrow> x = Some (Some y)"
by(cases x) simp_all
lemma Some_eq_join_option [simp]: "Some y = join_option x \<longleftrightarrow> x = Some (Some y)"
by(cases x) auto
lemma join_option_eq_None: "join_option x = None \<longleftrightarrow> x = None \<or> x = Some None"
by(cases x) simp_all
lemma None_eq_join_option: "None = join_option x \<longleftrightarrow> x = None \<or> x = Some None"
by(cases x) auto
subsubsection \<open>Zip on options\<close>
function zip_option :: "'a option \<Rightarrow> 'b option \<Rightarrow> ('a \<times> 'b) option"
where
"zip_option (Some x) (Some y) = Some (x, y)"
| "zip_option _ None = None"
| "zip_option None _ = None"
by pat_completeness auto
termination by lexicographic_order
lemma zip_option_eq_Some_iff [iff]:
"zip_option x y = Some (a, b) \<longleftrightarrow> x = Some a \<and> y = Some b"
by(cases "(x, y)" rule: zip_option.cases) simp_all
lemma set_zip_option [simp]:
"set_option (zip_option x y) = set_option x \<times> set_option y"
by auto
lemma zip_map_option1: "zip_option (map_option f x) y = map_option (apfst f) (zip_option x y)"
by(cases "(x, y)" rule: zip_option.cases) simp_all
lemma zip_map_option2: "zip_option x (map_option g y) = map_option (apsnd g) (zip_option x y)"
by(cases "(x, y)" rule: zip_option.cases) simp_all
lemma map_zip_option:
"map_option (map_prod f g) (zip_option x y) = zip_option (map_option f x) (map_option g y)"
by(simp add: zip_map_option1 zip_map_option2 option.map_comp apfst_def apsnd_def o_def prod.map_comp)
lemma zip_conv_bind_option:
"zip_option x y = Option.bind x (\<lambda>x. Option.bind y (\<lambda>y. Some (x, y)))"
by(cases "(x, y)" rule: zip_option.cases) simp_all
lemma zip_option_parametric [transfer_rule]:
includes lifting_syntax shows
"(rel_option R ===> rel_option Q ===> rel_option (rel_prod R Q)) zip_option zip_option"
unfolding zip_conv_bind_option[abs_def] by transfer_prover
lemma rel_option_eqI [simp]: "rel_option (=) x x"
by(simp add: option.rel_eq)
subsubsection \<open>Binary supremum on @{typ "'a option"}\<close>
primrec sup_option :: "'a option \<Rightarrow> 'a option \<Rightarrow> 'a option"
where
"sup_option x None = x"
| "sup_option x (Some y) = (Some y)"
lemma sup_option_idem [simp]: "sup_option x x = x"
by(cases x) simp_all
lemma sup_option_assoc: "sup_option (sup_option x y) z = sup_option x (sup_option y z)"
by(cases z) simp_all
lemma sup_option_left_idem: "sup_option x (sup_option x y) = sup_option x y"
by(rewrite sup_option_assoc[symmetric])(simp)
lemmas sup_option_ai = sup_option_assoc sup_option_left_idem
lemma sup_option_None [simp]: "sup_option None y = y"
by(cases y) simp_all
subsubsection \<open>Restriction on @{typ "'a option"}\<close>
primrec (transfer) enforce_option :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'a option" where
"enforce_option P (Some x) = (if P x then Some x else None)"
| "enforce_option P None = None"
lemma set_enforce_option [simp]: "set_option (enforce_option P x) = {a \<in> set_option x. P a}"
by(cases x) auto
lemma enforce_map_option: "enforce_option P (map_option f x) = map_option f (enforce_option (P \<circ> f) x)"
by(cases x) auto
lemma enforce_bind_option [simp]:
"enforce_option P (Option.bind x f) = Option.bind x (enforce_option P \<circ> f)"
by(cases x) auto
lemma enforce_option_alt_def:
"enforce_option P x = Option.bind x (\<lambda>a. Option.bind (assert_option (P a)) (\<lambda>_ :: unit. Some a))"
by(cases x) simp_all
lemma enforce_option_eq_None_iff [simp]:
"enforce_option P x = None \<longleftrightarrow> (\<forall>a. x = Some a \<longrightarrow> \<not> P a)"
by(cases x) auto
lemma enforce_option_eq_Some_iff [simp]:
"enforce_option P x = Some y \<longleftrightarrow> x = Some y \<and> P y"
by(cases x) auto
lemma Some_eq_enforce_option_iff [simp]:
"Some y = enforce_option P x \<longleftrightarrow> x = Some y \<and> P y"
by(cases x) auto
lemma enforce_option_top [simp]: "enforce_option \<top> = id"
by(rule ext; rename_tac x; case_tac x; simp)
lemma enforce_option_K_True [simp]: "enforce_option (\<lambda>_. True) x = x"
by(cases x) simp_all
lemma enforce_option_bot [simp]: "enforce_option \<bottom> = (\<lambda>_. None)"
by(simp add: fun_eq_iff)
lemma enforce_option_K_False [simp]: "enforce_option (\<lambda>_. False) x = None"
by simp
lemma enforce_pred_id_option: "pred_option P x \<Longrightarrow> enforce_option P x = x"
by(cases x) auto
subsubsection \<open>Maps\<close>
lemma map_add_apply: "(m1 ++ m2) x = sup_option (m1 x) (m2 x)"
by(simp add: map_add_def split: option.split)
lemma map_le_map_upd2: "\<lbrakk> f \<subseteq>\<^sub>m g; \<And>y'. f x = Some y' \<Longrightarrow> y' = y \<rbrakk> \<Longrightarrow> f \<subseteq>\<^sub>m g(x \<mapsto> y)"
by(cases "x \<in> dom f")(auto simp add: map_le_def Ball_def)
lemma eq_None_iff_not_dom: "f x = None \<longleftrightarrow> x \<notin> dom f"
by auto
lemma card_ran_le_dom: "finite (dom m) \<Longrightarrow> card (ran m) \<le> card (dom m)"
by(simp add: ran_alt_def card_image_le)
lemma dom_subset_ran_iff:
assumes "finite (ran m)"
shows "dom m \<subseteq> ran m \<longleftrightarrow> dom m = ran m"
proof
assume le: "dom m \<subseteq> ran m"
then have "card (dom m) \<le> card (ran m)" by(simp add: card_mono assms)
moreover have "card (ran m) \<le> card (dom m)" by(simp add: finite_subset[OF le assms] card_ran_le_dom)
ultimately show "dom m = ran m" using card_subset_eq[OF assms le] by simp
qed simp
text \<open>
We need a polymorphic constant for the empty map such that \<open>transfer_prover\<close>
can use a custom transfer rule for @{const Map.empty}
\<close>
definition Map_empty where [simp]: "Map_empty \<equiv> Map.empty"
lemma map_le_Some1D: "\<lbrakk> m \<subseteq>\<^sub>m m'; m x = Some y \<rbrakk> \<Longrightarrow> m' x = Some y"
by(auto simp add: map_le_def Ball_def)
lemma map_le_fun_upd2: "\<lbrakk> f \<subseteq>\<^sub>m g; x \<notin> dom f \<rbrakk> \<Longrightarrow> f \<subseteq>\<^sub>m g(x := y)"
by(auto simp add: map_le_def)
lemma map_eqI: "\<forall>x\<in>dom m \<union> dom m'. m x = m' x \<Longrightarrow> m = m'"
by(auto simp add: fun_eq_iff domIff intro: option.expand)
subsection \<open>Countable\<close>
lemma countable_lfp:
assumes step: "\<And>Y. countable Y \<Longrightarrow> countable (F Y)"
and cont: "Order_Continuity.sup_continuous F"
shows "countable (lfp F)"
by(subst sup_continuous_lfp[OF cont])(simp add: countable_funpow[OF step])
lemma countable_lfp_apply:
assumes step: "\<And>Y x. (\<And>x. countable (Y x)) \<Longrightarrow> countable (F Y x)"
and cont: "Order_Continuity.sup_continuous F"
shows "countable (lfp F x)"
proof -
{ fix n
have "\<And>x. countable ((F ^^ n) bot x)"
by(induct n)(auto intro: step) }
thus ?thesis using cont by(simp add: sup_continuous_lfp)
qed
subsection \<open> Extended naturals \<close>
lemma idiff_enat_eq_enat_iff: "x - enat n = enat m \<longleftrightarrow> (\<exists>k. x = enat k \<and> k - n = m)"
by (cases x) simp_all
lemma eSuc_SUP: "A \<noteq> {} \<Longrightarrow> eSuc (\<Squnion> (f ` A)) = (\<Squnion>x\<in>A. eSuc (f x))"
by (subst eSuc_Sup) (simp_all add: image_comp)
lemma ereal_of_enat_1: "ereal_of_enat 1 = ereal 1"
by (simp add: one_enat_def)
lemma ennreal_real_conv_ennreal_of_enat: "ennreal (real n) = ennreal_of_enat n"
by (simp add: ennreal_of_nat_eq_real_of_nat)
lemma enat_add_sub_same2: "b \<noteq> \<infinity> \<Longrightarrow> a + b - b = (a :: enat)"
by (cases a; cases b) simp_all
lemma enat_sub_add: "y \<le> x \<Longrightarrow> x - y + z = x + z - (y :: enat)"
by (cases x; cases y; cases z) simp_all
lemma SUP_enat_eq_0_iff [simp]: "\<Squnion> (f ` A) = (0 :: enat) \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
by (simp add: bot_enat_def [symmetric])
lemma SUP_enat_add_left:
assumes "I \<noteq> {}"
shows "(SUP i\<in>I. f i + c :: enat) = (SUP i\<in>I. f i) + c" (is "?lhs = ?rhs")
proof(cases "c", rule antisym)
case (enat n)
show "?lhs \<le> ?rhs" by(auto 4 3 intro: SUP_upper intro: SUP_least)
have "(SUP i\<in>I. f i) \<le> ?lhs - c" using enat
by(auto simp add: enat_add_sub_same2 intro!: SUP_least order_trans[OF _ SUP_upper[THEN enat_minus_mono1]])
note add_right_mono[OF this, of c]
also have "\<dots> + c \<le> ?lhs" using assms
by(subst enat_sub_add)(auto intro: SUP_upper2 simp add: enat_add_sub_same2 enat)
finally show "?rhs \<le> ?lhs" .
qed(simp add: assms SUP_constant)
lemma SUP_enat_add_right:
assumes "I \<noteq> {}"
shows "(SUP i\<in>I. c + f i :: enat) = c + (SUP i\<in>I. f i)"
using SUP_enat_add_left[OF assms, of f c]
by(simp add: add.commute)
lemma iadd_SUP_le_iff: "n + (SUP x\<in>A. f x :: enat) \<le> y \<longleftrightarrow> (if A = {} then n \<le> y else \<forall>x\<in>A. n + f x \<le> y)"
by(simp add: bot_enat_def SUP_enat_add_right[symmetric] SUP_le_iff)
lemma SUP_iadd_le_iff: "(SUP x\<in>A. f x :: enat) + n \<le> y \<longleftrightarrow> (if A = {} then n \<le> y else \<forall>x\<in>A. f x + n \<le> y)"
using iadd_SUP_le_iff[of n f A y] by(simp add: add.commute)
subsection \<open>Extended non-negative reals\<close>
lemma (in finite_measure) nn_integral_indicator_neq_infty:
"f -` A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A (f x) \<partial>M) \<noteq> \<infinity>"
unfolding ennreal_indicator[symmetric]
apply(rule integrableD)
apply(rule integrable_const_bound[where B=1])
apply(simp_all add: indicator_vimage[symmetric])
done
lemma (in finite_measure) nn_integral_indicator_neq_top:
"f -` A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A (f x) \<partial>M) \<noteq> \<top>"
by(drule nn_integral_indicator_neq_infty) simp
lemma nn_integral_indicator_map:
assumes [measurable]: "f \<in> measurable M N" "{x\<in>space N. P x} \<in> sets N"
shows "(\<integral>\<^sup>+x. indicator {x\<in>space N. P x} (f x) \<partial>M) = emeasure M {x\<in>space M. P (f x)}"
using assms(1)[THEN measurable_space]
by (subst nn_integral_indicator[symmetric])
(auto intro!: nn_integral_cong split: split_indicator simp del: nn_integral_indicator)
subsection \<open>BNF material\<close>
lemma transp_rel_fun: "\<lbrakk> is_equality Q; transp R \<rbrakk> \<Longrightarrow> transp (rel_fun Q R)"
by(rule transpI)(auto dest: transpD rel_funD simp add: is_equality_def)
lemma rel_fun_inf: "inf (rel_fun Q R) (rel_fun Q R') = rel_fun Q (inf R R')"
by(rule antisym)(auto elim: rel_fun_mono dest: rel_funD)
lemma reflp_fun1: includes lifting_syntax shows "\<lbrakk> is_equality A; reflp B \<rbrakk> \<Longrightarrow> reflp (A ===> B)"
by(simp add: reflp_def rel_fun_def is_equality_def)
lemma type_copy_id': "type_definition (\<lambda>x. x) (\<lambda>x. x) UNIV"
by unfold_locales simp_all
lemma type_copy_id: "type_definition id id UNIV"
by(simp add: id_def type_copy_id')
lemma GrpE [cases pred]:
assumes "BNF_Def.Grp A f x y"
obtains (Grp) "y = f x" "x \<in> A"
using assms
by(simp add: Grp_def)
lemma rel_fun_Grp_copy_Abs:
includes lifting_syntax
assumes "type_definition Rep Abs A"
shows "rel_fun (BNF_Def.Grp A Abs) (BNF_Def.Grp B g) = BNF_Def.Grp {f. f ` A \<subseteq> B} (Rep ---> g)"
proof -
interpret type_definition Rep Abs A by fact
show ?thesis
by(auto simp add: rel_fun_def Grp_def fun_eq_iff Abs_inverse Rep_inverse intro!: Rep)
qed
lemma rel_set_Grp:
"rel_set (BNF_Def.Grp A f) = BNF_Def.Grp {B. B \<subseteq> A} (image f)"
by(auto simp add: rel_set_def BNF_Def.Grp_def fun_eq_iff)
lemma rel_set_comp_Grp:
"rel_set R = (BNF_Def.Grp {x. x \<subseteq> {(x, y). R x y}} ((`) fst))\<inverse>\<inverse> OO BNF_Def.Grp {x. x \<subseteq> {(x, y). R x y}} ((`) snd)"
apply(auto 4 4 del: ext intro!: ext simp add: BNF_Def.Grp_def intro!: rel_setI intro: rev_bexI)
apply(simp add: relcompp_apply)
subgoal for A B
apply(rule exI[where x="A \<times> B \<inter> {(x, y). R x y}"])
apply(auto 4 3 dest: rel_setD1 rel_setD2 intro: rev_image_eqI)
done
done
lemma Domainp_Grp: "Domainp (BNF_Def.Grp A f) = (\<lambda>x. x \<in> A)"
by(auto simp add: fun_eq_iff Grp_def)
lemma pred_prod_conj [simp]:
shows pred_prod_conj1: "\<And>P Q R. pred_prod (\<lambda>x. P x \<and> Q x) R = (\<lambda>x. pred_prod P R x \<and> pred_prod Q R x)"
and pred_prod_conj2: "\<And>P Q R. pred_prod P (\<lambda>x. Q x \<and> R x) = (\<lambda>x. pred_prod P Q x \<and> pred_prod P R x)"
by(auto simp add: pred_prod.simps)
lemma pred_sum_conj [simp]:
shows pred_sum_conj1: "\<And>P Q R. pred_sum (\<lambda>x. P x \<and> Q x) R = (\<lambda>x. pred_sum P R x \<and> pred_sum Q R x)"
and pred_sum_conj2: "\<And>P Q R. pred_sum P (\<lambda>x. Q x \<and> R x) = (\<lambda>x. pred_sum P Q x \<and> pred_sum P R x)"
by(auto simp add: pred_sum.simps fun_eq_iff)
lemma pred_list_conj [simp]: "list_all (\<lambda>x. P x \<and> Q x) = (\<lambda>x. list_all P x \<and> list_all Q x)"
by(auto simp add: list_all_def)
lemma pred_prod_top [simp]:
"pred_prod (\<lambda>_. True) (\<lambda>_. True) = (\<lambda>_. True)"
by(simp add: pred_prod.simps fun_eq_iff)
lemma rel_fun_conversep: includes lifting_syntax shows
"(A^--1 ===> B^--1) = (A ===> B)^--1"
by(auto simp add: rel_fun_def fun_eq_iff)
lemma left_unique_Grp [iff]:
"left_unique (BNF_Def.Grp A f) \<longleftrightarrow> inj_on f A"
unfolding Grp_def left_unique_def by(auto simp add: inj_on_def)
lemma right_unique_Grp [simp, intro!]: "right_unique (BNF_Def.Grp A f)"
by(simp add: Grp_def right_unique_def)
lemma bi_unique_Grp [iff]:
"bi_unique (BNF_Def.Grp A f) \<longleftrightarrow> inj_on f A"
by(simp add: bi_unique_alt_def)
lemma left_total_Grp [iff]:
"left_total (BNF_Def.Grp A f) \<longleftrightarrow> A = UNIV"
by(auto simp add: left_total_def Grp_def)
lemma right_total_Grp [iff]:
"right_total (BNF_Def.Grp A f) \<longleftrightarrow> f ` A = UNIV"
by(auto simp add: right_total_def BNF_Def.Grp_def image_def)
lemma bi_total_Grp [iff]:
"bi_total (BNF_Def.Grp A f) \<longleftrightarrow> A = UNIV \<and> surj f"
by(auto simp add: bi_total_alt_def)
lemma left_unique_vimage2p [simp]:
"\<lbrakk> left_unique P; inj f \<rbrakk> \<Longrightarrow> left_unique (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro left_unique_OO) simp_all
lemma right_unique_vimage2p [simp]:
"\<lbrakk> right_unique P; inj g \<rbrakk> \<Longrightarrow> right_unique (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro right_unique_OO) simp_all
lemma bi_unique_vimage2p [simp]:
"\<lbrakk> bi_unique P; inj f; inj g \<rbrakk> \<Longrightarrow> bi_unique (BNF_Def.vimage2p f g P)"
unfolding bi_unique_alt_def by simp
lemma left_total_vimage2p [simp]:
"\<lbrakk> left_total P; surj g \<rbrakk> \<Longrightarrow> left_total (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro left_total_OO) simp_all
lemma right_total_vimage2p [simp]:
"\<lbrakk> right_total P; surj f \<rbrakk> \<Longrightarrow> right_total (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro right_total_OO) simp_all
lemma bi_total_vimage2p [simp]:
"\<lbrakk> bi_total P; surj f; surj g \<rbrakk> \<Longrightarrow> bi_total (BNF_Def.vimage2p f g P)"
unfolding bi_total_alt_def by simp
lemma vimage2p_eq [simp]:
"inj f \<Longrightarrow> BNF_Def.vimage2p f f (=) = (=)"
by(auto simp add: vimage2p_def fun_eq_iff inj_on_def)
lemma vimage2p_conversep: "BNF_Def.vimage2p f g R^--1 = (BNF_Def.vimage2p g f R)^--1"
by(simp add: vimage2p_def fun_eq_iff)
lemma rel_fun_refl: "\<lbrakk> A \<le> (=); (=) \<le> B \<rbrakk> \<Longrightarrow> (=) \<le> rel_fun A B"
by(subst fun.rel_eq[symmetric])(rule fun_mono)
lemma rel_fun_mono_strong:
"\<lbrakk> rel_fun A B f g; A' \<le> A; \<And>x y. \<lbrakk> x \<in> f ` {x. Domainp A' x}; y \<in> g ` {x. Rangep A' x}; B x y \<rbrakk> \<Longrightarrow> B' x y \<rbrakk> \<Longrightarrow> rel_fun A' B' f g"
by(auto simp add: rel_fun_def) fastforce
lemma rel_fun_refl_strong:
assumes "A \<le> (=)" "\<And>x. x \<in> f ` {x. Domainp A x} \<Longrightarrow> B x x"
shows "rel_fun A B f f"
proof -
have "rel_fun (=) (=) f f" by(simp add: rel_fun_eq)
then show ?thesis using assms(1)
by(rule rel_fun_mono_strong) (auto intro: assms(2))
qed
lemma Grp_iff: "BNF_Def.Grp B g x y \<longleftrightarrow> y = g x \<and> x \<in> B" by(simp add: Grp_def)
lemma Rangep_Grp: "Rangep (BNF_Def.Grp A f) = (\<lambda>x. x \<in> f ` A)"
by(auto simp add: fun_eq_iff Grp_iff)
lemma rel_fun_Grp:
"rel_fun (BNF_Def.Grp UNIV h)\<inverse>\<inverse> (BNF_Def.Grp A g) = BNF_Def.Grp {f. f ` range h \<subseteq> A} (map_fun h g)"
by(auto simp add: rel_fun_def fun_eq_iff Grp_iff)
subsection \<open>Transfer and lifting material\<close>
context includes lifting_syntax begin
lemma monotone_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_total A"
shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) monotone monotone"
unfolding monotone_def[abs_def] by transfer_prover
lemma fun_ord_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_total C"
shows "((A ===> B ===> (=)) ===> (C ===> A) ===> (C ===> B) ===> (=)) fun_ord fun_ord"
unfolding fun_ord_def[abs_def] by transfer_prover
lemma Plus_parametric [transfer_rule]:
"(rel_set A ===> rel_set B ===> rel_set (rel_sum A B)) (<+>) (<+>)"
unfolding Plus_def[abs_def] by transfer_prover
lemma pred_fun_parametric [transfer_rule]:
assumes [transfer_rule]: "bi_total A"
shows "((A ===> (=)) ===> (B ===> (=)) ===> (A ===> B) ===> (=)) pred_fun pred_fun"
unfolding pred_fun_def by(transfer_prover)
lemma rel_fun_eq_OO: "((=) ===> A) OO ((=) ===> B) = ((=) ===> A OO B)"
by(clarsimp simp add: rel_fun_def fun_eq_iff relcompp.simps) metis
end
lemma Quotient_set_rel_eq:
includes lifting_syntax
assumes "Quotient R Abs Rep T"
shows "(rel_set T ===> rel_set T ===> (=)) (rel_set R) (=)"
proof(rule rel_funI iffI)+
fix A B C D
assume AB: "rel_set T A B" and CD: "rel_set T C D"
have *: "\<And>x y. R x y = (T x (Abs x) \<and> T y (Abs y) \<and> Abs x = Abs y)"
"\<And>a b. T a b \<Longrightarrow> Abs a = b"
using assms unfolding Quotient_alt_def by simp_all
{ assume [simp]: "B = D"
thus "rel_set R A C"
by(auto 4 4 intro!: rel_setI dest: rel_setD1[OF AB, simplified] rel_setD2[OF AB, simplified] rel_setD2[OF CD] rel_setD1[OF CD] simp add: * elim!: rev_bexI)
next
assume AC: "rel_set R A C"
show "B = D"
apply safe
apply(drule rel_setD2[OF AB], erule bexE)
apply(drule rel_setD1[OF AC], erule bexE)
apply(drule rel_setD1[OF CD], erule bexE)
apply(simp add: *)
apply(drule rel_setD2[OF CD], erule bexE)
apply(drule rel_setD2[OF AC], erule bexE)
apply(drule rel_setD1[OF AB], erule bexE)
apply(simp add: *)
done
}
qed
lemma Domainp_eq: "Domainp (=) = (\<lambda>_. True)"
by(simp add: Domainp.simps fun_eq_iff)
lemma rel_fun_eq_onpI: "eq_onp (pred_fun P Q) f g \<Longrightarrow> rel_fun (eq_onp P) (eq_onp Q) f g"
by(auto simp add: eq_onp_def rel_fun_def)
lemma bi_unique_eq_onp: "bi_unique (eq_onp P)"
by(simp add: bi_unique_def eq_onp_def)
lemma rel_fun_eq_conversep: includes lifting_syntax shows "(A\<inverse>\<inverse> ===> (=)) = (A ===> (=))\<inverse>\<inverse>"
by(auto simp add: fun_eq_iff rel_fun_def)
lemma rel_fun_comp:
"\<And>f g h. rel_fun A B (f \<circ> g) h = rel_fun A (\<lambda>x. B (f x)) g h"
"\<And>f g h. rel_fun A B f (g \<circ> h) = rel_fun A (\<lambda>x y. B x (g y)) f h"
by(auto simp add: rel_fun_def)
lemma rel_fun_map_fun1: "rel_fun (BNF_Def.Grp UNIV h)\<inverse>\<inverse> A f g \<Longrightarrow> rel_fun (=) A (map_fun h id f) g"
by(auto simp add: rel_fun_def Grp_def)
lemma map_fun2_id: "map_fun f g x = g \<circ> map_fun f id x"
by(simp add: map_fun_def o_assoc)
lemma map_fun_id2_in: "map_fun g h f = map_fun g id (h \<circ> f)"
by(simp add: map_fun_def)
lemma Domainp_rel_fun_le: "Domainp (rel_fun A B) \<le> pred_fun (Domainp A) (Domainp B)"
by(auto dest: rel_funD)
definition rel_witness_fun :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'd) \<times> ('c \<Rightarrow> 'e) \<Rightarrow> ('b \<Rightarrow> 'd \<times> 'e)" where
"rel_witness_fun A A' = (\<lambda>(f, g) b. (f (THE a. A a b), g (THE c. A' b c)))"
lemma
assumes fg: "rel_fun (A OO A') B f g"
and A: "left_unique A" "right_total A"
and A': "right_unique A'" "left_total A'"
shows rel_witness_fun1: "rel_fun A (\<lambda>x (x', y). x = x' \<and> B x' y) f (rel_witness_fun A A' (f, g))"
and rel_witness_fun2: "rel_fun A' (\<lambda>(x, y') y. y = y' \<and> B x y') (rel_witness_fun A A' (f, g)) g"
proof (goal_cases)
case 1
have "A x y \<Longrightarrow> f x = f (THE a. A a y) \<and> B (f (THE a. A a y)) (g (The (A' y)))" for x y
by(rule left_totalE[OF A'(2)]; erule meta_allE[of _ y]; erule exE; frule (1) fg[THEN rel_funD, OF relcomppI])
(auto intro!: arg_cong[where f=f] arg_cong[where f=g] rel_funI the_equality the_equality[symmetric] dest: left_uniqueD[OF A(1)] right_uniqueD[OF A'(1)] elim!: arg_cong2[where f=B, THEN iffD2, rotated -1])
with 1 show ?case by(clarsimp simp add: rel_fun_def rel_witness_fun_def)
next
case 2
have "A' x y \<Longrightarrow> g y = g (The (A' x)) \<and> B (f (THE a. A a x)) (g (The (A' x)))" for x y
by(rule right_totalE[OF A(2), of x]; frule (1) fg[THEN rel_funD, OF relcomppI])
(auto intro!: arg_cong[where f=f] arg_cong[where f=g] rel_funI the_equality the_equality[symmetric] dest: left_uniqueD[OF A(1)] right_uniqueD[OF A'(1)] elim!: arg_cong2[where f=B, THEN iffD2, rotated -1])
with 2 show ?case by(clarsimp simp add: rel_fun_def rel_witness_fun_def)
qed
lemma rel_witness_fun_eq [simp]: "rel_witness_fun (=) (=) (f, g) = (\<lambda>x. (f x, g x))"
by(simp add: rel_witness_fun_def)
subsection \<open>Arithmetic\<close>
lemma abs_diff_triangle_ineq2: "\<bar>a - b :: _ :: ordered_ab_group_add_abs\<bar> \<le> \<bar>a - c\<bar> + \<bar>c - b\<bar>"
by(rule order_trans[OF _ abs_diff_triangle_ineq]) simp
lemma (in ordered_ab_semigroup_add) add_left_mono_trans:
"\<lbrakk> x \<le> a + b; b \<le> c \<rbrakk> \<Longrightarrow> x \<le> a + c"
by(erule order_trans)(rule add_left_mono)
lemma of_nat_le_one_cancel_iff [simp]:
fixes n :: nat shows "real n \<le> 1 \<longleftrightarrow> n \<le> 1"
by linarith
lemma (in linordered_semidom) mult_right_le: "c \<le> 1 \<Longrightarrow> 0 \<le> a \<Longrightarrow> c * a \<le> a"
by(subst mult.commute)(rule mult_left_le)
subsection \<open>Chain-complete partial orders and \<open>partial_function\<close>\<close>
lemma fun_ordD: "fun_ord ord f g \<Longrightarrow> ord (f x) (g x)"
by(simp add: fun_ord_def)
lemma parallel_fixp_induct_strong:
assumes ccpo1: "class.ccpo luba orda (mk_less orda)"
and ccpo2: "class.ccpo lubb ordb (mk_less ordb)"
and adm: "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>x. P (fst x) (snd x))"
and f: "monotone orda orda f"
and g: "monotone ordb ordb g"
and bot: "P (luba {}) (lubb {})"
and step: "\<And>x y. \<lbrakk> orda x (ccpo.fixp luba orda f); ordb y (ccpo.fixp lubb ordb g); P x y \<rbrakk> \<Longrightarrow> P (f x) (g y)"
shows "P (ccpo.fixp luba orda f) (ccpo.fixp lubb ordb g)"
proof -
let ?P="\<lambda>x y. orda x (ccpo.fixp luba orda f) \<and> ordb y (ccpo.fixp lubb ordb g) \<and> P x y"
show ?thesis using ccpo1 ccpo2 _ f g
proof(rule parallel_fixp_induct[where P="?P", THEN conjunct2, THEN conjunct2])
note [cont_intro] =
admissible_leI[OF ccpo1] ccpo.mcont_const[OF ccpo1]
admissible_leI[OF ccpo2] ccpo.mcont_const[OF ccpo2]
show "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (\<lambda>xy. ?P (fst xy) (snd xy))"
using adm by simp
show "?P (luba {}) (lubb {})" using bot by(auto intro: ccpo.ccpo_Sup_least ccpo1 ccpo2 chain_empty)
show "?P (f x) (g y)" if "?P x y" for x y using that
apply(subst ccpo.fixp_unfold[OF ccpo1 f])
apply(subst ccpo.fixp_unfold[OF ccpo2 g])
apply(auto intro: step monotoneD[OF f] monotoneD[OF g])
done
qed
qed
lemma parallel_fixp_induct_strong_uc:
assumes a: "partial_function_definitions orda luba"
and b: "partial_function_definitions ordb lubb"
and F: "\<And>x. monotone (fun_ord orda) orda (\<lambda>f. U1 (F (C1 f)) x)"
and G: "\<And>y. monotone (fun_ord ordb) ordb (\<lambda>g. U2 (G (C2 g)) y)"
and eq1: "f \<equiv> C1 (ccpo.fixp (fun_lub luba) (fun_ord orda) (\<lambda>f. U1 (F (C1 f))))"
and eq2: "g \<equiv> C2 (ccpo.fixp (fun_lub lubb) (fun_ord ordb) (\<lambda>g. U2 (G (C2 g))))"
and inverse: "\<And>f. U1 (C1 f) = f"
and inverse2: "\<And>g. U2 (C2 g) = g"
and adm: "ccpo.admissible (prod_lub (fun_lub luba) (fun_lub lubb)) (rel_prod (fun_ord orda) (fun_ord ordb)) (\<lambda>x. P (fst x) (snd x))"
and bot: "P (\<lambda>_. luba {}) (\<lambda>_. lubb {})"
and step: "\<And>f' g'. \<lbrakk> \<And>x. orda (U1 f' x) (U1 f x); \<And>y. ordb (U2 g' y) (U2 g y); P (U1 f') (U2 g') \<rbrakk> \<Longrightarrow> P (U1 (F f')) (U2 (G g'))"
shows "P (U1 f) (U2 g)"
apply(unfold eq1 eq2 inverse inverse2)
apply(rule parallel_fixp_induct_strong[OF partial_function_definitions.ccpo[OF a] partial_function_definitions.ccpo[OF b] adm])
using F apply(simp add: monotone_def fun_ord_def)
using G apply(simp add: monotone_def fun_ord_def)
apply(simp add: fun_lub_def bot)
apply(rule step; simp add: inverse inverse2 eq1 eq2 fun_ordD)
done
lemmas parallel_fixp_induct_strong_1_1 = parallel_fixp_induct_strong_uc[
of _ _ _ _ "\<lambda>x. x" _ "\<lambda>x. x" "\<lambda>x. x" _ "\<lambda>x. x",
OF _ _ _ _ _ _ refl refl]
lemmas parallel_fixp_induct_strong_2_2 = parallel_fixp_induct_strong_uc[
of _ _ _ _ "case_prod" _ "curry" "case_prod" _ "curry",
where P="\<lambda>f g. P (curry f) (curry g)",
unfolded case_prod_curry curry_case_prod curry_K,
OF _ _ _ _ _ _ refl refl,
split_format (complete), unfolded prod.case]
for P
lemma fixp_induct_option': \<comment> \<open>Stronger induction rule\<close>
fixes F :: "'c \<Rightarrow> 'c" and
U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a option" and
C :: "('b \<Rightarrow> 'a option) \<Rightarrow> 'c" and
P :: "'b \<Rightarrow> 'a \<Rightarrow> bool"
assumes mono: "\<And>x. mono_option (\<lambda>f. U (F (C f)) x)"
assumes eq: "f \<equiv> C (ccpo.fixp (fun_lub (flat_lub None)) (fun_ord option_ord) (\<lambda>f. U (F (C f))))"
assumes inverse2: "\<And>f. U (C f) = f"
assumes step: "\<And>g x y. \<lbrakk> \<And>x y. U g x = Some y \<Longrightarrow> P x y; U (F g) x = Some y; \<And>x. option_ord (U g x) (U f x) \<rbrakk> \<Longrightarrow> P x y"
assumes defined: "U f x = Some y"
shows "P x y"
using step defined option.fixp_strong_induct_uc[of U F C, OF mono eq inverse2 option_admissible, of P]
unfolding fun_lub_def flat_lub_def fun_ord_def
by(simp (no_asm_use)) blast
declaration \<open>Partial_Function.init "option'" @{term option.fixp_fun}
@{term option.mono_body} @{thm option.fixp_rule_uc} @{thm option.fixp_induct_uc}
(SOME @{thm fixp_induct_option'})\<close>
lemma bot_fun_least [simp]: "(\<lambda>_. bot :: 'a :: order_bot) \<le> x"
by(fold bot_fun_def) simp
lemma fun_ord_conv_rel_fun: "fun_ord = rel_fun (=)"
by(simp add: fun_ord_def fun_eq_iff rel_fun_def)
inductive finite_chains :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
for ord
where finite_chainsI: "(\<And>Y. Complete_Partial_Order.chain ord Y \<Longrightarrow> finite Y) \<Longrightarrow> finite_chains ord"
lemma finite_chainsD: "\<lbrakk> finite_chains ord; Complete_Partial_Order.chain ord Y \<rbrakk> \<Longrightarrow> finite Y"
by(rule finite_chains.cases)
lemma finite_chains_flat_ord [simp, intro!]: "finite_chains (flat_ord x)"
proof
fix Y
assume chain: "Complete_Partial_Order.chain (flat_ord x) Y"
show "finite Y"
proof(cases "\<exists>y \<in> Y. y \<noteq> x")
case True
then obtain y where y: "y \<in> Y" and yx: "y \<noteq> x" by blast
hence "Y \<subseteq> {x, y}" by(auto dest: chainD[OF chain] simp add: flat_ord_def)
thus ?thesis by(rule finite_subset) simp
next
case False
hence "Y \<subseteq> {x}" by auto
thus ?thesis by(rule finite_subset) simp
qed
qed
lemma mcont_finite_chains:
assumes finite: "finite_chains ord"
and mono: "monotone ord ord' f"
and ccpo: "class.ccpo lub ord (mk_less ord)"
and ccpo': "class.ccpo lub' ord' (mk_less ord')"
shows "mcont lub ord lub' ord' f"
proof(intro mcontI contI)
fix Y
assume chain: "Complete_Partial_Order.chain ord Y" and Y: "Y \<noteq> {}"
from finite chain have fin: "finite Y" by(rule finite_chainsD)
from ccpo chain fin Y have lub: "lub Y \<in> Y" by(rule ccpo.in_chain_finite)
interpret ccpo': ccpo lub' ord' "mk_less ord'" by(rule ccpo')
have chain': "Complete_Partial_Order.chain ord' (f ` Y)" using chain
by(rule chain_imageI)(rule monotoneD[OF mono])
have "ord' (f (lub Y)) (lub' (f ` Y))" using chain'
by(rule ccpo'.ccpo_Sup_upper)(simp add: lub)
moreover
have "ord' (lub' (f ` Y)) (f (lub Y))" using chain'
by(rule ccpo'.ccpo_Sup_least)(blast intro: monotoneD[OF mono] ccpo.ccpo_Sup_upper[OF ccpo chain])
- ultimately show "f (lub Y) = lub' (f ` Y)" by(rule ccpo'.antisym)
+ ultimately show "f (lub Y) = lub' (f ` Y)" by(rule ccpo'.order.antisym)
qed(fact mono)
lemma rel_fun_curry: includes lifting_syntax shows
"(A ===> B ===> C) f g \<longleftrightarrow> (rel_prod A B ===> C) (case_prod f) (case_prod g)"
by(auto simp add: rel_fun_def)
lemma (in ccpo) Sup_image_mono:
assumes ccpo: "class.ccpo luba orda lessa"
and mono: "monotone orda (\<le>) f"
and chain: "Complete_Partial_Order.chain orda A"
and "A \<noteq> {}"
shows "Sup (f ` A) \<le> (f (luba A))"
proof(rule ccpo_Sup_least)
from chain show "Complete_Partial_Order.chain (\<le>) (f ` A)"
by(rule chain_imageI)(rule monotoneD[OF mono])
fix x
assume "x \<in> f ` A"
then obtain y where "x = f y" "y \<in> A" by blast
from \<open>y \<in> A\<close> have "orda y (luba A)" by(rule ccpo.ccpo_Sup_upper[OF ccpo chain])
hence "f y \<le> f (luba A)" by(rule monotoneD[OF mono])
thus "x \<le> f (luba A)" using \<open>x = f y\<close> by simp
qed
lemma (in ccpo) admissible_le_mono:
assumes "monotone (\<le>) (\<le>) f"
shows "ccpo.admissible Sup (\<le>) (\<lambda>x. x \<le> f x)"
proof(rule ccpo.admissibleI)
fix Y
assume chain: "Complete_Partial_Order.chain (\<le>) Y"
and Y: "Y \<noteq> {}"
and le [rule_format]: "\<forall>x\<in>Y. x \<le> f x"
have "\<Squnion>Y \<le> \<Squnion>(f ` Y)" using chain
by(rule ccpo_Sup_least)(rule order_trans[OF le]; blast intro!: ccpo_Sup_upper chain_imageI[OF chain] intro: monotoneD[OF assms])
also have "\<dots> \<le> f (\<Squnion>Y)"
by(rule Sup_image_mono[OF _ assms chain Y, where lessa="(<)"]) unfold_locales
finally show "\<Squnion>Y \<le> \<dots>" .
qed
lemma (in ccpo) fixp_induct_strong2:
assumes adm: "ccpo.admissible Sup (\<le>) P"
and mono: "monotone (\<le>) (\<le>) f"
and bot: "P (\<Squnion>{})"
and step: "\<And>x. \<lbrakk> x \<le> ccpo_class.fixp f; x \<le> f x; P x \<rbrakk> \<Longrightarrow> P (f x)"
shows "P (ccpo_class.fixp f)"
proof(rule fixp_strong_induct[where P="\<lambda>x. x \<le> f x \<and> P x", THEN conjunct2])
show "ccpo.admissible Sup (\<le>) (\<lambda>x. x \<le> f x \<and> P x)"
using admissible_le_mono adm by(rule admissible_conj)(rule mono)
next
show "\<Squnion>{} \<le> f (\<Squnion>{}) \<and> P (\<Squnion>{})"
by(auto simp add: bot chain_empty intro: ccpo_Sup_least)
next
fix x
assume "x \<le> ccpo_class.fixp f" "x \<le> f x \<and> P x"
thus "f x \<le> f (f x) \<and> P (f x)"
by(auto dest: monotoneD[OF mono] intro: step)
qed(rule mono)
context partial_function_definitions begin
lemma fixp_induct_strong2_uc:
fixes F :: "'c \<Rightarrow> 'c"
and U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a"
and C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c"
and P :: "('b \<Rightarrow> 'a) \<Rightarrow> bool"
assumes mono: "\<And>x. mono_body (\<lambda>f. U (F (C f)) x)"
and eq: "f \<equiv> C (fixp_fun (\<lambda>f. U (F (C f))))"
and inverse: "\<And>f. U (C f) = f"
and adm: "ccpo.admissible lub_fun le_fun P"
and bot: "P (\<lambda>_. lub {})"
and step: "\<And>f'. \<lbrakk> le_fun (U f') (U f); le_fun (U f') (U (F f')); P (U f') \<rbrakk> \<Longrightarrow> P (U (F f'))"
shows "P (U f)"
unfolding eq inverse
apply (rule ccpo.fixp_induct_strong2[OF ccpo adm])
apply (insert mono, auto simp: monotone_def fun_ord_def bot fun_lub_def)[2]
apply (rule_tac f'5="C x" in step)
apply (simp_all add: inverse eq)
done
end
lemmas parallel_fixp_induct_2_4 = parallel_fixp_induct_uc[
of _ _ _ _ "case_prod" _ "curry" "\<lambda>f. case_prod (case_prod (case_prod f))" _ "\<lambda>f. curry (curry (curry f))",
where P="\<lambda>f g. P (curry f) (curry (curry (curry g)))",
unfolded case_prod_curry curry_case_prod curry_K,
OF _ _ _ _ _ _ refl refl]
for P
lemma (in ccpo) fixp_greatest:
assumes f: "monotone (\<le>) (\<le>) f"
and ge: "\<And>y. f y \<le> y \<Longrightarrow> x \<le> y"
shows "x \<le> ccpo.fixp Sup (\<le>) f"
by(rule ge)(simp add: fixp_unfold[OF f, symmetric])
lemma fixp_rolling:
assumes "class.ccpo lub1 leq1 (mk_less leq1)"
and "class.ccpo lub2 leq2 (mk_less leq2)"
and f: "monotone leq1 leq2 f"
and g: "monotone leq2 leq1 g"
shows "ccpo.fixp lub1 leq1 (\<lambda>x. g (f x)) = g (ccpo.fixp lub2 leq2 (\<lambda>x. f (g x)))"
proof -
interpret c1: ccpo lub1 leq1 "mk_less leq1" by fact
interpret c2: ccpo lub2 leq2 "mk_less leq2" by fact
show ?thesis
- proof(rule c1.antisym)
+ proof(rule c1.order.antisym)
have fg: "monotone leq2 leq2 (\<lambda>x. f (g x))" using f g by(rule monotone2monotone) simp_all
have gf: "monotone leq1 leq1 (\<lambda>x. g (f x))" using g f by(rule monotone2monotone) simp_all
show "leq1 (c1.fixp (\<lambda>x. g (f x))) (g (c2.fixp (\<lambda>x. f (g x))))" using gf
by(rule c1.fixp_lowerbound)(subst (2) c2.fixp_unfold[OF fg], simp)
show "leq1 (g (c2.fixp (\<lambda>x. f (g x)))) (c1.fixp (\<lambda>x. g (f x)))" using gf
proof(rule c1.fixp_greatest)
fix u
assume u: "leq1 (g (f u)) u"
have "leq1 (g (c2.fixp (\<lambda>x. f (g x)))) (g (f u))"
by(intro monotoneD[OF g] c2.fixp_lowerbound[OF fg] monotoneD[OF f u])
then show "leq1 (g (c2.fixp (\<lambda>x. f (g x)))) u" using u by(rule c1.order_trans)
qed
qed
qed
lemma fixp_lfp_parametric_eq:
includes lifting_syntax
assumes f: "\<And>x. lfp.mono_body (\<lambda>f. F f x)"
and g: "\<And>x. lfp.mono_body (\<lambda>f. G f x)"
and param: "((A ===> (=)) ===> A ===> (=)) F G"
shows "(A ===> (=)) (lfp.fixp_fun F) (lfp.fixp_fun G)"
using f g
proof(rule parallel_fixp_induct_1_1[OF complete_lattice_partial_function_definitions complete_lattice_partial_function_definitions _ _ reflexive reflexive, where P="(A ===> (=))"])
show "ccpo.admissible (prod_lub lfp.lub_fun lfp.lub_fun) (rel_prod lfp.le_fun lfp.le_fun) (\<lambda>x. (A ===> (=)) (fst x) (snd x))"
unfolding rel_fun_def by simp
show "(A ===> (=)) (\<lambda>_. \<Squnion>{}) (\<lambda>_. \<Squnion>{})" by auto
show "(A ===> (=)) (F f) (G g)" if "(A ===> (=)) f g" for f g
using that by(rule rel_funD[OF param])
qed
lemma mono2mono_map_option[THEN option.mono2mono, simp, cont_intro]:
shows monotone_map_option: "monotone option_ord option_ord (map_option f)"
by(rule monotoneI)(auto simp add: flat_ord_def)
lemma mcont2mcont_map_option[THEN option.mcont2mcont, simp, cont_intro]:
shows mcont_map_option: "mcont (flat_lub None) option_ord (flat_lub None) option_ord (map_option f)"
by(rule mcont_finite_chains[OF _ _ flat_interpretation[THEN ccpo] flat_interpretation[THEN ccpo]]) simp_all
lemma mono2mono_set_option [THEN lfp.mono2mono]:
shows monotone_set_option: "monotone option_ord (\<subseteq>) set_option"
by(auto intro!: monotoneI simp add: option_ord_Some1_iff)
lemma mcont2mcont_set_option [THEN lfp.mcont2mcont, cont_intro, simp]:
shows mcont_set_option: "mcont (flat_lub None) option_ord Union (\<subseteq>) set_option"
by(rule mcont_finite_chains)(simp_all add: monotone_set_option ccpo option.partial_function_definitions_axioms)
lemma eadd_gfp_partial_function_mono [partial_function_mono]:
"\<lbrakk> monotone (fun_ord (\<ge>)) (\<ge>) f; monotone (fun_ord (\<ge>)) (\<ge>) g \<rbrakk>
\<Longrightarrow> monotone (fun_ord (\<ge>)) (\<ge>) (\<lambda>x. f x + g x :: enat)"
by(rule mono2mono_gfp_eadd)
lemma map_option_mono [partial_function_mono]:
"mono_option B \<Longrightarrow> mono_option (\<lambda>f. map_option g (B f))"
unfolding map_conv_bind_option by(rule bind_mono) simp_all
subsection \<open>Folding over finite sets\<close>
lemma (in comp_fun_commute) fold_invariant_remove [consumes 1, case_names start step]:
assumes fin: "finite A"
and start: "I A s"
and step: "\<And>x s A'. \<lbrakk> x \<in> A'; I A' s; A' \<subseteq> A \<rbrakk> \<Longrightarrow> I (A' - {x}) (f x s)"
shows "I {} (Finite_Set.fold f s A)"
proof -
define A' where "A' == A"
with fin start have "finite A'" "A' \<subseteq> A" "I A' s" by simp_all
thus "I {} (Finite_Set.fold f s A')"
proof(induction arbitrary: s)
case empty thus ?case by simp
next
case (insert x A')
let ?A' = "insert x A'"
have "x \<in> ?A'" "I ?A' s" "?A' \<subseteq> A" using insert by auto
hence "I (?A' - {x}) (f x s)" by(rule step)
with insert have "A' \<subseteq> A" "I A' (f x s)" by auto
hence "I {} (Finite_Set.fold f (f x s) A')" by(rule insert.IH)
thus ?case using insert by(simp add: fold_insert2 del: fold_insert)
qed
qed
lemma (in comp_fun_commute) fold_invariant_insert [consumes 1, case_names start step]:
assumes fin: "finite A"
and start: "I {} s"
and step: "\<And>x s A'. \<lbrakk> I A' s; x \<notin> A'; x \<in> A; A' \<subseteq> A \<rbrakk> \<Longrightarrow> I (insert x A') (f x s)"
shows "I A (Finite_Set.fold f s A)"
using fin start
proof(rule fold_invariant_remove[where I="\<lambda>A'. I (A - A')" and A=A and s=s, simplified])
fix x s A'
assume *: "x \<in> A'" "I (A - A') s" "A' \<subseteq> A"
hence "x \<notin> A - A'" "x \<in> A" "A - A' \<subseteq> A" by auto
with \<open>I (A - A') s\<close> have "I (insert x (A - A')) (f x s)" by(rule step)
also have "insert x (A - A') = A - (A' - {x})" using * by auto
finally show "I \<dots> (f x s)" .
qed
lemma (in comp_fun_idem) fold_set_union:
assumes "finite A" "finite B"
shows "Finite_Set.fold f z (A \<union> B) = Finite_Set.fold f (Finite_Set.fold f z A) B"
using assms(2,1) by induction simp_all
subsection \<open>Parametrisation of transfer rules\<close>
attribute_setup transfer_parametric = \<open>
Attrib.thm >> (fn parametricity =>
Thm.rule_attribute [] (fn context => fn transfer_rule =>
let
val ctxt = Context.proof_of context;
val thm' = Lifting_Term.parametrize_transfer_rule ctxt transfer_rule
in Lifting_Def.generate_parametric_transfer_rule ctxt thm' parametricity
end
handle Lifting_Term.MERGE_TRANSFER_REL msg => error (Pretty.string_of msg)
))
\<close> "combine transfer rule with parametricity theorem"
subsection \<open>Lists\<close>
lemma nth_eq_tlI: "xs ! n = z \<Longrightarrow> (x # xs) ! Suc n = z"
by simp
lemma list_all2_append':
"length us = length vs \<Longrightarrow> list_all2 P (xs @ us) (ys @ vs) \<longleftrightarrow> list_all2 P xs ys \<and> list_all2 P us vs"
by(auto simp add: list_all2_append1 list_all2_append2 dest: list_all2_lengthD)
definition disjointp :: "('a \<Rightarrow> bool) list \<Rightarrow> bool"
where "disjointp xs = disjoint_family_on (\<lambda>n. {x. (xs ! n) x}) {0..<length xs}"
lemma disjointpD:
"\<lbrakk> disjointp xs; (xs ! n) x; (xs ! m) x; n < length xs; m < length xs \<rbrakk> \<Longrightarrow> n = m"
by(auto 4 3 simp add: disjointp_def disjoint_family_on_def)
lemma disjointpD':
"\<lbrakk> disjointp xs; P x; Q x; xs ! n = P; xs ! m = Q; n < length xs; m < length xs \<rbrakk> \<Longrightarrow> n = m"
by(auto 4 3 simp add: disjointp_def disjoint_family_on_def)
lemma wf_strict_prefix: "wfP strict_prefix"
proof -
from wf have "wf (inv_image {(x, y). x < y} length)" by(rule wf_inv_image)
moreover have "{(x, y). strict_prefix x y} \<subseteq> inv_image {(x, y). x < y} length" by(auto intro: prefix_length_less)
ultimately show ?thesis unfolding wfP_def by(rule wf_subset)
qed
lemma strict_prefix_setD:
"strict_prefix xs ys \<Longrightarrow> set xs \<subseteq> set ys"
by(auto simp add: strict_prefix_def prefix_def)
subsubsection \<open>List of a given length\<close>
inductive_set nlists :: "'a set \<Rightarrow> nat \<Rightarrow> 'a list set" for A n
where nlists: "\<lbrakk> set xs \<subseteq> A; length xs = n \<rbrakk> \<Longrightarrow> xs \<in> nlists A n"
hide_fact (open) nlists
lemma nlists_alt_def: "nlists A n = {xs. set xs \<subseteq> A \<and> length xs = n}"
by(auto simp add: nlists.simps)
lemma nlists_empty: "nlists {} n = (if n = 0 then {[]} else {})"
by(auto simp add: nlists_alt_def)
lemma nlists_empty_gt0 [simp]: "n > 0 \<Longrightarrow> nlists {} n = {}"
by(simp add: nlists_empty)
lemma nlists_0 [simp]: "nlists A 0 = {[]}"
by(auto simp add: nlists_alt_def)
lemma Cons_in_nlists_Suc [simp]: "x # xs \<in> nlists A (Suc n) \<longleftrightarrow> x \<in> A \<and> xs \<in> nlists A n"
by(simp add: nlists_alt_def)
lemma Nil_in_nlists [simp]: "[] \<in> nlists A n \<longleftrightarrow> n = 0"
by(auto simp add: nlists_alt_def)
lemma Cons_in_nlists_iff: "x # xs \<in> nlists A n \<longleftrightarrow> (\<exists>n'. n = Suc n' \<and> x \<in> A \<and> xs \<in> nlists A n')"
by(cases n) simp_all
lemma in_nlists_Suc_iff: "xs \<in> nlists A (Suc n) \<longleftrightarrow> (\<exists>x xs'. xs = x # xs' \<and> x \<in> A \<and> xs' \<in> nlists A n)"
by(cases xs) simp_all
lemma nlists_Suc: "nlists A (Suc n) = (\<Union>x\<in>A. (#) x ` nlists A n)"
by(auto 4 3 simp add: in_nlists_Suc_iff intro: rev_image_eqI)
lemma replicate_in_nlists [simp, intro]: "x \<in> A \<Longrightarrow> replicate n x \<in> nlists A n"
by(simp add: nlists_alt_def set_replicate_conv_if)
lemma nlists_eq_empty_iff [simp]: "nlists A n = {} \<longleftrightarrow> n > 0 \<and> A = {}"
using replicate_in_nlists by(cases n)(auto)
lemma finite_nlists [simp]: "finite A \<Longrightarrow> finite (nlists A n)"
by(induction n)(simp_all add: nlists_Suc)
lemma finite_nlistsD:
assumes "finite (nlists A n)"
shows "finite A \<or> n = 0"
proof(rule disjCI)
assume "n \<noteq> 0"
then obtain n' where n: "n = Suc n'" by(cases n)auto
then have "A = hd ` nlists A n" by(auto 4 4 simp add: nlists_Suc intro: rev_image_eqI rev_bexI)
also have "finite \<dots>" using assms ..
finally show "finite A" .
qed
lemma finite_nlists_iff: "finite (nlists A n) \<longleftrightarrow> finite A \<or> n = 0"
by(auto dest: finite_nlistsD)
lemma card_nlists: "card (nlists A n) = card A ^ n"
proof(induction n)
case (Suc n)
have "card (\<Union>x\<in>A. (#) x ` nlists A n) = card A * card (nlists A n)"
proof(cases "finite A")
case True
then show ?thesis by(subst card_UN_disjoint)(auto simp add: card_image inj_on_def)
next
case False
hence "\<not> finite (\<Union>x\<in>A. (#) x ` nlists A n)"
unfolding nlists_Suc[symmetric] by(auto dest: finite_nlistsD)
then show ?thesis using False by simp
qed
then show ?case using Suc.IH by(simp add: nlists_Suc)
qed simp
lemma in_nlists_UNIV: "xs \<in> nlists UNIV n \<longleftrightarrow> length xs = n"
by(simp add: nlists_alt_def)
subsubsection \<open> The type of lists of a given length \<close>
typedef (overloaded) ('a, 'b :: len0) nlist = "nlists (UNIV :: 'a set) (LENGTH('b))"
proof
show "replicate LENGTH('b) undefined \<in> ?nlist" by simp
qed
setup_lifting type_definition_nlist
subsection \<open>Streams and infinite lists\<close>
primrec sprefix :: "'a list \<Rightarrow> 'a stream \<Rightarrow> bool" where
sprefix_Nil: "sprefix [] ys = True"
| sprefix_Cons: "sprefix (x # xs) ys \<longleftrightarrow> x = shd ys \<and> sprefix xs (stl ys)"
lemma sprefix_append: "sprefix (xs @ ys) zs \<longleftrightarrow> sprefix xs zs \<and> sprefix ys (sdrop (length xs) zs)"
by(induct xs arbitrary: zs) simp_all
lemma sprefix_stake_same [simp]: "sprefix (stake n xs) xs"
by(induct n arbitrary: xs) simp_all
lemma sprefix_same_imp_eq:
assumes "sprefix xs ys" "sprefix xs' ys"
and "length xs = length xs'"
shows "xs = xs'"
using assms(3,1,2) by(induct arbitrary: ys rule: list_induct2) auto
lemma sprefix_shift_same [simp]:
"sprefix xs (xs @- ys)"
by(induct xs) simp_all
lemma sprefix_shift [simp]:
"length xs \<le> length ys \<Longrightarrow> sprefix xs (ys @- zs) \<longleftrightarrow> prefix xs ys"
by(induct xs arbitrary: ys)(simp, case_tac ys, auto)
lemma prefixeq_stake2 [simp]: "prefix xs (stake n ys) \<longleftrightarrow> length xs \<le> n \<and> sprefix xs ys"
proof(induct xs arbitrary: n ys)
case (Cons x xs)
thus ?case by(cases ys n rule: stream.exhaust[case_product nat.exhaust]) auto
qed simp
lemma tlength_eq_infinity_iff: "tlength xs = \<infinity> \<longleftrightarrow> \<not> tfinite xs"
including tllist.lifting by transfer(simp add: llength_eq_infty_conv_lfinite)
subsection \<open>Monomorphic monads\<close>
context includes lifting_syntax begin
local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "monad")\<close>
definition bind_option :: "'m fail \<Rightarrow> 'a option \<Rightarrow> ('a \<Rightarrow> 'm) \<Rightarrow> 'm"
where "bind_option fail x f = (case x of None \<Rightarrow> fail | Some x' \<Rightarrow> f x')" for fail
simps_of_case bind_option_simps [simp]: bind_option_def
lemma bind_option_parametric [transfer_rule]:
"(M ===> rel_option B ===> (B ===> M) ===> M) bind_option bind_option"
unfolding bind_option_def by transfer_prover
lemma bind_option_K:
"\<And>monad. (x = None \<Longrightarrow> m = fail) \<Longrightarrow> bind_option fail x (\<lambda>_. m) = m"
by(cases x) simp_all
end
lemma bind_option_option [simp]: "monad.bind_option None = Option.bind"
by(simp add: monad.bind_option_def fun_eq_iff split: option.split)
context monad_fail_hom begin
lemma hom_bind_option: "h (monad.bind_option fail1 x f) = monad.bind_option fail2 x (h \<circ> f)"
by(cases x)(simp_all)
end
lemma bind_option_set [simp]: "monad.bind_option fail_set = (\<lambda>x f. \<Union> (f ` set_option x))"
by(simp add: monad.bind_option_def fun_eq_iff split: option.split)
lemma run_bind_option_stateT [simp]:
"\<And>more. run_state (monad.bind_option (fail_state fail) x f) s =
monad.bind_option fail x (\<lambda>y. run_state (f y) s)"
by(cases x) simp_all
lemma run_bind_option_envT [simp]:
"\<And>more. run_env (monad.bind_option (fail_env fail) x f) s =
monad.bind_option fail x (\<lambda>y. run_env (f y) s)"
by(cases x) simp_all
subsection \<open>Measures\<close>
declare sets_restrict_space_count_space [measurable_cong]
lemma (in sigma_algebra) sets_Collect_countable_Ex1:
"(\<And>i :: 'i :: countable. {x \<in> \<Omega>. P i x} \<in> M) \<Longrightarrow> {x \<in> \<Omega>. \<exists>!i. P i x} \<in> M"
using sets_Collect_countable_Ex1'[of "UNIV :: 'i set"] by simp
lemma pred_countable_Ex1 [measurable]:
"(\<And>i :: _ :: countable. Measurable.pred M (\<lambda>x. P i x))
\<Longrightarrow> Measurable.pred M (\<lambda>x. \<exists>!i. P i x)"
unfolding pred_def by(rule sets.sets_Collect_countable_Ex1)
lemma measurable_snd_count_space [measurable]:
"A \<subseteq> B \<Longrightarrow> snd \<in> measurable (M1 \<Otimes>\<^sub>M count_space A) (count_space B)"
by(auto simp add: measurable_def space_pair_measure snd_vimage_eq_Times Times_Int_Times)
lemma integrable_scale_measure [simp]:
"\<lbrakk> integrable M f; r < \<top> \<rbrakk> \<Longrightarrow> integrable (scale_measure r M) f"
for f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
by(auto simp add: integrable_iff_bounded nn_integral_scale_measure ennreal_mult_less_top)
lemma integral_scale_measure:
assumes "integrable M f" "r < \<top>"
shows "integral\<^sup>L (scale_measure r M) f = enn2real r * integral\<^sup>L M f"
using assms
apply(subst (1 2) real_lebesgue_integral_def)
apply(simp_all add: nn_integral_scale_measure ennreal_enn2real_if)
by(auto simp add: ennreal_mult_less_top ennreal_less_top_iff ennreal_mult_eq_top_iff enn2real_mult right_diff_distrib elim!: integrableE)
subsection \<open>Sequence space\<close>
lemma (in sequence_space) nn_integral_split:
assumes f[measurable]: "f \<in> borel_measurable S"
shows "(\<integral>\<^sup>+\<omega>. f \<omega> \<partial>S) = (\<integral>\<^sup>+\<omega>. (\<integral>\<^sup>+\<omega>'. f (comb_seq i \<omega> \<omega>') \<partial>S) \<partial>S)"
by (subst PiM_comb_seq[symmetric, where i=i])
(simp add: nn_integral_distr P.nn_integral_fst[symmetric])
lemma (in sequence_space) prob_Collect_split:
assumes f[measurable]: "{x\<in>space S. P x} \<in> sets S"
shows "\<P>(x in S. P x) = (\<integral>\<^sup>+x. \<P>(x' in S. P (comb_seq i x x')) \<partial>S)"
proof -
have "\<P>(x in S. P x) = (\<integral>\<^sup>+x. (\<integral>\<^sup>+x'. indicator {x\<in>space S. P x} (comb_seq i x x') \<partial>S) \<partial>S)"
using nn_integral_split[of "indicator {x\<in>space S. P x}"] by (auto simp: emeasure_eq_measure)
also have "\<dots> = (\<integral>\<^sup>+x. \<P>(x' in S. P (comb_seq i x x')) \<partial>S)"
by (intro nn_integral_cong) (auto simp: emeasure_eq_measure nn_integral_indicator_map)
finally show ?thesis .
qed
subsection \<open>Probability mass functions\<close>
lemma measure_map_pmf_conv_distr:
"measure_pmf (map_pmf f p) = distr (measure_pmf p) (count_space UNIV) f"
by(fact map_pmf_rep_eq)
abbreviation coin_pmf :: "bool pmf" where "coin_pmf \<equiv> pmf_of_set UNIV"
text \<open>The rule @{thm [source] rel_pmf_bindI} is not complete as a program logic.\<close>
notepad begin
define x where "x = pmf_of_set {True, False}"
define y where "y = pmf_of_set {True, False}"
define f where "f x = pmf_of_set {True, False}" for x :: bool
define g :: "bool \<Rightarrow> bool pmf" where "g = return_pmf"
define P :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "P = (=)"
have "rel_pmf P (bind_pmf x f) (bind_pmf y g)"
by(simp add: P_def f_def[abs_def] g_def y_def bind_return_pmf' pmf.rel_eq)
have "\<not> R x y" if "\<And>x y. R x y \<Longrightarrow> rel_pmf P (f x) (g y)" for R x y
\<comment> \<open>Only the empty relation satisfies @{thm [source] rel_pmf_bindI}'s second premise.\<close>
proof
assume "R x y"
hence "rel_pmf P (f x) (g y)" by(rule that)
thus False by(auto simp add: P_def f_def g_def rel_pmf_return_pmf2)
qed
define R where "R x y = False" for x y :: bool
have "\<not> rel_pmf R x y" by(simp add: R_def[abs_def])
end
lemma pred_rel_pmf:
"\<lbrakk> pred_pmf P p; rel_pmf R p q \<rbrakk> \<Longrightarrow> pred_pmf (Imagep R P) q"
unfolding pred_pmf_def
apply(rule ballI)
apply(unfold rel_pmf.simps)
apply(erule exE conjE)+
apply hypsubst
apply(unfold pmf.set_map)
apply(erule imageE, hypsubst)
apply(drule bspec)
apply(erule rev_image_eqI)
apply(rule refl)
apply(erule Imagep.intros)
apply(erule allE)+
apply(erule mp)
apply(unfold prod.collapse)
apply assumption
done
lemma pmf_rel_mono': "\<lbrakk> rel_pmf P x y; P \<le> Q \<rbrakk> \<Longrightarrow> rel_pmf Q x y"
by(drule pmf.rel_mono) (auto)
lemma rel_pmf_eqI [simp]: "rel_pmf (=) x x"
by(simp add: pmf.rel_eq)
lemma rel_pmf_bind_reflI:
"(\<And>x. x \<in> set_pmf p \<Longrightarrow> rel_pmf R (f x) (g x))
\<Longrightarrow> rel_pmf R (bind_pmf p f) (bind_pmf p g)"
by(rule rel_pmf_bindI[where R="\<lambda>x y. x = y \<and> x \<in> set_pmf p"])(auto intro: rel_pmf_reflI)
lemma pmf_pred_mono_strong:
"\<lbrakk> pred_pmf P p; \<And>a. \<lbrakk> a \<in> set_pmf p; P a \<rbrakk> \<Longrightarrow> P' a \<rbrakk> \<Longrightarrow> pred_pmf P' p"
by(simp add: pred_pmf_def)
lemma rel_pmf_restrict_relpI [intro?]:
"\<lbrakk> rel_pmf R x y; pred_pmf P x; pred_pmf Q y \<rbrakk> \<Longrightarrow> rel_pmf (R \<upharpoonleft> P \<otimes> Q) x y"
by(erule pmf.rel_mono_strong)(simp add: pred_pmf_def)
lemma rel_pmf_restrict_relpE [elim?]:
assumes "rel_pmf (R \<upharpoonleft> P \<otimes> Q) x y"
obtains "rel_pmf R x y" "pred_pmf P x" "pred_pmf Q y"
proof
show "rel_pmf R x y" using assms by(auto elim!: pmf.rel_mono_strong)
have "pred_pmf (Domainp (R \<upharpoonleft> P \<otimes> Q)) x" using assms by(fold pmf.Domainp_rel) blast
then show "pred_pmf P x" by(rule pmf_pred_mono_strong)(blast dest!: restrict_relp_DomainpD)
have "pred_pmf (Domainp (R \<upharpoonleft> P \<otimes> Q)\<inverse>\<inverse>) y" using assms
by(fold pmf.Domainp_rel)(auto simp only: pmf.rel_conversep Domainp_conversep)
then show "pred_pmf Q y" by(rule pmf_pred_mono_strong)(auto dest!: restrict_relp_DomainpD)
qed
lemma rel_pmf_restrict_relp_iff:
"rel_pmf (R \<upharpoonleft> P \<otimes> Q) x y \<longleftrightarrow> rel_pmf R x y \<and> pred_pmf P x \<and> pred_pmf Q y"
by(blast intro: rel_pmf_restrict_relpI elim: rel_pmf_restrict_relpE)
lemma rel_pmf_OO_trans [trans]:
"\<lbrakk> rel_pmf R p q; rel_pmf S q r \<rbrakk> \<Longrightarrow> rel_pmf (R OO S) p r"
unfolding pmf.rel_compp by blast
lemma pmf_pred_map [simp]: "pred_pmf P (map_pmf f p) = pred_pmf (P \<circ> f) p"
by(simp add: pred_pmf_def)
lemma pred_pmf_bind [simp]: "pred_pmf P (bind_pmf p f) = pred_pmf (pred_pmf P \<circ> f) p"
by(simp add: pred_pmf_def)
lemma pred_pmf_return [simp]: "pred_pmf P (return_pmf x) = P x"
by(simp add: pred_pmf_def)
lemma pred_pmf_of_set [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> pred_pmf P (pmf_of_set A) = Ball A P"
by(simp add: pred_pmf_def)
lemma pred_pmf_of_multiset [simp]: "M \<noteq> {#} \<Longrightarrow> pred_pmf P (pmf_of_multiset M) = Ball (set_mset M) P"
by(simp add: pred_pmf_def)
lemma pred_pmf_cond [simp]:
"set_pmf p \<inter> A \<noteq> {} \<Longrightarrow> pred_pmf P (cond_pmf p A) = pred_pmf (\<lambda>x. x \<in> A \<longrightarrow> P x) p"
by(auto simp add: pred_pmf_def)
lemma pred_pmf_pair [simp]:
"pred_pmf P (pair_pmf p q) = pred_pmf (\<lambda>x. pred_pmf (P \<circ> Pair x) q) p"
by(simp add: pred_pmf_def)
lemma pred_pmf_join [simp]: "pred_pmf P (join_pmf p) = pred_pmf (pred_pmf P) p"
by(simp add: pred_pmf_def)
lemma pred_pmf_bernoulli [simp]: "\<lbrakk> 0 < p; p < 1 \<rbrakk> \<Longrightarrow> pred_pmf P (bernoulli_pmf p) = All P"
by(simp add: pred_pmf_def)
lemma pred_pmf_geometric [simp]: "\<lbrakk> 0 < p; p < 1 \<rbrakk> \<Longrightarrow> pred_pmf P (geometric_pmf p) = All P"
by(simp add: pred_pmf_def set_pmf_geometric)
lemma pred_pmf_poisson [simp]: "0 < rate \<Longrightarrow> pred_pmf P (poisson_pmf rate) = All P"
by(simp add: pred_pmf_def)
lemma pmf_rel_map_restrict_relp:
shows pmf_rel_map_restrict_relp1: "rel_pmf (R \<upharpoonleft> P \<otimes> Q) (map_pmf f p) = rel_pmf (R \<circ> f \<upharpoonleft> P \<circ> f \<otimes> Q) p"
and pmf_rel_map_restrict_relp2: "rel_pmf (R \<upharpoonleft> P \<otimes> Q) p (map_pmf g q) = rel_pmf ((\<lambda>x. R x \<circ> g) \<upharpoonleft> P \<otimes> Q \<circ> g) p q"
by(simp_all add: pmf.rel_map restrict_relp_def fun_eq_iff)
lemma pred_pmf_conj [simp]: "pred_pmf (\<lambda>x. P x \<and> Q x) = (\<lambda>x. pred_pmf P x \<and> pred_pmf Q x)"
by(auto simp add: pred_pmf_def)
lemma pred_pmf_top [simp]:
"pred_pmf (\<lambda>_. True) = (\<lambda>_. True)"
by(simp add: pred_pmf_def)
lemma rel_pmf_of_setI:
assumes A: "A \<noteq> {}" "finite A"
and B: "B \<noteq> {}" "finite B"
and card: "\<And>X. X \<subseteq> A \<Longrightarrow> card B * card X \<le> card A * card {y\<in>B. \<exists>x\<in>X. R x y}"
shows "rel_pmf R (pmf_of_set A) (pmf_of_set B)"
apply(rule rel_pmf_measureI)
using assms
apply(clarsimp simp add: measure_pmf_of_set card_gt_0_iff field_simps of_nat_mult[symmetric] simp del: of_nat_mult)
apply(subst mult.commute)
apply(erule meta_allE)
apply(erule meta_impE)
prefer 2
apply(erule order_trans)
apply(auto simp add: card_gt_0_iff intro: card_mono)
done
consts rel_witness_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<times> 'b pmf \<Rightarrow> ('a \<times> 'b) pmf"
specification (rel_witness_pmf)
set_rel_witness_pmf': "rel_pmf A (fst xy) (snd xy) \<Longrightarrow> set_pmf (rel_witness_pmf A xy) \<subseteq> {(a, b). A a b}"
map1_rel_witness_pmf': "rel_pmf A (fst xy) (snd xy) \<Longrightarrow> map_pmf fst (rel_witness_pmf A xy) = fst xy"
map2_rel_witness_pmf': "rel_pmf A (fst xy) (snd xy) \<Longrightarrow> map_pmf snd (rel_witness_pmf A xy) = snd xy"
apply(fold all_conj_distrib imp_conjR)
apply(rule choice allI)+
apply(unfold pmf.in_rel)
by blast
lemmas set_rel_witness_pmf = set_rel_witness_pmf'[of _ "(x, y)" for x y, simplified]
lemmas map1_rel_witness_pmf = map1_rel_witness_pmf'[of _ "(x, y)" for x y, simplified]
lemmas map2_rel_witness_pmf = map2_rel_witness_pmf'[of _ "(x, y)" for x y, simplified]
lemmas rel_witness_pmf = set_rel_witness_pmf map1_rel_witness_pmf map2_rel_witness_pmf
lemma rel_witness_pmf1:
assumes "rel_pmf A p q"
shows "rel_pmf (\<lambda>a (a', b). a = a' \<and> A a' b) p (rel_witness_pmf A (p, q))"
using map1_rel_witness_pmf[OF assms, symmetric]
unfolding pmf.rel_eq[symmetric] pmf.rel_map
by(rule pmf.rel_mono_strong)(auto dest: set_rel_witness_pmf[OF assms, THEN subsetD])
lemma rel_witness_pmf2:
assumes "rel_pmf A p q"
shows "rel_pmf (\<lambda>(a, b') b. b = b' \<and> A a b') (rel_witness_pmf A (p, q)) q"
using map2_rel_witness_pmf[OF assms]
unfolding pmf.rel_eq[symmetric] pmf.rel_map
by(rule pmf.rel_mono_strong)(auto dest: set_rel_witness_pmf[OF assms, THEN subsetD])
lemma cond_pmf_of_set:
assumes fin: "finite A" and nonempty: "A \<inter> B \<noteq> {}"
shows "cond_pmf (pmf_of_set A) B = pmf_of_set (A \<inter> B)" (is "?lhs = ?rhs")
proof(rule pmf_eqI)
from nonempty have A: "A \<noteq> {}" by auto
show "pmf ?lhs x = pmf ?rhs x" for x
by(subst pmf_cond; clarsimp simp add: fin A nonempty measure_pmf_of_set split: split_indicator)
qed
lemma pair_pmf_of_set:
assumes A: "finite A" "A \<noteq> {}"
and B: "finite B" "B \<noteq> {}"
shows "pair_pmf (pmf_of_set A) (pmf_of_set B) = pmf_of_set (A \<times> B)"
by(rule pmf_eqI)(clarsimp simp add: pmf_pair assms split: split_indicator)
lemma emeasure_cond_pmf:
fixes p A
defines "q \<equiv> cond_pmf p A"
assumes "set_pmf p \<inter> A \<noteq> {}"
shows "emeasure (measure_pmf q) B = emeasure (measure_pmf p) (A \<inter> B) / emeasure (measure_pmf p) A"
proof -
note [transfer_rule] = cond_pmf.transfer[OF assms(2), folded q_def]
interpret pmf_as_measure .
show ?thesis by transfer simp
qed
lemma measure_cond_pmf:
"measure (measure_pmf (cond_pmf p A)) B = measure (measure_pmf p) (A \<inter> B) / measure (measure_pmf p) A"
if "set_pmf p \<inter> A \<noteq> {}"
using emeasure_cond_pmf[OF that, of B] that
by(auto simp add: measure_pmf.emeasure_eq_measure measure_pmf_posI divide_ennreal)
lemma emeasure_measure_pmf_zero_iff: "emeasure (measure_pmf p) s = 0 \<longleftrightarrow> set_pmf p \<inter> s = {}" (is "?lhs = ?rhs")
proof -
have "?lhs \<longleftrightarrow> (AE x in measure_pmf p. x \<notin> s)"
by(subst AE_iff_measurable)(auto)
also have "\<dots> = ?rhs" by(auto simp add: AE_measure_pmf_iff)
finally show ?thesis .
qed
subsection \<open>Subprobability mass functions\<close>
lemma ord_spmf_return_spmf1: "ord_spmf R (return_spmf x) p \<longleftrightarrow> lossless_spmf p \<and> (\<forall>y\<in>set_spmf p. R x y)"
by(auto simp add: rel_pmf_return_pmf1 ord_option.simps in_set_spmf lossless_iff_set_pmf_None Ball_def) (metis option.exhaust)
lemma ord_spmf_conv:
"ord_spmf R = rel_spmf R OO ord_spmf (=)"
apply(subst pmf.rel_compp[symmetric])
apply(rule arg_cong[where f="rel_pmf"])
apply(rule ext)+
apply(auto elim!: ord_option.cases option.rel_cases intro: option.rel_intros)
done
lemma ord_spmf_expand:
"NO_MATCH (=) R \<Longrightarrow> ord_spmf R = rel_spmf R OO ord_spmf (=)"
by(rule ord_spmf_conv)
lemma ord_spmf_eqD_measure: "ord_spmf (=) p q \<Longrightarrow> measure (measure_spmf p) A \<le> measure (measure_spmf q) A"
by(drule ord_spmf_eqD_measure_spmf)(simp add: le_measure measure_spmf.emeasure_eq_measure)
lemma ord_spmf_measureD:
assumes "ord_spmf R p q"
shows "measure (measure_spmf p) A \<le> measure (measure_spmf q) {y. \<exists>x\<in>A. R x y}"
(is "?lhs \<le> ?rhs")
proof -
from assms obtain p' where *: "rel_spmf R p p'" and **: "ord_spmf (=) p' q"
by(auto simp add: ord_spmf_expand)
have "?lhs \<le> measure (measure_spmf p') {y. \<exists>x\<in>A. R x y}" using * by(rule rel_spmf_measureD)
also have "\<dots> \<le> ?rhs" using ** by(rule ord_spmf_eqD_measure)
finally show ?thesis .
qed
lemma ord_spmf_bind_pmfI1:
"(\<And>x. x \<in> set_pmf p \<Longrightarrow> ord_spmf R (f x) q) \<Longrightarrow> ord_spmf R (bind_pmf p f) q"
apply(rewrite at "ord_spmf _ _ \<hole>" bind_return_pmf[symmetric, where f="\<lambda>_ :: unit. q"])
apply(rule rel_pmf_bindI[where R="\<lambda>x y. x \<in> set_pmf p"])
apply(simp_all add: rel_pmf_return_pmf2)
done
lemma ord_spmf_bind_spmfI1:
"(\<And>x. x \<in> set_spmf p \<Longrightarrow> ord_spmf R (f x) q) \<Longrightarrow> ord_spmf R (bind_spmf p f) q"
unfolding bind_spmf_def by(rule ord_spmf_bind_pmfI1)(auto split: option.split simp add: in_set_spmf)
lemma spmf_of_set_empty: "spmf_of_set {} = return_pmf None"
by(simp add: spmf_of_set_def)
lemma rel_spmf_of_setI:
assumes card: "\<And>X. X \<subseteq> A \<Longrightarrow> card B * card X \<le> card A * card {y\<in>B. \<exists>x\<in>X. R x y}"
and eq: "(finite A \<and> A \<noteq> {}) \<longleftrightarrow> (finite B \<and> B \<noteq> {})"
shows "rel_spmf R (spmf_of_set A) (spmf_of_set B)"
using eq by(clarsimp simp add: spmf_of_set_def card rel_pmf_of_setI simp del: spmf_of_pmf_pmf_of_set cong: conj_cong)
lemmas map_bind_spmf = map_spmf_bind_spmf
lemma nn_integral_measure_spmf_conv_measure_pmf:
assumes [measurable]: "f \<in> borel_measurable (count_space UNIV)"
shows "nn_integral (measure_spmf p) f = nn_integral (restrict_space (measure_pmf p) (range Some)) (f \<circ> the)"
by(simp add: measure_spmf_def nn_integral_distr o_def)
lemma nn_integral_spmf_neq_infinity: "(\<integral>\<^sup>+ x. spmf p x \<partial>count_space UNIV) \<noteq> \<infinity>"
using nn_integral_measure_spmf[where f="\<lambda>_. 1", of p, symmetric] by simp
lemma return_pmf_bind_option:
"return_pmf (Option.bind x f) = bind_spmf (return_pmf x) (return_pmf \<circ> f)"
by(cases x) simp_all
lemma rel_spmf_pos_distr: "rel_spmf A OO rel_spmf B \<le> rel_spmf (A OO B)"
unfolding option.rel_compp pmf.rel_compp ..
lemma rel_spmf_OO_trans [trans]:
"\<lbrakk> rel_spmf R p q; rel_spmf S q r \<rbrakk> \<Longrightarrow> rel_spmf (R OO S) p r"
by(rule rel_spmf_pos_distr[THEN predicate2D]) auto
lemma map_spmf_eq_map_spmf_iff: "map_spmf f p = map_spmf g q \<longleftrightarrow> rel_spmf (\<lambda>x y. f x = g y) p q"
by(simp add: spmf_rel_eq[symmetric] spmf_rel_map)
lemma map_spmf_eq_map_spmfI: "rel_spmf (\<lambda>x y. f x = g y) p q \<Longrightarrow> map_spmf f p = map_spmf g q"
by(simp add: map_spmf_eq_map_spmf_iff)
lemma spmf_rel_mono_strong:
"\<lbrakk>rel_spmf A f g; \<And>x y. \<lbrakk> x \<in> set_spmf f; y \<in> set_spmf g; A x y \<rbrakk> \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_spmf B f g"
apply(erule pmf.rel_mono_strong)
apply(erule option.rel_mono_strong)
by(clarsimp simp add: in_set_spmf)
lemma set_spmf_eq_empty: "set_spmf p = {} \<longleftrightarrow> p = return_pmf None"
by auto (metis restrict_spmf_empty restrict_spmf_trivial)
lemma measure_pair_spmf_times:
"measure (measure_spmf (pair_spmf p q)) (A \<times> B) = measure (measure_spmf p) A * measure (measure_spmf q) B"
proof -
have "emeasure (measure_spmf (pair_spmf p q)) (A \<times> B) = (\<integral>\<^sup>+ x. ennreal (spmf (pair_spmf p q) x) * indicator (A \<times> B) x \<partial>count_space UNIV)"
by(simp add: nn_integral_spmf[symmetric] nn_integral_count_space_indicator)
also have "\<dots> = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. (ennreal (spmf p x) * indicator A x) * (ennreal (spmf q y) * indicator B y) \<partial>count_space UNIV) \<partial>count_space UNIV)"
by(subst nn_integral_fst_count_space[symmetric])(auto intro!: nn_integral_cong split: split_indicator simp add: ennreal_mult)
also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (spmf p x) * indicator A x * emeasure (measure_spmf q) B \<partial>count_space UNIV)"
by(simp add: nn_integral_cmult nn_integral_spmf[symmetric] nn_integral_count_space_indicator)
also have "\<dots> = emeasure (measure_spmf p) A * emeasure (measure_spmf q) B"
by(simp add: nn_integral_multc)(simp add: nn_integral_spmf[symmetric] nn_integral_count_space_indicator)
finally show ?thesis by(simp add: measure_spmf.emeasure_eq_measure ennreal_mult[symmetric])
qed
lemma lossless_spmfD_set_spmf_nonempty: "lossless_spmf p \<Longrightarrow> set_spmf p \<noteq> {}"
using set_pmf_not_empty[of p] by(auto simp add: set_spmf_def bind_UNION lossless_iff_set_pmf_None)
lemma set_spmf_return_pmf: "set_spmf (return_pmf x) = set_option x"
by(cases x) simp_all
lemma bind_spmf_pmf_assoc: "bind_spmf (bind_pmf p f) g = bind_pmf p (\<lambda>x. bind_spmf (f x) g)"
by(simp add: bind_spmf_def bind_assoc_pmf)
lemma bind_spmf_of_set: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> bind_spmf (spmf_of_set A) f = bind_pmf (pmf_of_set A) f"
by(simp add: spmf_of_set_def del: spmf_of_pmf_pmf_of_set)
lemma bind_spmf_map_pmf:
"bind_spmf (map_pmf f p) g = bind_pmf p (\<lambda>x. bind_spmf (return_pmf (f x)) g)"
by(simp add: map_pmf_def bind_spmf_def bind_assoc_pmf)
lemma rel_spmf_eqI [simp]: "rel_spmf (=) x x"
by(simp add: option.rel_eq)
lemma set_spmf_map_pmf: "set_spmf (map_pmf f p) = (\<Union>x\<in>set_pmf p. set_option (f x))" (* Move up *)
by(simp add: set_spmf_def bind_UNION)
lemma ord_spmf_return_spmf [simp]: "ord_spmf (=) (return_spmf x) p \<longleftrightarrow> p = return_spmf x"
proof -
have "p = return_spmf x \<Longrightarrow> ord_spmf (=) (return_spmf x) p" by simp
thus ?thesis
by (metis (no_types) ord_option_eq_simps(2) rel_pmf_return_pmf1 rel_pmf_return_pmf2 spmf.leq_antisym)
qed
declare
set_bind_spmf [simp]
set_spmf_return_pmf [simp]
lemma bind_spmf_pmf_commute:
"bind_spmf p (\<lambda>x. bind_pmf q (f x)) = bind_pmf q (\<lambda>y. bind_spmf p (\<lambda>x. f x y))"
unfolding bind_spmf_def
by(subst bind_commute_pmf)(auto intro: bind_pmf_cong[OF refl] split: option.split)
lemma return_pmf_map_option_conv_bind:
"return_pmf (map_option f x) = bind_spmf (return_pmf x) (return_spmf \<circ> f)"
by(cases x) simp_all
lemma lossless_return_pmf_iff [simp]: "lossless_spmf (return_pmf x) \<longleftrightarrow> x \<noteq> None"
by(cases x) simp_all
lemma lossless_map_pmf: "lossless_spmf (map_pmf f p) \<longleftrightarrow> (\<forall>x \<in> set_pmf p. f x \<noteq> None)"
using image_iff by(fastforce simp add: lossless_iff_set_pmf_None)
lemma bind_pmf_spmf_assoc:
"g None = return_pmf None
\<Longrightarrow> bind_pmf (bind_spmf p f) g = bind_spmf p (\<lambda>x. bind_pmf (f x) g)"
by(auto simp add: bind_spmf_def bind_assoc_pmf bind_return_pmf fun_eq_iff intro!: arg_cong2[where f=bind_pmf] split: option.split)
abbreviation pred_spmf :: "('a \<Rightarrow> bool) \<Rightarrow> 'a spmf \<Rightarrow> bool"
where "pred_spmf P \<equiv> pred_pmf (pred_option P)"
lemma pred_spmf_def: "pred_spmf P p \<longleftrightarrow> (\<forall>x\<in>set_spmf p. P x)"
by(auto simp add: pred_pmf_def pred_option_def set_spmf_def)
lemma spmf_pred_mono_strong:
"\<lbrakk> pred_spmf P p; \<And>a. \<lbrakk> a \<in> set_spmf p; P a \<rbrakk> \<Longrightarrow> P' a \<rbrakk> \<Longrightarrow> pred_spmf P' p"
by(simp add: pred_spmf_def)
lemma spmf_Domainp_rel: "Domainp (rel_spmf R) = pred_spmf (Domainp R)"
by(simp add: pmf.Domainp_rel option.Domainp_rel)
lemma rel_spmf_restrict_relpI [intro?]:
"\<lbrakk> rel_spmf R p q; pred_spmf P p; pred_spmf Q q \<rbrakk> \<Longrightarrow> rel_spmf (R \<upharpoonleft> P \<otimes> Q) p q"
by(erule spmf_rel_mono_strong)(simp add: pred_spmf_def)
lemma rel_spmf_restrict_relpE [elim?]:
assumes "rel_spmf (R \<upharpoonleft> P \<otimes> Q) x y"
obtains "rel_spmf R x y" "pred_spmf P x" "pred_spmf Q y"
proof
show "rel_spmf R x y" using assms by(auto elim!: spmf_rel_mono_strong)
have "pred_spmf (Domainp (R \<upharpoonleft> P \<otimes> Q)) x" using assms by(fold spmf_Domainp_rel) blast
then show "pred_spmf P x" by(rule spmf_pred_mono_strong)(blast dest!: restrict_relp_DomainpD)
have "pred_spmf (Domainp (R \<upharpoonleft> P \<otimes> Q)\<inverse>\<inverse>) y" using assms
by(fold spmf_Domainp_rel)(auto simp only: spmf_rel_conversep Domainp_conversep)
then show "pred_spmf Q y" by(rule spmf_pred_mono_strong)(auto dest!: restrict_relp_DomainpD)
qed
lemma rel_spmf_restrict_relp_iff:
"rel_spmf (R \<upharpoonleft> P \<otimes> Q) x y \<longleftrightarrow> rel_spmf R x y \<and> pred_spmf P x \<and> pred_spmf Q y"
by(blast intro: rel_spmf_restrict_relpI elim: rel_spmf_restrict_relpE)
lemma spmf_pred_map: "pred_spmf P (map_spmf f p) = pred_spmf (P \<circ> f) p"
by(simp)
lemma pred_spmf_bind [simp]: "pred_spmf P (bind_spmf p f) = pred_spmf (pred_spmf P \<circ> f) p"
by(simp add: pred_spmf_def bind_UNION)
lemma pred_spmf_return: "pred_spmf P (return_spmf x) = P x"
by simp
lemma pred_spmf_return_pmf_None: "pred_spmf P (return_pmf None)"
by simp
lemma pred_spmf_spmf_of_pmf [simp]: "pred_spmf P (spmf_of_pmf p) = pred_pmf P p"
unfolding pred_spmf_def by(simp add: pred_pmf_def)
lemma pred_spmf_of_set [simp]: "pred_spmf P (spmf_of_set A) = (finite A \<longrightarrow> Ball A P)"
by(auto simp add: pred_spmf_def set_spmf_of_set)
lemma pred_spmf_assert_spmf [simp]: "pred_spmf P (assert_spmf b) = (b \<longrightarrow> P ())"
by(cases b) simp_all
lemma pred_spmf_pair [simp]:
"pred_spmf P (pair_spmf p q) = pred_spmf (\<lambda>x. pred_spmf (P \<circ> Pair x) q) p"
by(simp add: pred_spmf_def)
lemma set_spmf_try [simp]:
"set_spmf (try_spmf p q) = set_spmf p \<union> (if lossless_spmf p then {} else set_spmf q)"
by(auto simp add: try_spmf_def set_spmf_bind_pmf in_set_spmf lossless_iff_set_pmf_None split: option.splits)(metis option.collapse)
lemma try_spmf_bind_out1:
"(\<And>x. lossless_spmf (f x)) \<Longrightarrow> bind_spmf (TRY p ELSE q) f = TRY (bind_spmf p f) ELSE (bind_spmf q f)"
apply(clarsimp simp add: bind_spmf_def try_spmf_def bind_assoc_pmf bind_return_pmf intro!: bind_pmf_cong[OF refl] split: option.split)
apply(rewrite in "\<hole> = _" bind_return_pmf'[symmetric])
apply(rule bind_pmf_cong[OF refl])
apply(clarsimp split: option.split simp add: lossless_iff_set_pmf_None)
done
lemma pred_spmf_try [simp]:
"pred_spmf P (try_spmf p q) = (pred_spmf P p \<and> (\<not> lossless_spmf p \<longrightarrow> pred_spmf P q))"
by(auto simp add: pred_spmf_def)
lemma pred_spmf_cond [simp]:
"pred_spmf P (cond_spmf p A) = pred_spmf (\<lambda>x. x \<in> A \<longrightarrow> P x) p"
by(auto simp add: pred_spmf_def)
lemma spmf_rel_map_restrict_relp:
shows spmf_rel_map_restrict_relp1: "rel_spmf (R \<upharpoonleft> P \<otimes> Q) (map_spmf f p) = rel_spmf (R \<circ> f \<upharpoonleft> P \<circ> f \<otimes> Q) p"
and spmf_rel_map_restrict_relp2: "rel_spmf (R \<upharpoonleft> P \<otimes> Q) p (map_spmf g q) = rel_spmf ((\<lambda>x. R x \<circ> g) \<upharpoonleft> P \<otimes> Q \<circ> g) p q"
by(simp_all add: spmf_rel_map restrict_relp_def)
lemma pred_spmf_conj: "pred_spmf (\<lambda>x. P x \<and> Q x) = (\<lambda>x. pred_spmf P x \<and> pred_spmf Q x)"
by simp
lemma spmf_of_pmf_parametric [transfer_rule]:
includes lifting_syntax shows
"(rel_pmf A ===> rel_spmf A) spmf_of_pmf spmf_of_pmf"
unfolding spmf_of_pmf_def[abs_def] by transfer_prover
lemma mono2mono_return_pmf[THEN spmf.mono2mono, simp, cont_intro]: (* Move to SPMF *)
shows monotone_return_pmf: "monotone option_ord (ord_spmf (=)) return_pmf"
by(rule monotoneI)(auto simp add: flat_ord_def)
lemma mcont2mcont_return_pmf[THEN spmf.mcont2mcont, simp, cont_intro]: (* Move to SPMF *)
shows mcont_return_pmf: "mcont (flat_lub None) option_ord lub_spmf (ord_spmf (=)) return_pmf"
by(rule mcont_finite_chains[OF _ _ flat_interpretation[THEN ccpo] ccpo_spmf]) simp_all
lemma pred_spmf_top: (* Move up *)
"pred_spmf (\<lambda>_. True) = (\<lambda>_. True)"
by(simp)
lemma rel_spmf_restrict_relpI' [intro?]:
"\<lbrakk> rel_spmf (\<lambda>x y. P x \<longrightarrow> Q y \<longrightarrow> R x y) p q; pred_spmf P p; pred_spmf Q q \<rbrakk> \<Longrightarrow> rel_spmf (R \<upharpoonleft> P \<otimes> Q) p q"
by(erule spmf_rel_mono_strong)(simp add: pred_spmf_def)
lemma set_spmf_map_pmf_MATCH [simp]:
assumes "NO_MATCH (map_option g) f"
shows "set_spmf (map_pmf f p) = (\<Union>x\<in>set_pmf p. set_option (f x))"
by(rule set_spmf_map_pmf)
lemma rel_spmf_bindI':
"\<lbrakk> rel_spmf A p q; \<And>x y. \<lbrakk> A x y; x \<in> set_spmf p; y \<in> set_spmf q \<rbrakk> \<Longrightarrow> rel_spmf B (f x) (g y) \<rbrakk>
\<Longrightarrow> rel_spmf B (p \<bind> f) (q \<bind> g)"
apply(rule rel_spmf_bindI[where R="\<lambda>x y. A x y \<and> x \<in> set_spmf p \<and> y \<in> set_spmf q"])
apply(erule spmf_rel_mono_strong; simp)
apply simp
done
definition rel_witness_spmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a spmf \<times> 'b spmf \<Rightarrow> ('a \<times> 'b) spmf" where
"rel_witness_spmf A = map_pmf rel_witness_option \<circ> rel_witness_pmf (rel_option A)"
lemma assumes "rel_spmf A p q"
shows rel_witness_spmf1: "rel_spmf (\<lambda>a (a', b). a = a' \<and> A a' b) p (rel_witness_spmf A (p, q))"
and rel_witness_spmf2: "rel_spmf (\<lambda>(a, b') b. b = b' \<and> A a b') (rel_witness_spmf A (p, q)) q"
by(auto simp add: pmf.rel_map rel_witness_spmf_def intro: pmf.rel_mono_strong[OF rel_witness_pmf1[OF assms]] rel_witness_option1 pmf.rel_mono_strong[OF rel_witness_pmf2[OF assms]] rel_witness_option2)
lemma weight_assert_spmf [simp]: "weight_spmf (assert_spmf b) = indicator {True} b"
by(simp split: split_indicator)
definition enforce_spmf :: "('a \<Rightarrow> bool) \<Rightarrow> 'a spmf \<Rightarrow> 'a spmf" where
"enforce_spmf P = map_pmf (enforce_option P)"
lemma enforce_spmf_parametric [transfer_rule]: includes lifting_syntax shows
"((A ===> (=)) ===> rel_spmf A ===> rel_spmf A) enforce_spmf enforce_spmf"
unfolding enforce_spmf_def by transfer_prover
lemma enforce_return_spmf [simp]:
"enforce_spmf P (return_spmf x) = (if P x then return_spmf x else return_pmf None)"
by(simp add: enforce_spmf_def)
lemma enforce_return_pmf_None [simp]:
"enforce_spmf P (return_pmf None) = return_pmf None"
by(simp add: enforce_spmf_def)
lemma enforce_map_spmf:
"enforce_spmf P (map_spmf f p) = map_spmf f (enforce_spmf (P \<circ> f) p)"
by(simp add: enforce_spmf_def pmf.map_comp o_def enforce_map_option)
lemma enforce_bind_spmf [simp]:
"enforce_spmf P (bind_spmf p f) = bind_spmf p (enforce_spmf P \<circ> f)"
by(auto simp add: enforce_spmf_def bind_spmf_def map_bind_pmf intro!: bind_pmf_cong split: option.split)
lemma set_enforce_spmf [simp]: "set_spmf (enforce_spmf P p) = {a \<in> set_spmf p. P a}"
by(auto simp add: enforce_spmf_def in_set_spmf)
lemma enforce_spmf_alt_def:
"enforce_spmf P p = bind_spmf p (\<lambda>a. bind_spmf (assert_spmf (P a)) (\<lambda>_ :: unit. return_spmf a))"
by(auto simp add: enforce_spmf_def assert_spmf_def map_pmf_def bind_spmf_def bind_return_pmf intro!: bind_pmf_cong split: option.split)
lemma bind_enforce_spmf [simp]:
"bind_spmf (enforce_spmf P p) f = bind_spmf p (\<lambda>x. if P x then f x else return_pmf None)"
by(auto simp add: enforce_spmf_alt_def assert_spmf_def intro!: bind_spmf_cong)
lemma weight_enforce_spmf:
"weight_spmf (enforce_spmf P p) = weight_spmf p - measure (measure_spmf p) {x. \<not> P x}" (is "?lhs = ?rhs")
proof -
have "?lhs = LINT x|measure_spmf p. indicator {x. P x} x"
by(auto simp add: enforce_spmf_alt_def weight_bind_spmf o_def simp del: Bochner_Integration.integral_indicator intro!: Bochner_Integration.integral_cong split: split_indicator)
also have "\<dots> = ?rhs"
by(subst measure_spmf.finite_measure_Diff[symmetric])(auto simp add: space_measure_spmf intro!: arg_cong2[where f=measure])
finally show ?thesis .
qed
lemma lossless_enforce_spmf [simp]:
"lossless_spmf (enforce_spmf P p) \<longleftrightarrow> lossless_spmf p \<and> set_spmf p \<subseteq> {x. P x}"
by(auto simp add: enforce_spmf_alt_def)
lemma enforce_spmf_top [simp]: "enforce_spmf \<top> = id"
by(simp add: enforce_spmf_def)
lemma enforce_spmf_K_True [simp]: "enforce_spmf (\<lambda>_. True) p = p"
using enforce_spmf_top[THEN fun_cong, of p] by(simp add: top_fun_def)
lemma enforce_spmf_bot [simp]: "enforce_spmf \<bottom> = (\<lambda>_. return_pmf None)"
by(simp add: enforce_spmf_def fun_eq_iff)
lemma enforce_spmf_K_False [simp]: "enforce_spmf (\<lambda>_. False) p = return_pmf None"
using enforce_spmf_bot[THEN fun_cong, of p] by(simp add: bot_fun_def)
lemma enforce_pred_id_spmf: "enforce_spmf P p = p" if "pred_spmf P p"
proof -
have "enforce_spmf P p = map_pmf id p" using that
by(auto simp add: enforce_spmf_def enforce_pred_id_option simp del: map_pmf_id intro!: pmf.map_cong_pred[OF refl] elim!: pmf_pred_mono_strong)
then show ?thesis by simp
qed
lemma map_the_spmf_of_pmf [simp]: "map_pmf the (spmf_of_pmf p) = p"
by(simp add: spmf_of_pmf_def pmf.map_comp o_def)
lemma bind_bind_conv_pair_spmf:
"bind_spmf p (\<lambda>x. bind_spmf q (f x)) = bind_spmf (pair_spmf p q) (\<lambda>(x, y). f x y)"
by(simp add: pair_spmf_alt_def)
lemma cond_spmf_spmf_of_set:
"cond_spmf (spmf_of_set A) B = spmf_of_set (A \<inter> B)" if "finite A"
by(rule spmf_eqI)(auto simp add: spmf_of_set measure_spmf_of_set that split: split_indicator)
lemma pair_spmf_of_set:
"pair_spmf (spmf_of_set A) (spmf_of_set B) = spmf_of_set (A \<times> B)"
by(rule spmf_eqI)(clarsimp simp add: spmf_of_set card_cartesian_product split: split_indicator)
lemma emeasure_cond_spmf:
"emeasure (measure_spmf (cond_spmf p A)) B = emeasure (measure_spmf p) (A \<inter> B) / emeasure (measure_spmf p) A"
apply(clarsimp simp add: cond_spmf_def emeasure_measure_spmf_conv_measure_pmf emeasure_measure_pmf_zero_iff set_pmf_Int_Some split!: if_split)
apply blast
apply(subst (asm) emeasure_cond_pmf)
by(auto simp add: set_pmf_Int_Some image_Int)
lemma measure_cond_spmf:
"measure (measure_spmf (cond_spmf p A)) B = measure (measure_spmf p) (A \<inter> B) / measure (measure_spmf p) A"
apply(clarsimp simp add: cond_spmf_def measure_measure_spmf_conv_measure_pmf measure_pmf_zero_iff set_pmf_Int_Some split!: if_split)
apply(subst (asm) measure_cond_pmf)
by(auto simp add: image_Int set_pmf_Int_Some)
lemma lossless_cond_spmf [simp]: "lossless_spmf (cond_spmf p A) \<longleftrightarrow> set_spmf p \<inter> A \<noteq> {}"
by(clarsimp simp add: cond_spmf_def lossless_iff_set_pmf_None set_pmf_Int_Some)
lemma measure_spmf_eq_density: "measure_spmf p = density (count_space UNIV) (spmf p)"
by(rule measure_eqI)(simp_all add: emeasure_density nn_integral_spmf[symmetric] nn_integral_count_space_indicator)
lemma integral_measure_spmf:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
assumes A: "finite A"
shows "(\<And>a. a \<in> set_spmf M \<Longrightarrow> f a \<noteq> 0 \<Longrightarrow> a \<in> A) \<Longrightarrow> (LINT x|measure_spmf M. f x) = (\<Sum>a\<in>A. spmf M a *\<^sub>R f a)"
unfolding measure_spmf_eq_density
apply (simp add: integral_density)
apply (subst lebesgue_integral_count_space_finite_support)
by (auto intro!: finite_subset[OF _ \<open>finite A\<close>] sum.mono_neutral_left simp: spmf_eq_0_set_spmf)
lemma image_set_spmf_eq:
"f ` set_spmf p = g ` set_spmf q" if "ASSUMPTION (map_spmf f p = map_spmf g q)"
using that[unfolded ASSUMPTION_def, THEN arg_cong[where f=set_spmf]] by simp
lemma map_spmf_const: "map_spmf (\<lambda>_. x) p = scale_spmf (weight_spmf p) (return_spmf x)"
by(simp add: map_spmf_conv_bind_spmf bind_spmf_const)
lemma cond_return_pmf [simp]: "cond_pmf (return_pmf x) A = return_pmf x" if "x \<in> A"
using that by(intro pmf_eqI)(auto simp add: pmf_cond split: split_indicator)
lemma cond_return_spmf [simp]: "cond_spmf (return_spmf x) A = (if x \<in> A then return_spmf x else return_pmf None)"
by(simp add: cond_spmf_def)
lemma measure_range_Some_eq_weight:
"measure (measure_pmf p) (range Some) = weight_spmf p"
by (simp add: measure_measure_spmf_conv_measure_pmf space_measure_spmf)
lemma restrict_spmf_eq_return_pmf_None [simp]:
"restrict_spmf p A = return_pmf None \<longleftrightarrow> set_spmf p \<inter> A = {}"
by(auto 4 3 simp add: restrict_spmf_def map_pmf_eq_return_pmf_iff bind_UNION in_set_spmf bind_eq_None_conv option.the_def dest: bspec split: if_split_asm option.split_asm)
definition mk_lossless :: "'a spmf \<Rightarrow> 'a spmf" where
"mk_lossless p = scale_spmf (inverse (weight_spmf p)) p"
lemma mk_lossless_idem [simp]: "mk_lossless (mk_lossless p) = mk_lossless p"
by(simp add: mk_lossless_def weight_scale_spmf min_def max_def inverse_eq_divide)
lemma mk_lossless_return [simp]: "mk_lossless (return_pmf x) = return_pmf x"
by(cases x)(simp_all add: mk_lossless_def)
lemma mk_lossless_map [simp]: "mk_lossless (map_spmf f p) = map_spmf f (mk_lossless p)"
by(simp add: mk_lossless_def map_scale_spmf)
lemma spmf_mk_lossless [simp]: "spmf (mk_lossless p) x = spmf p x / weight_spmf p"
by(simp add: mk_lossless_def spmf_scale_spmf inverse_eq_divide max_def)
lemma set_spmf_mk_lossless [simp]: "set_spmf (mk_lossless p) = set_spmf p"
by(simp add: mk_lossless_def set_scale_spmf measure_spmf_zero_iff zero_less_measure_iff)
lemma mk_lossless_lossless [simp]: "lossless_spmf p \<Longrightarrow> mk_lossless p = p"
by(simp add: mk_lossless_def lossless_weight_spmfD)
lemma mk_lossless_eq_return_pmf_None [simp]: "mk_lossless p = return_pmf None \<longleftrightarrow> p = return_pmf None"
proof -
have aux: "weight_spmf p = 0 \<Longrightarrow> spmf p i = 0" for i
by(rule antisym, rule order_trans[OF spmf_le_weight]) (auto intro!: order_trans[OF spmf_le_weight])
have[simp]: " spmf (scale_spmf (inverse (weight_spmf p)) p) = spmf (return_pmf None) \<Longrightarrow> spmf p i = 0" for i
by(drule fun_cong[where x=i]) (auto simp add: aux spmf_scale_spmf max_def)
show ?thesis by(auto simp add: mk_lossless_def intro: spmf_eqI)
qed
lemma return_pmf_None_eq_mk_lossless [simp]: "return_pmf None = mk_lossless p \<longleftrightarrow> p = return_pmf None"
by(metis mk_lossless_eq_return_pmf_None)
lemma mk_lossless_spmf_of_set [simp]: "mk_lossless (spmf_of_set A) = spmf_of_set A"
by(simp add: spmf_of_set_def del: spmf_of_pmf_pmf_of_set)
lemma weight_mk_lossless: "weight_spmf (mk_lossless p) = (if p = return_pmf None then 0 else 1)"
by(simp add: mk_lossless_def weight_scale_spmf min_def max_def inverse_eq_divide weight_spmf_eq_0)
lemma mk_lossless_parametric [transfer_rule]: includes lifting_syntax shows
"(rel_spmf A ===> rel_spmf A) mk_lossless mk_lossless"
by(simp add: mk_lossless_def rel_fun_def rel_spmf_weightD rel_spmf_scaleI)
lemma rel_spmf_mk_losslessI:
"rel_spmf A p q \<Longrightarrow> rel_spmf A (mk_lossless p) (mk_lossless q)"
by(rule mk_lossless_parametric[THEN rel_funD])
lemma rel_spmf_restrict_spmfI:
"rel_spmf (\<lambda>x y. (x \<in> A \<and> y \<in> B \<and> R x y) \<or> x \<notin> A \<and> y \<notin> B) p q
\<Longrightarrow> rel_spmf R (restrict_spmf p A) (restrict_spmf q B)"
by(auto simp add: restrict_spmf_def pmf.rel_map elim!: option.rel_cases pmf.rel_mono_strong)
lemma cond_spmf_alt: "cond_spmf p A = mk_lossless (restrict_spmf p A)"
proof(cases "set_spmf p \<inter> A = {}")
case True
then show ?thesis by(simp add: cond_spmf_def measure_spmf_zero_iff)
next
case False
show ?thesis
by(rule spmf_eqI)(simp add: False cond_spmf_def pmf_cond set_pmf_Int_Some image_iff measure_measure_spmf_conv_measure_pmf[symmetric] spmf_scale_spmf max_def inverse_eq_divide)
qed
lemma cond_spmf_bind:
"cond_spmf (bind_spmf p f) A = mk_lossless (p \<bind> (\<lambda>x. f x \<upharpoonleft> A))"
by(simp add: cond_spmf_alt restrict_bind_spmf scale_bind_spmf)
lemma cond_spmf_UNIV [simp]: "cond_spmf p UNIV = mk_lossless p"
by(clarsimp simp add: cond_spmf_alt)
lemma cond_pmf_singleton:
"cond_pmf p A = return_pmf x" if "set_pmf p \<inter> A = {x}"
proof -
have[simp]: "set_pmf p \<inter> A = {x} \<Longrightarrow> x \<in> A \<Longrightarrow> measure_pmf.prob p A = pmf p x"
by(auto simp add: measure_pmf_single[symmetric] AE_measure_pmf_iff intro!: measure_pmf.finite_measure_eq_AE)
have "pmf (cond_pmf p A) i = pmf (return_pmf x) i" for i
using that by(auto simp add: pmf_cond measure_pmf_zero_iff pmf_eq_0_set_pmf split: split_indicator)
then show ?thesis by(rule pmf_eqI)
qed
definition cond_spmf_fst :: "('a \<times> 'b) spmf \<Rightarrow> 'a \<Rightarrow> 'b spmf" where
"cond_spmf_fst p a = map_spmf snd (cond_spmf p ({a} \<times> UNIV))"
lemma cond_spmf_fst_return_spmf [simp]:
"cond_spmf_fst (return_spmf (x, y)) x = return_spmf y"
by(simp add: cond_spmf_fst_def)
lemma cond_spmf_fst_map_Pair [simp]: "cond_spmf_fst (map_spmf (Pair x) p) x = mk_lossless p"
by(clarsimp simp add: cond_spmf_fst_def spmf.map_comp o_def)
lemma cond_spmf_fst_map_Pair' [simp]: "cond_spmf_fst (map_spmf (\<lambda>y. (x, f y)) p) x = map_spmf f (mk_lossless p)"
by(subst spmf.map_comp[where f="Pair x", symmetric, unfolded o_def]) simp
lemma cond_spmf_fst_eq_return_None [simp]: "cond_spmf_fst p x = return_pmf None \<longleftrightarrow> x \<notin> fst ` set_spmf p"
by(auto 4 4 simp add: cond_spmf_fst_def map_pmf_eq_return_pmf_iff in_set_spmf[symmetric] dest: bspec[where x="Some _"] intro: ccontr rev_image_eqI)
lemma cond_spmf_fst_map_Pair1:
"cond_spmf_fst (map_spmf (\<lambda>x. (f x, g x)) p) (f x) = return_spmf (g (inv_into (set_spmf p) f (f x)))"
if "x \<in> set_spmf p" "inj_on f (set_spmf p)"
proof -
let ?foo="\<lambda>y. map_option (\<lambda>x. (f x, g x)) -` Some ` ({f y} \<times> UNIV)"
have[simp]: "y \<in> set_spmf p \<Longrightarrow> f x = f y \<Longrightarrow> set_pmf p \<inter> (?foo y) \<noteq> {}" for y
by(auto simp add: vimage_def image_def in_set_spmf)
have[simp]: "y \<in> set_spmf p \<Longrightarrow> f x = f y \<Longrightarrow> map_spmf snd (map_spmf (\<lambda>x. (f x, g x)) (cond_pmf p (?foo y))) = return_spmf (g x)" for y
using that by(subst cond_pmf_singleton[where x="Some x"]) (auto simp add: in_set_spmf elim: inj_onD)
show ?thesis
using that
by(auto simp add: cond_spmf_fst_def cond_spmf_def)
(erule notE, subst cond_map_pmf, simp_all)
qed
lemma lossless_cond_spmf_fst [simp]: "lossless_spmf (cond_spmf_fst p x) \<longleftrightarrow> x \<in> fst ` set_spmf p"
by(auto simp add: cond_spmf_fst_def intro: rev_image_eqI)
lemma cond_spmf_fst_inverse:
"bind_spmf (map_spmf fst p) (\<lambda>x. map_spmf (Pair x) (cond_spmf_fst p x)) = p"
(is "?lhs = ?rhs")
proof(rule spmf_eqI)
fix i :: "'a \<times> 'b"
have *: "({x} \<times> UNIV \<inter> (Pair x \<circ> snd) -` {i}) = (if x = fst i then {i} else {})" for x by(cases i)auto
have "spmf ?lhs i = LINT x|measure_spmf (map_spmf fst p). spmf (map_spmf (Pair x \<circ> snd) (cond_spmf p ({x} \<times> UNIV))) i"
by(auto simp add: spmf_bind spmf.map_comp[symmetric] cond_spmf_fst_def intro!: integral_cong_AE)
also have "\<dots> = LINT x|measure_spmf (map_spmf fst p). measure (measure_spmf (cond_spmf p ({x} \<times> UNIV))) ((Pair x \<circ> snd) -` {i})"
by(rule integral_cong_AE)(auto simp add: spmf_map)
also have "\<dots> = LINT x|measure_spmf (map_spmf fst p). measure (measure_spmf p) ({x} \<times> UNIV \<inter> (Pair x \<circ> snd) -` {i}) /
measure (measure_spmf p) ({x} \<times> UNIV)"
by(rule integral_cong_AE; clarsimp simp add: measure_cond_spmf)
also have "\<dots> = spmf (map_spmf fst p) (fst i) * spmf p i / measure (measure_spmf p) ({fst i} \<times> UNIV)"
by(simp add: * if_distrib[where f="measure (measure_spmf _)"] cong: if_cong)
(subst integral_measure_spmf[where A="{fst i}"]; auto split: if_split_asm simp add: spmf_conv_measure_spmf)
also have "\<dots> = spmf p i"
by(clarsimp simp add: spmf_map vimage_fst)(metis (no_types, lifting) Int_insert_left_if1 in_set_spmf_iff_spmf insertI1 insert_UNIV insert_absorb insert_not_empty measure_spmf_zero_iff mem_Sigma_iff prod.collapse)
finally show "spmf ?lhs i = spmf ?rhs i" .
qed
subsubsection \<open>Embedding of @{typ "'a option"} into @{typ "'a spmf"}\<close>
text \<open>This theoretically follows from the embedding between @{typ "_ id"} into @{typ "_ prob"} and the isomorphism
between @{typ "(_, _ prob) optionT"} and @{typ "_ spmf"}, but we would only get the monomorphic
version via this connection. So we do it directly.
\<close>
lemma bind_option_spmf_monad [simp]: "monad.bind_option (return_pmf None) x = bind_spmf (return_pmf x)"
by(cases x)(simp_all add: fun_eq_iff)
locale option_to_spmf begin
text \<open>
We have to get the embedding into the lifting package such that we can use the parametrisation of transfer rules.
\<close>
definition the_pmf :: "'a pmf \<Rightarrow> 'a" where "the_pmf p = (THE x. p = return_pmf x)"
lemma the_pmf_return [simp]: "the_pmf (return_pmf x) = x"
by(simp add: the_pmf_def)
lemma type_definition_option_spmf: "type_definition return_pmf the_pmf {x. \<exists>y :: 'a option. x = return_pmf y}"
by unfold_locales(auto)
context begin
private setup_lifting type_definition_option_spmf
abbreviation cr_spmf_option where "cr_spmf_option \<equiv> cr_option"
abbreviation pcr_spmf_option where "pcr_spmf_option \<equiv> pcr_option"
lemmas Quotient_spmf_option = Quotient_option
and cr_spmf_option_def = cr_option_def
and pcr_spmf_option_bi_unique = option.bi_unique
and Domainp_pcr_spmf_option = option.domain
and Domainp_pcr_spmf_option_eq = option.domain_eq
and Domainp_pcr_spmf_option_par = option.domain_par
and Domainp_pcr_spmf_option_left_total = option.domain_par_left_total
and pcr_spmf_option_left_unique = option.left_unique
and pcr_spmf_option_cr_eq = option.pcr_cr_eq
and pcr_spmf_option_return_pmf_transfer = option.rep_transfer
and pcr_spmf_option_right_total = option.right_total
and pcr_spmf_option_right_unique = option.right_unique
and pcr_spmf_option_def = pcr_option_def
bundle spmf_option_lifting = [[Lifting.lifting_restore_internal "Misc_CryptHOL.option.lifting"]]
end
context includes lifting_syntax begin
lemma return_option_spmf_transfer [transfer_parametric return_spmf_parametric, transfer_rule]:
"((=) ===> cr_spmf_option) return_spmf Some"
by(rule rel_funI)(simp add: cr_spmf_option_def)
lemma map_option_spmf_transfer [transfer_parametric map_spmf_parametric, transfer_rule]:
"(((=) ===> (=)) ===> cr_spmf_option ===> cr_spmf_option) map_spmf map_option"
unfolding rel_fun_eq by(auto simp add: rel_fun_def cr_spmf_option_def)
lemma fail_option_spmf_transfer [transfer_parametric return_spmf_None_parametric, transfer_rule]:
"cr_spmf_option (return_pmf None) None"
by(simp add: cr_spmf_option_def)
lemma bind_option_spmf_transfer [transfer_parametric bind_spmf_parametric, transfer_rule]:
"(cr_spmf_option ===> ((=) ===> cr_spmf_option) ===> cr_spmf_option) bind_spmf Option.bind"
apply(clarsimp simp add: rel_fun_def cr_spmf_option_def)
subgoal for x f g by(cases x; simp)
done
lemma set_option_spmf_transfer [transfer_parametric set_spmf_parametric, transfer_rule]:
"(cr_spmf_option ===> rel_set (=)) set_spmf set_option"
by(clarsimp simp add: rel_fun_def cr_spmf_option_def rel_set_eq)
lemma rel_option_spmf_transfer [transfer_parametric rel_spmf_parametric, transfer_rule]:
"(((=) ===> (=) ===> (=)) ===> cr_spmf_option ===> cr_spmf_option ===> (=)) rel_spmf rel_option"
unfolding rel_fun_eq by(simp add: rel_fun_def cr_spmf_option_def)
end
end
locale option_le_spmf begin
text \<open>
Embedding where only successful computations in the option monad are related to Dirac spmf.
\<close>
definition cr_option_le_spmf :: "'a option \<Rightarrow> 'a spmf \<Rightarrow> bool"
where "cr_option_le_spmf x p \<longleftrightarrow> ord_spmf (=) (return_pmf x) p"
context includes lifting_syntax begin
lemma return_option_le_spmf_transfer [transfer_rule]:
"((=) ===> cr_option_le_spmf) (\<lambda>x. x) return_pmf"
by(rule rel_funI)(simp add: cr_option_le_spmf_def ord_option_reflI)
lemma map_option_le_spmf_transfer [transfer_rule]:
"(((=) ===> (=)) ===> cr_option_le_spmf ===> cr_option_le_spmf) map_option map_spmf"
unfolding rel_fun_eq
apply(clarsimp simp add: rel_fun_def cr_option_le_spmf_def rel_pmf_return_pmf1 ord_option_map1 ord_option_map2)
subgoal for f x p y by(cases x; simp add: ord_option_reflI)
done
lemma bind_option_le_spmf_transfer [transfer_rule]:
"(cr_option_le_spmf ===> ((=) ===> cr_option_le_spmf) ===> cr_option_le_spmf) Option.bind bind_spmf"
apply(clarsimp simp add: rel_fun_def cr_option_le_spmf_def)
subgoal for x p f g by(cases x; auto 4 3 simp add: rel_pmf_return_pmf1 set_pmf_bind_spmf)
done
end
end
interpretation rel_spmf_characterisation by unfold_locales(rule rel_pmf_measureI)
lemma if_distrib_bind_spmf1 [if_distribs]:
"bind_spmf (if b then x else y) f = (if b then bind_spmf x f else bind_spmf y f)"
by simp
lemma if_distrib_bind_spmf2 [if_distribs]:
"bind_spmf x (\<lambda>y. if b then f y else g y) = (if b then bind_spmf x f else bind_spmf x g)"
by simp
lemma rel_spmf_if_distrib [if_distribs]:
"rel_spmf R (if b then x else y) (if b then x' else y') \<longleftrightarrow>
(b \<longrightarrow> rel_spmf R x x') \<and> (\<not> b \<longrightarrow> rel_spmf R y y')"
by(simp)
lemma if_distrib_map_spmf [if_distribs]:
"map_spmf f (if b then p else q) = (if b then map_spmf f p else map_spmf f q)"
by simp
lemma if_distrib_restrict_spmf1 [if_distribs]:
"restrict_spmf (if b then p else q) A = (if b then restrict_spmf p A else restrict_spmf q A)"
by simp
end
diff --git a/thys/Deriving/Comparator_Generator/Comparator.thy b/thys/Deriving/Comparator_Generator/Comparator.thy
--- a/thys/Deriving/Comparator_Generator/Comparator.thy
+++ b/thys/Deriving/Comparator_Generator/Comparator.thy
@@ -1,141 +1,154 @@
(* Title: Deriving class instances for datatypes
Author: Christian Sternagel and René Thiemann <christian.sternagel|rene.thiemann@uibk.ac.at>
Maintainer: Christian Sternagel and René Thiemann
License: LGPL
*)
section \<open>Comparisons\<close>
subsection \<open>Comparators and Linear Orders\<close>
theory Comparator
imports Main
begin
text \<open>Instead of having to define a strict and a weak linear order, @{term "((<), (\<le>))"},
one can alternative use a comparator to define the linear order, which may deliver
three possible outcomes when comparing two values.\<close>
datatype order = Eq | Lt | Gt
type_synonym 'a comparator = "'a \<Rightarrow> 'a \<Rightarrow> order"
text \<open>In the following, we provide the obvious definitions how to switch between
linear orders and comparators.\<close>
definition lt_of_comp :: "'a comparator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"lt_of_comp acomp x y = (case acomp x y of Lt \<Rightarrow> True | _ \<Rightarrow> False)"
definition le_of_comp :: "'a comparator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"le_of_comp acomp x y = (case acomp x y of Gt \<Rightarrow> False | _ \<Rightarrow> True)"
definition comp_of_ords :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a comparator" where
"comp_of_ords le lt x y = (if lt x y then Lt else if le x y then Eq else Gt)"
lemma comp_of_ords_of_le_lt[simp]: "comp_of_ords (le_of_comp c) (lt_of_comp c) = c"
by (intro ext, auto simp: comp_of_ords_def le_of_comp_def lt_of_comp_def split: order.split)
lemma lt_of_comp_of_ords: "lt_of_comp (comp_of_ords le lt) = lt"
by (intro ext, auto simp: comp_of_ords_def le_of_comp_def lt_of_comp_def split: order.split)
lemma le_of_comp_of_ords_gen: "(\<And> x y. lt x y \<Longrightarrow> le x y) \<Longrightarrow> le_of_comp (comp_of_ords le lt) = le"
by (intro ext, auto simp: comp_of_ords_def le_of_comp_def lt_of_comp_def split: order.split)
lemma le_of_comp_of_ords_linorder: assumes "class.linorder le lt"
shows "le_of_comp (comp_of_ords le lt) = le"
proof -
interpret linorder le lt by fact
show ?thesis by (rule le_of_comp_of_ords_gen) simp
qed
fun invert_order:: "order \<Rightarrow> order" where
"invert_order Lt = Gt" |
"invert_order Gt = Lt" |
"invert_order Eq = Eq"
locale comparator =
fixes comp :: "'a comparator"
assumes sym: "invert_order (comp x y) = comp y x"
and weak_eq: "comp x y = Eq \<Longrightarrow> x = y"
- and trans: "comp x y = Lt \<Longrightarrow> comp y z = Lt \<Longrightarrow> comp x z = Lt"
+ and comp_trans: "comp x y = Lt \<Longrightarrow> comp y z = Lt \<Longrightarrow> comp x z = Lt"
begin
lemma eq: "(comp x y = Eq) = (x = y)"
proof
assume "x = y"
with sym [of y y] show "comp x y = Eq" by (cases "comp x y") auto
qed (rule weak_eq)
lemma comp_same [simp]:
"comp x x = Eq"
by (simp add: eq)
abbreviation "lt \<equiv> lt_of_comp comp"
abbreviation "le \<equiv> le_of_comp comp"
-lemma linorder: "class.linorder le lt"
+sublocale ordering le lt
proof
note [simp] = lt_of_comp_def le_of_comp_def
fix x y z :: 'a
- show "lt x y = (le x y \<and> \<not> le y x)"
- using sym [of x y] by (cases "comp x y") (simp_all)
- show "le x y \<or> le y x"
- using sym [of x y] by (cases "comp x y") (simp_all)
show "le x x" using eq [of x x] by (simp)
+ show "le x y \<Longrightarrow> le y z \<Longrightarrow> le x z"
+ by (cases "comp x y" "comp y z" rule: order.exhaust [case_product order.exhaust])
+ (auto dest: comp_trans simp: eq)
show "le x y \<Longrightarrow> le y x \<Longrightarrow> x = y"
using sym [of x y] by (cases "comp x y") (simp_all add: eq)
- show "le x y \<Longrightarrow> le y z \<Longrightarrow> le x z"
- by (cases "comp x y" "comp y z" rule: order.exhaust [case_product order.exhaust])
- (auto dest: trans simp: eq)
+ show "lt x y \<longleftrightarrow> le x y \<and> x \<noteq> y"
+ using eq [of x y] by (cases "comp x y") (simp_all)
+qed
+
+lemma linorder: "class.linorder le lt"
+proof (rule class.linorder.intro)
+ interpret order le lt
+ using ordering_axioms by (rule ordering_orderI)
+ show \<open>class.order le lt\<close>
+ by (fact order_axioms)
+ show \<open>class.linorder_axioms le\<close>
+ proof
+ note [simp] = lt_of_comp_def le_of_comp_def
+ fix x y :: 'a
+ show "le x y \<or> le y x"
+ using sym [of x y] by (cases "comp x y") (simp_all)
+ qed
qed
sublocale linorder le lt
by (rule linorder)
lemma Gt_lt_conv: "comp x y = Gt \<longleftrightarrow> lt y x"
unfolding lt_of_comp_def sym[of x y, symmetric]
by (cases "comp x y", auto)
lemma Lt_lt_conv: "comp x y = Lt \<longleftrightarrow> lt x y"
unfolding lt_of_comp_def by (cases "comp x y", auto)
lemma eq_Eq_conv: "comp x y = Eq \<longleftrightarrow> x = y"
by (rule eq)
lemma nGt_le_conv: "comp x y \<noteq> Gt \<longleftrightarrow> le x y"
unfolding le_of_comp_def by (cases "comp x y", auto)
lemma nLt_le_conv: "comp x y \<noteq> Lt \<longleftrightarrow> le y x"
unfolding le_of_comp_def sym[of x y, symmetric] by (cases "comp x y", auto)
lemma nEq_neq_conv: "comp x y \<noteq> Eq \<longleftrightarrow> x \<noteq> y"
using eq_Eq_conv[of x y] by simp
lemmas le_lt_convs = nLt_le_conv nGt_le_conv Gt_lt_conv Lt_lt_conv eq_Eq_conv nEq_neq_conv
lemma two_comparisons_into_case_order:
"(if le x y then (if x = y then P else Q) else R) = (case_order P Q R (comp x y))"
"(if le x y then (if y = x then P else Q) else R) = (case_order P Q R (comp x y))"
"(if le x y then (if le y x then P else Q) else R) = (case_order P Q R (comp x y))"
"(if le x y then (if lt x y then Q else P) else R) = (case_order P Q R (comp x y))"
"(if lt x y then Q else (if le x y then P else R)) = (case_order P Q R (comp x y))"
"(if lt x y then Q else (if lt y x then R else P)) = (case_order P Q R (comp x y))"
"(if lt x y then Q else (if x = y then P else R)) = (case_order P Q R (comp x y))"
"(if lt x y then Q else (if y = x then P else R)) = (case_order P Q R (comp x y))"
"(if x = y then P else (if lt y x then R else Q)) = (case_order P Q R (comp x y))"
"(if x = y then P else (if lt x y then Q else R)) = (case_order P Q R (comp x y))"
"(if x = y then P else (if le y x then R else Q)) = (case_order P Q R (comp x y))"
"(if x = y then P else (if le x y then Q else R)) = (case_order P Q R (comp x y))"
by (auto simp: le_lt_convs split: order.splits)
end
lemma comp_of_ords: assumes "class.linorder le lt"
shows "comparator (comp_of_ords le lt)"
proof -
interpret linorder le lt by fact
show ?thesis
by (unfold_locales, auto simp: comp_of_ords_def split: if_splits)
qed
definition (in linorder) comparator_of :: "'a comparator" where
"comparator_of x y = (if x < y then Lt else if x = y then Eq else Gt)"
lemma comparator_of: "comparator comparator_of"
by unfold_locales (auto split: if_splits simp: comparator_of_def)
end
diff --git a/thys/Deriving/Comparator_Generator/Comparator_Generator.thy b/thys/Deriving/Comparator_Generator/Comparator_Generator.thy
--- a/thys/Deriving/Comparator_Generator/Comparator_Generator.thy
+++ b/thys/Deriving/Comparator_Generator/Comparator_Generator.thy
@@ -1,251 +1,251 @@
(* Title: Deriving class instances for datatypes
Author: Christian Sternagel and René Thiemann <christian.sternagel|rene.thiemann@uibk.ac.at>
Maintainer: Christian Sternagel and René Thiemann
License: LGPL
*)
section \<open>Generating Comparators\<close>
theory Comparator_Generator
imports
"../Generator_Aux"
"../Derive_Manager"
Comparator
begin
typedecl ('a,'b,'c,'z)type
text \<open>In the following, we define a generator which for a given datatype @{typ "('a,'b,'c,'z)type"}
constructs a comparator of type
@{typ "'a comparator \<Rightarrow> 'b comparator \<Rightarrow> 'c comparator \<Rightarrow> 'z comparator \<Rightarrow> ('a,'b,'c,'z)type"}.
To this end, we first compare the index of the constructors, then for equal constructors, we
compare the arguments recursively and combine the results lexicographically.\<close>
hide_type "type"
subsection \<open>Lexicographic combination of @{typ order}\<close>
fun comp_lex :: "order list \<Rightarrow> order"
where
"comp_lex (c # cs) = (case c of Eq \<Rightarrow> comp_lex cs | _ \<Rightarrow> c)" |
"comp_lex [] = Eq"
subsection \<open>Improved code for non-lazy languages\<close>
text \<open>The following equations will eliminate all occurrences of @{term comp_lex}
in the generated code of the comparators.\<close>
lemma comp_lex_unfolds:
"comp_lex [] = Eq"
"comp_lex [c] = c"
"comp_lex (c # d # cs) = (case c of Eq \<Rightarrow> comp_lex (d # cs) | z \<Rightarrow> z)"
by (cases c, auto)+
subsection \<open>Pointwise properties for equality, symmetry, and transitivity\<close>
text \<open>The pointwise properties are important during inductive proofs of soundness of comparators.
They are defined in a way that are combinable with @{const comp_lex}.\<close>
lemma comp_lex_eq: "comp_lex os = Eq \<longleftrightarrow> (\<forall> ord \<in> set os. ord = Eq)"
by (induct os) (auto split: order.splits)
definition trans_order :: "order \<Rightarrow> order \<Rightarrow> order \<Rightarrow> bool" where
"trans_order x y z \<longleftrightarrow> x \<noteq> Gt \<longrightarrow> y \<noteq> Gt \<longrightarrow> z \<noteq> Gt \<and> ((x = Lt \<or> y = Lt) \<longrightarrow> z = Lt)"
lemma trans_orderI:
"(x \<noteq> Gt \<Longrightarrow> y \<noteq> Gt \<Longrightarrow> z \<noteq> Gt \<and> ((x = Lt \<or> y = Lt) \<longrightarrow> z = Lt)) \<Longrightarrow> trans_order x y z"
by (simp add: trans_order_def)
lemma trans_orderD:
assumes "trans_order x y z" and "x \<noteq> Gt" and "y \<noteq> Gt"
shows "z \<noteq> Gt" and "x = Lt \<or> y = Lt \<Longrightarrow> z = Lt"
using assms by (auto simp: trans_order_def)
lemma All_less_Suc:
"(\<forall>i < Suc x. P i) \<longleftrightarrow> P 0 \<and> (\<forall>i < x. P (Suc i))"
using less_Suc_eq_0_disj by force
lemma comp_lex_trans:
assumes "length xs = length ys"
and "length ys = length zs"
and "\<forall> i < length zs. trans_order (xs ! i) (ys ! i) (zs ! i)"
shows "trans_order (comp_lex xs) (comp_lex ys) (comp_lex zs)"
using assms
proof (induct xs ys zs rule: list_induct3)
case (Cons x xs y ys z zs)
then show ?case
by (intro trans_orderI)
(cases x y z rule: order.exhaust [case_product order.exhaust order.exhaust],
auto simp: All_less_Suc dest: trans_orderD)
qed (simp add: trans_order_def)
lemma comp_lex_sym:
assumes "length xs = length ys"
and "\<forall> i < length ys. invert_order (xs ! i) = ys ! i"
shows "invert_order (comp_lex xs) = comp_lex ys"
using assms by (induct xs ys rule: list_induct2, simp, case_tac x) fastforce+
declare comp_lex.simps [simp del]
definition peq_comp :: "'a comparator \<Rightarrow> 'a \<Rightarrow> bool"
where
"peq_comp acomp x \<longleftrightarrow> (\<forall> y. acomp x y = Eq \<longleftrightarrow> x = y)"
lemma peq_compD: "peq_comp acomp x \<Longrightarrow> acomp x y = Eq \<longleftrightarrow> x = y"
unfolding peq_comp_def by auto
lemma peq_compI: "(\<And> y. acomp x y = Eq \<longleftrightarrow> x = y) \<Longrightarrow> peq_comp acomp x"
unfolding peq_comp_def by auto
definition psym_comp :: "'a comparator \<Rightarrow> 'a \<Rightarrow> bool" where
"psym_comp acomp x \<longleftrightarrow> (\<forall> y. invert_order (acomp x y) = (acomp y x))"
lemma psym_compD:
assumes "psym_comp acomp x"
shows "invert_order (acomp x y) = (acomp y x)"
using assms unfolding psym_comp_def by blast+
lemma psym_compI:
assumes "\<And> y. invert_order (acomp x y) = (acomp y x)"
shows "psym_comp acomp x"
using assms unfolding psym_comp_def by blast
definition ptrans_comp :: "'a comparator \<Rightarrow> 'a \<Rightarrow> bool" where
"ptrans_comp acomp x \<longleftrightarrow> (\<forall> y z. trans_order (acomp x y) (acomp y z) (acomp x z))"
lemma ptrans_compD:
assumes "ptrans_comp acomp x"
shows "trans_order (acomp x y) (acomp y z) (acomp x z)"
using assms unfolding ptrans_comp_def by blast+
lemma ptrans_compI:
assumes "\<And> y z. trans_order (acomp x y) (acomp y z) (acomp x z)"
shows "ptrans_comp acomp x"
using assms unfolding ptrans_comp_def by blast
subsection \<open>Separate properties of comparators\<close>
definition eq_comp :: "'a comparator \<Rightarrow> bool" where
"eq_comp acomp \<longleftrightarrow> (\<forall> x. peq_comp acomp x)"
lemma eq_compD2: "eq_comp acomp \<Longrightarrow> peq_comp acomp x"
unfolding eq_comp_def by blast
lemma eq_compI2: "(\<And> x. peq_comp acomp x) \<Longrightarrow> eq_comp acomp"
unfolding eq_comp_def by blast
definition trans_comp :: "'a comparator \<Rightarrow> bool" where
"trans_comp acomp \<longleftrightarrow> (\<forall> x. ptrans_comp acomp x)"
lemma trans_compD2: "trans_comp acomp \<Longrightarrow> ptrans_comp acomp x"
unfolding trans_comp_def by blast
lemma trans_compI2: "(\<And> x. ptrans_comp acomp x) \<Longrightarrow> trans_comp acomp"
unfolding trans_comp_def by blast
definition sym_comp :: "'a comparator \<Rightarrow> bool" where
"sym_comp acomp \<longleftrightarrow> (\<forall> x. psym_comp acomp x)"
lemma sym_compD2:
"sym_comp acomp \<Longrightarrow> psym_comp acomp x"
unfolding sym_comp_def by blast
lemma sym_compI2: "(\<And> x. psym_comp acomp x) \<Longrightarrow> sym_comp acomp"
unfolding sym_comp_def by blast
lemma eq_compD: "eq_comp acomp \<Longrightarrow> acomp x y = Eq \<longleftrightarrow> x = y"
by (rule peq_compD[OF eq_compD2])
lemma eq_compI: "(\<And> x y. acomp x y = Eq \<longleftrightarrow> x = y) \<Longrightarrow> eq_comp acomp"
by (intro eq_compI2 peq_compI)
lemma trans_compD: "trans_comp acomp \<Longrightarrow> trans_order (acomp x y) (acomp y z) (acomp x z)"
by (rule ptrans_compD[OF trans_compD2])
lemma trans_compI: "(\<And> x y z. trans_order (acomp x y) (acomp y z) (acomp x z)) \<Longrightarrow> trans_comp acomp"
by (intro trans_compI2 ptrans_compI)
lemma sym_compD:
"sym_comp acomp \<Longrightarrow> invert_order (acomp x y) = (acomp y x)"
by (rule psym_compD[OF sym_compD2])
lemma sym_compI: "(\<And> x y. invert_order (acomp x y) = (acomp y x)) \<Longrightarrow> sym_comp acomp"
by (intro sym_compI2 psym_compI)
lemma eq_sym_trans_imp_comparator:
assumes "eq_comp acomp" and "sym_comp acomp" and "trans_comp acomp"
shows "comparator acomp"
proof
fix x y z
show "invert_order (acomp x y) = acomp y x"
using sym_compD [OF \<open>sym_comp acomp\<close>] .
{
assume "acomp x y = Eq"
with eq_compD [OF \<open>eq_comp acomp\<close>]
show "x = y" by blast
}
{
assume "acomp x y = Lt" and "acomp y z = Lt"
with trans_orderD [OF trans_compD [OF \<open>trans_comp acomp\<close>], of x y z]
show "acomp x z = Lt" by auto
}
qed
lemma comparator_imp_eq_sym_trans:
assumes "comparator acomp"
shows "eq_comp acomp" "sym_comp acomp" "trans_comp acomp"
proof -
interpret comparator acomp by fact
show "eq_comp acomp" using eq by (intro eq_compI, auto)
show "sym_comp acomp" using sym by (intro sym_compI, auto)
show "trans_comp acomp"
proof (intro trans_compI trans_orderI)
fix x y z
assume "acomp x y \<noteq> Gt" "acomp y z \<noteq> Gt"
thus "acomp x z \<noteq> Gt \<and> (acomp x y = Lt \<or> acomp y z = Lt \<longrightarrow> acomp x z = Lt)"
- using trans [of x y z] and eq [of x y] and eq [of y z]
+ using comp_trans [of x y z] and eq [of x y] and eq [of y z]
by (cases "acomp x y" "acomp y z" rule: order.exhaust [case_product order.exhaust]) auto
qed
qed
context
fixes acomp :: "'a comparator"
assumes c: "comparator acomp"
begin
lemma comp_to_psym_comp: "psym_comp acomp x"
using comparator_imp_eq_sym_trans[OF c]
by (intro sym_compD2)
lemma comp_to_peq_comp: "peq_comp acomp x"
using comparator_imp_eq_sym_trans [OF c]
by (intro eq_compD2)
lemma comp_to_ptrans_comp: "ptrans_comp acomp x"
using comparator_imp_eq_sym_trans [OF c]
by (intro trans_compD2)
end
subsection \<open>Auxiliary Lemmas for Comparator Generator\<close>
lemma forall_finite: "(\<forall> i < (0 :: nat). P i) = True"
"(\<forall> i < Suc 0. P i) = P 0"
"(\<forall> i < Suc (Suc x). P i) = (P 0 \<and> (\<forall> i < Suc x. P (Suc i)))"
by (auto, case_tac i, auto)
lemma trans_order_different:
"trans_order a b Lt"
"trans_order Gt b c"
"trans_order a Gt c"
by (intro trans_orderI, auto)+
lemma length_nth_simps:
"length [] = 0" "length (x # xs) = Suc (length xs)"
"(x # xs) ! 0 = x" "(x # xs) ! (Suc n) = xs ! n" by auto
subsection \<open>The Comparator Generator\<close>
ML_file \<open>comparator_generator.ML\<close>
end
diff --git a/thys/Dijkstra_Shortest_Path/Weight.thy b/thys/Dijkstra_Shortest_Path/Weight.thy
--- a/thys/Dijkstra_Shortest_Path/Weight.thy
+++ b/thys/Dijkstra_Shortest_Path/Weight.thy
@@ -1,182 +1,182 @@
section "Weights for Dijkstra's Algorithm"
theory Weight
imports Complex_Main
begin
text \<open>
In this theory, we set up a type class for weights, and
a typeclass for weights with an infinity element. The latter
one is used internally in Dijkstra's algorithm.
Moreover, we provide a datatype that adds an infinity element to a given
base type.
\<close>
subsection \<open>Type Classes Setup\<close>
class weight = ordered_ab_semigroup_add + comm_monoid_add + linorder
begin
lemma add_nonneg_nonneg [simp]:
assumes "0 \<le> a" and "0 \<le> b" shows "0 \<le> a + b"
proof -
have "0 + 0 \<le> a + b"
using assms by (rule add_mono)
then show ?thesis by simp
qed
lemma add_nonpos_nonpos[simp]:
assumes "a \<le> 0" and "b \<le> 0" shows "a + b \<le> 0"
proof -
have "a + b \<le> 0 + 0"
using assms by (rule add_mono)
then show ?thesis by simp
qed
lemma add_nonneg_eq_0_iff:
assumes x: "0 \<le> x" and y: "0 \<le> y"
shows "x + y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
- by (metis add.comm_neutral add.left_neutral add_left_mono antisym x y)
+ by (metis local.add_0_left local.add_0_right local.add_left_mono local.antisym_conv x y)
lemma add_incr: "0\<le>b \<Longrightarrow> a \<le> a+b"
by (metis add.comm_neutral add_left_mono)
lemma add_incr_left[simp, intro!]: "0\<le>b \<Longrightarrow> a \<le> b + a"
by (metis add_incr add.commute)
lemma sum_not_less[simp, intro!]:
"0\<le>b \<Longrightarrow> \<not> (a+b < a)"
"0\<le>a \<Longrightarrow> \<not> (a+b < b)"
apply (metis add_incr less_le_not_le)
apply (metis add_incr_left less_le_not_le)
done
end
instance nat :: weight ..
instance int :: weight ..
instance rat :: weight ..
instance real :: weight ..
term top
class top_weight = order_top + weight +
assumes inf_add_right[simp]: "a + top = top"
begin
lemma inf_add_left[simp]: "top + a = top"
by (metis add.commute inf_add_right)
lemmas [simp] = top_unique less_top[symmetric]
lemma not_less_inf[simp]:
"\<not> (a < top) \<longleftrightarrow> a=top"
by simp
end
subsection \<open>Adding Infinity\<close>
text \<open>
We provide a standard way to add an infinity element to any type.
\<close>
datatype 'a infty = Infty | Num 'a
primrec val where "val (Num d) = d"
lemma num_val_iff[simp]: "e\<noteq>Infty \<Longrightarrow> Num (val e) = e" by (cases e) auto
type_synonym NatB = "nat infty"
instantiation infty :: (weight) top_weight
begin
definition "(0::'a infty) == Num 0"
definition "top \<equiv> Infty"
fun less_eq_infty where
"less_eq Infty (Num _) \<longleftrightarrow> False" |
"less_eq _ Infty \<longleftrightarrow> True" |
"less_eq (Num a) (Num b) \<longleftrightarrow> a\<le>b"
lemma [simp]: "Infty\<le>a \<longleftrightarrow> a=Infty"
by (cases a) auto
fun less_infty where
"less Infty _ \<longleftrightarrow> False" |
"less (Num _) Infty \<longleftrightarrow> True" |
"less (Num a) (Num b) \<longleftrightarrow> a<b"
lemma [simp]: "less a Infty \<longleftrightarrow> a \<noteq> Infty"
by (cases a) auto
fun plus_infty where
"plus _ Infty = Infty" |
"plus Infty _ = Infty" |
"plus (Num a) (Num b) = Num (a+b)"
lemma [simp]: "plus Infty a = Infty" by (cases a) simp_all
instance
apply (intro_classes)
apply (case_tac [!] x) [4]
apply simp_all
apply (case_tac [!] y) [3]
apply (simp_all add: less_le_not_le)
apply (case_tac z)
apply (simp_all add: top_infty_def zero_infty_def)
apply (case_tac [!] a) [4]
apply simp_all
apply (case_tac [!] b) [3]
apply (simp_all add: ac_simps)
apply (case_tac [!] c) [2]
apply (simp_all add: ac_simps add_right_mono)
apply (case_tac "(x,y)" rule: less_eq_infty.cases)
apply (simp_all add: linear)
done
end
subsubsection \<open>Unboxing\<close>
text \<open>Conversion between the constants defined by the
typeclass, and the concrete functions on the @{typ "'a infty"} type.
\<close>
lemma infty_inf_unbox:
"Num a \<noteq> top"
"top \<noteq> Num a"
"Infty = top"
by (auto simp add: top_infty_def)
lemma infty_ord_unbox:
"Num a \<le> Num b \<longleftrightarrow> a \<le> b"
"Num a < Num b \<longleftrightarrow> a < b"
by auto
lemma infty_plus_unbox:
"Num a + Num b = Num (a+b)"
by (auto)
lemma infty_zero_unbox:
"Num a = 0 \<longleftrightarrow> a = 0"
"Num 0 = 0"
by (auto simp: zero_infty_def)
lemmas infty_unbox =
infty_inf_unbox infty_zero_unbox infty_ord_unbox infty_plus_unbox
lemma inf_not_zero[simp]:
"top\<noteq>(0::_ infty)" "(0::_ infty)\<noteq>top"
apply (unfold zero_infty_def top_infty_def)
apply auto
done
lemma num_val_iff'[simp]: "e\<noteq>top \<Longrightarrow> Num (val e) = e"
by (cases e) (auto simp add: infty_unbox)
lemma infty_neE:
"\<lbrakk>a\<noteq>Infty; \<And>d. a=Num d \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
"\<lbrakk>a\<noteq>top; \<And>d. a=Num d \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (case_tac [!] a) (auto simp add: infty_unbox)
end
diff --git a/thys/Formal_SSA/Construct_SSA_notriv.thy b/thys/Formal_SSA/Construct_SSA_notriv.thy
--- a/thys/Formal_SSA/Construct_SSA_notriv.thy
+++ b/thys/Formal_SSA/Construct_SSA_notriv.thy
@@ -1,865 +1,865 @@
(* Title: Construct_SSA_notriv_code.thy
Author: Sebastian Ullrich, Denis Lohner
*)
subsection \<open>Inductive Removal of Trivial Phi Functions\<close>
theory Construct_SSA_notriv
imports SSA_CFG Minimality "HOL-Library.While_Combinator"
begin
locale CFG_SSA_Transformed_notriv_base = CFG_SSA_Transformed_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses" phis var
for
\<alpha>e :: "'g \<Rightarrow> ('node::linorder \<times> 'edgeD \<times> 'node) set" and
\<alpha>n :: "'g \<Rightarrow> 'node list" and
invar :: "'g \<Rightarrow> bool" and
inEdges' :: "'g \<Rightarrow> 'node \<Rightarrow> ('node \<times> 'edgeD) list" and
Entry::"'g \<Rightarrow> 'node" and
oldDefs :: "'g \<Rightarrow> 'node \<Rightarrow> 'var::linorder set" and
oldUses :: "'g \<Rightarrow> 'node \<Rightarrow> 'var set" and
"defs" :: "'g \<Rightarrow> 'node \<Rightarrow> 'val::linorder set" and
"uses" :: "'g \<Rightarrow> 'node \<Rightarrow> 'val set" and
phis :: "'g \<Rightarrow> ('node, 'val) phis" and
var :: "'g \<Rightarrow> 'val \<Rightarrow> 'var" +
fixes chooseNext_all :: "('node \<Rightarrow> 'val set) \<Rightarrow> ('node, 'val) phis \<Rightarrow> 'g \<Rightarrow> ('node \<times> 'val)"
begin
abbreviation "chooseNext g \<equiv> snd (chooseNext_all (uses g) (phis g) g)"
abbreviation "chooseNext' g \<equiv> chooseNext_all (uses g) (phis g) g"
definition "substitution g \<equiv> THE v'. isTrivialPhi g (chooseNext g) v'"
definition "substNext g \<equiv> \<lambda>v. if v = chooseNext g then substitution g else v"
definition[simp]: "uses' g n \<equiv> substNext g ` uses g n"
definition[simp]: "phis' g x \<equiv> case x of (n,v) \<Rightarrow> if v = chooseNext g
then None
else map_option (map (substNext g)) (phis g (n,v))"
end
locale CFG_SSA_Transformed_notriv = CFG_SSA_Transformed \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses" phis var
+ CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses" phis var chooseNext_all
for
\<alpha>e :: "'g \<Rightarrow> ('node::linorder \<times> 'edgeD \<times> 'node) set" and
\<alpha>n :: "'g \<Rightarrow> 'node list" and
invar :: "'g \<Rightarrow> bool" and
inEdges' :: "'g \<Rightarrow> 'node \<Rightarrow> ('node \<times> 'edgeD) list" and
Entry::"'g \<Rightarrow> 'node" and
"oldDefs" :: "'g \<Rightarrow> 'node \<Rightarrow> 'var::linorder set" and
"oldUses" :: "'g \<Rightarrow> 'node \<Rightarrow> 'var set" and
"defs" :: "'g \<Rightarrow> 'node \<Rightarrow> 'val::linorder set" and
"uses" :: "'g \<Rightarrow> 'node \<Rightarrow> 'val set" and
phis :: "'g \<Rightarrow> ('node, 'val) phis" and
var :: "'g \<Rightarrow> 'val \<Rightarrow> 'var" and
chooseNext_all :: "('node \<Rightarrow> 'val set) \<Rightarrow> ('node, 'val) phis \<Rightarrow> 'g \<Rightarrow> ('node \<times> 'val)" +
assumes chooseNext_all: "CFG_SSA_Transformed \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses defs u p var \<Longrightarrow>
CFG_SSA_wf_base.redundant \<alpha>n inEdges' defs u p g \<Longrightarrow>
chooseNext_all (u g) (p g) g \<in> dom (p g) \<and>
CFG_SSA_wf_base.trivial \<alpha>n inEdges' defs u p g (snd (chooseNext_all (u g) (p g) g))"
begin
lemma chooseNext':"redundant g \<Longrightarrow> chooseNext' g \<in> dom (phis g) \<and> trivial g (chooseNext g)"
by (rule chooseNext_all, unfold_locales)
lemma chooseNext: "redundant g \<Longrightarrow> chooseNext g \<in> allVars g \<and> trivial g (chooseNext g)"
by (drule chooseNext', auto simp: trivial_in_allVars)
lemmas chooseNext_in_allVars[simp] = chooseNext[THEN conjunct1]
lemma isTrivialPhi_det: "trivial g v \<Longrightarrow> \<exists>!v'. isTrivialPhi g v v'"
proof(rule ex_ex1I)
fix v' v''
assume "isTrivialPhi g v v'" "isTrivialPhi g v v''"
from this[unfolded isTrivialPhi_def, THEN conjunct2] show "v' = v''" by (auto simp:isTrivialPhi_def doubleton_eq_iff split:option.splits)
qed (auto simp: trivial_def)
lemma trivialPhi_strict_dom:
assumes[simp]: "v \<in> allVars g" and triv: "isTrivialPhi g v v'"
shows "strict_def_dom g v' v"
proof
let ?n = "defNode g v"
let ?n' = "defNode g v'"
from triv obtain vs where vs: "phi g v = Some vs" "(set vs = {v'} \<or> set vs = {v,v'})" by (auto simp:isTrivialPhi_def split:option.splits)
hence "?n \<noteq> Entry g" by auto
have other_preds_dominated: "\<And>m. m \<in> set (old.predecessors g ?n) \<Longrightarrow> v' \<notin> phiUses g m \<Longrightarrow> old.dominates g ?n m"
proof-
fix m
assume m: "m \<in> set (old.predecessors g ?n)" "v' \<notin> phiUses g m"
hence[simp]: "m \<in> set (\<alpha>n g)" by auto
show "old.dominates g ?n m"
proof (cases "v \<in> phiUses g m")
case True
hence "v \<in> allUses g m" by simp
thus ?thesis by (rule allUses_dominated) simp_all
next
case False
with vs have "v' \<in> phiUses g m" by - (rule phiUses_exI[OF m(1)], auto simp:phi_def)
with m(2) show ?thesis by simp
qed
qed
show "?n' \<noteq> ?n"
proof (rule notI)
assume asm: "?n' = ?n"
have "\<And>m. m \<in> set (old.predecessors g ?n) \<Longrightarrow> v' \<in> phiUses g m \<Longrightarrow> old.dominates g ?n m"
proof-
fix m
assume "m \<in> set (old.predecessors g ?n)" "v' \<in> phiUses g m"
hence "old.dominates g ?n' m" by - (rule allUses_dominated, auto)
thus "?thesis m" by (simp add:asm)
qed
with non_dominated_predecessor[of ?n g] other_preds_dominated \<open>?n \<noteq> Entry g\<close> show False by auto
qed
show "old.dominates g ?n' ?n"
proof
fix ns
assume asm: "g \<turnstile> Entry g-ns\<rightarrow>?n"
from \<open>?n \<noteq> Entry g\<close> obtain m ns'
where ns': "g \<turnstile> Entry g-ns'\<rightarrow>m" "m \<in> set (old.predecessors g ?n)" "?n \<notin> set ns'" "set ns' \<subseteq> set ns"
by - (rule old.simple_path2_unsnoc[OF asm], auto)
hence[simp]: "m \<in> set (\<alpha>n g)" by auto
from ns' have "\<not>old.dominates g ?n m" by (auto elim:old.dominatesE)
with other_preds_dominated[of m] ns'(2) have "v' \<in> phiUses g m" by auto
hence "old.dominates g ?n' m" by - (rule allUses_dominated, auto)
with ns'(1) have "?n' \<in> set ns'" by - (erule old.dominatesE)
with ns'(4) show "?n' \<in> set ns" by auto
qed auto
qed
lemma isTrivialPhi_asymmetric:
assumes "isTrivialPhi g a b"
and "isTrivialPhi g b a"
shows "False"
using assms
proof -
from \<open>isTrivialPhi g a b\<close>
have "b \<in> allVars g"
unfolding isTrivialPhi_def
by (fastforce intro!: phiArg_in_allVars simp: phiArg_def split: option.splits)
from \<open>isTrivialPhi g b a\<close>
have "a \<in> allVars g"
unfolding isTrivialPhi_def
by (fastforce intro!: phiArg_in_allVars simp: phiArg_def split: option.splits)
from trivialPhi_strict_dom [OF \<open>a \<in> allVars g\<close> assms(1)]
trivialPhi_strict_dom [OF \<open>b \<in> allVars g\<close> assms(2)]
show ?thesis by blast
qed
lemma substitution[intro]: "redundant g \<Longrightarrow> isTrivialPhi g (chooseNext g) (substitution g)"
unfolding substitution_def by (rule theI', rule isTrivialPhi_det, simp add: chooseNext)
lemma trivialPhi_in_allVars[simp]:
assumes "isTrivialPhi g v v'" and[simp]: "v \<in> allVars g"
shows "v' \<in> allVars g"
proof-
from assms(1) have "phiArg g v v'"
unfolding phiArg_def
by (auto simp:isTrivialPhi_def split:option.splits)
thus "v' \<in> allVars g" by - (rule phiArg_in_allVars, auto)
qed
lemma substitution_in_allVars[simp]:
assumes "redundant g"
shows "substitution g \<in> allVars g"
using assms by - (rule trivialPhi_in_allVars, auto)
lemma defs_uses_disjoint_inv:
assumes[simp]: "n \<in> set (\<alpha>n g)" "redundant g"
shows "defs g n \<inter> uses' g n = {}"
proof (rule equals0I)
fix v'
assume asm: "v' \<in> defs g n \<inter> uses' g n"
then obtain v where v: "v \<in> uses g n" "v' = substNext g v" and v': "v' \<in> defs g n" by auto
show False
proof (cases "v = chooseNext g")
case False
thus ?thesis using v v' defs_uses_disjoint[of n g] by (auto simp:substNext_def split:if_split_asm)
next
case [simp]: True
from v' have n_defNode: "n = defNode g v'" by - (rule defNode_eq[symmetric], auto)
from v(1) have[simp]: "v \<in> allVars g" by - (rule allUses_in_allVars[where n=n], auto)
let ?n' = "defNode g v"
have "old.strict_dom g n ?n'"
by (simp only:n_defNode v(2), rule trivialPhi_strict_dom, auto simp:substNext_def)
moreover from v(1) have "old.dominates g ?n' n" by - (rule allUses_dominated, auto)
ultimately show False by auto
qed
qed
end
context CFG_SSA_wf
begin
inductive liveVal' :: "'g \<Rightarrow> 'val list \<Rightarrow> bool"
for g :: 'g
where
liveSimple': "\<lbrakk>n \<in> set (\<alpha>n g); val \<in> uses g n\<rbrakk> \<Longrightarrow> liveVal' g [val]"
| livePhi': "\<lbrakk>liveVal' g (v#vs); phiArg g v v'\<rbrakk> \<Longrightarrow> liveVal' g (v'#v#vs)"
lemma liveVal'_suffix:
assumes "liveVal' g vs" "suffix vs' vs" "vs' \<noteq> []"
shows "liveVal' g vs'"
using assms proof induction
case (liveSimple' n v)
from liveSimple'.prems have "vs' = [v]"
- by (metis append_Nil butlast.simps(2) suffixI suffix_order.antisym suffix_unsnoc)
+ by (metis append_Nil butlast.simps(2) suffixI suffix_order.order_antisym suffix_unsnoc)
with liveSimple'.hyps show ?case by (auto intro: liveVal'.liveSimple')
next
case (livePhi' v vs v')
show ?case
proof (cases "vs' = v' # v # vs")
case True
with livePhi' show ?thesis by - (auto intro: liveVal'.livePhi')
next
case False
with livePhi'.prems have "suffix vs' (v#vs)"
by (metis list.sel(3) self_append_conv2 suffixI suffix_take tl_append2)
with livePhi'.prems(2) show ?thesis by - (rule livePhi'.IH)
qed
qed
lemma liveVal'I:
assumes "liveVal g v"
obtains vs where "liveVal' g (v#vs)"
using assms proof induction
case (liveSimple n v)
thus thesis by - (rule liveSimple(3), rule liveSimple')
next
case (livePhi v v')
show thesis
proof (rule livePhi.IH)
fix vs
assume asm: "liveVal' g (v#vs)"
show thesis
proof (cases "v' \<in> set (v#vs)")
case False
with livePhi.hyps asm show thesis by - (rule livePhi.prems, rule livePhi')
next
case True
then obtain vs' where "suffix (v'#vs') (v#vs)"
by - (drule split_list_last, auto simp: Sublist.suffix_def)
with asm show thesis by - (rule livePhi.prems, rule liveVal'_suffix, simp_all)
qed
qed
qed
lemma liveVal'D:
assumes "liveVal' g vs" "vs = v#vs'"
shows "liveVal g v"
using assms proof (induction arbitrary: v vs')
case (liveSimple' n vs)
thus ?case by - (rule liveSimple, auto)
next
case (livePhi' v\<^sub>2 vs v')
thus ?case by - (rule livePhi, auto)
qed
end
locale CFG_SSA_step = CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses" phis var chooseNext_all
for
\<alpha>e :: "'g \<Rightarrow> ('node::linorder \<times> 'edgeD \<times> 'node) set" and
\<alpha>n :: "'g \<Rightarrow> 'node list" and
invar :: "'g \<Rightarrow> bool" and
inEdges' :: "'g \<Rightarrow> 'node \<Rightarrow> ('node \<times> 'edgeD) list" and
Entry::"'g \<Rightarrow> 'node" and
"oldDefs" :: "'g \<Rightarrow> 'node \<Rightarrow> 'var::linorder set" and
"oldUses" :: "'g \<Rightarrow> 'node \<Rightarrow> 'var set" and
"defs" :: "'g \<Rightarrow> 'node \<Rightarrow> 'val::linorder set" and
"uses" :: "'g \<Rightarrow> 'node \<Rightarrow> 'val set" and
phis :: "'g \<Rightarrow> ('node, 'val) phis" and
var :: "'g \<Rightarrow> 'val \<Rightarrow> 'var" and
chooseNext_all :: "('node \<Rightarrow> 'val set) \<Rightarrow> ('node, 'val) phis \<Rightarrow> 'g \<Rightarrow> ('node \<times> 'val)" and
g :: "'g" +
assumes redundant[simp]: "redundant g"
begin
abbreviation "u_g \<equiv> uses(g:=uses' g)"
abbreviation "p_g \<equiv> phis(g:=phis' g)"
sublocale step: CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u_g p_g var chooseNext_all .
lemma simpleDefs_phiDefs_disjoint_inv:
assumes "n \<in> set (\<alpha>n g)"
shows "defs g n \<inter> step.phiDefs g n = {}"
using simpleDefs_phiDefs_disjoint[OF assms]
by (auto simp: phiDefs_def step.phiDefs_def dom_def split:option.splits)
lemma allDefs_disjoint_inv:
assumes "n \<in> set (\<alpha>n g)" "m \<in> set (\<alpha>n g)" "n \<noteq> m"
shows "step.allDefs g n \<inter> step.allDefs g m = {}"
using allDefs_disjoint[OF assms]
by (auto simp: CFG_SSA_defs step.CFG_SSA_defs dom_def split:option.splits)
lemma phis_finite_inv:
shows "finite (dom (phis' g))"
using phis_finite[of g] by - (rule finite_subset, auto split:if_split_asm)
lemma phis_wf_inv:
assumes "phis' g (n, v) = Some args"
shows "length (old.predecessors g n) = length args"
using phis_wf[of g] assms by (auto split:if_split_asm)
sublocale step: CFG_SSA \<alpha>e \<alpha>n invar inEdges' Entry "defs" u_g p_g
apply unfold_locales
apply (rename_tac g')
apply (case_tac "g'=g")
apply (simp add:defs_uses_disjoint_inv[simplified])
apply (simp add:defs_uses_disjoint)
apply (rule defs_finite)
apply (auto simp: uses_in_\<alpha>n split: if_split_asm)[1]
apply (rename_tac g' n)
apply (case_tac "g'=g")
apply simp
apply simp
apply (rule invar)
apply (rename_tac g')
apply (case_tac "g'=g")
apply (simp add:phis_finite_inv)
apply (simp add:phis_finite)
apply (auto simp: phis_in_\<alpha>n split: if_split_asm)[1]
apply (rename_tac g' n v args)
apply (case_tac "g'=g")
apply (simp add:phis_wf_inv)
apply (simp add:phis_wf)
apply (rename_tac g')
apply (case_tac "g'=g")
apply (simp add: simpleDefs_phiDefs_disjoint_inv)
apply (simp add: simpleDefs_phiDefs_disjoint[unfolded CFG_SSA_defs] step.CFG_SSA_defs )
apply (rename_tac g' m)
apply (case_tac "g'=g")
apply (simp add: allDefs_disjoint_inv)
apply (simp add: allDefs_disjoint[unfolded CFG_SSA_defs] step.CFG_SSA_defs)
done
lemma allUses_narrows:
assumes "n \<in> set (\<alpha>n g)"
shows "step.allUses g n \<subseteq> substNext g ` allUses g n"
proof-
have "\<And>n' v' z b. phis g (n', v') = Some z \<Longrightarrow> (n, b) \<in> set (zip (old.predecessors g n') z) \<Longrightarrow> b \<notin> phiUses g n \<Longrightarrow> b \<in> uses g n"
proof-
fix n' v' z b
assume "(n, b) \<in> set (zip (old.predecessors g n') (z :: 'val list))"
with assms(1) have "n' \<in> set (\<alpha>n g)" by auto
thus "phis g (n', v') = Some z \<Longrightarrow> (n, b) \<in> set (zip (old.predecessors g n') z) \<Longrightarrow> b \<notin> phiUses g n \<Longrightarrow> b \<in> uses g n" by (auto intro:phiUsesI)
qed
thus ?thesis by (auto simp:step.allUses_def allUses_def zip_map2 intro!:imageI elim!:step.phiUsesE phiUsesE split:if_split_asm)
qed
lemma allDefs_narrows[simp]: "v \<in> step.allDefs g n \<Longrightarrow> v \<in> allDefs g n"
by (auto simp:step.allDefs_def step.phiDefs_def phiDefs_def allDefs_def split:if_split_asm)
lemma allUses_def_ass_inv:
assumes "v' \<in> step.allUses g n" "n \<in> set (\<alpha>n g)"
shows "step.defAss g n v'"
proof (rule step.defAssI)
fix ns
assume asm: "g \<turnstile> Entry g-ns\<rightarrow>n"
from assms obtain v where v': "v' = substNext g v" and[simp]: "v \<in> allUses g n"
using allUses_narrows by auto
with assms(2) have[simp]: "v \<in> allVars g" by - (rule allUses_in_allVars)
have[simp]: "v' \<in> allVars g" by (simp add:substNext_def v')
let ?n\<^sub>v = "defNode g v"
let ?n\<^sub>v\<^sub>' = "defNode g v'"
from assms(2) asm have 1: "?n\<^sub>v \<in> set ns" using allUses_def_ass[of v g n] by (simp add:defAss_defNode)
then obtain ns\<^sub>v where ns\<^sub>v: "prefix (ns\<^sub>v@[?n\<^sub>v]) ns" by (rule prefix_split_first)
with asm have 2: "g \<turnstile> Entry g-ns\<^sub>v@[?n\<^sub>v]\<rightarrow>?n\<^sub>v" by auto
show "\<exists>n \<in> set ns. v' \<in> step.allDefs g n"
proof (cases "v = chooseNext g")
case True
hence dom: "strict_def_dom g v' v" using substitution[of g] by - (rule trivialPhi_strict_dom, simp_all add:substNext_def v')
hence[simp]: "v' \<noteq> v" by auto
have "v' \<in> allDefs g ?n\<^sub>v\<^sub>'" by simp
hence "v' \<in> step.allDefs g ?n\<^sub>v\<^sub>'" unfolding step.allDefs_def step.phiDefs_def allDefs_def phiDefs_def by (auto simp:True[symmetric])
moreover have "?n\<^sub>v\<^sub>' \<in> set ns"
proof-
from dom have "def_dominates g v' v" by auto
hence "?n\<^sub>v\<^sub>' \<in> set (ns\<^sub>v@[?n\<^sub>v])" using 2 by -(erule old.dominatesE)
with ns\<^sub>v show ?thesis by auto
qed
ultimately show ?thesis by auto
next
case [simp]: False
have[simp]: "v' = v" by (simp add:v' substNext_def)
have "v \<in> allDefs g ?n\<^sub>v" by simp
thus ?thesis by - (rule bexI[of _ ?n\<^sub>v\<^sub>'], auto simp:allDefs_def step.allDefs_def step.phiDefs_def 1 phiDefs_def)
qed
qed
lemma Entry_no_phis_inv: "phis' g (Entry g, v) = None"
by (simp add:Entry_no_phis)
sublocale step: CFG_SSA_wf \<alpha>e \<alpha>n invar inEdges' Entry "defs" u_g p_g
apply unfold_locales
apply (rename_tac g' n)
apply (case_tac "g'=g")
apply (simp add:allUses_def_ass_inv)
apply (simp add:allUses_def_ass[unfolded CFG_SSA_defs, simplified] step.CFG_SSA_defs)
apply (rename_tac g' v)
apply (case_tac "g'=g")
apply (simp add:Entry_no_phis_inv)
apply (simp)
done
lemma chooseNext_eliminated: "chooseNext g \<notin> step.allDefs g (defNode g (chooseNext g))"
proof-
let ?v = "chooseNext g"
let ?n = "defNode g ?v"
from chooseNext[OF redundant] have "?v \<in> phiDefs g ?n" "?n \<in> set (\<alpha>n g)"
by (auto simp: trivial_def isTrivialPhi_def phiDefs_def phi_def split: option.splits)
hence "?v \<notin> defs g ?n" using simpleDefs_phiDefs_disjoint[of ?n g] by auto
thus ?thesis by (auto simp:step.allDefs_def step.phiDefs_def)
qed
lemma oldUses_inv:
assumes "n \<in> set (\<alpha>n g)"
shows "oldUses g n = var g ` u_g g n"
proof-
have "var g (substitution g) = var g (chooseNext g)" using substitution[of g]
by - (rule phiArg_same_var, auto simp: isTrivialPhi_def phiArg_def split: option.splits)
thus ?thesis using assms by (auto simp: substNext_def oldUses_def image_Un)
qed
lemma conventional_inv:
assumes "g \<turnstile> n-ns\<rightarrow>m" "n \<notin> set (tl ns)" "v \<in> step.allDefs g n" "v \<in> step.allUses g m" "x \<in> set (tl ns)" "v' \<in> step.allDefs g x"
shows "var g v' \<noteq> var g v"
proof-
from assms(1,3) have[simp]: "n = defNode g v" "v \<in> allDefs g n" by - (rule defNode_eq[symmetric], auto)
from assms(1) have[simp]: "m \<in> set (\<alpha>n g)" by auto
from assms(4) obtain v\<^sub>0 where v\<^sub>0: "v = substNext g v\<^sub>0" "v\<^sub>0 \<in> allUses g m" using allUses_narrows[of m] by auto
hence[simp]: "v\<^sub>0 \<in> allVars g" using assms(1) by auto
let ?n\<^sub>0 = "defNode g v\<^sub>0"
show ?thesis
proof (cases "v\<^sub>0 = chooseNext g")
case False
with v\<^sub>0 have "v = v\<^sub>0" by (simp add:substNext_def split:if_split_asm)
with assms v\<^sub>0 show ?thesis by - (rule conventional, auto)
next
case True
hence dom: "strict_def_dom g v v\<^sub>0" using substitution[of g] by - (rule trivialPhi_strict_dom, simp_all add:substNext_def v\<^sub>0)
from v\<^sub>0(2) have "old.dominates g ?n\<^sub>0 m" using assms(1) by - (rule allUses_dominated, auto)
with assms(1) dom have "?n\<^sub>0 \<in> set ns" by - (rule old.dominates_mid, auto)
with assms(1) obtain ns\<^sub>1 ns\<^sub>3 ns\<^sub>2 where
ns: "ns = ns\<^sub>1@ns\<^sub>3@ns\<^sub>2" and
ns\<^sub>1: "g \<turnstile> n-ns\<^sub>1@[?n\<^sub>0]\<rightarrow>?n\<^sub>0" "?n\<^sub>0 \<notin> set ns\<^sub>1" and
ns\<^sub>3: "g \<turnstile> ?n\<^sub>0-ns\<^sub>3\<rightarrow>?n\<^sub>0" and
ns\<^sub>2: "g \<turnstile> ?n\<^sub>0-?n\<^sub>0#ns\<^sub>2\<rightarrow>m" "?n\<^sub>0 \<notin> set ns\<^sub>2" by (rule old.path2_split_first_last)
have[simp]: "ns\<^sub>1 \<noteq> []"
proof
assume "ns\<^sub>1 = []"
hence "?n\<^sub>0 = n" "hd ns = n" using assms(1) ns\<^sub>3 by (auto simp:ns old.path2_def)
thus False by (metis \<open>n = defNode g v\<close> dom)
qed
hence "length (ns\<^sub>1@[?n\<^sub>0]) \<ge> 2" by (cases ns\<^sub>1, auto)
with ns\<^sub>1 have 1: "g \<turnstile> n-ns\<^sub>1\<rightarrow>last ns\<^sub>1" "last ns\<^sub>1 \<in> set (old.predecessors g ?n\<^sub>0)" by - (erule old.path2_unsnoc, simp, simp, erule old.path2_unsnoc, auto)
from \<open>v\<^sub>0 = chooseNext g\<close> v\<^sub>0 have triv: "isTrivialPhi g v\<^sub>0 v" using substitution[of g] by (auto simp:substNext_def)
then obtain vs where vs: "phi g v\<^sub>0 = Some vs" "set vs = {v\<^sub>0,v} \<or> set vs = {v}" by (auto simp:isTrivialPhi_def split:option.splits)
hence[simp]: "var g v\<^sub>0 = var g v" by - (rule phiArg_same_var[symmetric], auto simp: phiArg_def)
have[simp]: "v \<in> phiUses g (last ns\<^sub>1)"
proof-
from vs ns\<^sub>1 1 have "v \<in> phiUses g (last ns\<^sub>1) \<or> v\<^sub>0 \<in> phiUses g (last ns\<^sub>1)" by - (rule phiUses_exI[of "last ns\<^sub>1" g ?n\<^sub>0 v\<^sub>0 vs], auto simp:phi_def)
moreover have "v\<^sub>0 \<notin> phiUses g (last ns\<^sub>1)"
proof
assume asm: "v\<^sub>0 \<in> phiUses g (last ns\<^sub>1)"
from True have "last ns\<^sub>1 \<in> set ns\<^sub>1" by - (rule last_in_set, auto)
hence "last ns\<^sub>1 \<in> set (\<alpha>n g)" by - (rule old.path2_in_\<alpha>n[OF ns\<^sub>1(1)], auto)
with asm ns\<^sub>1 have "old.dominates g ?n\<^sub>0 (last ns\<^sub>1)" by - (rule allUses_dominated, auto)
moreover have "strict_def_dom g v v\<^sub>0" using triv by - (rule trivialPhi_strict_dom, auto)
ultimately have "?n\<^sub>0 \<in> set ns\<^sub>1" using 1(1) by - (rule old.dominates_mid, auto)
with ns\<^sub>1(2) show False ..
qed
ultimately show ?thesis by simp
qed
show ?thesis
proof (cases "x \<in> set (tl ns\<^sub>1)")
case True
thus ?thesis using assms(2,3,6) by - (rule conventional[where x=x, OF 1(1)], auto simp:ns)
next
case False
show ?thesis
proof (cases "var g v' = var g v\<^sub>0")
case [simp]: True
{
assume asm: "x \<in> set ns\<^sub>3"
with assms(6)[THEN allDefs_narrows] have[simp]: "x = defNode g v'"
using ns\<^sub>3 by - (rule defNode_eq[symmetric], auto)
{
assume "v' = v\<^sub>0"
hence False using assms(6) \<open>v\<^sub>0 = chooseNext g\<close> simpleDefs_phiDefs_disjoint[of x g] vs(1)
by (auto simp: step.allDefs_def step.phiDefs_def)
}
moreover {
assume "v' \<noteq> v\<^sub>0"
hence "x \<noteq> ?n\<^sub>0" using allDefs_var_disjoint[OF _ assms(6)[THEN allDefs_narrows], of v\<^sub>0]
by auto
from ns\<^sub>3 asm ns obtain ns\<^sub>3 where ns\<^sub>3: "g \<turnstile> ?n\<^sub>0-ns\<^sub>3\<rightarrow>?n\<^sub>0" "?n\<^sub>0 \<notin> set (tl (butlast ns\<^sub>3))" "x \<in> set ns\<^sub>3" "set ns\<^sub>3 \<subseteq> set (tl ns)"
by - (rule old.path2_simple_loop, auto)
with \<open>x \<noteq> ?n\<^sub>0\<close> have "length ns\<^sub>3 > 1"
by (metis empty_iff graph_path_base.path2_def hd_Cons_tl insert_iff length_greater_0_conv length_tl list.set(1) list.set(2) zero_less_diff)
with ns\<^sub>3 obtain ns' m where ns': "g \<turnstile> ?n\<^sub>0-ns'\<rightarrow>m" "m \<in> set (old.predecessors g ?n\<^sub>0)" "ns' = butlast ns\<^sub>3"
by - (rule old.path2_unsnoc, auto)
with vs ns\<^sub>3 have "v \<in> phiUses g m \<or> v\<^sub>0 \<in> phiUses g m"
by - (rule phiUses_exI[of m g ?n\<^sub>0 v\<^sub>0 vs], auto simp:phi_def)
moreover {
assume "v \<in> phiUses g m"
have "var g v\<^sub>0 \<noteq> var g v"
proof (rule conventional)
show "g \<turnstile> n-ns\<^sub>1 @ ns'\<rightarrow>m" using old.path2_app'[OF ns\<^sub>1(1) ns'(1)] by simp
have "n \<notin> set (tl ns\<^sub>1)" using ns assms(2) by auto
moreover have "n \<notin> set ns'" using ns'(3) ns\<^sub>3(4) assms(2) by (auto dest: in_set_butlastD)
ultimately show "n \<notin> set (tl (ns\<^sub>1 @ ns'))" by simp
show "v \<in> allDefs g n" using \<open>v \<in> allDefs g n\<close> .
show "?n\<^sub>0 \<in> set (tl (ns\<^sub>1 @ ns'))" using ns'(1) by (auto simp: old.path2_def)
qed (auto simp: \<open>v \<in> phiUses g m\<close>)
hence False by simp
}
moreover {
assume "v\<^sub>0 \<in> phiUses g m"
moreover from ns\<^sub>3(1,3) \<open>x \<noteq> ?n\<^sub>0\<close> \<open>length ns\<^sub>3 > 1\<close> have "x \<in> set (tl (butlast ns\<^sub>3))"
by (cases ns\<^sub>3, auto simp: old.path2_def intro: in_set_butlastI)
ultimately have "var g v' \<noteq> var g v\<^sub>0"
using assms(6)[THEN allDefs_narrows] ns\<^sub>3(2,3) ns'(3) by - (rule conventional[OF ns'(1)], auto)
hence False by simp
}
ultimately have False by auto
}
ultimately have False by auto
}
moreover {
assume asm: "x \<notin> set ns\<^sub>3"
have "var g v' \<noteq> var g v\<^sub>0"
proof (cases "x = ?n\<^sub>0")
case True
moreover have "v\<^sub>0 \<notin> step.allDefs g ?n\<^sub>0" by (auto simp:\<open>v\<^sub>0 = chooseNext g\<close> chooseNext_eliminated)
ultimately show ?thesis using assms(6) vs(1) by - (rule allDefs_var_disjoint[of x g], auto)
next
case False
with \<open>x \<notin> set (tl ns\<^sub>1)\<close> assms(5) asm have "x \<in> set ns\<^sub>2" by (auto simp:ns)
thus ?thesis using assms(2,6) v\<^sub>0(2) ns\<^sub>2(2) by - (rule conventional[OF ns\<^sub>2(1), where x=x], auto simp:ns)
qed
}
ultimately show ?thesis by auto
qed auto
qed
qed
qed
lemma[simp]: "var g (substNext g v) = var g v"
using substitution[OF redundant]
by (auto simp:substNext_def isTrivialPhi_def phi_def split:option.splits)
lemma phis_same_var_inv:
assumes "phis' g (n,v) = Some vs" "v' \<in> set vs"
shows "var g v' = var g v"
proof-
from assms obtain vs\<^sub>0 v\<^sub>0 where 1: "phis g (n,v) = Some vs\<^sub>0" "v\<^sub>0 \<in> set vs\<^sub>0" "v' = substNext g v\<^sub>0" by (auto split:if_split_asm)
hence "var g v\<^sub>0 = var g v" by auto
with 1 show ?thesis by auto
qed
lemma allDefs_var_disjoint_inv: "\<lbrakk>n \<in> set (\<alpha>n g); v \<in> step.allDefs g n; v' \<in> step.allDefs g n; v \<noteq> v'\<rbrakk> \<Longrightarrow> var g v' \<noteq> var g v"
using allDefs_var_disjoint
by (auto simp:step.allDefs_def)
lemma step_CFG_SSA_Transformed_notriv: "CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses defs u_g p_g var chooseNext_all"
apply unfold_locales
apply (rule oldDefs_def)
apply (rename_tac g')
apply (case_tac "g'=g")
apply (simp add:oldUses_inv)
apply (simp add:oldUses_def)
apply (rename_tac g' n ns m v x v')
apply (case_tac "g'=g")
apply (simp add:conventional_inv)
apply (simp add:conventional[unfolded CFG_SSA_defs, simplified] step.CFG_SSA_defs)
apply (rename_tac g' n v vs v')
apply (case_tac "g'=g")
apply (simp add:phis_same_var_inv)
apply (simp add:phis_same_var)
apply (rename_tac g' v v')
apply (case_tac "g'=g")
apply (simp add:allDefs_var_disjoint_inv)
apply (simp add:allDefs_var_disjoint[unfolded allDefs_def phiDefs_def, simplified] step.allDefs_def step.phiDefs_def)
by (rule chooseNext_all)
sublocale step: CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u_g p_g var chooseNext_all
by (rule step_CFG_SSA_Transformed_notriv)
lemma step_defNode: "v \<in> allVars g \<Longrightarrow> v \<noteq> chooseNext g \<Longrightarrow> step.defNode g v = defNode g v"
by (auto simp: step.CFG_SSA_wf_defs dom_def CFG_SSA_wf_defs)
lemma step_phi: "v \<in> allVars g \<Longrightarrow> v \<noteq> chooseNext g \<Longrightarrow> step.phi g v = map_option (map (substNext g)) (phi g v)"
by (auto simp: step.phi_def step_defNode phi_def)
lemma liveVal'_inv:
assumes "liveVal' g (v#vs)" "v \<noteq> chooseNext g"
obtains vs' where "step.liveVal' g (v#vs')"
using assms proof (induction "length vs" arbitrary: v vs rule: nat_less_induct)
case (1 vs v)
from "1.prems"(2) show thesis
proof cases
case (liveSimple' n)
with "1.prems"(3) show thesis by - (rule "1.prems"(1), rule step.liveSimple', auto simp: substNext_def)
next
case (livePhi' v' vs')
from this(2) have[simp]: "v' \<in> allVars g" by - (drule liveVal'D, rule, rule liveVal_in_allVars)
show thesis
proof (cases "chooseNext g = v'")
case False
show thesis
proof (rule "1.hyps"[rule_format, of "length vs'" vs' v'])
fix vs'\<^sub>2
assume asm: "step.liveVal' g (v'#vs'\<^sub>2)"
have "step.phiArg g v' v" using livePhi'(3) False "1.prems"(3) by (auto simp: step.phiArg_def phiArg_def step_phi substNext_def)
thus thesis by - (rule "1.prems"(1), rule step.livePhi', rule asm)
qed (auto simp: livePhi' False[symmetric])
next
case [simp]: True
with "1.prems"(3) have[simp]: "v \<noteq> v'" by simp
from True have "trivial g v'" using chooseNext[OF redundant] by auto
with \<open>phiArg g v' v\<close> have "isTrivialPhi g v' v" by (auto simp: phiArg_def trivial_def isTrivialPhi_def)
hence[simp]: "substitution g = v" unfolding substitution_def
by - (rule the1_equality, auto intro!: isTrivialPhi_det[unfolded trivial_def])
obtain vs'\<^sub>2 where vs'\<^sub>2: "suffix (v'#vs'\<^sub>2) (v'#vs')" "v' \<notin> set vs'\<^sub>2"
using split_list_last[of v' "v'#vs'"] by (auto simp: Sublist.suffix_def)
with \<open>liveVal' g (v'#vs')\<close> have "liveVal' g (v'#vs'\<^sub>2)" by - (rule liveVal'_suffix, simp_all)
thus thesis
proof (cases rule: liveVal'.cases)
case (liveSimple' n)
hence "v \<in> uses' g n" by (auto simp: substNext_def)
with liveSimple' show thesis by - (rule "1.prems"(1), rule step.liveSimple', auto)
next
case (livePhi' v'' vs'')
from this(2) have[simp]: "v'' \<in> allVars g" by - (drule liveVal'D, rule, rule liveVal_in_allVars)
from vs'\<^sub>2(2) livePhi'(1) have[simp]: "v'' \<noteq> v'" by auto
show thesis
proof (rule "1.hyps"[rule_format, of "length vs''" vs'' v''])
show "length vs'' < length vs" using \<open>vs = v'#vs'\<close> livePhi'(1) vs'\<^sub>2(1)[THEN suffix_ConsD2]
by (auto simp: Sublist.suffix_def)
next
fix vs''\<^sub>2
assume asm: "step.liveVal' g (v''#vs''\<^sub>2)"
from livePhi' \<open>phiArg g v' v\<close> have "step.phiArg g v'' v" by (auto simp: phiArg_def step.phiArg_def step_phi substNext_def)
thus thesis by - (rule "1.prems"(1), rule step.livePhi', rule asm)
qed (auto simp: livePhi'(2))
qed
qed
qed
qed
lemma liveVal_inv:
assumes "liveVal g v" "v \<noteq> chooseNext g"
shows "step.liveVal g v"
apply (rule liveVal'I[OF assms(1)])
apply (erule liveVal'_inv[OF _ assms(2)])
apply (rule step.liveVal'D)
by simp_all
lemma pruned_inv:
assumes "pruned g"
shows "step.pruned g"
proof (rule step.pruned_def[THEN iffD2, rule_format])
fix n v
assume "v \<in> step.phiDefs g n" and[simp]: "n \<in> set (\<alpha>n g)"
hence "v \<in> phiDefs g n" "v \<noteq> chooseNext g" by (auto simp: step.CFG_SSA_defs CFG_SSA_defs split: if_split_asm)
hence "liveVal g v" using assms by (auto simp: pruned_def)
thus "step.liveVal g v" using \<open>v \<noteq> chooseNext g\<close> by (rule liveVal_inv)
qed
end
context CFG_SSA_Transformed_notriv_base
begin
abbreviation "inst g u p \<equiv> CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses defs (uses(g:=u)) (phis(g:=p)) var chooseNext_all"
abbreviation "inst' g \<equiv> \<lambda>(u,p). inst g u p"
interpretation uninst: CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u p var chooseNext_all
for u and p
by unfold_locales
definition "cond g \<equiv> \<lambda>(u,p). uninst.redundant (uses(g:=u)) (phis(g:=p)) g"
definition "step g \<equiv> \<lambda>(u,p). (uninst.uses' (uses(g:=u)) (phis(g:=p)) g,
uninst.phis' (uses(g:=u)) (phis(g:=p)) g)"
definition[code]: "substAll g \<equiv> while (cond g) (step g) (uses g,phis g)"
definition[code]: "uses'_all g \<equiv> fst (substAll g)"
definition[code]: "phis'_all g \<equiv> snd (substAll g)"
lemma uninst_allVars_simps [simp]:
"uninst.allVars u (\<lambda>_. p g) g = uninst.allVars u p g"
"uninst.allVars (\<lambda>_. u g) p g = uninst.allVars u p g"
"uninst.allVars (uses(g:=u g)) p g = uninst.allVars u p g"
"uninst.allVars u (phis(g:=p g)) g = uninst.allVars u p g"
unfolding uninst.allVars_def uninst.allDefs_def uninst.allUses_def uninst.phiDefs_def uninst.phiUses_def
by simp_all
lemma uninst_trivial_simps [simp]:
"uninst.trivial u (\<lambda>_. p g) g = uninst.trivial u p g"
"uninst.trivial (\<lambda>_. u g) p g = uninst.trivial u p g"
"uninst.trivial (uses(g:=u g)) p g = uninst.trivial u p g"
"uninst.trivial u (phis(g:=p g)) g = uninst.trivial u p g"
unfolding uninst.trivial_def [abs_def] uninst.isTrivialPhi_def uninst.phi_def uninst.defNode_code
uninst.allDefs_def uninst.phiDefs_def
by simp_all
end
context CFG_SSA_Transformed_notriv
begin
declare fun_upd_apply[simp del] fun_upd_same[simp]
lemma substAll_wf:
assumes[simp]: "redundant g"
shows "card (dom (phis' g)) < card (dom (phis g))"
proof (rule psubset_card_mono)
let ?v = "chooseNext g"
from chooseNext[of g] obtain n where "(n,?v) \<in> dom (phis g)" by (auto simp: trivial_def isTrivialPhi_def phi_def split:option.splits)
moreover have "(n,?v) \<notin> dom (phis' g)" by auto
ultimately have "dom (phis' g) \<noteq> dom (phis g)" by auto
thus "dom (phis' g) \<subset> dom (phis g)" by (auto split:if_split_asm)
qed (rule phis_finite)
lemma step_preserves_inst:
assumes "inst' g (u,p)"
and "CFG_SSA_wf_base.redundant \<alpha>n inEdges' defs (uses(g:=u)) (phis(g:=p)) g"
shows "inst' g (step g (u,p))"
proof-
from assms(1) interpret i: CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses(g:=u)" "phis(g:=p)" var
by simp
from assms(2) interpret step: CFG_SSA_step \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses(g:=u)" "phis(g:=p)" var chooseNext_all
by unfold_locales
show ?thesis using step.step_CFG_SSA_Transformed_notriv[simplified] by (simp add: step_def)
qed
lemma substAll:
assumes "P (uses g, phis g)"
assumes "\<And>x. P x \<Longrightarrow> inst' g x \<Longrightarrow> cond g x \<Longrightarrow> P (step g x)"
assumes "\<And>x. P x \<Longrightarrow> inst' g x \<Longrightarrow> \<not>cond g x \<Longrightarrow> Q (fst x) (snd x)"
shows "inst g (uses'_all g) (phis'_all g)" "Q (uses'_all g) (phis'_all g)"
proof-
note uses'_def[simp del]
note phis'_def[simp del]
have 2: "\<And>f x. f x = f (fst x, snd x)" by simp
have "inst' g (substAll g) \<and> Q (uses'_all g) (phis'_all g)" unfolding substAll_def uses'_all_def phis'_all_def
apply (rule while_rule[where P="\<lambda>x. inst' g x \<and> P x"])
apply (rule conjI)
apply (simp, unfold_locales)
apply (rule assms(1))
apply (rule conjI)
apply (clarsimp simp: cond_def step_def)
apply (rule step_preserves_inst [unfolded step_def, simplified], assumption+)
proof-
show "wf {(y,x). (inst' g x \<and> cond g x) \<and> y = step g x}"
apply (rule wf_if_measure[where f="\<lambda>(u,p). card (dom p)"])
apply (clarsimp simp: cond_def step_def split:prod.split)
proof-
fix u p
assume "inst g u p"
then interpret i: CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses(g:=u)" "phis(g:=p)" by simp
assume "i.redundant g"
thus "card (dom (i.phis' g)) < card (dom p)" by (rule i.substAll_wf[of g, simplified])
qed
qed (auto intro: assms(2,3))
thus "inst g (uses'_all g) (phis'_all g)" "Q (uses'_all g) (phis'_all g)"
by (auto simp: uses'_all_def phis'_all_def)
qed
sublocale notriv: CFG_SSA_Transformed \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" uses'_all phis'_all
proof-
interpret ssa: CFG_SSA \<alpha>e \<alpha>n invar inEdges' Entry "defs" uses'_all phis'_all
proof
fix g
interpret i: CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses(g:=uses'_all g)" "phis(g:=phis'_all g)" var
by (rule substAll, auto)
interpret uninst: CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u p var chooseNext_all
for u and p
by unfold_locales
fix n v args m
show "finite (defs g n)" by (rule defs_finite)
show "v \<in> uses'_all g n \<Longrightarrow> n \<in> set (\<alpha>n g)" by (rule i.uses_in_\<alpha>n[of _ g, simplified])
show "finite (uses'_all g n)" by (rule i.uses_finite[of g, simplified])
show "invar g" by (rule invar)
show "finite (dom (phis'_all g))" by (rule i.phis_finite[of g, simplified])
show "phis'_all g (n, v) = Some args \<Longrightarrow> n \<in> set (\<alpha>n g)" using i.phis_in_\<alpha>n[of g] by simp
show "phis'_all g (n, v) = Some args \<Longrightarrow> length (old.predecessors g n) = length args" using i.phis_wf[of g] by simp
show "n \<in> set (\<alpha>n g) \<Longrightarrow> defs g n \<inter> uninst.phiDefs phis'_all g n = {}" using i.simpleDefs_phiDefs_disjoint[of n g] by (simp add: uninst.CFG_SSA_defs)
show "n \<in> set (\<alpha>n g) \<Longrightarrow> m \<in> set (\<alpha>n g) \<Longrightarrow> n \<noteq> m \<Longrightarrow> uninst.allDefs phis'_all g n \<inter> uninst.allDefs phis'_all g m = {}"
using i.allDefs_disjoint[of n g] by (simp add: uninst.CFG_SSA_defs)
show "n \<in> set (\<alpha>n g) \<Longrightarrow> defs g n \<inter> uses'_all g n = {}" using i.defs_uses_disjoint[of n g] by simp
qed
interpret uninst: CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u p var chooseNext_all
for u and p
by unfold_locales
show "CFG_SSA_Transformed \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses defs uses'_all phis'_all var"
proof
fix g n v ns m x v' vs
interpret i: CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses(g:=uses'_all g)" "phis(g:=phis'_all g)" var
by (rule substAll, auto)
show "oldDefs g n = var g ` defs g n" by (rule oldDefs_def)
show "n \<in> set (\<alpha>n g) \<Longrightarrow> oldUses g n = var g ` uses'_all g n" using i.oldUses_def[of n g] by simp
show "v \<in> ssa.allUses g n \<Longrightarrow> n \<in> set (\<alpha>n g) \<Longrightarrow> ssa.defAss g n v" using i.allUses_def_ass[of v g n] by (simp add: uninst.CFG_SSA_defs)
show "old.path2 g n ns m \<Longrightarrow> n \<notin> set (tl ns) \<Longrightarrow> v \<in> ssa.allDefs g n \<Longrightarrow> v \<in> ssa.allUses g m \<Longrightarrow> x \<in> set (tl ns) \<Longrightarrow> v' \<in> ssa.allDefs g x \<Longrightarrow> var g v' \<noteq> var g v" using i.conventional[of g n ns m v x v'] by (simp add: uninst.CFG_SSA_defs)
show "phis'_all g (n, v) = Some vs \<Longrightarrow> v' \<in> set vs \<Longrightarrow> var g v' = var g v" using i.phis_same_var[of g n v] by simp
show "n \<in> set (\<alpha>n g) \<Longrightarrow> v \<in> ssa.allDefs g n \<Longrightarrow> v' \<in> ssa.allDefs g n \<Longrightarrow> v \<noteq> v' \<Longrightarrow> var g v' \<noteq> var g v" using i.allDefs_var_disjoint by (simp add: uninst.CFG_SSA_defs)
show "phis'_all g (Entry g, v) = None" using i.Entry_no_phis[of g v] by simp
qed
qed
theorem not_redundant: "\<not>notriv.redundant g"
proof-
interpret uninst: CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u p var chooseNext_all
for u and p
by unfold_locales
have 1: "\<And>u p. uninst.redundant (uses(g:=u g)) (phis(g:=p g)) g \<longleftrightarrow> uninst.redundant u p g"
by (simp add: uninst.CFG_SSA_wf_defs)
show ?thesis
by (rule substAll(2)[where Q="\<lambda>u p. \<not>uninst.redundant (uses(g:=u)) (phis(g:=p)) g" and P="\<lambda>_. True" and g=g, simplified cond_def substAll_def 1], auto)
qed
corollary minimal: "old.reducible g \<Longrightarrow> notriv.cytronMinimal g"
by (erule notriv.reducible_nonredundant_imp_minimal, rule not_redundant)
theorem pruned_invariant:
assumes "pruned g"
shows "notriv.pruned g"
proof-
{
fix u p
assume "inst g u p"
then interpret i: CFG_SSA_Transformed_notriv \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses(g:=u)" "phis(g:=p)" var chooseNext_all
by simp
assume "i.redundant g"
then interpret i: CFG_SSA_step \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" "uses(g:=u)" "phis(g:=p)" var chooseNext_all g
by unfold_locales
interpret uninst: CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u p var chooseNext_all
for u and p
by unfold_locales
assume "i.pruned g"
hence "uninst.pruned (uses(g:=i.uses' g)) (phis(g:=i.phis' g)) g"
by (rule i.pruned_inv[simplified])
}
note 1 = this
interpret uninst: CFG_SSA_Transformed_notriv_base \<alpha>e \<alpha>n invar inEdges' Entry oldDefs oldUses "defs" u p var chooseNext_all
for u and p
by unfold_locales
have 2: "\<And>u u' p p' g. uninst.pruned (u'(g:=u g)) (p'(g:=p g)) g \<longleftrightarrow> uninst.pruned u p g"
by (clarsimp simp: uninst.CFG_SSA_wf_defs)
from 1 assms show ?thesis
by - (rule substAll(2)[where P="\<lambda>(u,p). uninst.pruned (uses(g:=u)) (phis(g:=p)) g" and Q="\<lambda>u p. uninst.pruned (uses(g:=u)) (phis(g:=p)) g" and g=g, simplified 2],
auto simp: cond_def step_def)
qed
end
end
diff --git a/thys/Formal_SSA/FormalSSA_Misc.thy b/thys/Formal_SSA/FormalSSA_Misc.thy
--- a/thys/Formal_SSA/FormalSSA_Misc.thy
+++ b/thys/Formal_SSA/FormalSSA_Misc.thy
@@ -1,275 +1,275 @@
(* Title: FormalSSA_Misc.thy
Author: Sebastian Ullrich, Denis Lohner
*)
section \<open>Prelude\<close>
subsection \<open>Miscellaneous Lemmata\<close>
theory FormalSSA_Misc
imports Main "HOL-Library.Sublist"
begin
lemma length_1_last_hd: "length ns = 1 \<Longrightarrow> last ns = hd ns"
by (metis One_nat_def append.simps(1) append_butlast_last_id diff_Suc_Suc diff_zero length_0_conv length_butlast list.sel(1) zero_neq_one)
lemma not_in_butlast[simp]: "\<lbrakk>x \<in> set ys; x \<notin> set (butlast ys)\<rbrakk> \<Longrightarrow> x = last ys"
apply (cases "ys = []")
apply simp
apply (subst(asm) append_butlast_last_id[symmetric])
by (simp_all del:append_butlast_last_id)
lemma in_set_butlastI: "x \<in> set xs \<Longrightarrow> x \<noteq> last xs \<Longrightarrow> x \<in> set (butlast xs)"
by (metis append_butlast_last_id append_is_Nil_conv list.distinct(1) rotate1.simps(2) set_ConsD set_rotate1 split_list)
lemma butlast_strict_prefix: "xs \<noteq> [] \<Longrightarrow> strict_prefix (butlast xs) xs"
by (metis append_butlast_last_id strict_prefixI')
lemma set_tl: "set (tl xs) \<subseteq> set xs"
by (metis set_mono_suffix suffix_tl)
lemma in_set_tlD[elim]: "x \<in> set (tl xs) \<Longrightarrow> x \<in> set xs"
using set_tl[of xs] by auto
lemma suffix_unsnoc:
assumes "suffix xs ys" "xs \<noteq> []"
obtains x where "xs = butlast xs@[x]" "ys = butlast ys@[x]"
by (metis append_butlast_last_id append_is_Nil_conv assms(1) assms(2) last_appendR suffix_def)
lemma prefix_split_first:
assumes "x \<in> set xs"
obtains as where "prefix (as@[x]) xs" and "x \<notin> set as"
proof atomize_elim
from assms obtain as bs where "xs = as@x#bs \<and> x \<notin> set as" by (atomize_elim, rule split_list_first)
thus "\<exists>as. prefix (as@[x]) xs \<and> x \<notin> set as" by -(rule exI[where x = as], auto)
qed
lemma in_prefix[elim]:
assumes "prefix xs ys" and "x \<in> set xs"
shows "x \<in> set ys"
using assms by (auto elim!:prefixE)
lemma strict_prefix_butlast:
assumes "prefix xs (butlast ys)" "ys \<noteq> []"
shows "strict_prefix xs ys"
using assms unfolding append_butlast_last_id[symmetric] by (auto simp add:less_le butlast_strict_prefix prefix_order.le_less_trans)
lemma prefix_tl_subset: "prefix xs ys \<Longrightarrow> set (tl xs) \<subseteq> set (tl ys)"
- by (metis Nil_tl prefix_bot.bot.extremum prefix_def set_mono_prefix tl_append2)
+ by (metis Nil_tl prefix_bot.extremum prefix_def set_mono_prefix tl_append2)
lemma suffix_tl_subset: "suffix xs ys \<Longrightarrow> set (tl xs) \<subseteq> set (tl ys)"
by (metis append_Nil suffix_def set_mono_suffix suffix_tl suffix_order.order_trans tl_append2)
lemma set_tl_append': "set (tl (xs @ ys)) \<subseteq> set (tl xs) \<union> set ys"
by (metis list.sel(2) order_refl set_append set_mono_suffix suffix_tl tl_append2)
lemma last_in_tl: "length xs > 1 \<Longrightarrow> last xs \<in> set (tl xs)"
by (metis One_nat_def diff_Suc_Suc last_in_set last_tl length_tl less_numeral_extra(4) list.size(3) zero_less_diff)
lemma concat_join: "xs \<noteq> [] \<Longrightarrow> ys \<noteq> [] \<Longrightarrow> last xs = hd ys \<Longrightarrow> butlast xs@ys = xs@tl ys"
by (induction xs, auto)
lemma fold_induct[case_names Nil Cons]: "P s \<Longrightarrow> (\<And>x s. x \<in> set xs \<Longrightarrow> P s \<Longrightarrow> P (f x s)) \<Longrightarrow> P (fold f xs s)"
by (rule fold_invariant [where Q = "\<lambda>x. x \<in> set xs"]) simp
lemma fold_union_elem:
assumes "x \<in> fold (\<union>) xss xs"
obtains ys where "x \<in> ys" "ys \<in> set xss \<union> {xs}"
using assms
by (induction rule:fold_induct) auto
lemma fold_union_elemI:
assumes "x \<in> ys" "ys \<in> set xss \<union> {xs}"
shows "x \<in> fold (\<union>) xss xs"
using assms
by (metis Sup_empty Sup_insert Sup_set_fold Un_insert_right UnionI ccpo_Sup_singleton fold_simps(2) list.simps(15))
lemma fold_union_elemI':
assumes "x \<in> xs \<or> (\<exists>xs \<in> set xss. x \<in> xs)"
shows "x \<in> fold (\<union>) xss xs"
using assms
using fold_union_elemI by fastforce
lemma fold_union_finite[intro!]:
assumes "finite xs" "\<forall>xs \<in> set xss. finite xs"
shows "finite (fold (\<union>) xss xs)"
using assms by - (rule fold_invariant, auto)
lemma in_set_zip_map:
assumes "(x,y) \<in> set (zip xs (map f ys))"
obtains y' where "(x,y') \<in> set (zip xs ys)" "f y' = y"
proof-
from assms obtain i where "x = xs ! i" "y = map f ys ! i" "i < length xs" "i < length ys" by (auto simp:set_zip)
thus thesis by - (rule that[of "ys ! i"], auto simp:set_zip)
qed
lemma dom_comp_subset: "g ` dom (f \<circ> g) \<subseteq> dom f"
by (auto simp add:dom_def)
lemma finite_dom_comp:
assumes "finite (dom f)" "inj_on g (dom (f \<circ> g))"
shows "finite (dom (f \<circ> g))"
proof (rule finite_imageD)
have "g ` dom (f \<circ> g) \<subseteq> dom f" by auto
with assms(1) show "finite (g ` dom (f \<circ> g))" by - (rule finite_subset)
qed (simp add:assms(2))
lemma the1_list: "\<exists>!x \<in> set xs. P x \<Longrightarrow> (THE x. x \<in> set xs \<and> P x) = hd (filter P xs)"
proof (induction xs)
case (Cons y xs)
let ?Q = "\<lambda>xs x. x \<in> set xs \<and> P x"
from Cons.prems obtain x where x: "?Q (y#xs) x" by auto
have "x = hd (filter P (y#xs))"
proof (cases "x=y")
case True
with x show ?thesis by auto
next
case False
with Cons.prems x have 1: "\<exists>!x. x \<in> set xs \<and> P x" by auto
hence "(THE x. x \<in> set xs \<and> P x) = x" using x False by - (rule the1_equality, auto)
also from 1 have "(THE x. x \<in> set xs \<and> P x) = hd (filter P xs)" by (rule Cons.IH)
finally show ?thesis using False x Cons.prems by auto
qed
thus ?case using x by - (rule the1_equality[OF Cons.prems], auto)
qed auto
lemma set_zip_leftI:
assumes "length xs = length ys"
assumes "y \<in> set ys"
obtains x where "(x,y) \<in> set (zip xs ys)"
proof-
from assms(2) obtain i where "y = ys ! i" "i < length ys" by (metis in_set_conv_nth)
with assms(1) show thesis by - (rule that[of "xs ! i"], auto simp add:set_zip)
qed
lemma butlast_idx:
assumes "y \<in> set (butlast xs)"
obtains i where "xs ! i = y" "i < length xs - 1"
apply atomize_elim
using assms proof (induction xs arbitrary:y)
case (Cons x xs)
from Cons.prems have[simp]: "xs \<noteq> []" by (simp split:if_split_asm)
show ?case
proof (cases "y = x")
case True
show ?thesis by (rule exI[where x=0], simp_all add:True)
next
case False
with Cons.prems have "y \<in> set (butlast xs)" by simp
from Cons.IH[OF this] obtain i where "y = xs ! i" and "i < length xs - 1" by auto
thus ?thesis by - (rule exI[where x="Suc i"], simp)
qed
qed simp
lemma butlast_idx':
assumes "xs ! i = y" "i < length xs - 1" "length xs > 1"
shows "y \<in> set (butlast xs)"
using assms proof (induction xs arbitrary:i)
case (Cons x xs)
show ?case
proof (cases i)
case 0
with Cons.prems(1,3) show ?thesis by simp
next
case (Suc j)
with Cons.prems(1)[symmetric] Cons.prems(2,3) have "y \<in> set (butlast xs)" by - (rule Cons.IH, auto)
with Cons.prems(3) show ?thesis by simp
qed
qed simp
lemma card_eq_1_singleton:
assumes "card A = 1"
obtains x where "A = {x}"
using assms[simplified] by - (drule card_eq_SucD, auto)
lemma set_take_two:
assumes "card A \<ge> 2"
obtains x y where "x \<in> A" "y \<in> A" "x \<noteq> y"
proof-
from assms obtain k where "card A = Suc (Suc k)"
by (auto simp: le_iff_add)
from card_eq_SucD[OF this] obtain x B where x: "A = insert x B" "x \<notin> B" "card B = Suc k" by auto
from card_eq_SucD[OF this(3)] obtain y where y: "y \<in> B" by auto
from x y show ?thesis by - (rule that[of x y], auto)
qed
lemma singleton_list_hd_last: "length xs = 1 \<Longrightarrow> hd xs = last xs"
by (metis One_nat_def impossible_Cons last.simps length_0_conv less_nat_zero_code list.sel(1) nat_less_le neq_Nil_conv not_less_eq_eq)
lemma distinct_hd_tl: "distinct xs \<Longrightarrow> hd xs \<notin> set (tl xs)"
by (metis distinct.simps(2) hd_Cons_tl in_set_member list.sel(2) member_rec(2))
lemma set_mono_strict_prefix: "strict_prefix xs ys \<Longrightarrow> set xs \<subseteq> set (butlast ys)"
by (metis append_butlast_last_id strict_prefixE strict_prefix_simps(1) prefix_snoc set_mono_prefix)
lemma set_butlast_distinct: "distinct xs \<Longrightarrow> set (butlast xs) \<inter> {last xs} = {}"
by (metis append_butlast_last_id butlast.simps(1) distinct_append inf_bot_right inf_commute list.set(1) set_simps(2))
lemma disjoint_elem[elim]: "A \<inter> B = {} \<Longrightarrow> x \<in> A \<Longrightarrow> x \<notin> B" by auto
lemma prefix_butlastD[elim]: "prefix xs (butlast ys) \<Longrightarrow> prefix xs ys"
using strict_prefix_butlast by fastforce
lemma butlast_prefix: "prefix xs ys \<Longrightarrow> prefix (butlast xs) (butlast ys)"
by (induction xs ys rule: list_induct2'; auto)
lemma hd_in_butlast: "length xs > 1 \<Longrightarrow> hd xs \<in> set (butlast xs)"
by (metis butlast.simps(2) dual_order.strict_iff_order hd_Cons_tl hd_in_set length_greater_0_conv length_tl less_le_trans list.distinct(1) list.sel(1) neq0_conv zero_less_diff)
lemma nonsimple_length_gt_1: "xs \<noteq> [] \<Longrightarrow> hd xs \<noteq> last xs \<Longrightarrow> length xs > 1"
by (metis length_0_conv less_one nat_neq_iff singleton_list_hd_last)
lemma set_hd_tl: "xs \<noteq> [] \<Longrightarrow> set [hd xs] \<union> set (tl xs) = set xs"
by (metis inf_sup_aci(5) rotate1_hd_tl set_append set_rotate1)
lemma fold_update_conv:
"fold (\<lambda>n m. m(n \<mapsto> g n)) xs m x =
(if (x \<in> set xs) then Some (g x) else m x)"
by (induction xs arbitrary: m) auto
lemmas removeAll_le = length_removeAll_less_eq
lemmas removeAll_less [intro] = length_removeAll_less
lemma removeAll_induct:
assumes "\<And>xs. (\<And>x. x \<in> set xs \<Longrightarrow> P (removeAll x xs)) \<Longrightarrow> P xs"
shows "P xs"
by (induct xs rule:length_induct, rule assms) auto
lemma The_Min: "Ex1 P \<Longrightarrow> The P = Min {x. P x}"
apply (rule the_equality)
apply (metis (mono_tags) Min.infinite Min_in Min_singleton all_not_in_conv finite_subset insert_iff mem_Collect_eq subsetI)
by (metis (erased, hide_lams) Least_Min Least_equality Set.set_insert ex_in_conv finite.emptyI finite_insert insert_iff mem_Collect_eq order_refl)
lemma The_Max: "Ex1 P \<Longrightarrow> The P = Max {x. P x}"
apply (rule the_equality)
apply (metis (mono_tags) Max.infinite Max_in Max_singleton all_not_in_conv finite_subset insert_iff mem_Collect_eq subsetI)
by (metis Max_singleton Min_singleton Nitpick.Ex1_unfold The_Min the_equality)
lemma set_sorted_list_of_set_remove [simp]:
"set (sorted_list_of_set (Set.remove x A)) = Set.remove x (set (sorted_list_of_set A))"
unfolding Set.remove_def
by (cases "finite A"; simp)
lemma set_minus_one: "\<lbrakk>v \<noteq> v'; v' \<in> set vs\<rbrakk> \<Longrightarrow> set vs - {v'} \<subseteq> {v} \<longleftrightarrow> set vs = {v'} \<or> set vs = {v,v'}"
by auto
lemma set_single_hd: "set vs = {v} \<Longrightarrow> hd vs = v"
by (induction vs; auto)
lemma set_double_filter_hd: "\<lbrakk> set vs = {v,v'}; v \<noteq> v' \<rbrakk> \<Longrightarrow> hd [v'\<leftarrow>vs . v' \<noteq> v] = v'"
apply (induction vs)
apply simp
apply auto
apply (case_tac "v \<in> set vs")
prefer 2
apply (subgoal_tac "set vs = {v'}")
prefer 2
apply fastforce
apply (clarsimp simp: set_single_hd)
by fastforce
lemma map_option_the: "x = map_option f y \<Longrightarrow> x \<noteq> None \<Longrightarrow> the x = f (the y)"
by (auto simp: map_option_case split: option.split prod.splits)
end
diff --git a/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy b/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy
--- a/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy
+++ b/thys/Functional_Ordered_Resolution_Prover/Deterministic_FO_Ordered_Resolution_Prover.thy
@@ -1,1753 +1,1753 @@
(* Title: A Deterministic Ordered Resolution Prover for First-Order Clauses
Author: Jasmin Blanchette <j.c.blanchette at vu.nl>, 2017
Author: Anders Schlichtkrull <andschl at dtu.dk>, 2018
Maintainer: Anders Schlichtkrull <andschl at dtu.dk>
*)
section \<open>A Deterministic Ordered Resolution Prover for First-Order Clauses\<close>
text \<open>
The \<open>deterministic_RP\<close> prover introduced below is a deterministic program that works on finite
lists, committing to a strategy for assigning priorities to clauses. However, it is not fully
executable: It abstracts over operations on atoms and employs logical specifications instead of
executable functions for auxiliary notions.
\<close>
theory Deterministic_FO_Ordered_Resolution_Prover
imports
Polynomial_Factorization.Missing_List
Weighted_FO_Ordered_Resolution_Prover
Lambda_Free_RPOs.Lambda_Free_Util
begin
subsection \<open>Library\<close>
lemma apfst_fst_snd: "apfst f x = (f (fst x), snd x)"
by (rule apfst_conv[of _ "fst x" "snd x" for x, unfolded prod.collapse])
lemma apfst_comp_rpair_const: "apfst f \<circ> (\<lambda>x. (x, y)) = (\<lambda>x. (x, y)) \<circ> f"
by (simp add: comp_def)
(* TODO: Move to Isabelle's "List.thy"? *)
lemma length_remove1_less[termination_simp]: "x \<in> set xs \<Longrightarrow> length (remove1 x xs) < length xs"
by (induct xs) auto
lemma map_filter_neq_eq_filter_map:
"map f (filter (\<lambda>y. f x \<noteq> f y) xs) = filter (\<lambda>z. f x \<noteq> z) (map f xs)"
by (induct xs) auto
lemma mset_map_remdups_gen:
"mset (map f (remdups_gen f xs)) = mset (remdups_gen (\<lambda>x. x) (map f xs))"
by (induct f xs rule: remdups_gen.induct) (auto simp: map_filter_neq_eq_filter_map)
lemma mset_remdups_gen_ident: "mset (remdups_gen (\<lambda>x. x) xs) = mset_set (set xs)"
proof -
have "f = (\<lambda>x. x) \<Longrightarrow> mset (remdups_gen f xs) = mset_set (set xs)" for f
proof (induct f xs rule: remdups_gen.induct)
case (2 f x xs)
note ih = this(1) and f = this(2)
show ?case
unfolding f remdups_gen.simps ih[OF f, unfolded f] mset.simps
by (metis finite_set list.simps(15) mset_set.insert_remove removeAll_filter_not_eq
remove_code(1) remove_def)
qed simp
then show ?thesis
by simp
qed
(* TODO: Move to Isabelle? *)
lemma funpow_fixpoint: "f x = x \<Longrightarrow> (f ^^ n) x = x"
by (induct n) auto
lemma rtranclp_imp_eq_image: "(\<forall>x y. R x y \<longrightarrow> f x = f y) \<Longrightarrow> R\<^sup>*\<^sup>* x y \<Longrightarrow> f x = f y"
by (erule rtranclp.induct) auto
lemma tranclp_imp_eq_image: "(\<forall>x y. R x y \<longrightarrow> f x = f y) \<Longrightarrow> R\<^sup>+\<^sup>+ x y \<Longrightarrow> f x = f y"
by (erule tranclp.induct) auto
subsection \<open>Prover\<close>
type_synonym 'a lclause = "'a literal list"
type_synonym 'a dclause = "'a lclause \<times> nat"
type_synonym 'a dstate = "'a dclause list \<times> 'a dclause list \<times> 'a dclause list \<times> nat"
locale deterministic_FO_resolution_prover =
weighted_FO_resolution_prover_with_size_timestamp_factors S subst_atm id_subst comp_subst
renamings_apart atm_of_atms mgu less_atm size_atm timestamp_factor size_factor
for
S :: "('a :: wellorder) clause \<Rightarrow> 'a clause" and
subst_atm :: "'a \<Rightarrow> 's \<Rightarrow> 'a" and
id_subst :: "'s" and
comp_subst :: "'s \<Rightarrow> 's \<Rightarrow> 's" and
renamings_apart :: "'a literal multiset list \<Rightarrow> 's list" and
atm_of_atms :: "'a list \<Rightarrow> 'a" and
mgu :: "'a set set \<Rightarrow> 's option" and
less_atm :: "'a \<Rightarrow> 'a \<Rightarrow> bool" and
size_atm :: "'a \<Rightarrow> nat" and
timestamp_factor :: nat and
size_factor :: nat +
assumes
S_empty: "S C = {#}"
begin
lemma less_atm_irrefl: "\<not> less_atm A A"
using ex_ground_subst less_atm_ground less_atm_stable unfolding is_ground_subst_def by blast
fun wstate_of_dstate :: "'a dstate \<Rightarrow> 'a wstate" where
"wstate_of_dstate (N, P, Q, n) =
(mset (map (apfst mset) N), mset (map (apfst mset) P), mset (map (apfst mset) Q), n)"
fun state_of_dstate :: "'a dstate \<Rightarrow> 'a state" where
"state_of_dstate (N, P, Q, _) =
(set (map (mset \<circ> fst) N), set (map (mset \<circ> fst) P), set (map (mset \<circ> fst) Q))"
abbreviation clss_of_dstate :: "'a dstate \<Rightarrow> 'a clause set" where
"clss_of_dstate St \<equiv> clss_of_state (state_of_dstate St)"
fun is_final_dstate :: "'a dstate \<Rightarrow> bool" where
"is_final_dstate (N, P, Q, n) \<longleftrightarrow> N = [] \<and> P = []"
declare is_final_dstate.simps [simp del]
abbreviation rtrancl_weighted_RP (infix "\<leadsto>\<^sub>w\<^sup>*" 50) where
"(\<leadsto>\<^sub>w\<^sup>*) \<equiv> (\<leadsto>\<^sub>w)\<^sup>*\<^sup>*"
abbreviation trancl_weighted_RP (infix "\<leadsto>\<^sub>w\<^sup>+" 50) where
"(\<leadsto>\<^sub>w\<^sup>+) \<equiv> (\<leadsto>\<^sub>w)\<^sup>+\<^sup>+"
definition is_tautology :: "'a lclause \<Rightarrow> bool" where
"is_tautology C \<longleftrightarrow> (\<exists>A \<in> set (map atm_of C). Pos A \<in> set C \<and> Neg A \<in> set C)"
definition subsume :: "'a lclause list \<Rightarrow> 'a lclause \<Rightarrow> bool" where
"subsume Ds C \<longleftrightarrow> (\<exists>D \<in> set Ds. subsumes (mset D) (mset C))"
definition strictly_subsume :: "'a lclause list \<Rightarrow> 'a lclause \<Rightarrow> bool" where
"strictly_subsume Ds C \<longleftrightarrow> (\<exists>D \<in> set Ds. strictly_subsumes (mset D) (mset C))"
definition is_reducible_on :: "'a literal \<Rightarrow> 'a lclause \<Rightarrow> 'a literal \<Rightarrow> 'a lclause \<Rightarrow> bool" where
"is_reducible_on M D L C \<longleftrightarrow> subsumes (mset D + {#- M#}) (mset C + {#L#})"
definition is_reducible_lit :: "'a lclause list \<Rightarrow> 'a lclause \<Rightarrow> 'a literal \<Rightarrow> bool" where
"is_reducible_lit Ds C L \<longleftrightarrow>
(\<exists>D \<in> set Ds. \<exists>L' \<in> set D. \<exists>\<sigma>. - L = L' \<cdot>l \<sigma> \<and> mset (remove1 L' D) \<cdot> \<sigma> \<subseteq># mset C)"
primrec reduce :: "'a lclause list \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause" where
"reduce _ _ [] = []"
| "reduce Ds C (L # C') =
(if is_reducible_lit Ds (C @ C') L then reduce Ds C C' else L # reduce Ds (L # C) C')"
abbreviation is_irreducible :: "'a lclause list \<Rightarrow> 'a lclause \<Rightarrow> bool" where
"is_irreducible Ds C \<equiv> reduce Ds [] C = C"
abbreviation is_reducible :: "'a lclause list \<Rightarrow> 'a lclause \<Rightarrow> bool" where
"is_reducible Ds C \<equiv> reduce Ds [] C \<noteq> C"
definition reduce_all :: "'a lclause \<Rightarrow> 'a dclause list \<Rightarrow> 'a dclause list" where
"reduce_all D = map (apfst (reduce [D] []))"
fun reduce_all2 :: "'a lclause \<Rightarrow> 'a dclause list \<Rightarrow> 'a dclause list \<times> 'a dclause list" where
"reduce_all2 _ [] = ([], [])"
| "reduce_all2 D (Ci # Cs) =
(let
(C, i) = Ci;
C' = reduce [D] [] C
in
(if C' = C then apsnd else apfst) (Cons (C', i)) (reduce_all2 D Cs))"
fun remove_all :: "'b list \<Rightarrow> 'b list \<Rightarrow> 'b list" where
"remove_all xs [] = xs"
| "remove_all xs (y # ys) = (if y \<in> set xs then remove_all (remove1 y xs) ys else remove_all xs ys)"
lemma remove_all_mset_minus: "mset ys \<subseteq># mset xs \<Longrightarrow> mset (remove_all xs ys) = mset xs - mset ys"
proof (induction ys arbitrary: xs)
case (Cons y ys)
show ?case
proof (cases "y \<in> set xs")
case y_in: True
then have subs: "mset ys \<subseteq># mset (remove1 y xs)"
using Cons(2) by (simp add: insert_subset_eq_iff)
show ?thesis
using y_in Cons subs by auto
next
case False
then show ?thesis
using Cons by auto
qed
qed auto
definition resolvent :: "'a lclause \<Rightarrow> 'a \<Rightarrow>'a lclause \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause" where
"resolvent D A CA Ls =
map (\<lambda>M. M \<cdot>l (the (mgu {insert A (atms_of (mset Ls))}))) (remove_all CA Ls @ D)"
definition resolvable :: "'a \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause \<Rightarrow> bool" where
"resolvable A D CA Ls \<longleftrightarrow>
(let \<sigma> = (mgu {insert A (atms_of (mset Ls))}) in
\<sigma> \<noteq> None
\<and> Ls \<noteq> []
\<and> maximal_wrt (A \<cdot>a the \<sigma>) ((add_mset (Neg A) (mset D)) \<cdot> the \<sigma>)
\<and> strictly_maximal_wrt (A \<cdot>a the \<sigma>) ((mset CA - mset Ls) \<cdot> the \<sigma>)
\<and> (\<forall>L \<in> set Ls. is_pos L))"
definition resolve_on :: "'a \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause list" where
"resolve_on A D CA = map (resolvent D A CA) (filter (resolvable A D CA) (subseqs CA))"
definition resolve :: "'a lclause \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause list" where
"resolve C D =
concat (map (\<lambda>L.
(case L of
Pos A \<Rightarrow> []
| Neg A \<Rightarrow>
if maximal_wrt A (mset D) then
resolve_on A (remove1 L D) C
else
[])) D)"
definition resolve_rename :: "'a lclause \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause list" where
"resolve_rename C D =
(let \<sigma>s = renamings_apart [mset D, mset C] in
resolve (map (\<lambda>L. L \<cdot>l last \<sigma>s) C) (map (\<lambda>L. L \<cdot>l hd \<sigma>s) D))"
definition resolve_rename_either_way :: "'a lclause \<Rightarrow> 'a lclause \<Rightarrow> 'a lclause list" where
"resolve_rename_either_way C D = resolve_rename C D @ resolve_rename D C"
fun select_min_weight_clause :: "'a dclause \<Rightarrow> 'a dclause list \<Rightarrow> 'a dclause" where
"select_min_weight_clause Ci [] = Ci"
| "select_min_weight_clause Ci (Dj # Djs) =
select_min_weight_clause
(if weight (apfst mset Dj) < weight (apfst mset Ci) then Dj else Ci) Djs"
lemma select_min_weight_clause_in: "select_min_weight_clause P0 P \<in> set (P0 # P)"
by (induct P arbitrary: P0) auto
function remdups_clss :: "'a dclause list \<Rightarrow> 'a dclause list" where
"remdups_clss [] = []"
| "remdups_clss (Ci # Cis) =
(let
Ci' = select_min_weight_clause Ci Cis
in
Ci' # remdups_clss (filter (\<lambda>(D, _). mset D \<noteq> mset (fst Ci')) (Ci # Cis)))"
by pat_completeness auto
termination
apply (relation "measure length")
apply (rule wf_measure)
by (metis (mono_tags) in_measure length_filter_less prod.case_eq_if select_min_weight_clause_in)
declare remdups_clss.simps(2) [simp del]
fun deterministic_RP_step :: "'a dstate \<Rightarrow> 'a dstate" where
"deterministic_RP_step (N, P, Q, n) =
(if \<exists>Ci \<in> set (P @ Q). fst Ci = [] then
([], [], remdups_clss P @ Q, n + length (remdups_clss P))
else
(case N of
[] \<Rightarrow>
(case P of
[] \<Rightarrow> (N, P, Q, n)
| P0 # P' \<Rightarrow>
let
(C, i) = select_min_weight_clause P0 P';
N = map (\<lambda>D. (D, n)) (remdups_gen mset (resolve_rename C C
@ concat (map (resolve_rename_either_way C \<circ> fst) Q)));
P = filter (\<lambda>(D, j). mset D \<noteq> mset C) P;
Q = (C, i) # Q;
n = Suc n
in
(N, P, Q, n))
| (C, i) # N \<Rightarrow>
let
C = reduce (map fst (P @ Q)) [] C
in
if C = [] then
([], [], [([], i)], Suc n)
else if is_tautology C \<or> subsume (map fst (P @ Q)) C then
(N, P, Q, n)
else
let
P = reduce_all C P;
(back_to_P, Q) = reduce_all2 C Q;
P = back_to_P @ P;
Q = filter (Not \<circ> strictly_subsume [C] \<circ> fst) Q;
P = filter (Not \<circ> strictly_subsume [C] \<circ> fst) P;
P = (C, i) # P
in
(N, P, Q, n)))"
declare deterministic_RP_step.simps [simp del]
partial_function (option) deterministic_RP :: "'a dstate \<Rightarrow> 'a lclause list option" where
"deterministic_RP St =
(if is_final_dstate St then
let (_, _, Q, _) = St in Some (map fst Q)
else
deterministic_RP (deterministic_RP_step St))"
lemma is_final_dstate_imp_not_weighted_RP: "is_final_dstate St \<Longrightarrow> \<not> wstate_of_dstate St \<leadsto>\<^sub>w St'"
using wrp.final_weighted_RP
by (cases St) (auto intro: wrp.final_weighted_RP simp: is_final_dstate.simps)
lemma is_final_dstate_funpow_imp_deterministic_RP_neq_None:
"is_final_dstate ((deterministic_RP_step ^^ k) St) \<Longrightarrow> deterministic_RP St \<noteq> None"
proof (induct k arbitrary: St)
case (Suc k)
note ih = this(1) and final_Sk = this(2)[simplified, unfolded funpow_swap1]
show ?case
using ih[OF final_Sk] by (subst deterministic_RP.simps) (simp add: prod.case_eq_if)
qed (subst deterministic_RP.simps, simp add: prod.case_eq_if)
lemma is_reducible_lit_mono_cls:
"mset C \<subseteq># mset C' \<Longrightarrow> is_reducible_lit Ds C L \<Longrightarrow> is_reducible_lit Ds C' L"
- unfolding is_reducible_lit_def by (blast intro: subset_mset.order.trans)
+ unfolding is_reducible_lit_def by (blast intro: subset_mset.trans)
lemma is_reducible_lit_mset_iff:
"mset C = mset C' \<Longrightarrow> is_reducible_lit Ds C' L \<longleftrightarrow> is_reducible_lit Ds C L"
by (metis is_reducible_lit_mono_cls subset_mset.order_refl)
lemma is_reducible_lit_remove1_Cons_iff:
assumes "L \<in> set C'"
shows "is_reducible_lit Ds (C @ remove1 L (M # C')) L \<longleftrightarrow>
is_reducible_lit Ds (M # C @ remove1 L C') L"
using assms by (subst is_reducible_lit_mset_iff, auto)
lemma reduce_mset_eq: "mset C = mset C' \<Longrightarrow> reduce Ds C E = reduce Ds C' E"
proof (induct E arbitrary: C C')
case (Cons L E)
note ih = this(1) and mset_eq = this(2)
have
mset_lc_eq: "mset (L # C) = mset (L # C')" and
mset_ce_eq: "mset (C @ E) = mset (C' @ E)"
using mset_eq by simp+
show ?case
using ih[OF mset_eq] ih[OF mset_lc_eq] by (simp add: is_reducible_lit_mset_iff[OF mset_ce_eq])
qed simp
lemma reduce_rotate[simp]: "reduce Ds (C @ [L]) E = reduce Ds (L # C) E"
by (rule reduce_mset_eq) simp
lemma mset_reduce_subset: "mset (reduce Ds C E) \<subseteq># mset E"
by (induct E arbitrary: C) (auto intro: subset_mset_imp_subset_add_mset)
lemma reduce_idem: "reduce Ds C (reduce Ds C E) = reduce Ds C E"
by (induct E arbitrary: C)
(auto intro!: mset_reduce_subset
dest!: is_reducible_lit_mono_cls[of "C @ reduce Ds (L # C) E" "C @ E" Ds L for L E C,
rotated])
lemma is_reducible_lit_imp_is_reducible:
"L \<in> set C' \<Longrightarrow> is_reducible_lit Ds (C @ remove1 L C') L \<Longrightarrow> reduce Ds C C' \<noteq> C'"
proof (induct C' arbitrary: C)
case (Cons M C')
note ih = this(1) and l_in = this(2) and l_red = this(3)
show ?case
proof (cases "is_reducible_lit Ds (C @ C') M")
case True
then show ?thesis
by simp (metis mset.simps(2) mset_reduce_subset multi_self_add_other_not_self
subset_mset.eq_iff subset_mset_imp_subset_add_mset)
next
case m_irred: False
have
"L \<in> set C'" and
"is_reducible_lit Ds (M # C @ remove1 L C') L"
using l_in l_red m_irred is_reducible_lit_remove1_Cons_iff by auto
then show ?thesis
by (simp add: ih[of "M # C"] m_irred)
qed
qed simp
lemma is_reducible_imp_is_reducible_lit:
"reduce Ds C C' \<noteq> C' \<Longrightarrow> \<exists>L \<in> set C'. is_reducible_lit Ds (C @ remove1 L C') L"
proof (induct C' arbitrary: C)
case (Cons M C')
note ih = this(1) and mc'_red = this(2)
show ?case
proof (cases "is_reducible_lit Ds (C @ C') M")
case m_irred: False
show ?thesis
using ih[of "M # C"] mc'_red[simplified, simplified m_irred, simplified] m_irred
is_reducible_lit_remove1_Cons_iff
by auto
qed simp
qed simp
lemma is_irreducible_iff_nexists_is_reducible_lit:
"reduce Ds C C' = C' \<longleftrightarrow> \<not> (\<exists>L \<in> set C'. is_reducible_lit Ds (C @ remove1 L C') L)"
using is_reducible_imp_is_reducible_lit is_reducible_lit_imp_is_reducible by blast
lemma is_irreducible_mset_iff: "mset E = mset E' \<Longrightarrow> reduce Ds C E = E \<longleftrightarrow> reduce Ds C E' = E'"
unfolding is_irreducible_iff_nexists_is_reducible_lit
by (metis (full_types) is_reducible_lit_mset_iff mset_remove1 set_mset_mset union_code)
lemma select_min_weight_clause_min_weight:
assumes "Ci = select_min_weight_clause P0 P"
shows "weight (apfst mset Ci) = Min ((weight \<circ> apfst mset) ` set (P0 # P))"
using assms
proof (induct P arbitrary: P0 Ci)
case (Cons P1 P)
note ih = this(1) and ci = this(2)
show ?case
proof (cases "weight (apfst mset P1) < weight (apfst mset P0)")
case True
then have min: "Min ((weight \<circ> apfst mset) ` set (P0 # P1 # P)) =
Min ((weight \<circ> apfst mset) ` set (P1 # P))"
by (simp add: min_def)
show ?thesis
unfolding min by (rule ih[of Ci P1]) (simp add: ih[of Ci P1] ci True)
next
case False
have "Min ((weight \<circ> apfst mset) ` set (P0 # P1 # P)) =
Min ((weight \<circ> apfst mset) ` set (P1 # P0 # P))"
by (rule arg_cong[of _ _ Min]) auto
then have min: "Min ((weight \<circ> apfst mset) ` set (P0 # P1 # P)) =
Min ((weight \<circ> apfst mset) ` set (P0 # P))"
by (simp add: min_def) (use False eq_iff in fastforce)
show ?thesis
unfolding min by (rule ih[of Ci P0]) (simp add: ih[of Ci P1] ci False)
qed
qed simp
lemma remdups_clss_Nil_iff: "remdups_clss Cs = [] \<longleftrightarrow> Cs = []"
by (cases Cs, simp, hypsubst, subst remdups_clss.simps(2), simp add: Let_def)
lemma empty_N_if_Nil_in_P_or_Q:
assumes nil_in: "[] \<in> fst ` set (P @ Q)"
shows "wstate_of_dstate (N, P, Q, n) \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ([], P, Q, n)"
proof (induct N)
case ih: (Cons N0 N)
have "wstate_of_dstate (N0 # N, P, Q, n) \<leadsto>\<^sub>w wstate_of_dstate (N, P, Q, n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.forward_subsumption[of "{#}" "mset (map (apfst mset) P)" "mset (map (apfst mset) Q)"
"mset (fst N0)" "mset (map (apfst mset) N)" "snd N0" n]])
(use nil_in in \<open>force simp: image_def apfst_fst_snd\<close>)+
then show ?case
using ih by (rule converse_rtranclp_into_rtranclp)
qed simp
lemma remove_strictly_subsumed_clauses_in_P:
assumes
c_in: "C \<in> fst ` set N" and
p_nsubs: "\<forall>D \<in> fst ` set P. \<not> strictly_subsume [C] D"
shows "wstate_of_dstate (N, P @ P', Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, P @ filter (Not \<circ> strictly_subsume [C] \<circ> fst) P', Q, n)"
using p_nsubs
proof (induct "length P'" arbitrary: P P' rule: less_induct)
case less
note ih = this(1) and p_nsubs = this(2)
show ?case
proof (cases "length P'")
case Suc
let ?Dj = "hd P'"
let ?P'' = "tl P'"
have p': "P' = hd P' # tl P'"
using Suc by (metis length_Suc_conv list.distinct(1) list.exhaust_sel)
show ?thesis
proof (cases "strictly_subsume [C] (fst ?Dj)")
case subs: True
have p_filtered: "{#(E, k) \<in># image_mset (apfst mset) (mset P). E \<noteq> mset (fst ?Dj)#} =
image_mset (apfst mset) (mset P)"
by (rule filter_mset_cong[OF refl, of _ _ "\<lambda>_. True", simplified],
use subs p_nsubs in \<open>auto simp: strictly_subsume_def\<close>)
have "{#(E, k) \<in># image_mset (apfst mset) (mset P'). E \<noteq> mset (fst ?Dj)#} =
{#(E, k) \<in># image_mset (apfst mset) (mset ?P''). E \<noteq> mset (fst ?Dj)#}"
by (subst (2) p') (simp add: case_prod_beta)
also have "\<dots> =
image_mset (apfst mset) (mset (filter (\<lambda>(E, l). mset E \<noteq> mset (fst ?Dj)) ?P''))"
by (auto simp: image_mset_filter_swap[symmetric] case_prod_beta)
finally have p'_filtered:
"{#(E, k) \<in># image_mset (apfst mset) (mset P'). E \<noteq> mset (fst ?Dj)#} =
image_mset (apfst mset) (mset (filter (\<lambda>(E, l). mset E \<noteq> mset (fst ?Dj)) ?P''))"
.
have "wstate_of_dstate (N, P @ P', Q, n)
\<leadsto>\<^sub>w wstate_of_dstate (N, P @ filter (\<lambda>(E, l). mset E \<noteq> mset (fst ?Dj)) ?P'', Q, n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.backward_subsumption_P[of "mset C" "mset (map (apfst mset) N)" "mset (fst ?Dj)"
"mset (map (apfst mset) (P @ P'))" "mset (map (apfst mset) Q)" n]],
use c_in subs in \<open>auto simp add: p_filtered p'_filtered arg_cong[OF p', of set]
strictly_subsume_def\<close>)
also have "\<dots>
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, P @ filter (Not \<circ> strictly_subsume [C] \<circ> fst) P', Q, n)"
apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w\<^sup>*)", OF _ _
ih[of "filter (\<lambda>(E, l). mset E \<noteq> mset (fst ?Dj)) ?P''" P]])
apply simp_all
apply (subst (3) p')
using subs
apply (simp add: case_prod_beta)
apply (rule arg_cong[of _ _ "\<lambda>x. image_mset (apfst mset) x"])
apply (metis (no_types, hide_lams) strictly_subsume_def)
apply (subst (3) p')
apply (subst list.size)
apply (metis (no_types, lifting) less_Suc0 less_add_same_cancel1 linorder_neqE_nat
not_add_less1 sum_length_filter_compl trans_less_add1)
using p_nsubs by fast
ultimately show ?thesis
by (rule converse_rtranclp_into_rtranclp)
next
case nsubs: False
show ?thesis
apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w\<^sup>*)", OF _ _
ih[of ?P'' "P @ [?Dj]"]])
using nsubs p_nsubs
apply (simp_all add: arg_cong[OF p', of mset] arg_cong[OF p', of "filter f" for f])
apply (subst (1 2) p')
by simp
qed
qed simp
qed
lemma remove_strictly_subsumed_clauses_in_Q:
assumes c_in: "C \<in> fst ` set N"
shows "wstate_of_dstate (N, P, Q @ Q', n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, P, Q @ filter (Not \<circ> strictly_subsume [C] \<circ> fst) Q', n)"
proof (induct Q' arbitrary: Q)
case ih: (Cons Dj Q')
have "wstate_of_dstate (N, P, Q @ Dj # Q', n) \<leadsto>\<^sub>w\<^sup>*
wstate_of_dstate (N, P, Q @ filter (Not \<circ> strictly_subsume [C] \<circ> fst) [Dj] @ Q', n)"
proof (cases "strictly_subsume [C] (fst Dj)")
case subs: True
have "wstate_of_dstate (N, P, Q @ Dj # Q', n) \<leadsto>\<^sub>w wstate_of_dstate (N, P, Q @ Q', n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.backward_subsumption_Q[of "mset C" "mset (map (apfst mset) N)" "mset (fst Dj)"
"mset (map (apfst mset) P)" "mset (map (apfst mset) (Q @ Q'))" "snd Dj" n]])
(use c_in subs in \<open>auto simp: apfst_fst_snd strictly_subsume_def\<close>)
then show ?thesis
by auto
qed simp
then show ?case
using ih[of "Q @ filter (Not \<circ> strictly_subsume [C] \<circ> fst) [Dj]"] by force
qed simp
lemma reduce_clause_in_P:
assumes
c_in: "C \<in> fst ` set N" and
p_irred: "\<forall>(E, k) \<in> set (P @ P'). k > j \<longrightarrow> is_irreducible [C] E"
shows "wstate_of_dstate (N, P @ (D @ D', j) # P', Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, P @ (D @ reduce [C] D D', j) # P', Q, n)"
proof (induct D' arbitrary: D)
case ih: (Cons L D')
show ?case
proof (cases "is_reducible_lit [C] (D @ D') L")
case l_red: True
then obtain L' :: "'a literal" and \<sigma> :: 's where
l'_in: "L' \<in> set C" and
not_l: "- L = L' \<cdot>l \<sigma>" and
subs: "mset (remove1 L' C) \<cdot> \<sigma> \<subseteq># mset (D @ D')"
unfolding is_reducible_lit_def by force
have ldd'_red: "is_reducible [C] (L # D @ D')"
apply (rule is_reducible_lit_imp_is_reducible)
using l_red by auto
have lt_imp_neq: "\<forall>(E, k) \<in> set (P @ P'). j < k \<longrightarrow> mset E \<noteq> mset (L # D @ D')"
using p_irred ldd'_red is_irreducible_mset_iff by fast
have "wstate_of_dstate (N, P @ (D @ L # D', j) # P', Q, n)
\<leadsto>\<^sub>w wstate_of_dstate (N, P @ (D @ D', j) # P', Q, n)"
apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.backward_reduction_P[of "mset C - {#L'#}" L' "mset (map (apfst mset) N)" L \<sigma>
"mset (D @ D')" "mset (map (apfst mset) (P @ P'))" j "mset (map (apfst mset) Q)" n]])
using l'_in not_l subs c_in lt_imp_neq by (simp_all add: case_prod_beta) force+
then show ?thesis
using ih[of D] l_red by simp
next
case False
then show ?thesis
using ih[of "D @ [L]"] by simp
qed
qed simp
lemma reduce_clause_in_Q:
assumes
c_in: "C \<in> fst ` set N" and
p_irred: "\<forall>(E, k) \<in> set P. k > j \<longrightarrow> is_irreducible [C] E" and
d'_red: "reduce [C] D D' \<noteq> D'"
shows "wstate_of_dstate (N, P, Q @ (D @ D', j) # Q', n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, (D @ reduce [C] D D', j) # P, Q @ Q', n)"
using d'_red
proof (induct D' arbitrary: D)
case (Cons L D')
note ih = this(1) and ld'_red = this(2)
then show ?case
proof (cases "is_reducible_lit [C] (D @ D') L")
case l_red: True
then obtain L' :: "'a literal" and \<sigma> :: 's where
l'_in: "L' \<in> set C" and
not_l: "- L = L' \<cdot>l \<sigma>" and
subs: "mset (remove1 L' C) \<cdot> \<sigma> \<subseteq># mset (D @ D')"
unfolding is_reducible_lit_def by force
have "wstate_of_dstate (N, P, Q @ (D @ L # D', j) # Q', n)
\<leadsto>\<^sub>w wstate_of_dstate (N, (D @ D', j) # P, Q @ Q', n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.backward_reduction_Q[of "mset C - {#L'#}" L' "mset (map (apfst mset) N)" L \<sigma>
"mset (D @ D')" "mset (map (apfst mset) P)" "mset (map (apfst mset) (Q @ Q'))" j n]],
use l'_in not_l subs c_in in auto)
then show ?thesis
using l_red p_irred reduce_clause_in_P[OF c_in, of "[]" P j D D' "Q @ Q'" n] by simp
next
case l_nred: False
then have d'_red: "reduce [C] (D @ [L]) D' \<noteq> D'"
using ld'_red by simp
show ?thesis
using ih[OF d'_red] l_nred by simp
qed
qed simp
lemma reduce_clauses_in_P:
assumes
c_in: "C \<in> fst ` set N" and
p_irred: "\<forall>(E, k) \<in> set P. is_irreducible [C] E"
shows "wstate_of_dstate (N, P @ P', Q, n) \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, P @ reduce_all C P', Q, n)"
unfolding reduce_all_def
using p_irred
proof (induct "length P'" arbitrary: P P')
case (Suc l)
note ih = this(1) and suc_l = this(2) and p_irred = this(3)
have p'_nnil: "P' \<noteq> []"
using suc_l by auto
define j :: nat where
"j = Max (snd ` set P')"
obtain Dj :: "'a dclause" where
dj_in: "Dj \<in> set P'" and
snd_dj: "snd Dj = j"
using Max_in[of "snd ` set P'", unfolded image_def, simplified]
by (metis image_def j_def length_Suc_conv list.set_intros(1) suc_l)
have "\<forall>k \<in> snd ` set P'. k \<le> j"
unfolding j_def using p'_nnil by simp
then have j_max: "\<forall>(E, k) \<in> set P'. j \<ge> k"
unfolding image_def by fastforce
obtain P1' P2' :: "'a dclause list" where
p': "P' = P1' @ Dj # P2'"
using split_list[OF dj_in] by blast
have "wstate_of_dstate (N, P @ P1' @ Dj # P2', Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, P @ P1' @ apfst (reduce [C] []) Dj # P2', Q, n)"
unfolding append_assoc[symmetric]
apply (subst (1 2) surjective_pairing[of Dj, unfolded snd_dj])
apply (simp only: apfst_conv)
apply (rule reduce_clause_in_P[of _ _ _ _ _ "[]", unfolded append_Nil, OF c_in])
using p_irred j_max[unfolded p'] by (force simp: case_prod_beta)
moreover have "wstate_of_dstate (N, P @ P1' @ apfst (reduce [C] []) Dj # P2', Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, P @ map (apfst (reduce [C] [])) (P1' @ Dj # P2'), Q, n)"
apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w\<^sup>*)", OF _ _
ih[of "P1' @ P2'" "apfst (reduce [C] []) Dj # P"]])
using suc_l reduce_idem p_irred unfolding p' by (auto simp: case_prod_beta)
ultimately show ?case
unfolding p' by simp
qed simp
lemma reduce_clauses_in_Q:
assumes
c_in: "C \<in> fst ` set N" and
p_irred: "\<forall>(E, k) \<in> set P. is_irreducible [C] E"
shows "wstate_of_dstate (N, P, Q @ Q', n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, fst (reduce_all2 C Q') @ P, Q @ snd (reduce_all2 C Q'), n)"
using p_irred
proof (induct Q' arbitrary: P Q)
case (Cons Dj Q')
note ih = this(1) and p_irred = this(2)
show ?case
proof (cases "is_irreducible [C] (fst Dj)")
case True
then show ?thesis
using ih[of _ "Q @ [Dj]"] p_irred by (simp add: case_prod_beta)
next
case d_red: False
have "wstate_of_dstate (N, P, Q @ Dj # Q', n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (N, (reduce [C] [] (fst Dj), snd Dj) # P, Q @ Q', n)"
using p_irred reduce_clause_in_Q[of _ _ P "snd Dj" "[]" _ Q Q' n, OF c_in _ d_red]
by (cases Dj) force
then show ?thesis
using ih[of "(reduce [C] [] (fst Dj), snd Dj) # P" Q] d_red p_irred reduce_idem
by (force simp: case_prod_beta)
qed
qed simp
lemma eligible_iff:
"eligible S \<sigma> As DA \<longleftrightarrow> As = [] \<or> length As = 1 \<and> maximal_wrt (hd As \<cdot>a \<sigma>) (DA \<cdot> \<sigma>)"
unfolding eligible.simps S_empty by (fastforce dest: hd_conv_nth)
lemma ord_resolve_one_side_prem:
"ord_resolve S CAs DA AAs As \<sigma> E \<Longrightarrow> length CAs = 1 \<and> length AAs = 1 \<and> length As = 1"
by (force elim!: ord_resolve.cases simp: eligible_iff)
lemma ord_resolve_rename_one_side_prem:
"ord_resolve_rename S CAs DA AAs As \<sigma> E \<Longrightarrow> length CAs = 1 \<and> length AAs = 1 \<and> length As = 1"
by (force elim!: ord_resolve_rename.cases dest: ord_resolve_one_side_prem)
abbreviation Bin_ord_resolve :: "'a clause \<Rightarrow> 'a clause \<Rightarrow> 'a clause set" where
"Bin_ord_resolve C D \<equiv> {E. \<exists>AA A \<sigma>. ord_resolve S [C] D [AA] [A] \<sigma> E}"
abbreviation Bin_ord_resolve_rename :: "'a clause \<Rightarrow> 'a clause \<Rightarrow> 'a clause set" where
"Bin_ord_resolve_rename C D \<equiv> {E. \<exists>AA A \<sigma>. ord_resolve_rename S [C] D [AA] [A] \<sigma> E}"
lemma resolve_on_eq_UNION_Bin_ord_resolve:
"mset ` set (resolve_on A D CA) =
{E. \<exists>AA \<sigma>. ord_resolve S [mset CA] ({#Neg A#} + mset D) [AA] [A] \<sigma> E}"
proof
{
fix E :: "'a literal list"
assume "E \<in> set (resolve_on A D CA)"
then have "E \<in> resolvent D A CA ` {Ls. subseq Ls CA \<and> resolvable A D CA Ls}"
unfolding resolve_on_def by simp
then obtain Ls where Ls_p: "resolvent D A CA Ls = E" "subseq Ls CA \<and> resolvable A D CA Ls"
by auto
define \<sigma> where "\<sigma> = the (mgu {insert A (atms_of (mset Ls))})"
then have \<sigma>_p:
"mgu {insert A (atms_of (mset Ls))} = Some \<sigma>"
"Ls \<noteq> []"
"eligible S \<sigma> [A] (add_mset (Neg A) (mset D))"
"strictly_maximal_wrt (A \<cdot>a \<sigma>) ((mset CA - mset Ls) \<cdot> \<sigma>)"
"\<forall>L \<in> set Ls. is_pos L"
using Ls_p unfolding resolvable_def unfolding Let_def eligible.simps using S_empty by auto
from \<sigma>_p have \<sigma>_p2: "the (mgu {insert A (atms_of (mset Ls))}) = \<sigma>"
by auto
have Ls_sub_CA: "mset Ls \<subseteq># mset CA"
using subseq_mset_subseteq_mset Ls_p by auto
then have "mset (resolvent D A CA Ls) = sum_list [mset CA - mset Ls] \<cdot> \<sigma> + mset D \<cdot> \<sigma>"
unfolding resolvent_def \<sigma>_p2 subst_cls_def using remove_all_mset_minus[of Ls CA] by auto
moreover
have "length [mset CA - mset Ls] = Suc 0"
by auto
moreover
have "\<forall>L \<in> set Ls. is_pos L"
using \<sigma>_p(5) list_all_iff[of is_pos] by auto
then have "{#Pos (atm_of x). x \<in># mset Ls#} = mset Ls"
by (induction Ls) auto
then have "mset CA = [mset CA - mset Ls] ! 0 + {#Pos (atm_of x). x \<in># mset Ls#}"
using Ls_sub_CA by auto
moreover
have "Ls \<noteq> []"
using \<sigma>_p by -
moreover
have "Some \<sigma> = mgu {insert A (atm_of ` set Ls)}"
using \<sigma>_p unfolding atms_of_def by auto
moreover
have "eligible S \<sigma> [A] (add_mset (Neg A) (mset D))"
using \<sigma>_p by -
moreover
have "strictly_maximal_wrt (A \<cdot>a \<sigma>) ([mset CA - mset Ls] ! 0 \<cdot> \<sigma>)"
using \<sigma>_p(4) by auto
moreover have "S (mset CA) = {#}"
by (simp add: S_empty)
ultimately have "\<exists>Cs. mset (resolvent D A CA Ls) = sum_list Cs \<cdot> \<sigma> + mset D \<cdot> \<sigma>
\<and> length Cs = Suc 0 \<and> mset CA = Cs ! 0 + {#Pos (atm_of x). x \<in># mset Ls#}
\<and> Ls \<noteq> [] \<and> Some \<sigma> = mgu {insert A (atm_of ` set Ls)}
\<and> eligible S \<sigma> [A] (add_mset (Neg A) (mset D)) \<and> strictly_maximal_wrt (A \<cdot>a \<sigma>) (Cs ! 0 \<cdot> \<sigma>)
\<and> S (mset CA) = {#}"
by blast
then have "ord_resolve S [mset CA] (add_mset (Neg A) (mset D)) [image_mset atm_of (mset Ls)] [A]
\<sigma> (mset (resolvent D A CA Ls))"
unfolding ord_resolve.simps by auto
then have "\<exists>AA \<sigma>. ord_resolve S [mset CA] (add_mset (Neg A) (mset D)) [AA] [A] \<sigma> (mset E)"
using Ls_p by auto
}
then show "mset ` set (resolve_on A D CA)
\<subseteq> {E. \<exists>AA \<sigma>. ord_resolve S [mset CA] ({#Neg A#} + mset D) [AA] [A] \<sigma> E}"
by auto
next
{
fix E AA \<sigma>
assume "ord_resolve S [mset CA] (add_mset (Neg A) (mset D)) [AA] [A] \<sigma> E"
then obtain Cs where res': "E = sum_list Cs \<cdot> \<sigma> + mset D \<cdot> \<sigma>"
"length Cs = Suc 0"
"mset CA = Cs ! 0 + poss AA"
"AA \<noteq> {#}"
"Some \<sigma> = mgu {insert A (set_mset AA)}"
"eligible S \<sigma> [A] (add_mset (Neg A) (mset D))"
"strictly_maximal_wrt (A \<cdot>a \<sigma>) (Cs ! 0 \<cdot> \<sigma>)"
"S (Cs ! 0 + poss AA) = {#}"
unfolding ord_resolve.simps by auto
moreover define C where "C = Cs ! 0"
ultimately have res:
"E = sum_list Cs \<cdot> \<sigma> + mset D \<cdot> \<sigma>"
"mset CA = C + poss AA"
"AA \<noteq> {#}"
"Some \<sigma> = mgu {insert A (set_mset AA)}"
"eligible S \<sigma> [A] (add_mset (Neg A) (mset D))"
"strictly_maximal_wrt (A \<cdot>a \<sigma>) (C \<cdot> \<sigma>)"
"S (C + poss AA) = {#}"
unfolding ord_resolve.simps by auto
from this(1) have
"E = C \<cdot> \<sigma> + mset D \<cdot> \<sigma>"
unfolding C_def using res'(2) by (cases Cs) auto
note res' = this res(2-7)
have "\<exists>Al. mset Al = AA \<and> subseq (map Pos Al) CA"
using res(2)
proof (induction CA arbitrary: AA C)
case Nil
then show ?case by auto
next
case (Cons L CA)
then show ?case
proof (cases "L \<in># poss AA ")
case True
then have pos_L: "is_pos L"
by auto
have rem: "\<And>A'. Pos A' \<in># poss AA \<Longrightarrow>
remove1_mset (Pos A') (C + poss AA) = C + poss (remove1_mset A' AA)"
by (induct AA) auto
have "mset CA = C + (poss (AA - {#atm_of L#}))"
using True Cons(2)
by (metis add_mset_remove_trivial rem literal.collapse(1) mset.simps(2) pos_L)
then have "\<exists>Al. mset Al = remove1_mset (atm_of L) AA \<and> subseq (map Pos Al) CA"
using Cons(1)[of _ "((AA - {#atm_of L#}))"] by metis
then obtain Al where
"mset Al = remove1_mset (atm_of L) AA \<and> subseq (map Pos Al) CA"
by auto
then have
"mset (atm_of L # Al) = AA" and
"subseq (map Pos (atm_of L # Al)) (L # CA)"
using True by (auto simp add: pos_L)
then show ?thesis
by blast
next
case False
then have "mset CA = remove1_mset L C + poss AA"
using Cons(2)
by (metis Un_iff add_mset_remove_trivial mset.simps(2) set_mset_union single_subset_iff
subset_mset.add_diff_assoc2 union_single_eq_member)
then have "\<exists>Al. mset Al = AA \<and> subseq (map Pos Al) CA"
using Cons(1)[of "C - {#L#}" AA] Cons(2) by auto
then show ?thesis
by auto
qed
qed
then obtain Al where Al_p: "mset Al = AA" "subseq (map Pos Al) CA"
by auto
define Ls :: "'a lclause" where "Ls = map Pos Al"
have diff: "mset CA - mset Ls = C"
unfolding Ls_def using res(2) Al_p(1) by auto
have ls_subq_ca: "subseq Ls CA"
unfolding Ls_def using Al_p by -
moreover
{
have "\<exists>y. mgu {insert A (atms_of (mset Ls))} = Some y"
unfolding Ls_def using res(4) Al_p by (metis atms_of_poss mset_map)
moreover have "Ls \<noteq> []"
using Al_p(1) Ls_def res'(3) by auto
moreover have \<sigma>_p: "the (mgu {insert A (set Al)}) = \<sigma>"
using res'(4) Al_p(1) by (metis option.sel set_mset_mset)
then have "eligible S (the (mgu {insert A (atms_of (mset Ls))})) [A]
(add_mset (Neg A) (mset D))"
unfolding Ls_def using res by auto
moreover have "strictly_maximal_wrt (A \<cdot>a the (mgu {insert A (atms_of (mset Ls))}))
((mset CA - mset Ls) \<cdot> the (mgu {insert A (atms_of (mset Ls))}))"
unfolding Ls_def using res \<sigma>_p Al_p by auto
moreover have "\<forall>L \<in> set Ls. is_pos L"
by (simp add: Ls_def)
ultimately have "resolvable A D CA Ls"
unfolding resolvable_def unfolding eligible.simps using S_empty by simp
}
moreover have ls_sub_ca: "mset Ls \<subseteq># mset CA"
using ls_subq_ca subseq_mset_subseteq_mset[of Ls CA] by simp
have "{#x \<cdot>l \<sigma>. x \<in># mset CA - mset Ls#} + {#M \<cdot>l \<sigma>. M \<in># mset D#} = C \<cdot> \<sigma> + mset D \<cdot> \<sigma>"
using diff unfolding subst_cls_def by simp
then have "{#x \<cdot>l \<sigma>. x \<in># mset CA - mset Ls#} + {#M \<cdot>l \<sigma>. M \<in># mset D#} = E"
using res'(1) by auto
then have "{#M \<cdot>l \<sigma>. M \<in># mset (remove_all CA Ls)#} + {#M \<cdot>l \<sigma> . M \<in># mset D#} = E"
using remove_all_mset_minus[of Ls CA] ls_sub_ca by auto
then have "mset (resolvent D A CA Ls) = E"
unfolding resolvable_def Let_def resolvent_def using Al_p(1) Ls_def atms_of_poss res'(4)
by (metis image_mset_union mset_append mset_map option.sel)
ultimately have "E \<in> mset ` set (resolve_on A D CA)"
unfolding resolve_on_def by auto
}
then show "{E. \<exists>AA \<sigma>. ord_resolve S [mset CA] ({#Neg A#} + mset D) [AA] [A] \<sigma> E}
\<subseteq> mset ` set (resolve_on A D CA)"
by auto
qed
lemma set_resolve_eq_UNION_set_resolve_on:
"set (resolve C D) =
(\<Union>L \<in> set D.
(case L of
Pos _ \<Rightarrow> {}
| Neg A \<Rightarrow> if maximal_wrt A (mset D) then set (resolve_on A (remove1 L D) C) else {}))"
unfolding resolve_def by (fastforce split: literal.splits if_splits)
lemma resolve_eq_Bin_ord_resolve: "mset ` set (resolve C D) = Bin_ord_resolve (mset C) (mset D)"
unfolding set_resolve_eq_UNION_set_resolve_on
apply (unfold image_UN literal.case_distrib if_distrib)
apply (subst resolve_on_eq_UNION_Bin_ord_resolve)
apply (rule order_antisym)
apply (force split: literal.splits if_splits)
apply (clarsimp split: literal.splits if_splits)
apply (rule_tac x = "Neg A" in bexI)
apply (rule conjI)
apply blast
apply clarify
apply (rule conjI)
apply clarify
apply (rule_tac x = AA in exI)
apply (rule_tac x = \<sigma> in exI)
apply (frule ord_resolve.simps[THEN iffD1])
apply force
apply (drule ord_resolve.simps[THEN iffD1])
apply (clarsimp simp: eligible_iff simp del: subst_cls_add_mset subst_cls_union)
apply (drule maximal_wrt_subst)
apply sat
apply (drule ord_resolve.simps[THEN iffD1])
using set_mset_mset by fastforce
lemma poss_in_map_clauseD:
"poss AA \<subseteq># map_clause f C \<Longrightarrow> \<exists>AA0. poss AA0 \<subseteq># C \<and> AA = {#f A. A \<in># AA0#}"
proof (induct AA arbitrary: C)
case (add A AA)
note ih = this(1) and aaa_sub = this(2)
have "Pos A \<in># map_clause f C"
using aaa_sub by auto
then obtain A0 where
pa0_in: "Pos A0 \<in># C" and
a: "A = f A0"
by clarify (metis literal.distinct(1) literal.exhaust literal.inject(1) literal.simps(9,10))
have "poss AA \<subseteq># map_clause f (C - {#Pos A0#})"
using pa0_in aaa_sub[unfolded a] by (simp add: image_mset_remove1_mset_if insert_subset_eq_iff)
then obtain AA0 where
paa0_sub: "poss AA0 \<subseteq># C - {#Pos A0#}" and
aa: "AA = image_mset f AA0"
using ih by meson
have "poss (add_mset A0 AA0) \<subseteq># C"
using pa0_in paa0_sub by (simp add: insert_subset_eq_iff)
moreover have "add_mset A AA = image_mset f (add_mset A0 AA0)"
unfolding a aa by simp
ultimately show ?case
by blast
qed simp
lemma poss_subset_filterD:
"poss AA \<subseteq># {#L \<cdot>l \<rho>. L \<in># mset C#} \<Longrightarrow> \<exists>AA0. poss AA0 \<subseteq># mset C \<and> AA = AA0 \<cdot>am \<rho>"
unfolding subst_atm_mset_def subst_lit_def by (rule poss_in_map_clauseD)
lemma neg_in_map_literalD: "Neg A \<in> map_literal f ` D \<Longrightarrow> \<exists>A0. Neg A0 \<in> D \<and> A = f A0"
unfolding image_def by (clarify, case_tac x, auto)
lemma neg_in_filterD: "Neg A \<in># {#L \<cdot>l \<rho>'. L \<in># mset D#} \<Longrightarrow> \<exists>A0. Neg A0 \<in># mset D \<and> A = A0 \<cdot>a \<rho>'"
unfolding subst_lit_def image_def by (rule neg_in_map_literalD) simp
lemma resolve_rename_eq_Bin_ord_resolve_rename:
"mset ` set (resolve_rename C D) = Bin_ord_resolve_rename (mset C) (mset D)"
proof (intro order_antisym subsetI)
let ?\<rho>s = "renamings_apart [mset D, mset C]"
define \<rho>' :: 's where
"\<rho>' = hd ?\<rho>s"
define \<rho> :: 's where
"\<rho> = last ?\<rho>s"
have tl_\<rho>s: "tl ?\<rho>s = [\<rho>]"
unfolding \<rho>_def
using renamings_apart_length Nitpick.size_list_simp(2) Suc_length_conv last.simps
by (smt length_greater_0_conv list.sel(3))
{
fix E
assume e_in: "E \<in> mset ` set (resolve_rename C D)"
from e_in obtain AA :: "'a multiset" and A :: 'a and \<sigma> :: 's where
aa_sub: "poss AA \<subseteq># mset C \<cdot> \<rho>" and
a_in: "Neg A \<in># mset D \<cdot> \<rho>'" and
res_e: "ord_resolve S [mset C \<cdot> \<rho>] {#L \<cdot>l \<rho>'. L \<in># mset D#} [AA] [A] \<sigma> E"
unfolding \<rho>'_def \<rho>_def
apply atomize_elim
using e_in unfolding resolve_rename_def Let_def resolve_eq_Bin_ord_resolve
apply clarsimp
apply (frule ord_resolve_one_side_prem)
apply (frule ord_resolve.simps[THEN iffD1])
apply (rule_tac x = AA in exI)
apply (clarsimp simp: subst_cls_def)
apply (rule_tac x = A in exI)
by (metis (full_types) Melem_subst_cls set_mset_mset subst_cls_def union_single_eq_member)
obtain AA0 :: "'a multiset" where
aa0_sub: "poss AA0 \<subseteq># mset C" and
aa: "AA = AA0 \<cdot>am \<rho>"
using aa_sub
apply atomize_elim
apply (rule ord_resolve.cases[OF res_e])
by (rule poss_subset_filterD[OF aa_sub[unfolded subst_cls_def]])
obtain A0 :: 'a where
a0_in: "Neg A0 \<in> set D" and
a: "A = A0 \<cdot>a \<rho>'"
apply atomize_elim
apply (rule ord_resolve.cases[OF res_e])
using neg_in_filterD[OF a_in[unfolded subst_cls_def]] by simp
show "E \<in> Bin_ord_resolve_rename (mset C) (mset D)"
unfolding ord_resolve_rename.simps
using res_e
apply clarsimp
apply (rule_tac x = AA0 in exI)
apply (intro conjI)
apply (rule aa0_sub)
apply (rule_tac x = A0 in exI)
apply (intro conjI)
apply (rule a0_in)
apply (rule_tac x = \<sigma> in exI)
unfolding aa a \<rho>'_def[symmetric] \<rho>_def[symmetric] tl_\<rho>s by (simp add: subst_cls_def)
}
{
fix E
assume e_in: "E \<in> Bin_ord_resolve_rename (mset C) (mset D)"
show "E \<in> mset ` set (resolve_rename C D)"
using e_in
unfolding resolve_rename_def Let_def resolve_eq_Bin_ord_resolve ord_resolve_rename.simps
apply clarsimp
apply (rule_tac x = "AA \<cdot>am \<rho>" in exI)
apply (rule_tac x = "A \<cdot>a \<rho>'" in exI)
apply (rule_tac x = \<sigma> in exI)
unfolding tl_\<rho>s \<rho>'_def \<rho>_def by (simp add: subst_cls_def subst_cls_lists_def)
}
qed
lemma bin_ord_FO_\<Gamma>_def:
"ord_FO_\<Gamma> S = {Infer {#CA#} DA E | CA DA AA A \<sigma> E. ord_resolve_rename S [CA] DA [AA] [A] \<sigma> E}"
unfolding ord_FO_\<Gamma>_def
apply (rule order.antisym)
apply clarify
apply (frule ord_resolve_rename_one_side_prem)
apply simp
apply (metis Suc_length_conv length_0_conv)
by blast
lemma ord_FO_\<Gamma>_side_prem: "\<gamma> \<in> ord_FO_\<Gamma> S \<Longrightarrow> side_prems_of \<gamma> = {#THE D. D \<in># side_prems_of \<gamma>#}"
unfolding bin_ord_FO_\<Gamma>_def by clarsimp
lemma ord_FO_\<Gamma>_infer_from_Collect_eq:
"{\<gamma> \<in> ord_FO_\<Gamma> S. infer_from (DD \<union> {C}) \<gamma> \<and> C \<in># prems_of \<gamma>} =
{\<gamma> \<in> ord_FO_\<Gamma> S. \<exists>D \<in> DD \<union> {C}. prems_of \<gamma> = {#C, D#}}"
unfolding infer_from_def
apply (rule set_eq_subset[THEN iffD2])
apply (rule conjI)
apply clarify
apply (subst (asm) (1 2) ord_FO_\<Gamma>_side_prem, assumption, assumption)
apply (subst (1) ord_FO_\<Gamma>_side_prem, assumption)
apply force
apply clarify
apply (subst (asm) (1) ord_FO_\<Gamma>_side_prem, assumption)
apply (subst (1 2) ord_FO_\<Gamma>_side_prem, assumption)
by force
lemma inferences_between_eq_UNION: "inference_system.inferences_between (ord_FO_\<Gamma> S) Q C =
inference_system.inferences_between (ord_FO_\<Gamma> S) {C} C
\<union> (\<Union>D \<in> Q. inference_system.inferences_between (ord_FO_\<Gamma> S) {D} C)"
unfolding ord_FO_\<Gamma>_infer_from_Collect_eq inference_system.inferences_between_def by auto
lemma concls_of_inferences_between_singleton_eq_Bin_ord_resolve_rename:
"concls_of (inference_system.inferences_between (ord_FO_\<Gamma> S) {D} C) =
Bin_ord_resolve_rename C C \<union> Bin_ord_resolve_rename C D \<union> Bin_ord_resolve_rename D C"
proof (intro order_antisym subsetI)
fix E
assume e_in: "E \<in> concls_of (inference_system.inferences_between (ord_FO_\<Gamma> S) {D} C)"
then show "E \<in> Bin_ord_resolve_rename C C \<union> Bin_ord_resolve_rename C D
\<union> Bin_ord_resolve_rename D C"
unfolding inference_system.inferences_between_def ord_FO_\<Gamma>_infer_from_Collect_eq
bin_ord_FO_\<Gamma>_def infer_from_def by (fastforce simp: add_mset_eq_add_mset)
qed (force simp: inference_system.inferences_between_def infer_from_def ord_FO_\<Gamma>_def)
lemma concls_of_inferences_between_eq_Bin_ord_resolve_rename:
"concls_of (inference_system.inferences_between (ord_FO_\<Gamma> S) Q C) =
Bin_ord_resolve_rename C C \<union> (\<Union>D \<in> Q. Bin_ord_resolve_rename C D \<union> Bin_ord_resolve_rename D C)"
by (subst inferences_between_eq_UNION)
(auto simp: image_Un image_UN concls_of_inferences_between_singleton_eq_Bin_ord_resolve_rename)
lemma resolve_rename_either_way_eq_congls_of_inferences_between:
"mset ` set (resolve_rename C C) \<union> (\<Union>D \<in> Q. mset ` set (resolve_rename_either_way C D)) =
concls_of (inference_system.inferences_between (ord_FO_\<Gamma> S) (mset ` Q) (mset C))"
by (simp add: resolve_rename_either_way_def image_Un resolve_rename_eq_Bin_ord_resolve_rename
concls_of_inferences_between_eq_Bin_ord_resolve_rename UN_Un_distrib)
lemma compute_inferences:
assumes
ci_in: "(C, i) \<in> set P" and
ci_min: "\<forall>(D, j) \<in># mset (map (apfst mset) P). weight (mset C, i) \<le> weight (D, j)"
shows
"wstate_of_dstate ([], P, Q, n) \<leadsto>\<^sub>w
wstate_of_dstate (map (\<lambda>D. (D, n)) (remdups_gen mset (resolve_rename C C @
concat (map (resolve_rename_either_way C \<circ> fst) Q))),
filter (\<lambda>(D, j). mset D \<noteq> mset C) P, (C, i) # Q, Suc n)"
(is "_ \<leadsto>\<^sub>w wstate_of_dstate (?N, _)")
proof -
have ms_ci_in: "(mset C, i) \<in># image_mset (apfst mset) (mset P)"
using ci_in by force
show ?thesis
apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.inference_computation[of "mset (map (apfst mset) P) - {#(mset C, i)#}" "mset C" i
"mset (map (apfst mset) ?N)" n "mset (map (apfst mset) Q)"]])
apply (simp add: add_mset_remove_trivial_eq[THEN iffD2, OF ms_ci_in, symmetric])
using ms_ci_in
apply (simp add: ci_in image_mset_remove1_mset_if)
apply (smt apfst_conv case_prodE case_prodI2 case_prod_conv filter_mset_cong
image_mset_filter_swap mset_filter)
apply (metis ci_min in_diffD)
apply (simp only: list.map_comp apfst_comp_rpair_const)
apply (simp only: list.map_comp[symmetric])
apply (subst mset_map)
apply (unfold mset_map_remdups_gen mset_remdups_gen_ident)
apply (subst image_mset_mset_set)
apply (simp add: inj_on_def)
apply (subst mset_set_eq_iff)
apply simp
apply (simp add: finite_ord_FO_resolution_inferences_between)
apply (rule arg_cong[of _ _ "\<lambda>N. (\<lambda>D. (D, n)) ` N"])
apply (simp only: map_concat list.map_comp image_comp)
using resolve_rename_either_way_eq_congls_of_inferences_between[of C "fst ` set Q", symmetric]
by (simp add: image_comp comp_def image_UN)
qed
lemma nonfinal_deterministic_RP_step:
assumes
nonfinal: "\<not> is_final_dstate St" and
step: "St' = deterministic_RP_step St"
shows "wstate_of_dstate St \<leadsto>\<^sub>w\<^sup>+ wstate_of_dstate St'"
proof -
obtain N P Q :: "'a dclause list" and n :: nat where
st: "St = (N, P, Q, n)"
by (cases St) blast
note step = step[unfolded st deterministic_RP_step.simps, simplified]
show ?thesis
proof (cases "\<exists>Ci \<in> set P \<union> set Q. fst Ci = []")
case nil_in: True
note step = step[simplified nil_in, simplified]
have nil_in': "[] \<in> fst ` set (P @ Q)"
using nil_in by (force simp: image_def)
have star: "[] \<in> fst ` set (P @ Q) \<Longrightarrow>
wstate_of_dstate (N, P, Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ([], [], remdups_clss P @ Q, n + length (remdups_clss P))"
proof (induct "length (remdups_clss P)" arbitrary: N P Q n)
case 0
note len_p = this(1) and nil_in' = this(2)
have p_nil: "P = []"
using len_p remdups_clss_Nil_iff by simp
have "wstate_of_dstate (N, [], Q, n) \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ([], [], Q, n)"
by (rule empty_N_if_Nil_in_P_or_Q[OF nil_in'[unfolded p_nil]])
then show ?case
unfolding p_nil by simp
next
case (Suc k)
note ih = this(1) and suc_k = this(2) and nil_in' = this(3)
have "P \<noteq> []"
using suc_k remdups_clss_Nil_iff by force
hence p_cons: "P = hd P # tl P"
by simp
obtain C :: "'a lclause" and i :: nat where
ci: "(C, i) = select_min_weight_clause (hd P) (tl P)"
by (metis prod.exhaust)
have ci_in: "(C, i) \<in> set P"
unfolding ci using p_cons select_min_weight_clause_in[of "hd P" "tl P"] by simp
have ci_min: "\<forall>(D, j) \<in># mset (map (apfst mset) P). weight (mset C, i) \<le> weight (D, j)"
by (subst p_cons) (simp add: select_min_weight_clause_min_weight[OF ci, simplified])
let ?P' = "filter (\<lambda>(D, j). mset D \<noteq> mset C) P"
have ms_p'_ci_q_eq: "mset (remdups_clss ?P' @ (C, i) # Q) = mset (remdups_clss P @ Q)"
apply (subst (2) p_cons)
apply (subst remdups_clss.simps(2))
by (auto simp: Let_def case_prod_beta p_cons[symmetric] ci[symmetric])
then have len_p: "length (remdups_clss P) = length (remdups_clss ?P') + 1"
by (smt Suc_eq_plus1_left add.assoc add_right_cancel length_Cons length_append
mset_eq_length)
have "wstate_of_dstate (N, P, Q, n) \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ([], P, Q, n)"
by (rule empty_N_if_Nil_in_P_or_Q[OF nil_in'])
also obtain N' :: "'a dclause list" where
"\<dots> \<leadsto>\<^sub>w wstate_of_dstate (N', ?P', (C, i) # Q, Suc n)"
by (atomize_elim, rule exI, rule compute_inferences[OF ci_in], use ci_min in fastforce)
also have "\<dots> \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ([], [], remdups_clss P @ Q, n + length (remdups_clss P))"
apply (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w\<^sup>*)", OF _ _
ih[of ?P' "(C, i) # Q" N' "Suc n"], OF refl])
using ms_p'_ci_q_eq suc_k nil_in' ci_in
apply (simp_all add: len_p)
apply (metis (no_types) apfst_conv image_mset_add_mset)
by force
finally show ?case
.
qed
show ?thesis
unfolding st step using star[OF nil_in'] nonfinal[unfolded st is_final_dstate.simps]
by cases simp_all
next
case nil_ni: False
note step = step[simplified nil_ni, simplified]
show ?thesis
proof (cases N)
case n_nil: Nil
note step = step[unfolded n_nil, simplified]
show ?thesis
proof (cases P)
case Nil
then have False
using n_nil nonfinal[unfolded st] by (simp add: is_final_dstate.simps)
then show ?thesis
using step by simp
next
case p_cons: (Cons P0 P')
note step = step[unfolded p_cons list.case, folded p_cons]
obtain C :: "'a lclause" and i :: nat where
ci: "(C, i) = select_min_weight_clause P0 P'"
by (metis prod.exhaust)
note step = step[unfolded select, simplified]
have ci_in: "(C, i) \<in> set P"
by (rule select_min_weight_clause_in[of P0 P', folded ci p_cons])
show ?thesis
unfolding st n_nil step p_cons[symmetric] ci[symmetric] prod.case
by (rule tranclp.r_into_trancl, rule compute_inferences[OF ci_in])
(simp add: select_min_weight_clause_min_weight[OF ci, simplified] p_cons)
qed
next
case n_cons: (Cons Ci N')
note step = step[unfolded n_cons, simplified]
obtain C :: "'a lclause" and i :: nat where
ci: "Ci = (C, i)"
by (cases Ci) simp
note step = step[unfolded ci, simplified]
define C' :: "'a lclause" where
"C' = reduce (map fst P @ map fst Q) [] C"
note step = step[unfolded ci C'_def[symmetric], simplified]
have "wstate_of_dstate ((E @ C, i) # N', P, Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ((E @ reduce (map fst P @ map fst Q) E C, i) # N', P, Q, n)" for E
unfolding C'_def
proof (induct C arbitrary: E)
case (Cons L C)
note ih = this(1)
show ?case
proof (cases "is_reducible_lit (map fst P @ map fst Q) (E @ C) L")
case l_red: True
then have red_lc:
"reduce (map fst P @ map fst Q) E (L # C) = reduce (map fst P @ map fst Q) E C"
by simp
obtain D D' :: "'a literal list" and L' :: "'a literal" and \<sigma> :: 's where
"D \<in> set (map fst P @ map fst Q)" and
"D' = remove1 L' D" and
"L' \<in> set D" and
"- L = L' \<cdot>l \<sigma>" and
"mset D' \<cdot> \<sigma> \<subseteq># mset (E @ C)"
using l_red unfolding is_reducible_lit_def comp_def by blast
then have \<sigma>:
"mset D' + {#L'#} \<in> set (map (mset \<circ> fst) (P @ Q))"
"- L = L' \<cdot>l \<sigma> \<and> mset D' \<cdot> \<sigma> \<subseteq># mset (E @ C)"
unfolding is_reducible_lit_def by (auto simp: comp_def)
have "wstate_of_dstate ((E @ L # C, i) # N', P, Q, n)
\<leadsto>\<^sub>w wstate_of_dstate ((E @ C, i) # N', P, Q, n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.forward_reduction[of "mset D'" L' "mset (map (apfst mset) P)"
"mset (map (apfst mset) Q)" L \<sigma> "mset (E @ C)" "mset (map (apfst mset) N')"
i n]])
(use \<sigma> in \<open>auto simp: comp_def\<close>)
then show ?thesis
unfolding red_lc using ih[of E] by (rule converse_rtranclp_into_rtranclp)
next
case False
then show ?thesis
using ih[of "L # E"] by simp
qed
qed simp
then have red_C:
"wstate_of_dstate ((C, i) # N', P, Q, n) \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', P, Q, n)"
unfolding C'_def by (metis self_append_conv2)
have proc_C: "wstate_of_dstate ((C', i) # N', P', Q', n')
\<leadsto>\<^sub>w wstate_of_dstate (N', (C', i) # P', Q', n')" for P' Q' n'
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.clause_processing[of "mset (map (apfst mset) N')" "mset C'" i
"mset (map (apfst mset) P')" "mset (map (apfst mset) Q')" n']],
simp+)
show ?thesis
proof (cases "C' = []")
case True
note c'_nil = this
note step = step[simplified c'_nil, simplified]
have
filter_p: "filter (Not \<circ> strictly_subsume [[]] \<circ> fst) P = []" and
filter_q: "filter (Not \<circ> strictly_subsume [[]] \<circ> fst) Q = []"
using nil_ni unfolding strictly_subsume_def filter_empty_conv find_None_iff by force+
note red_C[unfolded c'_nil]
also have "wstate_of_dstate (([], i) # N', P, Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (([], i) # N', [], Q, n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w\<^sup>*)", OF _ _
remove_strictly_subsumed_clauses_in_P[of "[]" _ "[]", unfolded append_Nil],
OF refl])
(auto simp: filter_p)
also have "\<dots> \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (([], i) # N', [], [], n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w\<^sup>*)", OF _ _
remove_strictly_subsumed_clauses_in_Q[of "[]" _ _ "[]", unfolded append_Nil],
OF refl])
(auto simp: filter_q)
also note proc_C[unfolded c'_nil, THEN tranclp.r_into_trancl[of "(\<leadsto>\<^sub>w)"]]
also have "wstate_of_dstate (N', [([], i)], [], n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ([], [([], i)], [], n)"
by (rule empty_N_if_Nil_in_P_or_Q) simp
also have "\<dots> \<leadsto>\<^sub>w wstate_of_dstate ([], [], [([], i)], Suc n)"
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.inference_computation[of "{#}" "{#}" i "{#}" n "{#}"]])
(auto simp: ord_FO_resolution_inferences_between_empty_empty)
finally show ?thesis
unfolding step st n_cons ci .
next
case c'_nnil: False
note step = step[simplified c'_nnil, simplified]
show ?thesis
proof (cases "is_tautology C' \<or> subsume (map fst P @ map fst Q) C'")
case taut_or_subs: True
note step = step[simplified taut_or_subs, simplified]
have "wstate_of_dstate ((C', i) # N', P, Q, n) \<leadsto>\<^sub>w wstate_of_dstate (N', P, Q, n)"
proof (cases "is_tautology C'")
case True
then obtain A :: 'a where
neg_a: "Neg A \<in> set C'" and pos_a: "Pos A \<in> set C'"
unfolding is_tautology_def by blast
show ?thesis
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.tautology_deletion[of A "mset C'" "mset (map (apfst mset) N')" i
"mset (map (apfst mset) P)" "mset (map (apfst mset) Q)" n]])
(use neg_a pos_a in simp_all)
next
case False
then have "subsume (map fst P @ map fst Q) C'"
using taut_or_subs by blast
then obtain D :: "'a lclause" where
d_in: "D \<in> set (map fst P @ map fst Q)" and
subs: "subsumes (mset D) (mset C')"
unfolding subsume_def by blast
show ?thesis
by (rule arg_cong2[THEN iffD1, of _ _ _ _ "(\<leadsto>\<^sub>w)", OF _ _
wrp.forward_subsumption[of "mset D" "mset (map (apfst mset) P)"
"mset (map (apfst mset) Q)" "mset C'" "mset (map (apfst mset) N')" i n]],
use d_in subs in \<open>auto simp: subsume_def\<close>)
qed
then show ?thesis
unfolding step st n_cons ci using red_C by (rule rtranclp_into_tranclp1[rotated])
next
case not_taut_or_subs: False
note step = step[simplified not_taut_or_subs, simplified]
define P' :: "('a literal list \<times> nat) list" where
"P' = reduce_all C' P"
obtain back_to_P Q' :: "'a dclause list" where
red_Q: "(back_to_P, Q') = reduce_all2 C' Q"
by (metis prod.exhaust)
note step = step[unfolded red_Q[symmetric], simplified]
define Q'' :: "('a literal list \<times> nat) list" where
"Q'' = filter (Not \<circ> strictly_subsume [C'] \<circ> fst) Q'"
define P'' :: "('a literal list \<times> nat) list" where
"P'' = filter (Not \<circ> strictly_subsume [C'] \<circ> fst) (back_to_P @ P')"
note step = step[unfolded P'_def[symmetric] Q''_def[symmetric] P''_def[symmetric],
simplified]
note red_C
also have "wstate_of_dstate ((C', i) # N', P, Q, n)
\<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', P', Q, n)"
unfolding P'_def by (rule reduce_clauses_in_P[of _ _ "[]", unfolded append_Nil]) simp+
also have "\<dots> \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', back_to_P @ P', Q', n)"
unfolding P'_def
by (rule reduce_clauses_in_Q[of C' _ _ "[]" Q, folded red_Q,
unfolded append_Nil prod.sel])
(auto intro: reduce_idem simp: reduce_all_def)
also have "\<dots> \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', back_to_P @ P', Q'', n)"
unfolding Q''_def
by (rule remove_strictly_subsumed_clauses_in_Q[of _ _ _ "[]", unfolded append_Nil])
simp
also have "\<dots> \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate ((C', i) # N', P'', Q'', n)"
unfolding P''_def
by (rule remove_strictly_subsumed_clauses_in_P[of _ _ "[]", unfolded append_Nil]) auto
also note proc_C[THEN tranclp.r_into_trancl[of "(\<leadsto>\<^sub>w)"]]
finally show ?thesis
unfolding step st n_cons ci P''_def by simp
qed
qed
qed
qed
qed
lemma final_deterministic_RP_step: "is_final_dstate St \<Longrightarrow> deterministic_RP_step St = St"
by (cases St) (auto simp: deterministic_RP_step.simps is_final_dstate.simps)
lemma deterministic_RP_SomeD:
assumes "deterministic_RP (N, P, Q, n) = Some R"
shows "\<exists>N' P' Q' n'. (\<exists>k. (deterministic_RP_step ^^ k) (N, P, Q, n) = (N', P', Q', n'))
\<and> is_final_dstate (N', P', Q', n') \<and> R = map fst Q'"
proof (induct rule: deterministic_RP.raw_induct[OF _ assms])
case (1 self_call St R)
note ih = this(1) and step = this(2)
obtain N P Q :: "'a dclause list" and n :: nat where
st: "St = (N, P, Q, n)"
by (cases St) blast
note step = step[unfolded st, simplified]
show ?case
proof (cases "is_final_dstate (N, P, Q, n)")
case True
then have "(deterministic_RP_step ^^ 0) (N, P, Q, n) = (N, P, Q, n)
\<and> is_final_dstate (N, P, Q, n) \<and> R = map fst Q"
using step by simp
then show ?thesis
unfolding st by blast
next
case nonfinal: False
note step = step[simplified nonfinal, simplified]
obtain N' P' Q' :: "'a dclause list" and n' k :: nat where
"(deterministic_RP_step ^^ k) (deterministic_RP_step (N, P, Q, n)) = (N', P', Q', n')" and
"is_final_dstate (N', P', Q', n')"
"R = map fst Q'"
using ih[OF step] by blast
then show ?thesis
unfolding st funpow_Suc_right[symmetric, THEN fun_cong, unfolded comp_apply] by blast
qed
qed
context
fixes
N0 :: "'a dclause list" and
n0 :: nat and
R :: "'a lclause list"
begin
abbreviation St0 :: "'a dstate" where
"St0 \<equiv> (N0, [], [], n0)"
abbreviation grounded_N0 where
"grounded_N0 \<equiv> grounding_of_clss (set (map (mset \<circ> fst) N0))"
abbreviation grounded_R :: "'a clause set" where
"grounded_R \<equiv> grounding_of_clss (set (map mset R))"
primcorec derivation_from :: "'a dstate \<Rightarrow> 'a dstate llist" where
"derivation_from St =
LCons St (if is_final_dstate St then LNil else derivation_from (deterministic_RP_step St))"
abbreviation Sts :: "'a dstate llist" where
"Sts \<equiv> derivation_from St0"
abbreviation wSts :: "'a wstate llist" where
"wSts \<equiv> lmap wstate_of_dstate Sts"
lemma full_deriv_wSts_trancl_weighted_RP: "full_chain (\<leadsto>\<^sub>w\<^sup>+) wSts"
proof -
have "Sts' = derivation_from St0' \<Longrightarrow> full_chain (\<leadsto>\<^sub>w\<^sup>+) (lmap wstate_of_dstate Sts')"
for St0' Sts'
proof (coinduction arbitrary: St0' Sts' rule: full_chain.coinduct)
case sts': full_chain
show ?case
proof (cases "is_final_dstate St0'")
case True
then have "ltl (lmap wstate_of_dstate Sts') = LNil"
unfolding sts' by simp
then have "lmap wstate_of_dstate Sts' = LCons (wstate_of_dstate St0') LNil"
unfolding sts' by (subst derivation_from.code, subst (asm) derivation_from.code, auto)
moreover have "\<And>St''. \<not> wstate_of_dstate St0' \<leadsto>\<^sub>w St''"
using True by (rule is_final_dstate_imp_not_weighted_RP)
ultimately show ?thesis
by (meson tranclpD)
next
case nfinal: False
have "lmap wstate_of_dstate Sts' =
LCons (wstate_of_dstate St0') (lmap wstate_of_dstate (ltl Sts'))"
unfolding sts' by (subst derivation_from.code) simp
moreover have "ltl Sts' = derivation_from (deterministic_RP_step St0')"
unfolding sts' using nfinal by (subst derivation_from.code) simp
moreover have "wstate_of_dstate St0' \<leadsto>\<^sub>w\<^sup>+ wstate_of_dstate (lhd (ltl Sts'))"
unfolding sts' using nonfinal_deterministic_RP_step[OF nfinal refl] nfinal
by (subst derivation_from.code) simp
ultimately show ?thesis
by fastforce
qed
qed
then show ?thesis
by blast
qed
lemmas deriv_wSts_trancl_weighted_RP = full_chain_imp_chain[OF full_deriv_wSts_trancl_weighted_RP]
definition sswSts :: "'a wstate llist" where
"sswSts = (SOME wSts'.
full_chain (\<leadsto>\<^sub>w) wSts' \<and> emb wSts wSts' \<and> lhd wSts' = lhd wSts \<and> llast wSts' = llast wSts)"
lemma sswSts:
"full_chain (\<leadsto>\<^sub>w) sswSts \<and> emb wSts sswSts \<and> lhd sswSts = lhd wSts \<and> llast sswSts = llast wSts"
unfolding sswSts_def
by (rule someI_ex[OF full_chain_tranclp_imp_exists_full_chain[OF
full_deriv_wSts_trancl_weighted_RP]])
lemmas full_deriv_sswSts_weighted_RP = sswSts[THEN conjunct1]
lemmas emb_sswSts = sswSts[THEN conjunct2, THEN conjunct1]
lemmas lfinite_sswSts_iff = emb_lfinite[OF emb_sswSts]
lemmas lhd_sswSts = sswSts[THEN conjunct2, THEN conjunct2, THEN conjunct1]
lemmas llast_sswSts = sswSts[THEN conjunct2, THEN conjunct2, THEN conjunct2]
lemmas deriv_sswSts_weighted_RP = full_chain_imp_chain[OF full_deriv_sswSts_weighted_RP]
lemma not_lnull_sswSts: "\<not> lnull sswSts"
using deriv_sswSts_weighted_RP by (cases rule: chain.cases) auto
lemma empty_ssgP0: "wrp.P_of_wstate (lhd sswSts) = {}"
unfolding lhd_sswSts by (subst derivation_from.code) simp
lemma empty_ssgQ0: "wrp.Q_of_wstate (lhd sswSts) = {}"
unfolding lhd_sswSts by (subst derivation_from.code) simp
lemmas sswSts_thms = full_deriv_sswSts_weighted_RP empty_ssgP0 empty_ssgQ0
abbreviation S_ssgQ :: "'a clause \<Rightarrow> 'a clause" where
"S_ssgQ \<equiv> wrp.S_gQ sswSts"
abbreviation ord_\<Gamma> :: "'a inference set" where
"ord_\<Gamma> \<equiv> ground_resolution_with_selection.ord_\<Gamma> S_ssgQ"
abbreviation Rf :: "'a clause set \<Rightarrow> 'a clause set" where
"Rf \<equiv> standard_redundancy_criterion.Rf"
abbreviation Ri :: "'a clause set \<Rightarrow> 'a inference set" where
"Ri \<equiv> standard_redundancy_criterion.Ri ord_\<Gamma>"
abbreviation saturated_upto :: "'a clause set \<Rightarrow> bool" where
"saturated_upto \<equiv> redundancy_criterion.saturated_upto ord_\<Gamma> Rf Ri"
context
assumes drp_some: "deterministic_RP St0 = Some R"
begin
lemma lfinite_Sts: "lfinite Sts"
proof (induct rule: deterministic_RP.raw_induct[OF _ drp_some])
case (1 self_call St St')
note ih = this(1) and step = this(2)
show ?case
using step by (subst derivation_from.code, auto intro: ih)
qed
lemma lfinite_wSts: "lfinite wSts"
by (rule lfinite_lmap[THEN iffD2, OF lfinite_Sts])
lemmas lfinite_sswSts = lfinite_sswSts_iff[THEN iffD2, OF lfinite_wSts]
theorem
deterministic_RP_saturated: "saturated_upto grounded_R" (is ?saturated) and
deterministic_RP_model: "I \<Turnstile>s grounded_N0 \<longleftrightarrow> I \<Turnstile>s grounded_R" (is ?model)
proof -
obtain N' P' Q' :: "'a dclause list" and n' k :: nat where
k_steps: "(deterministic_RP_step ^^ k) St0 = (N', P', Q', n')" (is "_ = ?Stk") and
final: "is_final_dstate (N', P', Q', n')" and
r: "R = map fst Q'"
using deterministic_RP_SomeD[OF drp_some] by blast
have wrp: "wstate_of_dstate St0 \<leadsto>\<^sub>w\<^sup>* wstate_of_dstate (llast Sts)"
using lfinite_chain_imp_rtranclp_lhd_llast
by (metis (no_types) deriv_sswSts_weighted_RP derivation_from.disc_iff derivation_from.simps(2)
lfinite_Sts lfinite_sswSts llast_lmap llist.map_sel(1) sswSts)
have last_sts: "llast Sts = ?Stk"
proof -
have "(deterministic_RP_step ^^ k') St0' = ?Stk \<Longrightarrow> llast (derivation_from St0') = ?Stk"
for St0' k'
proof (induct k' arbitrary: St0')
case 0
then show ?case
using final by (subst derivation_from.code) simp
next
case (Suc k')
note ih = this(1) and suc_k'_steps = this(2)
show ?case
proof (cases "is_final_dstate St0'")
case True
then show ?thesis
using ih[of "deterministic_RP_step St0'"] suc_k'_steps final_deterministic_RP_step
funpow_fixpoint[of deterministic_RP_step]
by auto
next
case False
then show ?thesis
using ih[of "deterministic_RP_step St0'"] suc_k'_steps
by (subst derivation_from.code) (simp add: llast_LCons funpow_swap1[symmetric])
qed
qed
then show ?thesis
using k_steps by blast
qed
have fin_gr_fgsts: "lfinite (lmap wrp.grounding_of_wstate sswSts)"
by (rule lfinite_lmap[THEN iffD2, OF lfinite_sswSts])
have lim_last: "Liminf_llist (lmap wrp.grounding_of_wstate sswSts) =
wrp.grounding_of_wstate (llast sswSts)"
unfolding lfinite_Liminf_llist[OF fin_gr_fgsts] llast_lmap[OF lfinite_sswSts not_lnull_sswSts]
using not_lnull_sswSts by simp
have gr_st0: "wrp.grounding_of_wstate (wstate_of_dstate St0) = grounded_N0"
unfolding comp_def by simp
have "?saturated \<and> ?model"
proof (cases "[] \<in> set R")
case True
then have emp_in: "{#} \<in> grounded_R"
unfolding grounding_of_clss_def grounding_of_cls_def by (auto intro: ex_ground_subst)
have "grounded_R \<subseteq> wrp.grounding_of_wstate (llast sswSts)"
unfolding r llast_sswSts
by (simp add: last_sts llast_lmap[OF lfinite_Sts] grounding_of_clss_def)
then have gr_last_st: "grounded_R \<subseteq> wrp.grounding_of_wstate (wstate_of_dstate (llast Sts))"
by (simp add: lfinite_Sts llast_lmap llast_sswSts)
have gr_r_fls: "\<not> I \<Turnstile>s grounded_R"
using emp_in unfolding true_clss_def by force
then have gr_last_fls: "\<not> I \<Turnstile>s wrp.grounding_of_wstate (wstate_of_dstate (llast Sts))"
using gr_last_st unfolding true_clss_def by auto
have ?saturated
unfolding wrp.ord_\<Gamma>_saturated_upto_def[OF sswSts_thms]
wrp.ord_\<Gamma>_contradiction_Rf[OF sswSts_thms emp_in] inference_system.inferences_from_def
by auto
moreover have ?model
unfolding gr_r_fls[THEN eq_False[THEN iffD2]]
by (rule rtranclp_imp_eq_image[of "(\<leadsto>\<^sub>w)" "\<lambda>St. I \<Turnstile>s wrp.grounding_of_wstate St", OF _ wrp,
unfolded gr_st0 gr_last_fls[THEN eq_False[THEN iffD2]]])
(use wrp.weighted_RP_model[OF sswSts_thms] in blast)
ultimately show ?thesis
by blast
next
case False
then have gr_last: "wrp.grounding_of_wstate (llast sswSts) = grounded_R"
using final unfolding r llast_sswSts
by (simp add: last_sts llast_lmap[OF lfinite_Sts] comp_def is_final_dstate.simps)
then have gr_last_st: "wrp.grounding_of_wstate (wstate_of_dstate (llast Sts)) = grounded_R"
by (simp add: lfinite_Sts llast_lmap llast_sswSts)
have ?saturated
using wrp.weighted_RP_saturated[OF sswSts_thms, unfolded gr_last lim_last] by auto
moreover have ?model
by (rule rtranclp_imp_eq_image[of "(\<leadsto>\<^sub>w)" "\<lambda>St. I \<Turnstile>s wrp.grounding_of_wstate St", OF _ wrp,
unfolded gr_st0 gr_last_st])
(use wrp.weighted_RP_model[OF sswSts_thms] in blast)
ultimately show ?thesis
by blast
qed
then show ?saturated and ?model
by blast+
qed
corollary deterministic_RP_refutation:
"\<not> satisfiable grounded_N0 \<longleftrightarrow> {#} \<in> grounded_R" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?rhs
then have "\<not> satisfiable grounded_R"
unfolding true_clss_def true_cls_def by force
then show ?lhs
using deterministic_RP_model[THEN iffD1] by blast
next
assume ?lhs
then have "\<not> satisfiable grounded_R"
using deterministic_RP_model[THEN iffD2] by blast
then show ?rhs
unfolding wrp.ord_\<Gamma>_saturated_upto_complete[OF sswSts_thms deterministic_RP_saturated] .
qed
end
context
assumes drp_none: "deterministic_RP St0 = None"
begin
theorem deterministic_RP_complete: "satisfiable grounded_N0"
proof (rule ccontr)
assume unsat: "\<not> satisfiable grounded_N0"
have unsat_wSts0: "\<not> satisfiable (wrp.grounding_of_wstate (lhd wSts))"
using unsat by (subst derivation_from.code) (simp add: comp_def)
have bot_in_ss: "{#} \<in> Q_of_state (wrp.Liminf_wstate sswSts)"
by (rule wrp.weighted_RP_complete[OF sswSts_thms unsat_wSts0[folded lhd_sswSts]])
have bot_in_lim: "{#} \<in> Q_of_state (wrp.Liminf_wstate wSts)"
proof (cases "lfinite Sts")
case fin: True
have "wrp.Liminf_wstate sswSts = wrp.Liminf_wstate wSts"
by (rule Liminf_state_fin, simp_all add: fin lfinite_sswSts_iff not_lnull_sswSts,
subst (1 2) llast_lmap,
simp_all add: lfinite_sswSts_iff fin not_lnull_sswSts llast_sswSts)
then show ?thesis
using bot_in_ss by simp
next
case False
then show ?thesis
using bot_in_ss Q_of_Liminf_state_inf[OF _ emb_lmap[OF emb_sswSts]] by auto
qed
then obtain k :: nat where
k_lt: "enat k < llength Sts" and
emp_in: "{#} \<in> wrp.Q_of_wstate (lnth wSts k)"
unfolding Liminf_state_def Liminf_llist_def by auto
have emp_in: "{#} \<in> Q_of_state (state_of_dstate ((deterministic_RP_step ^^ k) St0))"
proof -
have "enat k < llength Sts' \<Longrightarrow> Sts' = derivation_from St0' \<Longrightarrow>
{#} \<in> wrp.Q_of_wstate (lnth (lmap wstate_of_dstate Sts') k) \<Longrightarrow>
{#} \<in> Q_of_state (state_of_dstate ((deterministic_RP_step ^^ k) St0'))" for St0' Sts' k
proof (induction k arbitrary: St0' Sts')
case 0
then show ?case
by (subst (asm) derivation_from.code, cases St0', auto simp: comp_def)
next
case (Suc k)
note ih = this(1) and sk_lt = this(2) and sts' = this(3) and emp_in_sk = this(4)
have k_lt: "enat k < llength (ltl Sts')"
using sk_lt by (cases Sts') (auto simp: Suc_ile_eq)
moreover have "ltl Sts' = derivation_from (deterministic_RP_step St0')"
using sts' k_lt by (cases Sts') auto
moreover have "{#} \<in> wrp.Q_of_wstate (lnth (lmap wstate_of_dstate (ltl Sts')) k)"
using emp_in_sk k_lt by (cases Sts') auto
ultimately show ?case
using ih[of "ltl Sts'" "deterministic_RP_step St0'"] by (simp add: funpow_swap1)
qed
then show ?thesis
using k_lt emp_in by blast
qed
have "deterministic_RP St0 \<noteq> None"
by (rule is_final_dstate_funpow_imp_deterministic_RP_neq_None[of "Suc k"],
cases "(deterministic_RP_step ^^ k) St0",
use emp_in in \<open>force simp: deterministic_RP_step.simps is_final_dstate.simps\<close>)
then show False
using drp_none ..
qed
end
end
end
end
diff --git a/thys/Groebner_Bases/Groebner_PM.thy b/thys/Groebner_Bases/Groebner_PM.thy
--- a/thys/Groebner_Bases/Groebner_PM.thy
+++ b/thys/Groebner_Bases/Groebner_PM.thy
@@ -1,486 +1,488 @@
(* Author: Alexander Maletzky *)
theory Groebner_PM
imports Polynomials.MPoly_PM Reduced_GB
begin
text \<open>We prove results that hold specifically for Gr\"obner bases in polynomial rings, where the
polynomials really have @{emph \<open>indeterminates\<close>}.\<close>
context pm_powerprod
begin
lemmas finite_reduced_GB_Polys =
punit.finite_reduced_GB_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_is_reduced_GB_Polys =
punit.reduced_GB_is_reduced_GB_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_is_GB_Polys =
punit.reduced_GB_is_GB_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_is_auto_reduced_Polys =
punit.reduced_GB_is_auto_reduced_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_is_monic_set_Polys =
punit.reduced_GB_is_monic_set_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_nonzero_Polys =
punit.reduced_GB_nonzero_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_ideal_Polys =
punit.reduced_GB_pmdl_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_unique_Polys =
punit.reduced_GB_unique_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas reduced_GB_Polys =
punit.reduced_GB_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
lemmas ideal_eq_UNIV_iff_reduced_GB_eq_one_Polys =
ideal_eq_UNIV_iff_reduced_GB_eq_one_dgrad_p_set[simplified, OF dickson_grading_varnum, where m=0, simplified dgrad_p_set_varnum]
subsection \<open>Univariate Polynomials\<close>
lemma (in -) adds_univariate_linear:
assumes "finite X" and "card X \<le> 1" and "s \<in> .[X]" and "t \<in> .[X]"
obtains "s adds t" | "t adds s"
proof (cases "s adds t")
case True
thus ?thesis ..
next
case False
then obtain x where 1: "lookup t x < lookup s x" by (auto simp: adds_poly_mapping le_fun_def not_le)
hence "x \<in> keys s" by (simp add: in_keys_iff)
also from assms(3) have "\<dots> \<subseteq> X" by (rule PPsD)
finally have "x \<in> X" .
have "t adds s" unfolding adds_poly_mapping le_fun_def
proof
fix y
show "lookup t y \<le> lookup s y"
proof (cases "y \<in> keys t")
case True
also from assms(4) have "keys t \<subseteq> X" by (rule PPsD)
finally have "y \<in> X" .
with assms(1, 2) \<open>x \<in> X\<close> have "x = y" by (simp add: card_le_Suc0_iff_eq)
with 1 show ?thesis by simp
next
case False
thus ?thesis by (simp add: in_keys_iff)
qed
qed
thus ?thesis ..
qed
context
fixes X :: "'x set"
assumes fin_X: "finite X" and card_X: "card X \<le> 1"
begin
lemma ord_iff_adds_univariate:
assumes "s \<in> .[X]" and "t \<in> .[X]"
shows "s \<preceq> t \<longleftrightarrow> s adds t"
proof
assume "s \<preceq> t"
from fin_X card_X assms show "s adds t"
proof (rule adds_univariate_linear)
assume "t adds s"
hence "t \<preceq> s" by (rule ord_adds)
- with \<open>s \<preceq> t\<close> have "s = t" by (rule ordered_powerprod_lin.antisym)
+ with \<open>s \<preceq> t\<close> have "s = t"
+ by simp
thus ?thesis by simp
qed
qed (rule ord_adds)
lemma adds_iff_deg_le_univariate:
assumes "s \<in> .[X]" and "t \<in> .[X]"
shows "s adds t \<longleftrightarrow> deg_pm s \<le> deg_pm t"
proof
assume *: "deg_pm s \<le> deg_pm t"
from fin_X card_X assms show "s adds t"
proof (rule adds_univariate_linear)
assume "t adds s"
hence "t = s" using * by (rule adds_deg_pm_antisym)
thus ?thesis by simp
qed
qed (rule deg_pm_mono)
corollary ord_iff_deg_le_univariate: "s \<in> .[X] \<Longrightarrow> t \<in> .[X] \<Longrightarrow> s \<preceq> t \<longleftrightarrow> deg_pm s \<le> deg_pm t"
by (simp only: ord_iff_adds_univariate adds_iff_deg_le_univariate)
lemma poly_deg_univariate:
assumes "p \<in> P[X]"
shows "poly_deg p = deg_pm (lpp p)"
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
hence lp_in: "lpp p \<in> keys p" by (rule punit.lt_in_keys)
also from assms have "\<dots> \<subseteq> .[X]" by (rule PolysD)
finally have "lpp p \<in> .[X]" .
show ?thesis
proof (intro antisym poly_deg_leI)
fix t
assume "t \<in> keys p"
hence "t \<preceq> lpp p" by (rule punit.lt_max_keys)
moreover from \<open>t \<in> keys p\<close> \<open>keys p \<subseteq> .[X]\<close> have "t \<in> .[X]" ..
ultimately show "deg_pm t \<le> deg_pm (lpp p)" using \<open>lpp p \<in> .[X]\<close>
by (simp only: ord_iff_deg_le_univariate)
next
from lp_in show "deg_pm (lpp p) \<le> poly_deg p" by (rule poly_deg_max_keys)
qed
qed
lemma reduced_GB_univariate_cases:
assumes "F \<subseteq> P[X]"
obtains g where "g \<in> P[X]" and "g \<noteq> 0" and "lcf g = 1" and "punit.reduced_GB F = {g}" |
"punit.reduced_GB F = {}"
proof (cases "punit.reduced_GB F = {}")
case True
thus ?thesis ..
next
case False
let ?G = "punit.reduced_GB F"
from fin_X assms have ar: "punit.is_auto_reduced ?G" and "0 \<notin> ?G" and "?G \<subseteq> P[X]"
and m: "punit.is_monic_set ?G"
by (rule reduced_GB_is_auto_reduced_Polys, rule reduced_GB_nonzero_Polys, rule reduced_GB_Polys,
rule reduced_GB_is_monic_set_Polys)
from False obtain g where "g \<in> ?G" by blast
with \<open>0 \<notin> ?G\<close> \<open>?G \<subseteq> P[X]\<close> have "g \<noteq> 0" and "g \<in> P[X]" by blast+
from this(1) have lp_g: "lpp g \<in> keys g" by (rule punit.lt_in_keys)
also from \<open>g \<in> P[X]\<close> have "\<dots> \<subseteq> .[X]" by (rule PolysD)
finally have "lpp g \<in> .[X]" .
note \<open>g \<in> P[X]\<close> \<open>g \<noteq> 0\<close>
moreover from m \<open>g \<in> ?G\<close> \<open>g \<noteq> 0\<close> have "lcf g = 1" by (rule punit.is_monic_setD)
moreover have "?G = {g}"
proof
show "?G \<subseteq> {g}"
proof
fix g'
assume "g' \<in> ?G"
with \<open>0 \<notin> ?G\<close> \<open>?G \<subseteq> P[X]\<close> have "g' \<noteq> 0" and "g' \<in> P[X]" by blast+
from this(1) have lp_g': "lpp g' \<in> keys g'" by (rule punit.lt_in_keys)
also from \<open>g' \<in> P[X]\<close> have "\<dots> \<subseteq> .[X]" by (rule PolysD)
finally have "lpp g' \<in> .[X]" .
have "g' = g"
proof (rule ccontr)
assume "g' \<noteq> g"
with \<open>g \<in> ?G\<close> \<open>g' \<in> ?G\<close> have g: "g \<in> ?G - {g'}" and g': "g' \<in> ?G - {g}" by blast+
from fin_X card_X \<open>lpp g \<in> .[X]\<close> \<open>lpp g' \<in> .[X]\<close> show False
proof (rule adds_univariate_linear)
assume *: "lpp g adds lpp g'"
from ar \<open>g' \<in> ?G\<close> have "\<not> punit.is_red (?G - {g'}) g'" by (rule punit.is_auto_reducedD)
moreover from g \<open>g \<noteq> 0\<close> lp_g' * have "punit.is_red (?G - {g'}) g'"
by (rule punit.is_red_addsI[simplified])
ultimately show ?thesis ..
next
assume *: "lpp g' adds lpp g"
from ar \<open>g \<in> ?G\<close> have "\<not> punit.is_red (?G - {g}) g" by (rule punit.is_auto_reducedD)
moreover from g' \<open>g' \<noteq> 0\<close> lp_g * have "punit.is_red (?G - {g}) g"
by (rule punit.is_red_addsI[simplified])
ultimately show ?thesis ..
qed
qed
thus "g' \<in> {g}" by simp
qed
next
from \<open>g \<in> ?G\<close> show "{g} \<subseteq> ?G" by simp
qed
ultimately show ?thesis ..
qed
corollary deg_reduced_GB_univariate_le:
assumes "F \<subseteq> P[X]" and "f \<in> ideal F" and "f \<noteq> 0" and "g \<in> punit.reduced_GB F"
shows "poly_deg g \<le> poly_deg f"
using assms(1)
proof (rule reduced_GB_univariate_cases)
let ?G = "punit.reduced_GB F"
fix g'
assume "g' \<in> P[X]" and "g' \<noteq> 0" and G: "?G = {g'}"
from fin_X assms(1) have gb: "punit.is_Groebner_basis ?G" and "ideal ?G = ideal F"
and "?G \<subseteq> P[X]"
by (rule reduced_GB_is_GB_Polys, rule reduced_GB_ideal_Polys, rule reduced_GB_Polys)
from assms(2) this(2) have "f \<in> ideal ?G" by simp
with gb obtain g'' where "g'' \<in> ?G" and "lpp g'' adds lpp f"
using assms(3) by (rule punit.GB_adds_lt[simplified])
with assms(4) have "lpp g adds lpp f" by (simp add: G)
hence "deg_pm (lpp g) \<le> deg_pm (lpp f)" by (rule deg_pm_mono)
moreover from assms(4) \<open>?G \<subseteq> P[X]\<close> have "g \<in> P[X]" ..
ultimately have "poly_deg g \<le> deg_pm (lpp f)" by (simp only: poly_deg_univariate)
also from punit.lt_in_keys have "\<dots> \<le> poly_deg f" by (rule poly_deg_max_keys) fact
finally show ?thesis .
next
assume "punit.reduced_GB F = {}"
with assms(4) show ?thesis by simp
qed
end
subsection \<open>Homogeneity\<close>
lemma is_reduced_GB_homogeneous:
assumes "\<And>f. f \<in> F \<Longrightarrow> homogeneous f" and "punit.is_reduced_GB G" and "ideal G = ideal F"
and "g \<in> G"
shows "homogeneous g"
proof (rule homogeneousI)
fix s t
have 1: "deg_pm u = deg_pm (lpp g)" if "u \<in> keys g" for u
proof -
from assms(4) have "g \<in> ideal G" by (rule ideal.span_base)
hence "g \<in> ideal F" by (simp only: assms(3))
from that have "u \<in> Keys (hom_components g)" by (simp only: Keys_hom_components)
then obtain q where q: "q \<in> hom_components g" and "u \<in> keys q" by (rule in_KeysE)
from assms(1) \<open>g \<in> ideal F\<close> q have "q \<in> ideal F" by (rule homogeneous_ideal')
from assms(2) have "punit.is_Groebner_basis G" by (rule punit.reduced_GB_D1)
moreover from \<open>q \<in> ideal F\<close> have "q \<in> ideal G" by (simp only: assms(3))
moreover from q have "q \<noteq> 0" by (rule hom_components_nonzero)
ultimately obtain g' where "g' \<in> G" and "g' \<noteq> 0" and adds: "lpp g' adds lpp q"
by (rule punit.GB_adds_lt[simplified])
from \<open>q \<noteq> 0\<close> have "lpp q \<in> keys q" by (rule punit.lt_in_keys)
also from q have "\<dots> \<subseteq> Keys (hom_components g)" by (rule keys_subset_Keys)
finally have "lpp q \<in> keys g" by (simp only: Keys_hom_components)
with _ \<open>g' \<noteq> 0\<close> have red: "punit.is_red {g'} g"
using adds by (rule punit.is_red_addsI[simplified]) simp
from assms(2) have "punit.is_auto_reduced G" by (rule punit.reduced_GB_D2)
hence "\<not> punit.is_red (G - {g}) g" using assms(4) by (rule punit.is_auto_reducedD)
with red have "\<not> {g'} \<subseteq> G - {g}" using punit.is_red_subset by blast
with \<open>g' \<in> G\<close> have "g' = g" by simp
from \<open>lpp q \<in> keys g\<close> have "lpp q \<preceq> lpp g" by (rule punit.lt_max_keys)
moreover from adds have "lpp g \<preceq> lpp q"
unfolding \<open>g' = g\<close> by (rule punit.ord_adds_term[simplified])
- ultimately have eq: "lpp q = lpp g" by (rule ordered_powerprod_lin.antisym)
+ ultimately have eq: "lpp q = lpp g"
+ by simp
from q have "homogeneous q" by (rule hom_components_homogeneous)
hence "deg_pm u = deg_pm (lpp q)"
using \<open>u \<in> keys q\<close> \<open>lpp q \<in> keys q\<close> by (rule homogeneousD)
thus ?thesis by (simp only: eq)
qed
assume "s \<in> keys g"
hence 2: "deg_pm s = deg_pm (lpp g)" by (rule 1)
assume "t \<in> keys g"
hence "deg_pm t = deg_pm (lpp g)" by (rule 1)
with 2 show "deg_pm s = deg_pm t" by simp
qed
lemma lp_dehomogenize:
assumes "is_hom_ord x" and "homogeneous p"
shows "lpp (dehomogenize x p) = except (lpp p) {x}"
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
hence "lpp p \<in> keys p" by (rule punit.lt_in_keys)
with assms(2) have "except (lpp p) {x} \<in> keys (dehomogenize x p)"
by (rule keys_dehomogenizeI)
thus ?thesis
proof (rule punit.lt_eqI_keys)
fix t
assume "t \<in> keys (dehomogenize x p)"
then obtain s where "s \<in> keys p" and t: "t = except s {x}" by (rule keys_dehomogenizeE)
from this(1) have "s \<preceq> lpp p" by (rule punit.lt_max_keys)
moreover from assms(2) \<open>s \<in> keys p\<close> \<open>lpp p \<in> keys p\<close> have "deg_pm s = deg_pm (lpp p)"
by (rule homogeneousD)
ultimately show "t \<preceq> except (lpp p) {x}" using assms(1) by (simp add: t is_hom_ordD)
qed
qed
lemma isGB_dehomogenize:
assumes "is_hom_ord x" and "finite X" and "G \<subseteq> P[X]" and "punit.is_Groebner_basis G"
and "\<And>g. g \<in> G \<Longrightarrow> homogeneous g"
shows "punit.is_Groebner_basis (dehomogenize x ` G)"
using dickson_grading_varnum
proof (rule punit.isGB_I_adds_lt[simplified])
from assms(2) show "finite (X - {x})" by simp
next
have "dehomogenize x ` G \<subseteq> P[X - {x}]"
proof
fix g
assume "g \<in> dehomogenize x ` G"
then obtain g' where "g' \<in> G" and g: "g = dehomogenize x g'" ..
from this(1) assms(3) have "g' \<in> P[X]" ..
hence "indets g' \<subseteq> X" by (rule PolysD)
have "indets g \<subseteq> indets g' - {x}" by (simp only: g indets_dehomogenize)
also from \<open>indets g' \<subseteq> X\<close> subset_refl have "\<dots> \<subseteq> X - {x}" by (rule Diff_mono)
finally show "g \<in> P[X - {x}]" by (rule PolysI_alt)
qed
thus "dehomogenize x ` G \<subseteq> punit.dgrad_p_set (varnum (X - {x})) 0"
by (simp only: dgrad_p_set_varnum)
next
fix p
assume "p \<in> ideal (dehomogenize x ` G)"
then obtain G0 q where "G0 \<subseteq> dehomogenize x ` G" and "finite G0" and p: "p = (\<Sum>g\<in>G0. q g * g)"
by (rule ideal.spanE)
from this(1) obtain G' where "G' \<subseteq> G" and G0: "G0 = dehomogenize x ` G'"
and inj: "inj_on (dehomogenize x) G'" by (rule subset_imageE_inj)
define p' where "p' = (\<Sum>g\<in>G'. q (dehomogenize x g) * g)"
have "p' \<in> ideal G'" unfolding p'_def by (rule ideal.sum_in_spanI)
also from \<open>G' \<subseteq> G\<close> have "\<dots> \<subseteq> ideal G" by (rule ideal.span_mono)
finally have "p' \<in> ideal G" .
with assms(5) have "homogenize x p' \<in> ideal G" (is "?p \<in> _") by (rule homogeneous_ideal_homogenize)
assume "p \<in> punit.dgrad_p_set (varnum (X - {x})) 0"
hence "p \<in> P[X - {x}]" by (simp only: dgrad_p_set_varnum)
hence "indets p \<subseteq> X - {x}" by (rule PolysD)
hence "x \<notin> indets p" by blast
have "p = dehomogenize x p" by (rule sym) (simp add: \<open>x \<notin> indets p\<close>)
also from inj have "\<dots> = dehomogenize x (\<Sum>g\<in>G'. q (dehomogenize x g) * dehomogenize x g)"
by (simp add: p G0 sum.reindex)
also have "\<dots> = dehomogenize x ?p"
by (simp add: dehomogenize_sum dehomogenize_times p'_def)
finally have p: "p = dehomogenize x ?p" .
moreover assume "p \<noteq> 0"
ultimately have "?p \<noteq> 0" by (auto simp del: dehomogenize_homogenize)
with assms(4) \<open>?p \<in> ideal G\<close> obtain g where "g \<in> G" and "g \<noteq> 0" and adds: "lpp g adds lpp ?p"
by (rule punit.GB_adds_lt[simplified])
from this(1) have "homogeneous g" by (rule assms(5))
show "\<exists>g\<in>dehomogenize x ` G. g \<noteq> 0 \<and> lpp g adds lpp p"
proof (intro bexI conjI notI)
assume "dehomogenize x g = 0"
hence "g = 0" using \<open>homogeneous g\<close> by (rule dehomogenize_zeroD)
with \<open>g \<noteq> 0\<close> show False ..
next
from assms(1) \<open>homogeneous g\<close> have "lpp (dehomogenize x g) = except (lpp g) {x}"
by (rule lp_dehomogenize)
also from adds have "\<dots> adds except (lpp ?p) {x}"
by (simp add: adds_poly_mapping le_fun_def lookup_except)
also from assms(1) homogeneous_homogenize have "\<dots> = lpp (dehomogenize x ?p)"
by (rule lp_dehomogenize[symmetric])
finally show "lpp (dehomogenize x g) adds lpp p" by (simp only: p)
next
from \<open>g \<in> G\<close> show "dehomogenize x g \<in> dehomogenize x ` G" by (rule imageI)
qed
qed
end (* pm_powerprod *)
context extended_ord_pm_powerprod
begin
lemma extended_ord_lp:
assumes "None \<notin> indets p"
shows "restrict_indets_pp (extended_ord.lpp p) = lpp (restrict_indets p)"
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
hence "extended_ord.lpp p \<in> keys p" by (rule extended_ord.punit.lt_in_keys)
hence "restrict_indets_pp (extended_ord.lpp p) \<in> restrict_indets_pp ` keys p" by (rule imageI)
also from assms have eq: "\<dots> = keys (restrict_indets p)" by (rule keys_restrict_indets[symmetric])
finally show ?thesis
proof (rule punit.lt_eqI_keys[symmetric])
fix t
assume "t \<in> keys (restrict_indets p)"
then obtain s where "s \<in> keys p" and t: "t = restrict_indets_pp s" unfolding eq[symmetric] ..
from this(1) have "extended_ord s (extended_ord.lpp p)" by (rule extended_ord.punit.lt_max_keys)
thus "t \<preceq> restrict_indets_pp (extended_ord.lpp p)" by (auto simp: t extended_ord_def)
qed
qed
lemma restrict_indets_reduced_GB:
assumes "finite X" and "F \<subseteq> P[X]"
shows "punit.is_Groebner_basis (restrict_indets ` extended_ord.punit.reduced_GB (homogenize None ` extend_indets ` F))"
(is ?thesis1)
and "ideal (restrict_indets ` extended_ord.punit.reduced_GB (homogenize None ` extend_indets ` F)) = ideal F"
(is ?thesis2)
and "restrict_indets ` extended_ord.punit.reduced_GB (homogenize None ` extend_indets ` F) \<subseteq> P[X]"
(is ?thesis3)
proof -
let ?F = "homogenize None ` extend_indets ` F"
let ?G = "extended_ord.punit.reduced_GB ?F"
from assms(1) have "finite (insert None (Some ` X))" by simp
moreover have "?F \<subseteq> P[insert None (Some ` X)]"
proof
fix hf
assume "hf \<in> ?F"
then obtain f where "f \<in> F" and hf: "hf = homogenize None (extend_indets f)" by auto
from this(1) assms(2) have "f \<in> P[X]" ..
hence "indets f \<subseteq> X" by (rule PolysD)
hence "Some ` indets f \<subseteq> Some ` X" by (rule image_mono)
with indets_extend_indets[of f] have "indets (extend_indets f) \<subseteq> Some ` X" by blast
hence "insert None (indets (extend_indets f)) \<subseteq> insert None (Some ` X)" by blast
with indets_homogenize_subset have "indets hf \<subseteq> insert None (Some ` X)"
unfolding hf by (rule subset_trans)
thus "hf \<in> P[insert None (Some ` X)]" by (rule PolysI_alt)
qed
ultimately have G_sub: "?G \<subseteq> P[insert None (Some ` X)]"
and ideal_G: "ideal ?G = ideal ?F"
and GB_G: "extended_ord.punit.is_reduced_GB ?G"
by (rule extended_ord.reduced_GB_Polys, rule extended_ord.reduced_GB_ideal_Polys,
rule extended_ord.reduced_GB_is_reduced_GB_Polys)
show ?thesis3
proof
fix g
assume "g \<in> restrict_indets ` ?G"
then obtain g' where "g' \<in> ?G" and g: "g = restrict_indets g'" ..
from this(1) G_sub have "g' \<in> P[insert None (Some ` X)]" ..
hence "indets g' \<subseteq> insert None (Some ` X)" by (rule PolysD)
have "indets g \<subseteq> the ` (indets g' - {None})" by (simp only: g indets_restrict_indets_subset)
also from \<open>indets g' \<subseteq> insert None (Some ` X)\<close> have "\<dots> \<subseteq> X" by auto
finally show "g \<in> P[X]" by (rule PolysI_alt)
qed
from dickson_grading_varnum show ?thesis1
proof (rule punit.isGB_I_adds_lt[simplified])
from \<open>?thesis3\<close> show "restrict_indets ` ?G \<subseteq> punit.dgrad_p_set (varnum X) 0"
by (simp only: dgrad_p_set_varnum)
next
fix p :: "('a \<Rightarrow>\<^sub>0 nat) \<Rightarrow>\<^sub>0 'b"
assume "p \<noteq> 0"
assume "p \<in> ideal (restrict_indets ` ?G)"
hence "extend_indets p \<in> extend_indets ` ideal (restrict_indets ` ?G)" by (rule imageI)
also have "\<dots> \<subseteq> ideal (extend_indets ` restrict_indets ` ?G)" by (fact extend_indets_ideal_subset)
also have "\<dots> = ideal (dehomogenize None ` ?G)"
by (simp only: image_comp extend_indets_comp_restrict_indets)
finally have p_in_ideal: "extend_indets p \<in> ideal (dehomogenize None ` ?G)" .
assume "p \<in> punit.dgrad_p_set (varnum X) 0"
hence "p \<in> P[X]" by (simp only: dgrad_p_set_varnum)
have "extended_ord.punit.is_Groebner_basis (dehomogenize None ` ?G)"
using extended_ord_is_hom_ord \<open>finite (insert None (Some ` X))\<close> G_sub
proof (rule extended_ord.isGB_dehomogenize)
from GB_G show "extended_ord.punit.is_Groebner_basis ?G"
by (rule extended_ord.punit.reduced_GB_D1)
next
fix g
assume "g \<in> ?G"
with _ GB_G ideal_G show "homogeneous g"
proof (rule extended_ord.is_reduced_GB_homogeneous)
fix hf
assume "hf \<in> ?F"
then obtain f where "hf = homogenize None f" ..
thus "homogeneous hf" by (simp only: homogeneous_homogenize)
qed
qed
moreover note p_in_ideal
moreover from \<open>p \<noteq> 0\<close> have "extend_indets p \<noteq> 0" by simp
ultimately obtain g where g_in: "g \<in> dehomogenize None ` ?G" and "g \<noteq> 0"
and adds: "extended_ord.lpp g adds extended_ord.lpp (extend_indets p)"
by (rule extended_ord.punit.GB_adds_lt[simplified])
have "None \<notin> indets g"
proof
assume "None \<in> indets g"
moreover from g_in obtain g0 where "g = dehomogenize None g0" ..
ultimately show False using indets_dehomogenize[of None g0] by blast
qed
show "\<exists>g\<in>restrict_indets ` ?G. g \<noteq> 0 \<and> lpp g adds lpp p"
proof (intro bexI conjI notI)
have "lpp (restrict_indets g) = restrict_indets_pp (extended_ord.lpp g)"
by (rule sym, intro extended_ord_lp \<open>None \<notin> indets g\<close>)
also from adds have "\<dots> adds restrict_indets_pp (extended_ord.lpp (extend_indets p))"
by (simp add: adds_poly_mapping le_fun_def lookup_restrict_indets_pp)
also have "\<dots> = lpp (restrict_indets (extend_indets p))"
proof (intro extended_ord_lp notI)
assume "None \<in> indets (extend_indets p)"
thus False by (simp add: indets_extend_indets)
qed
also have "\<dots> = lpp p" by simp
finally show "lpp (restrict_indets g) adds lpp p" .
next
from g_in have "restrict_indets g \<in> restrict_indets ` dehomogenize None ` ?G" by (rule imageI)
also have "\<dots> = restrict_indets ` ?G" by (simp only: image_comp restrict_indets_comp_dehomogenize)
finally show "restrict_indets g \<in> restrict_indets ` ?G" .
next
assume "restrict_indets g = 0"
with \<open>None \<notin> indets g\<close> extend_restrict_indets have "g = 0" by fastforce
with \<open>g \<noteq> 0\<close> show False ..
qed
qed (fact assms(1))
from ideal_G show ?thesis2 by (rule ideal_restrict_indets)
qed
end
end (* theory *)
diff --git a/thys/Groebner_Macaulay/Hilbert_Function.thy b/thys/Groebner_Macaulay/Hilbert_Function.thy
--- a/thys/Groebner_Macaulay/Hilbert_Function.thy
+++ b/thys/Groebner_Macaulay/Hilbert_Function.thy
@@ -1,1388 +1,1388 @@
(* Author: Alexander Maletzky *)
section \<open>Direct Decompositions and Hilbert Functions\<close>
theory Hilbert_Function
imports Dube_Prelims Degree_Section "HOL-Library.Permutations"
begin
subsection \<open>Direct Decompositions\<close>
text \<open>The main reason for defining \<open>direct_decomp\<close> in terms of lists rather than sets is that
lemma \<open>direct_decomp_direct_decomp\<close> can be proved easier. At some point one could invest the time
to re-define \<open>direct_decomp\<close> in terms of sets (possibly adding a couple of further assumptions to
\<open>direct_decomp_direct_decomp\<close>).\<close>
definition direct_decomp :: "'a set \<Rightarrow> 'a::comm_monoid_add set list \<Rightarrow> bool"
where "direct_decomp A ss \<longleftrightarrow> bij_betw sum_list (listset ss) A"
lemma direct_decompI:
"inj_on sum_list (listset ss) \<Longrightarrow> sum_list ` listset ss = A \<Longrightarrow> direct_decomp A ss"
by (simp add: direct_decomp_def bij_betw_def)
lemma direct_decompI_alt:
"(\<And>qs. qs \<in> listset ss \<Longrightarrow> sum_list qs \<in> A) \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> \<exists>!qs\<in>listset ss. a = sum_list qs) \<Longrightarrow>
direct_decomp A ss"
by (auto simp: direct_decomp_def intro!: bij_betwI') blast
lemma direct_decompD:
assumes "direct_decomp A ss"
shows "qs \<in> listset ss \<Longrightarrow> sum_list qs \<in> A" and "inj_on sum_list (listset ss)"
and "sum_list ` listset ss = A"
using assms by (auto simp: direct_decomp_def bij_betw_def)
lemma direct_decompE:
assumes "direct_decomp A ss" and "a \<in> A"
obtains qs where "qs \<in> listset ss" and "a = sum_list qs"
using assms by (auto simp: direct_decomp_def bij_betw_def)
lemma direct_decomp_unique:
"direct_decomp A ss \<Longrightarrow> qs \<in> listset ss \<Longrightarrow> qs' \<in> listset ss \<Longrightarrow> sum_list qs = sum_list qs' \<Longrightarrow>
qs = qs'"
by (auto dest: direct_decompD simp: inj_on_def)
lemma direct_decomp_singleton: "direct_decomp A [A]"
proof (rule direct_decompI_alt)
fix qs
assume "qs \<in> listset [A]"
then obtain q where "q \<in> A" and "qs = [q]" by (rule listset_singletonE)
thus "sum_list qs \<in> A" by simp
next
fix a
assume "a \<in> A"
show "\<exists>!qs\<in>listset [A]. a = sum_list qs"
proof (intro ex1I conjI allI impI)
from \<open>a \<in> A\<close> refl show "[a] \<in> listset [A]" by (rule listset_singletonI)
next
fix qs
assume "qs \<in> listset [A] \<and> a = sum_list qs"
hence a: "a = sum_list qs" and "qs \<in> listset [A]" by simp_all
from this(2) obtain b where qs: "qs = [b]" by (rule listset_singletonE)
with a show "qs = [a]" by simp
qed simp_all
qed
(* TODO: Move. *)
lemma mset_bij:
assumes "bij_betw f {..<length xs} {..<length ys}" and "\<And>i. i < length xs \<Longrightarrow> xs ! i = ys ! f i"
shows "mset xs = mset ys"
proof -
from assms(1) have 1: "inj_on f {0..<length xs}" and 2: "f ` {0..<length xs} = {0..<length ys}"
by (simp_all add: bij_betw_def lessThan_atLeast0)
let ?f = "(!) ys \<circ> f"
have "xs = map ?f [0..<length xs]" unfolding list_eq_iff_nth_eq
proof (intro conjI allI impI)
fix i
assume "i < length xs"
hence "xs ! i = ys ! f i" by (rule assms(2))
also from \<open>i < length xs\<close> have "\<dots> = map ((!) ys \<circ> f) [0..<length xs] ! i" by simp
finally show "xs ! i = map ((!) ys \<circ> f) [0..<length xs] ! i" .
qed simp
hence "mset xs = mset (map ?f [0..<length xs])" by (rule arg_cong)
also have "\<dots> = image_mset ((!) ys) (image_mset f (mset_set {0..<length xs}))"
by (simp flip: image_mset.comp)
also from 1 have "\<dots> = image_mset ((!) ys) (mset_set {0..<length ys})"
by (simp add: image_mset_mset_set 2)
also have "\<dots> = mset (map ((!) ys) [0..<length ys])" by simp
finally show "mset xs = mset ys" by (simp only: map_nth)
qed
lemma direct_decomp_perm:
assumes "direct_decomp A ss1" and "mset ss1 = mset ss2"
shows "direct_decomp A ss2"
proof -
from assms(2) have len_ss1: "length ss1 = length ss2"
using mset_eq_length by blast
from assms(2) obtain f where \<open>f permutes {..<length ss2}\<close>
\<open>permute_list f ss2 = ss1\<close>
by (rule mset_eq_permutation)
then have f_bij: "bij_betw f {..<length ss2} {..<length ss1}"
and f: "\<And>i. i < length ss2 \<Longrightarrow> ss1 ! i = ss2 ! f i"
by (auto simp add: permutes_imp_bij permute_list_nth)
define g where "g = inv_into {..<length ss2} f"
from f_bij have g_bij: "bij_betw g {..<length ss1} {..<length ss2}"
unfolding g_def len_ss1 by (rule bij_betw_inv_into)
have f_g: "f (g i) = i" if "i < length ss1" for i
proof -
from that f_bij have "i \<in> f ` {..<length ss2}" by (simp add: bij_betw_def len_ss1)
thus ?thesis by (simp only: f_inv_into_f g_def)
qed
have g_f: "g (f i) = i" if "i < length ss2" for i
proof -
from f_bij have "inj_on f {..<length ss2}" by (simp only: bij_betw_def)
moreover from that have "i \<in> {..<length ss2}" by simp
ultimately show ?thesis by (simp add: g_def)
qed
have g: "ss2 ! i = ss1 ! g i" if "i < length ss1" for i
proof -
from that have "i \<in> {..<length ss2}" by (simp add: len_ss1)
hence "g i \<in> g ` {..<length ss2}" by (rule imageI)
also from g_bij have "\<dots> = {..<length ss2}" by (simp only: len_ss1 bij_betw_def)
finally have "g i < length ss2" by simp
hence "ss1 ! g i = ss2 ! f (g i)" by (rule f)
with that show ?thesis by (simp only: f_g)
qed
show ?thesis
proof (rule direct_decompI_alt)
fix qs2
assume "qs2 \<in> listset ss2"
then obtain qs1 where qs1_in: "qs1 \<in> listset ss1" and len_qs1: "length qs1 = length qs2"
and *: "\<And>i. i < length qs2 \<Longrightarrow> qs1 ! i = qs2 ! f i" using f_bij f by (rule listset_permE) blast+
from \<open>qs2 \<in> listset ss2\<close> have "length qs2 = length ss2" by (rule listsetD)
with f_bij have "bij_betw f {..<length qs1} {..<length qs2}" by (simp only: len_qs1 len_ss1)
hence "mset qs1 = mset qs2" using * by (rule mset_bij) (simp only: len_qs1)
hence "sum_list qs2 = sum_list qs1" by (simp flip: sum_mset_sum_list)
also from assms(1) qs1_in have "\<dots> \<in> A" by (rule direct_decompD)
finally show "sum_list qs2 \<in> A" .
next
fix a
assume "a \<in> A"
with assms(1) obtain qs where a: "a = sum_list qs" and qs_in: "qs \<in> listset ss1"
by (rule direct_decompE)
from qs_in obtain qs2 where qs2_in: "qs2 \<in> listset ss2" and len_qs2: "length qs2 = length qs"
and 1: "\<And>i. i < length qs \<Longrightarrow> qs2 ! i = qs ! g i" using g_bij g by (rule listset_permE) blast+
show "\<exists>!qs\<in>listset ss2. a = sum_list qs"
proof (intro ex1I conjI allI impI)
from qs_in have len_qs: "length qs = length ss1" by (rule listsetD)
with g_bij have g_bij2: "bij_betw g {..<length qs2} {..<length qs}" by (simp only: len_qs2 len_ss1)
hence "mset qs2 = mset qs" using 1 by (rule mset_bij) (simp only: len_qs2)
thus a2: "a = sum_list qs2" by (simp only: a flip: sum_mset_sum_list)
fix qs'
assume "qs' \<in> listset ss2 \<and> a = sum_list qs'"
hence qs'_in: "qs' \<in> listset ss2" and a': "a = sum_list qs'" by simp_all
from this(1) obtain qs1 where qs1_in: "qs1 \<in> listset ss1" and len_qs1: "length qs1 = length qs'"
and 2: "\<And>i. i < length qs' \<Longrightarrow> qs1 ! i = qs' ! f i" using f_bij f by (rule listset_permE) blast+
from \<open>qs' \<in> listset ss2\<close> have "length qs' = length ss2" by (rule listsetD)
with f_bij have "bij_betw f {..<length qs1} {..<length qs'}" by (simp only: len_qs1 len_ss1)
hence "mset qs1 = mset qs'" using 2 by (rule mset_bij) (simp only: len_qs1)
hence "sum_list qs1 = sum_list qs'" by (simp flip: sum_mset_sum_list)
hence "sum_list qs1 = sum_list qs" by (simp only: a flip: a')
with assms(1) qs1_in qs_in have "qs1 = qs" by (rule direct_decomp_unique)
show "qs' = qs2" unfolding list_eq_iff_nth_eq
proof (intro conjI allI impI)
from qs'_in have "length qs' = length ss2" by (rule listsetD)
thus eq: "length qs' = length qs2" by (simp only: len_qs2 len_qs len_ss1)
fix i
assume "i < length qs'"
hence "i < length qs2" by (simp only: eq)
hence "i \<in> {..<length qs2}" and "i < length qs" and "i < length ss1"
by (simp_all add: len_qs2 len_qs)
from this(1) have "g i \<in> g ` {..<length qs2}" by (rule imageI)
also from g_bij2 have "\<dots> = {..<length qs}" by (simp only: bij_betw_def)
finally have "g i < length qs'" by (simp add: eq len_qs2)
from \<open>i < length qs\<close> have "qs2 ! i = qs ! g i" by (rule 1)
also have "\<dots> = qs1 ! g i" by (simp only: \<open>qs1 = qs\<close>)
also from \<open>g i < length qs'\<close> have "\<dots> = qs' ! f (g i)" by (rule 2)
also from \<open>i < length ss1\<close> have "\<dots> = qs' ! i" by (simp only: f_g)
finally show "qs' ! i = qs2 ! i" by (rule sym)
qed
qed fact
qed
qed
lemma direct_decomp_split_map:
"direct_decomp A (map f ss) \<Longrightarrow> direct_decomp A (map f (filter P ss) @ map f (filter (- P) ss))"
proof (rule direct_decomp_perm)
show "mset (map f ss) = mset (map f (filter P ss) @ map f (filter (- P) ss))"
by simp (metis image_mset_union multiset_partition)
qed
lemmas direct_decomp_split = direct_decomp_split_map[where f=id, simplified]
lemma direct_decomp_direct_decomp:
assumes "direct_decomp A (s # ss)" and "direct_decomp s rs"
shows "direct_decomp A (ss @ rs)" (is "direct_decomp A ?ss")
proof (rule direct_decompI_alt)
fix qs
assume "qs \<in> listset ?ss"
then obtain qs1 qs2 where qs1: "qs1 \<in> listset ss" and qs2: "qs2 \<in> listset rs" and qs: "qs = qs1 @ qs2"
by (rule listset_appendE)
have "sum_list qs = sum_list ((sum_list qs2) # qs1)" by (simp add: qs add.commute)
also from assms(1) have "\<dots> \<in> A"
proof (rule direct_decompD)
from assms(2) qs2 have "sum_list qs2 \<in> s" by (rule direct_decompD)
thus "sum_list qs2 # qs1 \<in> listset (s # ss)" using qs1 refl by (rule listset_ConsI)
qed
finally show "sum_list qs \<in> A" .
next
fix a
assume "a \<in> A"
with assms(1) obtain qs1 where qs1_in: "qs1 \<in> listset (s # ss)" and a: "a = sum_list qs1"
by (rule direct_decompE)
from qs1_in obtain qs11 qs12 where "qs11 \<in> s" and qs12_in: "qs12 \<in> listset ss"
and qs1: "qs1 = qs11 # qs12" by (rule listset_ConsE)
from assms(2) this(1) obtain qs2 where qs2_in: "qs2 \<in> listset rs" and qs11: "qs11 = sum_list qs2"
by (rule direct_decompE)
let ?qs = "qs12 @ qs2"
show "\<exists>!qs\<in>listset ?ss. a = sum_list qs"
proof (intro ex1I conjI allI impI)
from qs12_in qs2_in refl show "?qs \<in> listset ?ss" by (rule listset_appendI)
show "a = sum_list ?qs" by (simp add: a qs1 qs11 add.commute)
fix qs0
assume "qs0 \<in> listset ?ss \<and> a = sum_list qs0"
hence qs0_in: "qs0 \<in> listset ?ss" and a2: "a = sum_list qs0" by simp_all
from this(1) obtain qs01 qs02 where qs01_in: "qs01 \<in> listset ss" and qs02_in: "qs02 \<in> listset rs"
and qs0: "qs0 = qs01 @ qs02" by (rule listset_appendE)
note assms(1)
moreover from _ qs01_in refl have "(sum_list qs02) # qs01 \<in> listset (s # ss)" (is "?qs' \<in> _")
proof (rule listset_ConsI)
from assms(2) qs02_in show "sum_list qs02 \<in> s" by (rule direct_decompD)
qed
moreover note qs1_in
moreover from a2 have "sum_list ?qs' = sum_list qs1" by (simp add: qs0 a add.commute)
ultimately have "?qs' = qs11 # qs12" unfolding qs1 by (rule direct_decomp_unique)
hence "qs11 = sum_list qs02" and 1: "qs01 = qs12" by simp_all
from this(1) have "sum_list qs02 = sum_list qs2" by (simp only: qs11)
with assms(2) qs02_in qs2_in have "qs02 = qs2" by (rule direct_decomp_unique)
thus "qs0 = qs12 @ qs2" by (simp only: 1 qs0)
qed
qed
lemma sum_list_map_times: "sum_list (map ((*) x) xs) = (x::'a::semiring_0) * sum_list xs"
by (induct xs) (simp_all add: algebra_simps)
lemma direct_decomp_image_times:
assumes "direct_decomp (A::'a::semiring_0 set) ss" and "\<And>a b. x * a = x * b \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> a = b"
shows "direct_decomp ((*) x ` A) (map ((`) ((*) x)) ss)" (is "direct_decomp ?A ?ss")
proof (rule direct_decompI_alt)
fix qs
assume "qs \<in> listset ?ss"
then obtain qs0 where qs0_in: "qs0 \<in> listset ss" and qs: "qs = map ((*) x) qs0"
by (rule listset_map_imageE)
have "sum_list qs = x * sum_list qs0" by (simp only: qs sum_list_map_times)
moreover from assms(1) qs0_in have "sum_list qs0 \<in> A" by (rule direct_decompD)
ultimately show "sum_list qs \<in> (*) x ` A" by (rule image_eqI)
next
fix a
assume "a \<in> ?A"
then obtain a' where "a' \<in> A" and a: "a = x * a'" ..
from assms(1) this(1) obtain qs' where qs'_in: "qs' \<in> listset ss" and a': "a' = sum_list qs'"
by (rule direct_decompE)
define qs where "qs = map ((*) x) qs'"
show "\<exists>!qs\<in>listset ?ss. a = sum_list qs"
proof (intro ex1I conjI allI impI)
from qs'_in qs_def show "qs \<in> listset ?ss" by (rule listset_map_imageI)
fix qs0
assume "qs0 \<in> listset ?ss \<and> a = sum_list qs0"
hence "qs0 \<in> listset ?ss" and a0: "a = sum_list qs0" by simp_all
from this(1) obtain qs1 where qs1_in: "qs1 \<in> listset ss" and qs0: "qs0 = map ((*) x) qs1"
by (rule listset_map_imageE)
show "qs0 = qs"
proof (cases "x = 0")
case True
from qs1_in have "length qs1 = length ss" by (rule listsetD)
moreover from qs'_in have "length qs' = length ss" by (rule listsetD)
ultimately show ?thesis by (simp add: qs_def qs0 list_eq_iff_nth_eq True)
next
case False
have "x * sum_list qs1 = a" by (simp only: a0 qs0 sum_list_map_times)
also have "\<dots> = x * sum_list qs'" by (simp only: a' a)
finally have "sum_list qs1 = sum_list qs'" using False by (rule assms(2))
with assms(1) qs1_in qs'_in have "qs1 = qs'" by (rule direct_decomp_unique)
thus ?thesis by (simp only: qs0 qs_def)
qed
qed (simp only: a a' qs_def sum_list_map_times)
qed
lemma direct_decomp_appendD:
assumes "direct_decomp A (ss1 @ ss2)"
shows "{} \<notin> set ss2 \<Longrightarrow> direct_decomp (sum_list ` listset ss1) ss1" (is "_ \<Longrightarrow> ?thesis1")
and "{} \<notin> set ss1 \<Longrightarrow> direct_decomp (sum_list ` listset ss2) ss2" (is "_ \<Longrightarrow> ?thesis2")
and "direct_decomp A [sum_list ` listset ss1, sum_list ` listset ss2]" (is "direct_decomp _ ?ss")
proof -
have rl: "direct_decomp (sum_list ` listset ts1) ts1"
if "direct_decomp A (ts1 @ ts2)" and "{} \<notin> set ts2" for ts1 ts2
proof (intro direct_decompI inj_onI refl)
fix qs1 qs2
assume qs1: "qs1 \<in> listset ts1" and qs2: "qs2 \<in> listset ts1"
assume eq: "sum_list qs1 = sum_list qs2"
from that(2) have "listset ts2 \<noteq> {}" by (simp add: listset_empty_iff)
then obtain qs3 where qs3: "qs3 \<in> listset ts2" by blast
note that(1)
moreover from qs1 qs3 refl have "qs1 @ qs3 \<in> listset (ts1 @ ts2)" by (rule listset_appendI)
moreover from qs2 qs3 refl have "qs2 @ qs3 \<in> listset (ts1 @ ts2)" by (rule listset_appendI)
moreover have "sum_list (qs1 @ qs3) = sum_list (qs2 @ qs3)" by (simp add: eq)
ultimately have "qs1 @ qs3 = qs2 @ qs3" by (rule direct_decomp_unique)
thus "qs1 = qs2" by simp
qed
{
assume "{} \<notin> set ss2"
with assms show ?thesis1 by (rule rl)
}
{
from assms have "direct_decomp A (ss2 @ ss1)"
by (rule direct_decomp_perm) simp
moreover assume "{} \<notin> set ss1"
ultimately show ?thesis2 by (rule rl)
}
show "direct_decomp A ?ss"
proof (rule direct_decompI_alt)
fix qs
assume "qs \<in> listset ?ss"
then obtain q1 q2 where q1: "q1 \<in> sum_list ` listset ss1" and q2: "q2 \<in> sum_list ` listset ss2"
and qs: "qs = [q1, q2]" by (rule listset_doubletonE)
from q1 obtain qs1 where qs1: "qs1 \<in> listset ss1" and q1: "q1 = sum_list qs1" ..
from q2 obtain qs2 where qs2: "qs2 \<in> listset ss2" and q2: "q2 = sum_list qs2" ..
from qs1 qs2 refl have "qs1 @ qs2 \<in> listset (ss1 @ ss2)" by (rule listset_appendI)
with assms have "sum_list (qs1 @ qs2) \<in> A" by (rule direct_decompD)
thus "sum_list qs \<in> A" by (simp add: qs q1 q2)
next
fix a
assume "a \<in> A"
with assms obtain qs0 where qs0_in: "qs0 \<in> listset (ss1 @ ss2)" and a: "a = sum_list qs0"
by (rule direct_decompE)
from this(1) obtain qs1 qs2 where qs1: "qs1 \<in> listset ss1" and qs2: "qs2 \<in> listset ss2"
and qs0: "qs0 = qs1 @ qs2" by (rule listset_appendE)
from qs1 have len_qs1: "length qs1 = length ss1" by (rule listsetD)
define qs where "qs = [sum_list qs1, sum_list qs2]"
show "\<exists>!qs\<in>listset ?ss. a = sum_list qs"
proof (intro ex1I conjI)
from qs1 have "sum_list qs1 \<in> sum_list ` listset ss1" by (rule imageI)
moreover from qs2 have "sum_list qs2 \<in> sum_list ` listset ss2" by (rule imageI)
ultimately show "qs \<in> listset ?ss" using qs_def by (rule listset_doubletonI)
fix qs'
assume "qs' \<in> listset ?ss \<and> a = sum_list qs'"
hence "qs' \<in> listset ?ss" and a': "a = sum_list qs'" by simp_all
from this(1) obtain q1 q2 where q1: "q1 \<in> sum_list ` listset ss1"
and q2: "q2 \<in> sum_list ` listset ss2" and qs': "qs' = [q1, q2]" by (rule listset_doubletonE)
from q1 obtain qs1' where qs1': "qs1' \<in> listset ss1" and q1: "q1 = sum_list qs1'" ..
from q2 obtain qs2' where qs2': "qs2' \<in> listset ss2" and q2: "q2 = sum_list qs2'" ..
from qs1' have len_qs1': "length qs1' = length ss1" by (rule listsetD)
note assms
moreover from qs1' qs2' refl have "qs1' @ qs2' \<in> listset (ss1 @ ss2)" by (rule listset_appendI)
moreover note qs0_in
moreover have "sum_list (qs1' @ qs2') = sum_list qs0" by (simp add: a' qs' flip: a q1 q2)
ultimately have "qs1' @ qs2' = qs0" by (rule direct_decomp_unique)
also have "\<dots> = qs1 @ qs2" by fact
finally show "qs' = qs" by (simp add: qs_def qs' q1 q2 len_qs1 len_qs1')
qed (simp add: qs_def a qs0)
qed
qed
lemma direct_decomp_Cons_zeroI:
assumes "direct_decomp A ss"
shows "direct_decomp A ({0} # ss)"
proof (rule direct_decompI_alt)
fix qs
assume "qs \<in> listset ({0} # ss)"
then obtain q qs' where "q \<in> {0}" and "qs' \<in> listset ss" and "qs = q # qs'"
by (rule listset_ConsE)
from this(1, 3) have "sum_list qs = sum_list qs'" by simp
also from assms \<open>qs' \<in> listset ss\<close> have "\<dots> \<in> A" by (rule direct_decompD)
finally show "sum_list qs \<in> A" .
next
fix a
assume "a \<in> A"
with assms obtain qs' where qs': "qs' \<in> listset ss" and a: "a = sum_list qs'"
by (rule direct_decompE)
define qs where "qs = 0 # qs'"
show "\<exists>!qs. qs \<in> listset ({0} # ss) \<and> a = sum_list qs"
proof (intro ex1I conjI)
from _ qs' qs_def show "qs \<in> listset ({0} # ss)" by (rule listset_ConsI) simp
next
fix qs0
assume "qs0 \<in> listset ({0} # ss) \<and> a = sum_list qs0"
hence "qs0 \<in> listset ({0} # ss)" and a0: "a = sum_list qs0" by simp_all
from this(1) obtain q0 qs0' where "q0 \<in> {0}" and qs0': "qs0' \<in> listset ss"
and qs0: "qs0 = q0 # qs0'" by (rule listset_ConsE)
from this(1, 3) have "sum_list qs0' = sum_list qs'" by (simp add: a0 flip: a)
with assms qs0' qs' have "qs0' = qs'" by (rule direct_decomp_unique)
with \<open>q0 \<in> {0}\<close> show "qs0 = qs" by (simp add: qs_def qs0)
qed (simp add: qs_def a)
qed
lemma direct_decomp_Cons_zeroD:
assumes "direct_decomp A ({0} # ss)"
shows "direct_decomp A ss"
proof -
have "direct_decomp {0} []" by (simp add: direct_decomp_def bij_betw_def)
with assms have "direct_decomp A (ss @ [])" by (rule direct_decomp_direct_decomp)
thus ?thesis by simp
qed
lemma direct_decomp_Cons_subsetI:
assumes "direct_decomp A (s # ss)" and "\<And>s0. s0 \<in> set ss \<Longrightarrow> 0 \<in> s0"
shows "s \<subseteq> A"
proof
fix x
assume "x \<in> s"
moreover from assms(2) have "map (\<lambda>_. 0) ss \<in> listset ss"
by (induct ss, auto simp del: listset.simps(2) intro: listset_ConsI)
ultimately have "x # (map (\<lambda>_. 0) ss) \<in> listset (s # ss)" using refl by (rule listset_ConsI)
with assms(1) have "sum_list (x # (map (\<lambda>_. 0) ss)) \<in> A" by (rule direct_decompD)
thus "x \<in> A" by simp
qed
lemma direct_decomp_Int_zero:
assumes "direct_decomp A ss" and "i < j" and "j < length ss" and "\<And>s. s \<in> set ss \<Longrightarrow> 0 \<in> s"
shows "ss ! i \<inter> ss ! j = {0}"
proof -
from assms(2, 3) have "i < length ss" by (rule less_trans)
hence i_in: "ss ! i \<in> set ss" by simp
from assms(3) have j_in: "ss ! j \<in> set ss" by simp
show ?thesis
proof
show "ss ! i \<inter> ss ! j \<subseteq> {0}"
proof
fix x
assume "x \<in> ss ! i \<inter> ss ! j"
hence x_i: "x \<in> ss ! i" and x_j: "x \<in> ss ! j" by simp_all
have 1: "(map (\<lambda>_. 0) ss)[k := y] \<in> listset ss" if "k < length ss" and "y \<in> ss ! k" for k y
using assms(4) that
proof (induct ss arbitrary: k)
case Nil
from Nil(2) show ?case by simp
next
case (Cons s ss)
have *: "\<And>s'. s' \<in> set ss \<Longrightarrow> 0 \<in> s'" by (rule Cons.prems) simp
show ?case
proof (cases k)
case k: 0
with Cons.prems(3) have "y \<in> s" by simp
moreover from * have "map (\<lambda>_. 0) ss \<in> listset ss"
by (induct ss) (auto simp del: listset.simps(2) intro: listset_ConsI)
moreover have "(map (\<lambda>_. 0) (s # ss))[k := y] = y # map (\<lambda>_. 0) ss" by (simp add: k)
ultimately show ?thesis by (rule listset_ConsI)
next
case k: (Suc k')
have "0 \<in> s" by (rule Cons.prems) simp
moreover from * have "(map (\<lambda>_. 0) ss)[k' := y] \<in> listset ss"
proof (rule Cons.hyps)
from Cons.prems(2) show "k' < length ss" by (simp add: k)
next
from Cons.prems(3) show "y \<in> ss ! k'" by (simp add: k)
qed
moreover have "(map (\<lambda>_. 0) (s # ss))[k := y] = 0 # (map (\<lambda>_. 0) ss)[k' := y]"
by (simp add: k)
ultimately show ?thesis by (rule listset_ConsI)
qed
qed
have 2: "sum_list ((map (\<lambda>_. 0) ss)[k := y]) = y" if "k < length ss" for k and y::'a
using that by (induct ss arbitrary: k) (auto simp: add_ac split: nat.split)
define qs1 where "qs1 = (map (\<lambda>_. 0) ss)[i := x]"
define qs2 where "qs2 = (map (\<lambda>_. 0) ss)[j := x]"
note assms(1)
moreover from \<open>i < length ss\<close> x_i have "qs1 \<in> listset ss" unfolding qs1_def by (rule 1)
moreover from assms(3) x_j have "qs2 \<in> listset ss" unfolding qs2_def by (rule 1)
thm sum_list_update
moreover from \<open>i < length ss\<close> assms(3) have "sum_list qs1 = sum_list qs2"
by (simp add: qs1_def qs2_def 2)
ultimately have "qs1 = qs2" by (rule direct_decomp_unique)
hence "qs1 ! i = qs2 ! i" by simp
with \<open>i < length ss\<close> assms(2, 3) show "x \<in> {0}" by (simp add: qs1_def qs2_def)
qed
next
from i_in have "0 \<in> ss ! i" by (rule assms(4))
moreover from j_in have "0 \<in> ss ! j" by (rule assms(4))
ultimately show "{0} \<subseteq> ss ! i \<inter> ss ! j" by simp
qed
qed
corollary direct_decomp_pairwise_zero:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> 0 \<in> s"
shows "pairwise (\<lambda>s1 s2. s1 \<inter> s2 = {0}) (set ss)"
proof (rule pairwiseI)
fix s1 s2
assume "s1 \<in> set ss"
then obtain i where "i < length ss" and s1: "s1 = ss ! i" by (metis in_set_conv_nth)
assume "s2 \<in> set ss"
then obtain j where "j < length ss" and s2: "s2 = ss ! j" by (metis in_set_conv_nth)
assume "s1 \<noteq> s2"
hence "i < j \<or> j < i" by (auto simp: s1 s2)
thus "s1 \<inter> s2 = {0}"
proof
assume "i < j"
with assms(1) show ?thesis unfolding s1 s2 using \<open>j < length ss\<close> assms(2)
by (rule direct_decomp_Int_zero)
next
assume "j < i"
with assms(1) have "s2 \<inter> s1 = {0}" unfolding s1 s2 using \<open>i < length ss\<close> assms(2)
by (rule direct_decomp_Int_zero)
thus ?thesis by (simp only: Int_commute)
qed
qed
corollary direct_decomp_repeated_eq_zero:
assumes "direct_decomp A ss" and "1 < count_list ss X" and "\<And>s. s \<in> set ss \<Longrightarrow> 0 \<in> s"
shows "X = {0}"
proof -
from assms(2) obtain i j where "i < j" and "j < length ss" and 1: "ss ! i = X" and 2: "ss ! j = X"
by (rule count_list_gr_1_E)
from assms(1) this(1, 2) assms(3) have "ss ! i \<inter> ss ! j = {0}" by (rule direct_decomp_Int_zero)
thus ?thesis by (simp add: 1 2)
qed
corollary direct_decomp_map_Int_zero:
assumes "direct_decomp A (map f ss)" and "s1 \<in> set ss" and "s2 \<in> set ss" and "s1 \<noteq> s2"
and "\<And>s. s \<in> set ss \<Longrightarrow> 0 \<in> f s"
shows "f s1 \<inter> f s2 = {0}"
proof -
from assms(2) obtain i where "i < length ss" and s1: "s1 = ss ! i" by (metis in_set_conv_nth)
from this(1) have i: "i < length (map f ss)" by simp
from assms(3) obtain j where "j < length ss" and s2: "s2 = ss ! j" by (metis in_set_conv_nth)
from this(1) have j: "j < length (map f ss)" by simp
have *: "0 \<in> s" if "s \<in> set (map f ss)" for s
proof -
from that obtain s' where "s' \<in> set ss" and s: "s = f s'" unfolding set_map ..
from this(1) show "0 \<in> s" unfolding s by (rule assms(5))
qed
show ?thesis
proof (rule linorder_cases)
assume "i < j"
with assms(1) have "(map f ss) ! i \<inter> (map f ss) ! j = {0}"
using j * by (rule direct_decomp_Int_zero)
with i j show ?thesis by (simp add: s1 s2)
next
assume "j < i"
with assms(1) have "(map f ss) ! j \<inter> (map f ss) ! i = {0}"
using i * by (rule direct_decomp_Int_zero)
with i j show ?thesis by (simp add: s1 s2 Int_commute)
next
assume "i = j"
with assms(4) show ?thesis by (simp add: s1 s2)
qed
qed
subsection \<open>Direct Decompositions and Vector Spaces\<close>
definition (in vector_space) is_basis :: "'b set \<Rightarrow> 'b set \<Rightarrow> bool"
where "is_basis V B \<longleftrightarrow> (B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> card B = dim V)"
definition (in vector_space) some_basis :: "'b set \<Rightarrow> 'b set"
where "some_basis V = Eps (local.is_basis V)"
hide_const (open) real_vector.is_basis real_vector.some_basis
context vector_space
begin
lemma dim_empty [simp]: "dim {} = 0"
using dim_span_eq_card_independent independent_empty by fastforce
lemma dim_zero [simp]: "dim {0} = 0"
using dim_span_eq_card_independent independent_empty by fastforce
lemma independent_UnI:
assumes "independent A" and "independent B" and "span A \<inter> span B = {0}"
shows "independent (A \<union> B)"
proof
from span_superset have "A \<inter> B \<subseteq> span A \<inter> span B" by blast
hence "A \<inter> B = {}" unfolding assms(3) using assms(1, 2) dependent_zero by blast
assume "dependent (A \<union> B)"
then obtain T u v where "finite T" and "T \<subseteq> A \<union> B" and eq: "(\<Sum>v\<in>T. u v *s v) = 0"
and "v \<in> T" and "u v \<noteq> 0" unfolding dependent_explicit by blast
define TA where "TA = T \<inter> A"
define TB where "TB = T \<inter> B"
from \<open>T \<subseteq> A \<union> B\<close> have T: "T = TA \<union> TB" by (auto simp: TA_def TB_def)
from \<open>finite T\<close> have "finite TA" and "TA \<subseteq> A" by (simp_all add: TA_def)
from \<open>finite T\<close> have "finite TB" and "TB \<subseteq> B" by (simp_all add: TB_def)
from \<open>A \<inter> B = {}\<close> \<open>TA \<subseteq> A\<close> this(2) have "TA \<inter> TB = {}" by blast
have "0 = (\<Sum>v\<in>TA \<union> TB. u v *s v)" by (simp only: eq flip: T)
also have "\<dots> = (\<Sum>v\<in>TA. u v *s v) + (\<Sum>v\<in>TB. u v *s v)" by (rule sum.union_disjoint) fact+
finally have "(\<Sum>v\<in>TA. u v *s v) = (\<Sum>v\<in>TB. (- u) v *s v)" (is "?x = ?y")
by (simp add: sum_negf eq_neg_iff_add_eq_0)
from \<open>finite TB\<close> \<open>TB \<subseteq> B\<close> have "?y \<in> span B" by (auto simp: span_explicit simp del: uminus_apply)
moreover from \<open>finite TA\<close> \<open>TA \<subseteq> A\<close> have "?x \<in> span A" by (auto simp: span_explicit)
ultimately have "?y \<in> span A \<inter> span B" by (simp add: \<open>?x = ?y\<close>)
hence "?x = 0" and "?y = 0" by (simp_all add: \<open>?x = ?y\<close> assms(3))
from \<open>v \<in> T\<close> have "v \<in> TA \<union> TB" by (simp only: T)
hence "u v = 0"
proof
assume "v \<in> TA"
with assms(1) \<open>finite TA\<close> \<open>TA \<subseteq> A\<close> \<open>?x = 0\<close> show "u v = 0" by (rule independentD)
next
assume "v \<in> TB"
with assms(2) \<open>finite TB\<close> \<open>TB \<subseteq> B\<close> \<open>?y = 0\<close> have "(- u) v = 0" by (rule independentD)
thus "u v = 0" by simp
qed
with \<open>u v \<noteq> 0\<close> show False ..
qed
lemma subspace_direct_decomp:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> subspace s"
shows "subspace A"
proof (rule subspaceI)
let ?qs = "map (\<lambda>_. 0) ss"
from assms(2) have "?qs \<in> listset ss"
by (induct ss) (auto simp del: listset.simps(2) dest: subspace_0 intro: listset_ConsI)
with assms(1) have "sum_list ?qs \<in> A" by (rule direct_decompD)
thus "0 \<in> A" by simp
next
fix p q
assume "p \<in> A"
with assms(1) obtain ps where ps: "ps \<in> listset ss" and p: "p = sum_list ps" by (rule direct_decompE)
assume "q \<in> A"
with assms(1) obtain qs where qs: "qs \<in> listset ss" and q: "q = sum_list qs" by (rule direct_decompE)
from ps qs have l: "length ps = length qs" by (simp only: listsetD)
from ps qs have "map2 (+) ps qs \<in> listset ss" (is "?qs \<in> _")
by (rule listset_closed_map2) (auto dest: assms(2) subspace_add)
with assms(1) have "sum_list ?qs \<in> A" by (rule direct_decompD)
thus "p + q \<in> A" using l by (simp only: p q sum_list_map2_plus)
next
fix c p
assume "p \<in> A"
with assms(1) obtain ps where "ps \<in> listset ss" and p: "p = sum_list ps" by (rule direct_decompE)
from this(1) have "map ((*s) c) ps \<in> listset ss" (is "?qs \<in> _")
by (rule listset_closed_map) (auto dest: assms(2) subspace_scale)
with assms(1) have "sum_list ?qs \<in> A" by (rule direct_decompD)
also have "sum_list ?qs = c *s sum_list ps" by (induct ps) (simp_all add: scale_right_distrib)
finally show "c *s p \<in> A" by (simp only: p)
qed
lemma is_basis_alt: "subspace V \<Longrightarrow> is_basis V B \<longleftrightarrow> (independent B \<and> span B = V)"
by (metis (full_types) is_basis_def dim_eq_card span_eq span_eq_iff)
lemma is_basis_finite: "is_basis V A \<Longrightarrow> is_basis V B \<Longrightarrow> finite A \<longleftrightarrow> finite B"
unfolding is_basis_def using independent_span_bound by auto
lemma some_basis_is_basis: "is_basis V (some_basis V)"
proof -
obtain B where "B \<subseteq> V" and "independent B" and "V \<subseteq> span B" and "card B = dim V"
by (rule basis_exists)
hence "is_basis V B" by (simp add: is_basis_def)
thus ?thesis unfolding some_basis_def by (rule someI)
qed
corollary
shows some_basis_subset: "some_basis V \<subseteq> V"
and independent_some_basis: "independent (some_basis V)"
and span_some_basis_supset: "V \<subseteq> span (some_basis V)"
and card_some_basis: "card (some_basis V) = dim V"
using some_basis_is_basis[of V] by (simp_all add: is_basis_def)
lemma some_basis_not_zero: "0 \<notin> some_basis V"
using independent_some_basis dependent_zero by blast
lemma span_some_basis: "subspace V \<Longrightarrow> span (some_basis V) = V"
by (simp add: span_subspace some_basis_subset span_some_basis_supset)
lemma direct_decomp_some_basis_pairwise_disjnt:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> subspace s"
shows "pairwise (\<lambda>s1 s2. disjnt (some_basis s1) (some_basis s2)) (set ss)"
proof (rule pairwiseI)
fix s1 s2
assume "s1 \<in> set ss" and "s2 \<in> set ss" and "s1 \<noteq> s2"
have "some_basis s1 \<inter> some_basis s2 \<subseteq> s1 \<inter> s2" using some_basis_subset by blast
also from direct_decomp_pairwise_zero have "\<dots> = {0}"
proof (rule pairwiseD)
fix s
assume "s \<in> set ss"
hence "subspace s" by (rule assms(2))
thus "0 \<in> s" by (rule subspace_0)
qed fact+
finally have "some_basis s1 \<inter> some_basis s2 \<subseteq> {0}" .
with some_basis_not_zero show "disjnt (some_basis s1) (some_basis s2)"
unfolding disjnt_def by blast
qed
lemma direct_decomp_span_some_basis:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> subspace s"
shows "span (\<Union>(some_basis ` set ss)) = A"
proof -
from assms(1) have eq0[symmetric]: "sum_list ` listset ss = A" by (rule direct_decompD)
show ?thesis unfolding eq0 using assms(2)
proof (induct ss)
case Nil
show ?case by simp
next
case (Cons s ss)
have "subspace s" by (rule Cons.prems) simp
hence eq1: "span (some_basis s) = s" by (rule span_some_basis)
have "\<And>s'. s' \<in> set ss \<Longrightarrow> subspace s'" by (rule Cons.prems) simp
hence eq2: "span (\<Union> (some_basis ` set ss)) = sum_list ` listset ss" by (rule Cons.hyps)
have "span (\<Union> (some_basis ` set (s # ss))) = {x + y |x y. x \<in> s \<and> y \<in> sum_list ` listset ss}"
by (simp add: span_Un eq1 eq2)
also have "\<dots> = sum_list ` listset (s # ss)" (is "?A = ?B")
proof
show "?A \<subseteq> ?B"
proof
fix a
assume "a \<in> ?A"
then obtain x y where "x \<in> s" and "y \<in> sum_list ` listset ss" and a: "a = x + y" by blast
from this(2) obtain qs where "qs \<in> listset ss" and y: "y = sum_list qs" ..
from \<open>x \<in> s\<close> this(1) refl have "x # qs \<in> listset (s # ss)" by (rule listset_ConsI)
hence "sum_list (x # qs) \<in> ?B" by (rule imageI)
also have "sum_list (x # qs) = a" by (simp add: a y)
finally show "a \<in> ?B" .
qed
next
show "?B \<subseteq> ?A"
proof
fix a
assume "a \<in> ?B"
then obtain qs' where "qs' \<in> listset (s # ss)" and a: "a = sum_list qs'" ..
from this(1) obtain x qs where "x \<in> s" and "qs \<in> listset ss" and qs': "qs' = x # qs"
by (rule listset_ConsE)
from this(2) have "sum_list qs \<in> sum_list ` listset ss" by (rule imageI)
moreover have "a = x + sum_list qs" by (simp add: a qs')
ultimately show "a \<in> ?A" using \<open>x \<in> s\<close> by blast
qed
qed
finally show ?case .
qed
qed
lemma direct_decomp_independent_some_basis:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> subspace s"
shows "independent (\<Union>(some_basis ` set ss))"
using assms
proof (induct ss arbitrary: A)
case Nil
from independent_empty show ?case by simp
next
case (Cons s ss)
have 1: "\<And>s'. s' \<in> set ss \<Longrightarrow> subspace s'" by (rule Cons.prems) simp
have "subspace s" by (rule Cons.prems) simp
hence "0 \<in> s" and eq1: "span (some_basis s) = s" by (rule subspace_0, rule span_some_basis)
from Cons.prems(1) have *: "direct_decomp A ([s] @ ss)" by simp
moreover from \<open>0 \<in> s\<close> have "{} \<notin> set [s]" by auto
ultimately have 2: "direct_decomp (sum_list ` listset ss) ss" by (rule direct_decomp_appendD)
hence eq2: "span (\<Union> (some_basis ` set ss)) = sum_list ` listset ss" using 1
by (rule direct_decomp_span_some_basis)
note independent_some_basis[of s]
moreover from 2 1 have "independent (\<Union> (some_basis ` set ss))" by (rule Cons.hyps)
moreover have "span (some_basis s) \<inter> span (\<Union> (some_basis ` set ss)) = {0}"
proof -
from * have "direct_decomp A [sum_list ` listset [s], sum_list ` listset ss]"
by (rule direct_decomp_appendD)
hence "direct_decomp A [s, sum_list ` listset ss]" by (simp add: image_image)
moreover have "0 < (1::nat)" by simp
moreover have "1 < length [s, sum_list ` listset ss]" by simp
ultimately have "[s, sum_list ` listset ss] ! 0 \<inter> [s, sum_list ` listset ss] ! 1 = {0}"
by (rule direct_decomp_Int_zero) (auto simp: \<open>0 \<in> s\<close> eq2[symmetric] span_zero)
thus ?thesis by (simp add: eq1 eq2)
qed
ultimately have "independent (some_basis s \<union> (\<Union> (some_basis ` set ss)))"
by (rule independent_UnI)
thus ?case by simp
qed
corollary direct_decomp_is_basis:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> subspace s"
shows "is_basis A (\<Union>(some_basis ` set ss))"
proof -
from assms have "subspace A" by (rule subspace_direct_decomp)
moreover from assms have "span (\<Union>(some_basis ` set ss)) = A"
by (rule direct_decomp_span_some_basis)
moreover from assms have "independent (\<Union>(some_basis ` set ss))"
by (rule direct_decomp_independent_some_basis)
ultimately show ?thesis by (simp add: is_basis_alt)
qed
lemma dim_direct_decomp:
assumes "direct_decomp A ss" and "finite B" and "A \<subseteq> span B" and "\<And>s. s \<in> set ss \<Longrightarrow> subspace s"
shows "dim A = (\<Sum>s\<in>set ss. dim s)"
proof -
from assms(1, 4) have "is_basis A (\<Union>(some_basis ` set ss))"
(is "is_basis A ?B") by (rule direct_decomp_is_basis)
hence "dim A = card ?B" and "independent ?B" and "?B \<subseteq> A" by (simp_all add: is_basis_def)
from this(3) assms(3) have "?B \<subseteq> span B" by (rule subset_trans)
with assms(2) \<open>independent ?B\<close> have "finite ?B" using independent_span_bound by blast
note \<open>dim A = card ?B\<close>
also from finite_set have "card ?B = (\<Sum>s\<in>set ss. card (some_basis s))"
proof (intro card_UN_disjoint ballI impI)
fix s
assume "s \<in> set ss"
with \<open>finite ?B\<close> show "finite (some_basis s)" by auto
next
fix s1 s2
have "pairwise (\<lambda>s t. disjnt (some_basis s) (some_basis t)) (set ss)"
using assms(1, 4) by (rule direct_decomp_some_basis_pairwise_disjnt)
moreover assume "s1 \<in> set ss" and "s2 \<in> set ss" and "s1 \<noteq> s2"
thm pairwiseD
ultimately have "disjnt (some_basis s1) (some_basis s2)" by (rule pairwiseD)
thus "some_basis s1 \<inter> some_basis s2 = {}" by (simp only: disjnt_def)
qed
also from refl card_some_basis have "\<dots> = (\<Sum>s\<in>set ss. dim s)" by (rule sum.cong)
finally show ?thesis .
qed
end (* vector_space *)
subsection \<open>Homogeneous Sets of Polynomials with Fixed Degree\<close>
lemma homogeneous_set_direct_decomp:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> homogeneous_set s"
shows "homogeneous_set A"
proof (rule homogeneous_setI)
fix a n
assume "a \<in> A"
with assms(1) obtain qs where "qs \<in> listset ss" and a: "a = sum_list qs" by (rule direct_decompE)
have "hom_component a n = hom_component (sum_list qs) n" by (simp only: a)
also have "\<dots> = sum_list (map (\<lambda>q. hom_component q n) qs)"
by (induct qs) (simp_all add: hom_component_plus)
also from assms(1) have "\<dots> \<in> A"
proof (rule direct_decompD)
show "map (\<lambda>q. hom_component q n) qs \<in> listset ss"
proof (rule listset_closed_map)
fix s q
assume "s \<in> set ss"
hence "homogeneous_set s" by (rule assms(2))
moreover assume "q \<in> s"
ultimately show "hom_component q n \<in> s" by (rule homogeneous_setD)
qed fact
qed
finally show "hom_component a n \<in> A" .
qed
definition hom_deg_set :: "nat \<Rightarrow> (('x \<Rightarrow>\<^sub>0 nat) \<Rightarrow>\<^sub>0 'a) set \<Rightarrow> (('x \<Rightarrow>\<^sub>0 nat) \<Rightarrow>\<^sub>0 'a::zero) set"
where "hom_deg_set z A = (\<lambda>a. hom_component a z) ` A"
lemma hom_deg_setD:
assumes "p \<in> hom_deg_set z A"
shows "homogeneous p" and "p \<noteq> 0 \<Longrightarrow> poly_deg p = z"
proof -
from assms obtain a where "a \<in> A" and p: "p = hom_component a z" unfolding hom_deg_set_def ..
show *: "homogeneous p" by (simp only: p homogeneous_hom_component)
assume "p \<noteq> 0"
hence "keys p \<noteq> {}" by simp
then obtain t where "t \<in> keys p" by blast
with * have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg)
moreover from \<open>t \<in> keys p\<close> have "deg_pm t = z" unfolding p by (rule keys_hom_componentD)
ultimately show "poly_deg p = z" by simp
qed
lemma zero_in_hom_deg_set:
assumes "0 \<in> A"
shows "0 \<in> hom_deg_set z A"
proof -
have "0 = hom_component 0 z" by simp
also from assms have "\<dots> \<in> hom_deg_set z A" unfolding hom_deg_set_def by (rule imageI)
finally show ?thesis .
qed
lemma hom_deg_set_closed_uminus:
assumes "\<And>a. a \<in> A \<Longrightarrow> - a \<in> A" and "p \<in> hom_deg_set z A"
shows "- p \<in> hom_deg_set z A"
proof -
from assms(2) obtain a where "a \<in> A" and p: "p = hom_component a z" unfolding hom_deg_set_def ..
from this(1) have "- a \<in> A" by (rule assms(1))
moreover have "- p = hom_component (- a) z" by (simp add: p)
ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI)
qed
lemma hom_deg_set_closed_plus:
assumes "\<And>a1 a2. a1 \<in> A \<Longrightarrow> a2 \<in> A \<Longrightarrow> a1 + a2 \<in> A"
and "p \<in> hom_deg_set z A" and "q \<in> hom_deg_set z A"
shows "p + q \<in> hom_deg_set z A"
proof -
from assms(2) obtain a1 where "a1 \<in> A" and p: "p = hom_component a1 z" unfolding hom_deg_set_def ..
from assms(3) obtain a2 where "a2 \<in> A" and q: "q = hom_component a2 z" unfolding hom_deg_set_def ..
from \<open>a1 \<in> A\<close> this(1) have "a1 + a2 \<in> A" by (rule assms(1))
moreover have "p + q = hom_component (a1 + a2) z" by (simp only: p q hom_component_plus)
ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI)
qed
lemma hom_deg_set_closed_minus:
assumes "\<And>a1 a2. a1 \<in> A \<Longrightarrow> a2 \<in> A \<Longrightarrow> a1 - a2 \<in> A"
and "p \<in> hom_deg_set z A" and "q \<in> hom_deg_set z A"
shows "p - q \<in> hom_deg_set z A"
proof -
from assms(2) obtain a1 where "a1 \<in> A" and p: "p = hom_component a1 z" unfolding hom_deg_set_def ..
from assms(3) obtain a2 where "a2 \<in> A" and q: "q = hom_component a2 z" unfolding hom_deg_set_def ..
from \<open>a1 \<in> A\<close> this(1) have "a1 - a2 \<in> A" by (rule assms(1))
moreover have "p - q = hom_component (a1 - a2) z" by (simp only: p q hom_component_minus)
ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI)
qed
lemma hom_deg_set_closed_scalar:
assumes "\<And>a. a \<in> A \<Longrightarrow> c \<cdot> a \<in> A" and "p \<in> hom_deg_set z A"
shows "(c::'a::semiring_0) \<cdot> p \<in> hom_deg_set z A"
proof -
from assms(2) obtain a where "a \<in> A" and p: "p = hom_component a z" unfolding hom_deg_set_def ..
from this(1) have "c \<cdot> a \<in> A" by (rule assms(1))
moreover have "c \<cdot> p = hom_component (c \<cdot> a) z"
by (simp add: p punit.map_scale_eq_monom_mult hom_component_monom_mult)
ultimately show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI)
qed
lemma hom_deg_set_closed_sum:
assumes "0 \<in> A" and "\<And>a1 a2. a1 \<in> A \<Longrightarrow> a2 \<in> A \<Longrightarrow> a1 + a2 \<in> A"
and "\<And>i. i \<in> I \<Longrightarrow> f i \<in> hom_deg_set z A"
shows "sum f I \<in> hom_deg_set z A"
using assms(3)
proof (induct I rule: infinite_finite_induct)
case (infinite I)
with assms(1) show ?case by (simp add: zero_in_hom_deg_set)
next
case empty
with assms(1) show ?case by (simp add: zero_in_hom_deg_set)
next
case (insert j I)
from insert.hyps(1, 2) have "sum f (insert j I) = f j + sum f I" by simp
also from assms(2) have "\<dots> \<in> hom_deg_set z A"
proof (intro hom_deg_set_closed_plus insert.hyps)
show "f j \<in> hom_deg_set z A" by (rule insert.prems) simp
next
fix i
assume "i \<in> I"
hence "i \<in> insert j I" by simp
thus "f i \<in> hom_deg_set z A" by (rule insert.prems)
qed
finally show ?case .
qed
lemma hom_deg_set_subset: "homogeneous_set A \<Longrightarrow> hom_deg_set z A \<subseteq> A"
by (auto dest: homogeneous_setD simp: hom_deg_set_def)
lemma Polys_closed_hom_deg_set:
assumes "A \<subseteq> P[X]"
shows "hom_deg_set z A \<subseteq> P[X]"
proof
fix p
assume "p \<in> hom_deg_set z A"
then obtain p' where "p' \<in> A" and p: "p = hom_component p' z" unfolding hom_deg_set_def ..
from this(1) assms have "p' \<in> P[X]" ..
have "keys p \<subseteq> keys p'" by (simp add: p keys_hom_component)
also from \<open>p' \<in> P[X]\<close> have "\<dots> \<subseteq> .[X]" by (rule PolysD)
finally show "p \<in> P[X]" by (rule PolysI)
qed
lemma hom_deg_set_alt_homogeneous_set:
assumes "homogeneous_set A"
shows "hom_deg_set z A = {p \<in> A. homogeneous p \<and> (p = 0 \<or> poly_deg p = z)}" (is "?A = ?B")
proof
show "?A \<subseteq> ?B"
proof
fix h
assume "h \<in> ?A"
also from assms have "\<dots> \<subseteq> A" by (rule hom_deg_set_subset)
finally show "h \<in> ?B" using \<open>h \<in> ?A\<close> by (auto dest: hom_deg_setD)
qed
next
show "?B \<subseteq> ?A"
proof
fix h
assume "h \<in> ?B"
hence "h \<in> A" and "homogeneous h" and "h = 0 \<or> poly_deg h = z" by simp_all
from this(3) show "h \<in> ?A"
proof
assume "h = 0"
with \<open>h \<in> A\<close> have "0 \<in> A" by simp
thus ?thesis unfolding \<open>h = 0\<close> by (rule zero_in_hom_deg_set)
next
assume "poly_deg h = z"
with \<open>homogeneous h\<close> have "h = hom_component h z" by (simp add: hom_component_of_homogeneous)
with \<open>h \<in> A\<close> show ?thesis unfolding hom_deg_set_def by (rule rev_image_eqI)
qed
qed
qed
lemma hom_deg_set_sum_list_listset:
assumes "A = sum_list ` listset ss"
shows "hom_deg_set z A = sum_list ` listset (map (hom_deg_set z) ss)" (is "?A = ?B")
proof
show "?A \<subseteq> ?B"
proof
fix h
assume "h \<in> ?A"
then obtain a where "a \<in> A" and h: "h = hom_component a z" unfolding hom_deg_set_def ..
from this(1) obtain qs where "qs \<in> listset ss" and a: "a = sum_list qs" unfolding assms ..
have "h = hom_component (sum_list qs) z" by (simp only: a h)
also have "\<dots> = sum_list (map (\<lambda>q. hom_component q z) qs)"
by (induct qs) (simp_all add: hom_component_plus)
also have "\<dots> \<in> ?B"
proof (rule imageI)
show "map (\<lambda>q. hom_component q z) qs \<in> listset (map (hom_deg_set z) ss)"
unfolding hom_deg_set_def using \<open>qs \<in> listset ss\<close> refl by (rule listset_map_imageI)
qed
finally show "h \<in> ?B" .
qed
next
show "?B \<subseteq> ?A"
proof
fix h
assume "h \<in> ?B"
then obtain qs where "qs \<in> listset (map (hom_deg_set z) ss)" and h: "h = sum_list qs" ..
from this(1) obtain qs' where "qs' \<in> listset ss" and qs: "qs = map (\<lambda>q. hom_component q z) qs'"
unfolding hom_deg_set_def by (rule listset_map_imageE)
have "h = sum_list (map (\<lambda>q. hom_component q z) qs')" by (simp only: h qs)
also have "\<dots> = hom_component (sum_list qs') z" by (induct qs') (simp_all add: hom_component_plus)
finally have "h = hom_component (sum_list qs') z" .
moreover have "sum_list qs' \<in> A" unfolding assms using \<open>qs' \<in> listset ss\<close> by (rule imageI)
ultimately show "h \<in> ?A" unfolding hom_deg_set_def by (rule image_eqI)
qed
qed
lemma direct_decomp_hom_deg_set:
assumes "direct_decomp A ss" and "\<And>s. s \<in> set ss \<Longrightarrow> homogeneous_set s"
shows "direct_decomp (hom_deg_set z A) (map (hom_deg_set z) ss)"
proof (rule direct_decompI)
from assms(1) have "sum_list ` listset ss = A" by (rule direct_decompD)
from this[symmetric] show "sum_list ` listset (map (hom_deg_set z) ss) = hom_deg_set z A"
by (simp only: hom_deg_set_sum_list_listset)
next
from assms(1) have "inj_on sum_list (listset ss)" by (rule direct_decompD)
moreover have "listset (map (hom_deg_set z) ss) \<subseteq> listset ss"
proof (rule listset_mono)
fix i
assume "i < length ss"
hence "map (hom_deg_set z) ss ! i = hom_deg_set z (ss ! i)" by simp
also from \<open>i < length ss\<close> have "\<dots> \<subseteq> ss ! i" by (intro hom_deg_set_subset assms(2) nth_mem)
finally show "map (hom_deg_set z) ss ! i \<subseteq> ss ! i" .
qed simp
ultimately show "inj_on sum_list (listset (map (hom_deg_set z) ss))" by (rule inj_on_subset)
qed
subsection \<open>Interpreting Polynomial Rings as Vector Spaces over the Coefficient Field\<close>
text \<open>There is no need to set up any further interpretation, since interpretation \<open>phull\<close> is exactly
what we need.\<close>
lemma subspace_ideal: "phull.subspace (ideal (F::('b::comm_powerprod \<Rightarrow>\<^sub>0 'a::field) set))"
using ideal.span_zero ideal.span_add
proof (rule phull.subspaceI)
fix c p
assume "p \<in> ideal F"
thus "c \<cdot> p \<in> ideal F" unfolding map_scale_eq_times by (rule ideal.span_scale)
qed
lemma subspace_Polys: "phull.subspace (P[X]::(('x \<Rightarrow>\<^sub>0 nat) \<Rightarrow>\<^sub>0 'a::field) set)"
using zero_in_Polys Polys_closed_plus Polys_closed_map_scale by (rule phull.subspaceI)
lemma subspace_hom_deg_set:
assumes "phull.subspace A"
shows "phull.subspace (hom_deg_set z A)" (is "phull.subspace ?A")
proof (rule phull.subspaceI)
from assms have "0 \<in> A" by (rule phull.subspace_0)
thus "0 \<in> ?A" by (rule zero_in_hom_deg_set)
next
fix p q
assume "p \<in> ?A" and "q \<in> ?A"
with phull.subspace_add show "p + q \<in> ?A" by (rule hom_deg_set_closed_plus) (rule assms)
next
fix c p
assume "p \<in> ?A"
with phull.subspace_scale show "c \<cdot> p \<in> ?A" by (rule hom_deg_set_closed_scalar) (rule assms)
qed
lemma hom_deg_set_Polys_eq_span:
"hom_deg_set z P[X] = phull.span (monomial (1::'a::field) ` deg_sect X z)" (is "?A = ?B")
proof
show "?A \<subseteq> ?B"
proof
fix p
assume "p \<in> ?A"
also from this have "\<dots> = {p \<in> P[X]. homogeneous p \<and> (p = 0 \<or> poly_deg p = z)}"
by (simp only: hom_deg_set_alt_homogeneous_set[OF homogeneous_set_Polys])
finally have "p \<in> P[X]" and "homogeneous p" and "p \<noteq> 0 \<Longrightarrow> poly_deg p = z" by simp_all
thus "p \<in> ?B"
proof (induct p rule: poly_mapping_plus_induct)
case 1
from phull.span_zero show ?case .
next
case (2 p c t)
let ?m = "monomial c t"
from 2(1) have "t \<in> keys ?m" by simp
hence "t \<in> keys (?m + p)" using 2(2) by (rule in_keys_plusI1)
hence "?m + p \<noteq> 0" by auto
hence "poly_deg (monomial c t + p) = z" by (rule 2)
from 2(4) have "keys (?m + p) \<subseteq> .[X]" by (rule PolysD)
with \<open>t \<in> keys (?m + p)\<close> have "t \<in> .[X]" ..
hence "?m \<in> P[X]" by (rule Polys_closed_monomial)
have "t \<in> deg_sect X z"
proof (rule deg_sectI)
from 2(5) \<open>t \<in> keys (?m + p)\<close> have "deg_pm t = poly_deg (?m + p)"
by (rule homogeneousD_poly_deg)
also have "\<dots> = z" by fact
finally show "deg_pm t = z" .
qed fact
hence "monomial 1 t \<in> monomial 1 ` deg_sect X z" by (rule imageI)
hence "monomial 1 t \<in> ?B" by (rule phull.span_base)
hence "c \<cdot> monomial 1 t \<in> ?B" by (rule phull.span_scale)
hence "?m \<in> ?B" by simp
moreover have "p \<in> ?B"
proof (rule 2)
from 2(4) \<open>?m \<in> P[X]\<close> have "(?m + p) - ?m \<in> P[X]" by (rule Polys_closed_minus)
thus "p \<in> P[X]" by simp
next
have 1: "deg_pm s = z" if "s \<in> keys p" for s
proof -
from that 2(2) have "s \<noteq> t" by blast
hence "s \<notin> keys ?m" by simp
with that have "s \<in> keys (?m + p)" by (rule in_keys_plusI2)
with 2(5) have "deg_pm s = poly_deg (?m + p)" by (rule homogeneousD_poly_deg)
also have "\<dots> = z" by fact
finally show ?thesis .
qed
show "homogeneous p" by (rule homogeneousI) (simp add: 1)
assume "p \<noteq> 0"
show "poly_deg p = z"
proof (rule antisym)
show "poly_deg p \<le> z" by (rule poly_deg_leI) (simp add: 1)
next
from \<open>p \<noteq> 0\<close> have "keys p \<noteq> {}" by simp
then obtain s where "s \<in> keys p" by blast
hence "z = deg_pm s" by (simp only: 1)
also from \<open>s \<in> keys p\<close> have "\<dots> \<le> poly_deg p" by (rule poly_deg_max_keys)
finally show "z \<le> poly_deg p" .
qed
qed
ultimately show ?case by (rule phull.span_add)
qed
qed
next
show "?B \<subseteq> ?A"
proof
fix p
assume "p \<in> ?B"
then obtain M u where "M \<subseteq> monomial 1 ` deg_sect X z" and "finite M" and p: "p = (\<Sum>m\<in>M. u m \<cdot> m)"
by (auto simp: phull.span_explicit)
from this(1) obtain T where "T \<subseteq> deg_sect X z" and M: "M = monomial 1 ` T"
and inj: "inj_on (monomial (1::'a)) T" by (rule subset_imageE_inj)
define c where "c = (\<lambda>t. u (monomial 1 t))"
from inj have "p = (\<Sum>t\<in>T. monomial (c t) t)" by (simp add: p M sum.reindex c_def)
also have "\<dots> \<in> ?A"
proof (intro hom_deg_set_closed_sum zero_in_Polys Polys_closed_plus)
fix t
assume "t \<in> T"
hence "t \<in> deg_sect X z" using \<open>T \<subseteq> deg_sect X z\<close> ..
hence "t \<in> .[X]" and eq: "deg_pm t = z" by (rule deg_sectD)+
from this(1) have "monomial (c t) t \<in> P[X]" (is "?m \<in> _") by (rule Polys_closed_monomial)
thus "?m \<in> ?A"
by (simp add: hom_deg_set_alt_homogeneous_set[OF homogeneous_set_Polys] poly_deg_monomial
monomial_0_iff eq)
qed
finally show "p \<in> ?A" .
qed
qed
subsection \<open>(Projective) Hilbert Function\<close>
interpretation phull: vector_space map_scale
apply standard
subgoal by (fact map_scale_distrib_left)
subgoal by (fact map_scale_distrib_right)
subgoal by (fact map_scale_assoc)
subgoal by (fact map_scale_one_left)
done
definition Hilbert_fun :: "(('x \<Rightarrow>\<^sub>0 nat) \<Rightarrow>\<^sub>0 'a::field) set \<Rightarrow> nat \<Rightarrow> nat"
where "Hilbert_fun A z = phull.dim (hom_deg_set z A)"
lemma Hilbert_fun_empty [simp]: "Hilbert_fun {} = 0"
by (rule ext) (simp add: Hilbert_fun_def hom_deg_set_def)
lemma Hilbert_fun_zero [simp]: "Hilbert_fun {0} = 0"
by (rule ext) (simp add: Hilbert_fun_def hom_deg_set_def)
lemma Hilbert_fun_direct_decomp:
assumes "finite X" and "A \<subseteq> P[X]" and "direct_decomp (A::(('x::countable \<Rightarrow>\<^sub>0 nat) \<Rightarrow>\<^sub>0 'a::field) set) ps"
and "\<And>s. s \<in> set ps \<Longrightarrow> homogeneous_set s" and "\<And>s. s \<in> set ps \<Longrightarrow> phull.subspace s"
shows "Hilbert_fun A z = (\<Sum>p\<in>set ps. Hilbert_fun p z)"
proof -
from assms(3, 4) have dd: "direct_decomp (hom_deg_set z A) (map (hom_deg_set z) ps)"
by (rule direct_decomp_hom_deg_set)
have "Hilbert_fun A z = phull.dim (hom_deg_set z A)" by (fact Hilbert_fun_def)
also from dd have "\<dots> = sum phull.dim (set (map (hom_deg_set z) ps))"
proof (rule phull.dim_direct_decomp)
from assms(1) have "finite (deg_sect X z)" by (rule finite_deg_sect)
thus "finite (monomial (1::'a) ` deg_sect X z)" by (rule finite_imageI)
next
from assms(2) have "hom_deg_set z A \<subseteq> hom_deg_set z P[X]"
unfolding hom_deg_set_def by (rule image_mono)
thus "hom_deg_set z A \<subseteq> phull.span (monomial 1 ` deg_sect X z)"
by (simp only: hom_deg_set_Polys_eq_span)
next
fix s
assume "s \<in> set (map (hom_deg_set z) ps)"
then obtain s' where "s' \<in> set ps" and s: "s = hom_deg_set z s'" unfolding set_map ..
from this(1) have "phull.subspace s'" by (rule assms(5))
thus "phull.subspace s" unfolding s by (rule subspace_hom_deg_set)
qed
also have "\<dots> = sum (phull.dim \<circ> hom_deg_set z) (set ps)" unfolding set_map using finite_set
proof (rule sum.reindex_nontrivial)
fix s1 s2
note dd
moreover assume "s1 \<in> set ps" and "s2 \<in> set ps" and "s1 \<noteq> s2"
moreover have "0 \<in> hom_deg_set z s" if "s \<in> set ps" for s
proof (rule zero_in_hom_deg_set)
from that have "phull.subspace s" by (rule assms(5))
thus "0 \<in> s" by (rule phull.subspace_0)
qed
ultimately have "hom_deg_set z s1 \<inter> hom_deg_set z s2 = {0}" by (rule direct_decomp_map_Int_zero)
moreover assume "hom_deg_set z s1 = hom_deg_set z s2"
ultimately show "phull.dim (hom_deg_set z s1) = 0" by simp
qed
also have "\<dots> = (\<Sum>p\<in>set ps. Hilbert_fun p z)" by (simp only: o_def Hilbert_fun_def)
finally show ?thesis .
qed
context pm_powerprod
begin
lemma image_lt_hom_deg_set:
assumes "homogeneous_set A"
shows "lpp ` (hom_deg_set z A - {0}) = {t \<in> lpp ` (A - {0}). deg_pm t = z}" (is "?B = ?A")
proof (intro set_eqI iffI)
fix t
assume "t \<in> ?A"
hence "t \<in> lpp ` (A - {0})" and deg_t[symmetric]: "deg_pm t = z" by simp_all
from this(1) obtain p where "p \<in> A - {0}" and t: "t = lpp p" ..
from this(1) have "p \<in> A" and "p \<noteq> 0" by simp_all
from this(1) have 1: "hom_component p z \<in> hom_deg_set z A" (is "?p \<in> _")
unfolding hom_deg_set_def by (rule imageI)
from \<open>p \<noteq> 0\<close> have "?p \<noteq> 0" and "lpp ?p = t" unfolding t deg_t by (rule hom_component_lpp)+
note this(2)[symmetric]
moreover from 1 \<open>?p \<noteq> 0\<close> have "?p \<in> hom_deg_set z A - {0}" by simp
ultimately show "t \<in> ?B" by (rule image_eqI)
next
fix t
assume "t \<in> ?B"
then obtain p where "p \<in> hom_deg_set z A - {0}" and t: "t = lpp p" ..
from this(1) have "p \<in> hom_deg_set z A" and "p \<noteq> 0" by simp_all
with assms have "p \<in> A" and "homogeneous p" and "poly_deg p = z"
by (simp_all add: hom_deg_set_alt_homogeneous_set)
from this(1) \<open>p \<noteq> 0\<close> have "p \<in> A - {0}" by simp
hence 1: "t \<in> lpp ` (A - {0})" using t by (rule rev_image_eqI)
from \<open>p \<noteq> 0\<close> have "t \<in> keys p" unfolding t by (rule punit.lt_in_keys)
with \<open>homogeneous p\<close> have "deg_pm t = poly_deg p" by (rule homogeneousD_poly_deg)
with 1 show "t \<in> ?A" by (simp add: \<open>poly_deg p = z\<close>)
qed
lemma Hilbert_fun_alt:
assumes "finite X" and "A \<subseteq> P[X]" and "phull.subspace A"
shows "Hilbert_fun A z = card (lpp ` (hom_deg_set z A - {0}))" (is "_ = card ?A")
proof -
have "?A \<subseteq> lpp ` (hom_deg_set z A - {0})" by simp
then obtain B where sub: "B \<subseteq> hom_deg_set z A - {0}" and eq1: "?A = lpp ` B"
and inj: "inj_on lpp B" by (rule subset_imageE_inj)
have "Hilbert_fun A z = phull.dim (hom_deg_set z A)" by (fact Hilbert_fun_def)
also have "\<dots> = card B"
proof (rule phull.dim_eq_card)
show "phull.span B = phull.span (hom_deg_set z A)"
proof
from sub have "B \<subseteq> hom_deg_set z A" by blast
thus "phull.span B \<subseteq> phull.span (hom_deg_set z A)" by (rule phull.span_mono)
next
from assms(3) have "phull.subspace (hom_deg_set z A)" by (rule subspace_hom_deg_set)
hence "phull.span (hom_deg_set z A) = hom_deg_set z A" by (simp only: phull.span_eq_iff)
also have "\<dots> \<subseteq> phull.span B"
proof (rule ccontr)
assume "\<not> hom_deg_set z A \<subseteq> phull.span B"
then obtain p0 where "p0 \<in> hom_deg_set z A - phull.span B" (is "_ \<in> ?B") by blast
note assms(1) this
moreover have "?B \<subseteq> P[X]"
proof (rule subset_trans)
from assms(2) show "hom_deg_set z A \<subseteq> P[X]" by (rule Polys_closed_hom_deg_set)
qed blast
ultimately obtain p where "p \<in> ?B" and p_min: "\<And>q. punit.ord_strict_p q p \<Longrightarrow> q \<notin> ?B"
by (rule punit.ord_p_minimum_dgrad_p_set[OF dickson_grading_varnum, where m=0,
simplified dgrad_p_set_varnum]) blast
from this(1) have "p \<in> hom_deg_set z A" and "p \<notin> phull.span B" by simp_all
from phull.span_zero this(2) have "p \<noteq> 0" by blast
with \<open>p \<in> hom_deg_set z A\<close> have "p \<in> hom_deg_set z A - {0}" by simp
hence "lpp p \<in> lpp ` (hom_deg_set z A - {0})" by (rule imageI)
also have "\<dots> = lpp ` B" by (simp only: eq1)
finally obtain b where "b \<in> B" and eq2: "lpp p = lpp b" ..
from this(1) sub have "b \<in> hom_deg_set z A - {0}" ..
hence "b \<in> hom_deg_set z A" and "b \<noteq> 0" by simp_all
from this(2) have lcb: "punit.lc b \<noteq> 0" by (rule punit.lc_not_0)
from \<open>p \<noteq> 0\<close> have lcp: "punit.lc p \<noteq> 0" by (rule punit.lc_not_0)
from \<open>b \<in> B\<close> have "b \<in> phull.span B" by (rule phull.span_base)
hence "(punit.lc p / punit.lc b) \<cdot> b \<in> phull.span B" (is "?b \<in> _") by (rule phull.span_scale)
with \<open>p \<notin> phull.span B\<close> have "p - ?b \<noteq> 0" by auto
moreover from lcb lcp \<open>b \<noteq> 0\<close> have "lpp ?b = lpp p"
by (simp add: punit.map_scale_eq_monom_mult punit.lt_monom_mult eq2)
moreover from lcb have "punit.lc ?b = punit.lc p" by (simp add: punit.map_scale_eq_monom_mult)
ultimately have "lpp (p - ?b) \<prec> lpp p" by (rule punit.lt_minus_lessI)
hence "punit.ord_strict_p (p - ?b) p" by (rule punit.lt_ord_p)
hence "p - ?b \<notin> ?B" by (rule p_min)
hence "p - ?b \<notin> hom_deg_set z A \<or> p - ?b \<in> phull.span B" by simp
thus False
proof
assume *: "p - ?b \<notin> hom_deg_set z A"
from phull.subspace_scale have "?b \<in> hom_deg_set z A"
proof (rule hom_deg_set_closed_scalar)
show "phull.subspace A" by fact
next
show "b \<in> hom_deg_set z A" by fact
qed
with phull.subspace_diff \<open>p \<in> hom_deg_set z A\<close> have "p - ?b \<in> hom_deg_set z A"
by (rule hom_deg_set_closed_minus) (rule assms(3))
with * show ?thesis ..
next
assume "p - ?b \<in> phull.span B"
hence "p - ?b + ?b \<in> phull.span B" using \<open>?b \<in> phull.span B\<close> by (rule phull.span_add)
hence "p \<in> phull.span B" by simp
with \<open>p \<notin> phull.span B\<close> show ?thesis ..
qed
qed
finally show "phull.span (hom_deg_set z A) \<subseteq> phull.span B" .
qed
next
show "phull.independent B"
proof
assume "phull.dependent B"
then obtain B' u b' where "finite B'" and "B' \<subseteq> B" and "(\<Sum>b\<in>B'. u b \<cdot> b) = 0"
and "b' \<in> B'" and "u b' \<noteq> 0" unfolding phull.dependent_explicit by blast
define B0 where "B0 = {b \<in> B'. u b \<noteq> 0}"
have "B0 \<subseteq> B'" by (simp add: B0_def)
with \<open>finite B'\<close> have "(\<Sum>b\<in>B0. u b \<cdot> b) = (\<Sum>b\<in>B'. u b \<cdot> b)"
by (rule sum.mono_neutral_left) (simp add: B0_def)
also have "\<dots> = 0" by fact
finally have eq: "(\<Sum>b\<in>B0. u b \<cdot> b) = 0" .
define t where "t = ordered_powerprod_lin.Max (lpp ` B0)"
from \<open>b' \<in> B'\<close> \<open>u b' \<noteq> 0\<close> have "b' \<in> B0" by (simp add: B0_def)
hence "lpp b' \<in> lpp ` B0" by (rule imageI)
hence "lpp ` B0 \<noteq> {}" by blast
from \<open>B0 \<subseteq> B'\<close> \<open>finite B'\<close> have "finite B0" by (rule finite_subset)
hence "finite (lpp ` B0)" by (rule finite_imageI)
hence "t \<in> lpp ` B0" unfolding t_def using \<open>lpp ` B0 \<noteq> {}\<close>
by (rule ordered_powerprod_lin.Max_in)
then obtain b0 where "b0 \<in> B0" and t: "t = lpp b0" ..
note this(1)
moreover from \<open>B0 \<subseteq> B'\<close> \<open>B' \<subseteq> B\<close> have "B0 \<subseteq> B" by (rule subset_trans)
also have "\<dots> \<subseteq> hom_deg_set z A - {0}" by fact
finally have "b0 \<in> hom_deg_set z A - {0}" .
hence "b0 \<noteq> 0" by simp
hence "t \<in> keys b0" unfolding t by (rule punit.lt_in_keys)
have "lookup (\<Sum>b\<in>B0. u b \<cdot> b) t = (\<Sum>b\<in>B0. u b * lookup b t)" by (simp add: lookup_sum)
also from \<open>finite B0\<close> have "\<dots> = (\<Sum>b\<in>{b0}. u b * lookup b t)"
proof (rule sum.mono_neutral_right)
from \<open>b0 \<in> B0\<close> show "{b0} \<subseteq> B0" by simp
next
show "\<forall>b\<in>B0 - {b0}. u b * lookup b t = 0"
proof
fix b
assume "b \<in> B0 - {b0}"
hence "b \<in> B0" and "b \<noteq> b0" by simp_all
from this(1) have "lpp b \<in> lpp ` B0" by (rule imageI)
with \<open>finite (lpp ` B0)\<close> have "lpp b \<preceq> t" unfolding t_def
by (rule ordered_powerprod_lin.Max_ge)
have "t \<notin> keys b"
proof
assume "t \<in> keys b"
hence "t \<preceq> lpp b" by (rule punit.lt_max_keys)
with \<open>lpp b \<preceq> t\<close> have "lpp b = lpp b0"
- unfolding t by (rule ordered_powerprod_lin.antisym)
+ unfolding t by simp
from inj \<open>B0 \<subseteq> B\<close> have "inj_on lpp B0" by (rule inj_on_subset)
hence "b = b0" using \<open>lpp b = lpp b0\<close> \<open>b \<in> B0\<close> \<open>b0 \<in> B0\<close> by (rule inj_onD)
with \<open>b \<noteq> b0\<close> show False ..
qed
thus "u b * lookup b t = 0" by (simp add: in_keys_iff)
qed
qed
also from \<open>t \<in> keys b0\<close> \<open>b0 \<in> B0\<close> have "\<dots> \<noteq> 0" by (simp add: B0_def in_keys_iff)
finally show False by (simp add: eq)
qed
qed
also have "\<dots> = card ?A" unfolding eq1 using inj by (rule card_image[symmetric])
finally show ?thesis .
qed
end (* pm_powerprod *)
end (* theory *)
diff --git a/thys/Hybrid_Systems_VCs/ModalKleeneAlgebra/HS_VC_MKA.thy b/thys/Hybrid_Systems_VCs/ModalKleeneAlgebra/HS_VC_MKA.thy
--- a/thys/Hybrid_Systems_VCs/ModalKleeneAlgebra/HS_VC_MKA.thy
+++ b/thys/Hybrid_Systems_VCs/ModalKleeneAlgebra/HS_VC_MKA.thy
@@ -1,184 +1,184 @@
(* Title: Verification components with MKAs
Author: Jonathan Julián Huerta y Munive, 2020
Maintainer: Jonathan Julián Huerta y Munive <jonjulian23@gmail.com>
*)
section \<open> Verification components with MKA \<close>
text \<open> We use the forward box operator of antidomain Kleene algebras to derive rules for weakest
liberal preconditions (wlps) of regular programs. \<close>
theory HS_VC_MKA
imports "KAD.Modal_Kleene_Algebra"
begin
subsection \<open> Verification in AKA \<close>
text \<open>Here we derive verification components with weakest liberal preconditions based on
antidomain Kleene algebra \<close>
no_notation Range_Semiring.antirange_semiring_class.ars_r ("r")
and HOL.If ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10)
notation zero_class.zero ("0")
context antidomain_kleene_algebra
begin
\<comment> \<open> Skip \<close>
lemma "|1] x = d x"
using fbox_one .
\<comment> \<open> Abort \<close>
lemma "|0] q = 1"
using fbox_zero .
\<comment> \<open> Sequential composition \<close>
lemma "|x \<cdot> y] q = |x] |y] q"
using fbox_mult .
declare fbox_mult [simp]
\<comment> \<open> Nondeterministic choice \<close>
lemma "|x + y] q = |x] q \<cdot> |y] q"
using fbox_add2 .
lemma le_fbox_choice_iff: "d p \<le> |x + y]q \<longleftrightarrow> (d p \<le> |x]q) \<and> (d p \<le> |y]q)"
by (metis local.a_closure' local.ads_d_def local.dnsz.dom_glb_eq local.fbox_add2 local.fbox_def)
\<comment> \<open> Conditional statement \<close> (* by Victor Gomes, Georg Struth *)
definition aka_cond :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("if _ then _ else _" [64,64,64] 63)
where "if p then x else y = d p \<cdot> x + ad p \<cdot> y"
lemma fbox_export1: "ad p + |x] q = |d p \<cdot> x] q"
using a_d_add_closure addual.ars_r_def fbox_def fbox_mult by auto
lemma fbox_cond [simp]: "|if p then x else y] q = (ad p + |x] q) \<cdot> (d p + |y] q)"
using fbox_export1 local.ans_d_def local.fbox_mult
unfolding aka_cond_def ads_d_def fbox_def by auto
lemma fbox_cond2: "|if p then x else y] q = (d p \<cdot> |x] q) + (ad p \<cdot> |y] q)" (is "?lhs = ?d1 + ?d2")
proof -
have obs: "?lhs = d p \<cdot> ?lhs + ad p \<cdot> ?lhs"
by (metis (no_types, lifting) local.a_closure' local.a_de_morgan fbox_def ans_d_def
ads_d_def local.am2 local.am5_lem local.dka.dsg3 local.dka.dsr5)
have "d p \<cdot> ?lhs = d p \<cdot> |x] q \<cdot> (d p + d ( |y] q))"
using fbox_cond local.a_d_add_closure local.ads_d_def
local.ds.ddual.mult_assoc local.fbox_def by auto
also have "... = d p \<cdot> |x] q"
by (metis local.ads_d_def local.am2 local.dka.dns5 local.ds.ddual.mult_assoc local.fbox_def)
finally have "d p \<cdot> ?lhs = d p \<cdot> |x] q" .
moreover have "ad p \<cdot> ?lhs = ad p \<cdot> |y] q"
by (metis add_commute fbox_cond local.a_closure' local.a_mult_add ads_d_def ans_d_def
local.dnsz.dns5 local.ds.ddual.mult_assoc local.fbox_def)
ultimately show ?thesis
using obs by simp
qed
\<comment> \<open> While loop \<close> (* by Victor Gomes, Georg Struth *)
definition aka_whilei :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("while _ do _ inv _" [64,64,64] 63) where
"while t do x inv i = (d t \<cdot> x)\<^sup>\<star> \<cdot> ad t"
lemma fbox_frame: "d p \<cdot> x \<le> x \<cdot> d p \<Longrightarrow> d q \<le> |x] r \<Longrightarrow> d p \<cdot> d q \<le> |x] (d p \<cdot> d r)"
using dual.mult_isol_var fbox_add1 fbox_demodalisation3 fbox_simp by auto
lemma fbox_shunt: "d p \<cdot> d q \<le> |x] t \<longleftrightarrow> d p \<le> ad q + |x] t"
by (metis a_6 a_antitone' a_loc add_commute addual.ars_r_def am_d_def da_shunt2 fbox_def)
lemma fbox_export2: "|x] p \<le> |x \<cdot> ad q] (d p \<cdot> ad q)"
proof -
{fix t
have "d t \<cdot> x \<le> x \<cdot> d p \<Longrightarrow> d t \<cdot> x \<cdot> ad q \<le> x \<cdot> ad q \<cdot> d p \<cdot> ad q"
by (metis (full_types) a_comm_var a_mult_idem ads_d_def am2 ds.ddual.mult_assoc phl_export2)
hence "d t \<le> |x] p \<Longrightarrow> d t \<le> |x \<cdot> ad q] (d p \<cdot> ad q)"
by (metis a_closure' addual.ars_r_def ans_d_def dka.dsg3 ds.ddual.mult_assoc fbox_def fbox_demodalisation3)}
thus ?thesis
by (metis a_closure' addual.ars_r_def ans_d_def fbox_def order_refl)
qed
lemma fbox_while: "d p \<cdot> d t \<le> |x] p \<Longrightarrow> d p \<le> |(d t \<cdot> x)\<^sup>\<star> \<cdot> ad t] (d p \<cdot> ad t)"
proof -
assume "d p \<cdot> d t \<le> |x] p"
hence "d p \<le> |d t \<cdot> x] p"
by (simp add: fbox_export1 fbox_shunt)
hence "d p \<le> |(d t \<cdot> x)\<^sup>\<star>] p"
by (simp add: fbox_star_induct_var)
thus ?thesis
using order_trans fbox_export2 by presburger
qed
lemma fbox_whilei:
assumes "d p \<le> d i" and "d i \<cdot> ad t \<le> d q" and "d i \<cdot> d t \<le> |x] i"
shows "d p \<le> |while t do x inv i] q"
proof-
have "d i \<le> |(d t \<cdot> x)\<^sup>\<star> \<cdot> ad t] (d i \<cdot> ad t)"
using fbox_while assms by blast
also have "... \<le> |(d t \<cdot> x)\<^sup>\<star> \<cdot> ad t] q"
by (metis assms(2) local.dka.dom_iso local.dka.domain_invol local.fbox_iso)
finally show ?thesis
unfolding aka_whilei_def
using assms(1) local.dual_order.trans by blast
qed
lemma fbox_seq_var: "p \<le> |x] p' \<Longrightarrow> p' \<le> |y] q \<Longrightarrow> p \<le> |x \<cdot> y] q"
proof -
assume h1: "p \<le> |x] p'" and h2: "p' \<le> |y] q"
hence "|x] p' \<le> |x] |y] q"
by (metis ads_d_def fbox_antitone_var fbox_dom fbox_iso)
thus ?thesis
by (metis dual_order.trans fbox_mult h1)
qed
lemma fbox_whilei_break:
"d p \<le> |y] i \<Longrightarrow> d i \<cdot> ad t \<le> d q \<Longrightarrow> d i \<cdot> d t \<le> |x] i \<Longrightarrow> d p \<le> |y \<cdot> (while t do x inv i)] q"
apply (rule fbox_seq_var[OF _ fbox_whilei])
using fbox_simp by auto
\<comment> \<open> Finite iteration \<close>
definition aka_loopi :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("loop _ inv _ " [64,64] 63)
where "loop x inv i = x\<^sup>\<star>"
lemma "d p \<le> |x] p \<Longrightarrow> d p \<le> |x\<^sup>\<star>] p"
using fbox_star_induct_var .
lemma fbox_loopi: "p \<le> d i \<Longrightarrow> d i \<le> |x] i \<Longrightarrow> d i \<le> d q \<Longrightarrow> p \<le> |loop x inv i] q"
unfolding aka_loopi_def by (meson dual_order.trans fbox_iso fbox_star_induct_var)
lemma fbox_loopi_break:
"p \<le> |y] d i \<Longrightarrow> d i \<le> |x] i \<Longrightarrow> d i \<le> d q \<Longrightarrow> p \<le> |y \<cdot> (loop x inv i)] q"
by (rule fbox_seq_var, force) (rule fbox_loopi, auto)
\<comment> \<open> Invariants \<close>
lemma "p \<le> i \<Longrightarrow> i \<le> |x]i \<Longrightarrow> i \<le> q \<Longrightarrow> p \<le> |x]q"
by (metis local.ads_d_def local.dpdz.dom_iso local.dual_order.trans local.fbox_iso)
lemma "p \<le> d i \<Longrightarrow> d i \<le> |x]i \<Longrightarrow> i \<le> d q \<Longrightarrow> p \<le> |x]q"
- by (metis local.a_4 local.a_antitone' local.a_subid_aux2 ads_d_def local.antisym fbox_def
+ by (metis local.a_4 local.a_antitone' local.a_subid_aux2 ads_d_def order.antisym fbox_def
local.dka.dsg1 local.dual.mult_isol_var local.dual_order.trans local.order.refl)
lemma "(i \<le> |x] i) \<or> (j \<le> |x] j) \<Longrightarrow> (i + j) \<le> |x] (i + j)"
(*nitpick*)
oops
lemma "d i \<le> |x] i \<Longrightarrow> d j \<le> |x] j \<Longrightarrow> (d i + d j) \<le> |x] (d i + d j)"
by (metis (no_types, lifting) dual_order.trans fbox_simp fbox_subdist join.le_supE join.le_supI)
lemma plus_inv: "i \<le> |x] i \<Longrightarrow> j \<le> |x] j \<Longrightarrow> (i + j) \<le> |x] (i + j)"
by (metis ads_d_def dka.dsr5 fbox_simp fbox_subdist join.sup_mono order_trans)
lemma mult_inv: "d i \<le> |x] i \<Longrightarrow> d j \<le> |x] j \<Longrightarrow> (d i \<cdot> d j) \<le> |x] (d i \<cdot> d j)"
using fbox_demodalisation3 fbox_frame fbox_simp by auto
end
end
\ No newline at end of file
diff --git a/thys/KAD/Antidomain_Semiring.thy b/thys/KAD/Antidomain_Semiring.thy
--- a/thys/KAD/Antidomain_Semiring.thy
+++ b/thys/KAD/Antidomain_Semiring.thy
@@ -1,1262 +1,1262 @@
(* Title: Antidomain Semirings
Author: Victor B. F. Gomes, Walter Guttmann, Peter Höfner, Georg Struth, Tjark Weber
Maintainer: Walter Guttmann <walter.guttman at canterbury.ac.nz>
Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Antidomain Semirings\<close>
theory Antidomain_Semiring
imports Domain_Semiring
begin
subsection \<open>Antidomain Monoids\<close>
text \<open>We axiomatise antidomain monoids, using the axioms of~\cite{DesharnaisJipsenStruth}.\<close>
class antidomain_op =
fixes antidomain_op :: "'a \<Rightarrow> 'a" ("ad")
class antidomain_left_monoid = monoid_mult + antidomain_op +
assumes am1 [simp]: "ad x \<cdot> x = ad 1"
and am2: "ad x \<cdot> ad y = ad y \<cdot> ad x"
and am3 [simp]: "ad (ad x) \<cdot> x = x"
and am4 [simp]: "ad (x \<cdot> y) \<cdot> ad (x \<cdot> ad y) = ad x"
and am5 [simp]: "ad (x \<cdot> y) \<cdot> x \<cdot> ad y = ad (x \<cdot> y) \<cdot> x"
begin
no_notation domain_op ("d")
no_notation zero_class.zero ("0")
text \<open>We define a zero element and operations of domain and addition.\<close>
definition a_zero :: "'a" ("0") where
"0 = ad 1"
definition am_d :: "'a \<Rightarrow> 'a" ("d") where
"d x = ad (ad x)"
definition am_add_op :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oplus>" 65) where
"x \<oplus> y \<equiv> ad (ad x \<cdot> ad y)"
lemma a_d_zero [simp]: "ad x \<cdot> d x = 0"
by (metis am1 am2 a_zero_def am_d_def)
lemma a_d_one [simp]: "d x \<oplus> ad x = 1"
by (metis am1 am3 mult_1_right am_d_def am_add_op_def)
lemma n_annil [simp]: "0 \<cdot> x = 0"
proof -
have "0 \<cdot> x = d x \<cdot> ad x \<cdot> x"
by (simp add: a_zero_def am_d_def)
also have "... = d x \<cdot> 0"
by (metis am1 mult_assoc a_zero_def)
thus ?thesis
by (metis am1 am2 am3 mult_assoc a_zero_def)
qed
lemma a_mult_idem [simp]: "ad x \<cdot> ad x = ad x"
proof -
have "ad x \<cdot> ad x = ad (1 \<cdot> x) \<cdot> 1 \<cdot> ad x"
by simp
also have "... = ad (1 \<cdot> x) \<cdot> 1"
using am5 by blast
finally show ?thesis
by simp
qed
lemma a_add_idem [simp]: "ad x \<oplus> ad x = ad x"
by (metis am1 am3 am4 mult_1_right am_add_op_def)
text \<open>The next three axioms suffice to show that the domain elements form a Boolean algebra.\<close>
lemma a_add_comm: "x \<oplus> y = y \<oplus> x"
using am2 am_add_op_def by auto
lemma a_add_assoc: "x \<oplus> (y \<oplus> z) = (x \<oplus> y) \<oplus> z"
proof -
have "\<And>x y. ad x \<cdot> ad (x \<cdot> y) = ad x"
by (metis a_mult_idem am2 am4 mult_assoc)
thus ?thesis
by (metis a_add_comm am_add_op_def local.am3 local.am4 mult_assoc)
qed
lemma huntington [simp]: "ad (x \<oplus> y) \<oplus> ad (x \<oplus> ad y) = ad x"
using a_add_idem am_add_op_def by auto
lemma a_absorb1 [simp]: "(ad x \<oplus> ad y) \<cdot> ad x = ad x"
by (metis a_add_idem a_mult_idem am4 mult_assoc am_add_op_def)
lemma a_absorb2 [simp]: "ad x \<oplus> ad x \<cdot> ad y = ad x"
proof -
have "ad (ad x) \<cdot> ad (ad x \<cdot> ad y) = ad (ad x)"
by (metis (no_types) a_mult_idem local.am4 local.mult.semigroup_axioms semigroup.assoc)
then show ?thesis
using a_add_idem am_add_op_def by auto
qed
text \<open>The distributivity laws remain to be proved; our proofs follow those of Maddux~\cite{Maddux}.\<close>
lemma prod_split [simp]: "ad x \<cdot> ad y \<oplus> ad x \<cdot> d y = ad x"
using a_add_idem am_d_def am_add_op_def by auto
lemma sum_split [simp]: "(ad x \<oplus> ad y) \<cdot> (ad x \<oplus> d y) = ad x"
using a_add_idem am_d_def am_add_op_def by fastforce
lemma a_comp_simp [simp]: "(ad x \<oplus> ad y) \<cdot> d x = ad y \<cdot> d x"
proof -
have f1: "(ad x \<oplus> ad y) \<cdot> d x = ad (ad (ad x) \<cdot> ad (ad y)) \<cdot> ad (ad x) \<cdot> ad (ad (ad y))"
by (simp add: am_add_op_def am_d_def)
have f2: "ad y = ad (ad (ad y))"
using a_add_idem am_add_op_def by auto
have "ad y = ad (ad (ad x) \<cdot> ad (ad y)) \<cdot> ad y"
by (metis (no_types) a_absorb1 a_add_comm am_add_op_def)
then show ?thesis
using f2 f1 by (simp add: am_d_def local.am2 local.mult.semigroup_axioms semigroup.assoc)
qed
lemma a_distrib1: "ad x \<cdot> (ad y \<oplus> ad z) = ad x \<cdot> ad y \<oplus> ad x \<cdot> ad z"
proof -
have f1: "\<And>a. ad (ad (ad (a::'a)) \<cdot> ad (ad a)) = ad a"
using a_add_idem am_add_op_def by auto
have f2: "\<And>a aa. ad ((a::'a) \<cdot> aa) \<cdot> (a \<cdot> ad aa) = ad (a \<cdot> aa) \<cdot> a"
using local.am5 mult_assoc by auto
have f3: "\<And>a. ad (ad (ad (a::'a))) = ad a"
using f1 by simp
have "\<And>a. ad (a::'a) \<cdot> ad a = ad a"
by simp
then have "\<And>a aa. ad (ad (ad (a::'a) \<cdot> ad aa)) = ad aa \<cdot> ad a"
using f3 f2 by (metis (no_types) local.am2 local.am4 mult_assoc)
then have "ad x \<cdot> (ad y \<oplus> ad z) = ad x \<cdot> (ad y \<oplus> ad z) \<cdot> ad y \<oplus> ad x \<cdot> (ad y \<oplus> ad z) \<cdot> d y"
using am_add_op_def am_d_def local.am2 local.am4 by presburger
also have "... = ad x \<cdot> ad y \<oplus> ad x \<cdot> (ad y \<oplus> ad z) \<cdot> d y"
by (simp add: mult_assoc)
also have "... = ad x \<cdot> ad y \<oplus> ad x \<cdot> ad z \<cdot> d y"
by (simp add: mult_assoc)
also have "... = ad x \<cdot> ad y \<oplus> ad x \<cdot> ad y \<cdot> ad z \<oplus> ad x \<cdot> ad z \<cdot> d y"
by (metis a_add_idem a_mult_idem local.am4 mult_assoc am_add_op_def)
also have "... = ad x \<cdot> ad y \<oplus> (ad x \<cdot> ad z \<cdot> ad y \<oplus> ad x \<cdot> ad z \<cdot> d y)"
by (metis am2 mult_assoc a_add_assoc)
finally show ?thesis
by (metis a_add_idem a_mult_idem am4 am_d_def am_add_op_def)
qed
lemma a_distrib2: "ad x \<oplus> ad y \<cdot> ad z = (ad x \<oplus> ad y) \<cdot> (ad x \<oplus> ad z)"
proof -
have f1: "\<And>a aa ab. ad (ad (ad (a::'a) \<cdot> ad aa) \<cdot> ad (ad a \<cdot> ad ab)) = ad a \<cdot> ad (ad (ad aa) \<cdot> ad (ad ab))"
using a_distrib1 am_add_op_def by auto
have "\<And>a. ad (ad (ad (a::'a))) = ad a"
by (metis a_absorb2 a_mult_idem am_add_op_def)
then have "ad (ad (ad x) \<cdot> ad (ad y)) \<cdot> ad (ad (ad x) \<cdot> ad (ad z)) = ad (ad (ad x) \<cdot> ad (ad y \<cdot> ad z))"
using f1 by (metis (full_types))
then show ?thesis
by (simp add: am_add_op_def)
qed
lemma aa_loc [simp]: "d (x \<cdot> d y) = d (x \<cdot> y)"
proof -
have f1: "x \<cdot> d y \<cdot> y = x \<cdot> y"
by (metis am3 mult_assoc am_d_def)
have f2: "\<And>w z. ad (w \<cdot> z) \<cdot> (w \<cdot> ad z) = ad (w \<cdot> z) \<cdot> w"
by (metis am5 mult_assoc)
hence f3: "\<And>z. ad (x \<cdot> y) \<cdot> (x \<cdot> z) = ad (x \<cdot> y) \<cdot> (x \<cdot> (ad (ad (ad y) \<cdot> y) \<cdot> z))"
using f1 by (metis (no_types) mult_assoc am_d_def)
have "ad (x \<cdot> ad (ad y)) \<cdot> (x \<cdot> y) = 0" using f1
by (metis am1 mult_assoc n_annil a_zero_def am_d_def)
thus ?thesis
by (metis a_d_zero am_d_def f3 local.am1 local.am2 local.am3 local.am4)
qed
lemma a_loc [simp]: "ad (x \<cdot> d y) = ad (x \<cdot> y)"
proof -
have "\<And>a. ad (ad (ad (a::'a))) = ad a"
using am_add_op_def am_d_def prod_split by auto
then show ?thesis
by (metis (full_types) aa_loc am_d_def)
qed
lemma d_a_export [simp]: "d (ad x \<cdot> y) = ad x \<cdot> d y"
proof -
have f1: "\<And>a aa. ad ((a::'a) \<cdot> ad (ad aa)) = ad (a \<cdot> aa)"
using a_loc am_d_def by auto
have "\<And>a. ad (ad (a::'a) \<cdot> a) = 1"
using a_d_one am_add_op_def am_d_def by auto
then have "\<And>a aa. ad (ad (ad (a::'a) \<cdot> ad aa)) = ad a \<cdot> ad aa"
using f1 by (metis a_distrib2 am_add_op_def local.mult_1_left)
then show ?thesis
using f1 by (metis (no_types) am_d_def)
qed
text \<open>Every antidomain monoid is a domain monoid.\<close>
sublocale dm: domain_monoid am_d "(\<cdot>)" 1
apply (unfold_locales)
apply (simp add: am_d_def)
apply simp
using am_d_def d_a_export apply auto[1]
by (simp add: am_d_def local.am2)
lemma ds_ord_iso1: "x \<sqsubseteq> y \<Longrightarrow> z \<cdot> x \<sqsubseteq> z \<cdot> y"
(*nitpick [expect=genuine]*)
oops
lemma a_very_costrict: "ad x = 1 \<longleftrightarrow> x = 0"
proof
assume a: "ad x = 1"
hence "0 = ad x \<cdot> x"
using a_zero_def by force
thus "x = 0"
by (simp add: a)
next
assume "x = 0"
thus "ad x = 1"
using a_zero_def am_d_def dm.dom_one by auto
qed
lemma a_weak_loc: "x \<cdot> y = 0 \<longleftrightarrow> x \<cdot> d y = 0"
proof -
have "x \<cdot> y = 0 \<longleftrightarrow> ad (x \<cdot> y) = 1"
by (simp add: a_very_costrict)
also have "... \<longleftrightarrow> ad (x \<cdot> d y) = 1"
by simp
finally show ?thesis
using a_very_costrict by blast
qed
lemma a_closure [simp]: "d (ad x) = ad x"
using a_add_idem am_add_op_def am_d_def by auto
lemma a_d_mult_closure [simp]: "d (ad x \<cdot> ad y) = ad x \<cdot> ad y"
by simp
lemma kat_3': "d x \<cdot> y \<cdot> ad z = 0 \<Longrightarrow> d x \<cdot> y = d x \<cdot> y \<cdot> d z"
by (metis dm.dom_one local.am5 local.mult_1_left a_zero_def am_d_def)
lemma s4 [simp]: "ad x \<cdot> ad (ad x \<cdot> y) = ad x \<cdot> ad y"
proof -
have "\<And>a aa. ad (a::'a) \<cdot> ad (ad aa) = ad (ad (ad a \<cdot> aa))"
using am_d_def d_a_export by presburger
then have "\<And>a aa. ad (ad (a::'a)) \<cdot> ad aa = ad (ad (ad aa \<cdot> a))"
using local.am2 by presburger
then show ?thesis
by (metis a_comp_simp a_d_mult_closure am_add_op_def am_d_def local.am2)
qed
end
class antidomain_monoid = antidomain_left_monoid +
assumes am6 [simp]: "x \<cdot> ad 1 = ad 1"
begin
lemma kat_3_equiv: "d x \<cdot> y \<cdot> ad z = 0 \<longleftrightarrow> d x \<cdot> y = d x \<cdot> y \<cdot> d z"
apply standard
apply (metis kat_3')
by (simp add: mult_assoc a_zero_def am_d_def)
no_notation a_zero ("0")
no_notation am_d ("d")
end
subsection \<open>Antidomain Near-Semirings\<close>
text \<open>We define antidomain near-semirings. We do not consider units separately. The axioms are taken from~\cite{DesharnaisStruthAMAST}.\<close>
notation zero_class.zero ("0")
class antidomain_near_semiring = ab_near_semiring_one_zerol + antidomain_op + plus_ord +
assumes ans1 [simp]: "ad x \<cdot> x = 0"
and ans2 [simp]: "ad (x \<cdot> y) + ad (x \<cdot> ad (ad y)) = ad (x \<cdot> ad (ad y))"
and ans3 [simp]: "ad (ad x) + ad x = 1"
and ans4 [simp]: "ad (x + y) = ad x \<cdot> ad y"
begin
definition ans_d :: "'a \<Rightarrow> 'a" ("d") where
"d x = ad (ad x)"
lemma a_a_one [simp]: "d 1 = 1"
proof -
have "d 1 = d 1 + 0"
by simp
also have "... = d 1 + ad 1"
by (metis ans1 mult_1_right)
finally show ?thesis
by (simp add: ans_d_def)
qed
lemma a_very_costrict': "ad x = 1 \<longleftrightarrow> x = 0"
proof
assume "ad x = 1"
hence "x = ad x \<cdot> x"
by simp
thus "x = 0"
by auto
next
assume "x = 0"
hence "ad x = ad 0"
by blast
thus "ad x = 1"
by (metis a_a_one ans_d_def local.ans1 local.mult_1_right)
qed
lemma one_idem [simp]: "1 + 1 = 1"
proof -
have "1 + 1 = d 1 + d 1"
by simp
also have "... = ad (ad 1 \<cdot> 1) + ad (ad 1 \<cdot> d 1)"
using a_a_one ans_d_def by auto
also have "... = ad (ad 1 \<cdot> d 1)"
using ans_d_def local.ans2 by presburger
also have "... = ad (ad 1 \<cdot> 1)"
by simp
also have "... = d 1"
by (simp add: ans_d_def)
finally show ?thesis
by simp
qed
text \<open>Every antidomain near-semiring is automatically a dioid, and therefore ordered.\<close>
subclass near_dioid_one_zerol
proof
show "\<And>x. x + x = x"
proof -
fix x
have "x + x = 1 \<cdot> x + 1 \<cdot> x"
by simp
also have "... = (1 + 1) \<cdot> x"
using distrib_right' by presburger
finally show "x + x = x"
by simp
qed
qed
lemma d1_a [simp]: "d x \<cdot> x = x"
proof -
have "x = (d x + ad x) \<cdot> x"
by (simp add: ans_d_def)
also have "... = d x \<cdot> x + ad x \<cdot> x"
using distrib_right' by blast
also have "... = d x \<cdot> x + 0"
by simp
finally show ?thesis
by auto
qed
lemma a_comm: "ad x \<cdot> ad y = ad y \<cdot> ad x"
using add_commute ans4 by fastforce
lemma a_subid: "ad x \<le> 1"
using local.ans3 local.join.sup_ge2 by fastforce
lemma a_subid_aux1: "ad x \<cdot> y \<le> y"
using a_subid mult_isor by fastforce
lemma a_subdist: "ad (x + y) \<le> ad x"
by (metis a_subid_aux1 ans4 add_comm)
lemma a_antitone: "x \<le> y \<Longrightarrow> ad y \<le> ad x"
using a_subdist local.order_prop by auto
lemma a_mul_d [simp]: "ad x \<cdot> d x = 0"
by (metis a_comm ans_d_def local.ans1)
lemma a_gla1: "ad x \<cdot> y = 0 \<Longrightarrow> ad x \<le> ad y"
proof -
assume "ad x \<cdot> y = 0"
hence a: "ad x \<cdot> d y = 0"
by (metis a_subid a_very_costrict' ans_d_def local.ans2 local.join.sup.order_iff)
have "ad x = (d y + ad y ) \<cdot> ad x"
by (simp add: ans_d_def)
also have "... = d y \<cdot> ad x + ad y \<cdot> ad x"
using distrib_right' by blast
also have "... = ad x \<cdot> d y + ad x \<cdot> ad y"
using a_comm ans_d_def by auto
also have "... = ad x \<cdot> ad y"
by (simp add: a)
finally show "ad x \<le> ad y"
by (metis a_subid_aux1)
qed
lemma a_gla2: "ad x \<le> ad y \<Longrightarrow> ad x \<cdot> y = 0"
proof -
assume "ad x \<le> ad y"
hence "ad x \<cdot> y \<le> ad y \<cdot> y"
using mult_isor by blast
thus ?thesis
by (simp add: join.le_bot)
qed
lemma a2_eq [simp]: "ad (x \<cdot> d y) = ad (x \<cdot> y)"
-proof (rule antisym)
+proof (rule order.antisym)
show "ad (x \<cdot> y) \<le> ad (x \<cdot> d y)"
by (simp add: ans_d_def local.less_eq_def)
next
show "ad (x \<cdot> d y) \<le> ad (x \<cdot> y)"
by (metis a_gla1 a_mul_d ans1 d1_a mult_assoc)
qed
lemma a_export' [simp]: "ad (ad x \<cdot> y) = d x + ad y"
-proof (rule antisym)
+proof (rule order.antisym)
have "ad (ad x \<cdot> y) \<cdot> ad x \<cdot> d y = 0"
by (simp add: a_gla2 local.mult.semigroup_axioms semigroup.assoc)
hence a: "ad (ad x \<cdot> y) \<cdot> d y \<le> ad (ad x)"
by (metis a_comm a_gla1 ans4 mult_assoc ans_d_def)
have "ad (ad x \<cdot> y) = ad (ad x \<cdot> y) \<cdot> d y + ad (ad x \<cdot> y) \<cdot> ad y"
by (metis (no_types) add_commute ans3 ans4 distrib_right' mult_onel ans_d_def)
thus "ad (ad x \<cdot> y) \<le> d x + ad y"
by (metis a_subid_aux1 a join.sup_mono ans_d_def)
next
show "d x + ad y \<le> ad (ad x \<cdot> y)"
by (metis a2_eq a_antitone a_comm a_subid_aux1 join.sup_least ans_d_def)
qed
text \<open>Every antidomain near-semiring is a domain near-semiring.\<close>
sublocale dnsz: domain_near_semiring_one_zerol "(+)" "(\<cdot>)" 1 0 "ans_d" "(\<le>)" "(<)"
apply (unfold_locales)
apply simp
using a2_eq ans_d_def apply auto[1]
apply (simp add: a_subid ans_d_def local.join.sup_absorb2)
apply (simp add: ans_d_def)
apply (simp add: a_comm ans_d_def)
using a_a_one a_very_costrict' ans_d_def by force
lemma a_idem [simp]: "ad x \<cdot> ad x = ad x"
proof -
have "ad x = (d x + ad x ) \<cdot> ad x"
by (simp add: ans_d_def)
also have "... = d x \<cdot> ad x + ad x \<cdot> ad x"
using distrib_right' by blast
finally show ?thesis
by (simp add: ans_d_def)
qed
lemma a_3_var [simp]: "ad x \<cdot> ad y \<cdot> (x + y) = 0"
by (metis ans1 ans4)
lemma a_3 [simp]: "ad x \<cdot> ad y \<cdot> d (x + y) = 0"
by (metis a_mul_d ans4)
lemma a_closure' [simp]: "d (ad x) = ad x"
proof -
have "d (ad x) = ad (1 \<cdot> d x)"
by (simp add: ans_d_def)
also have "... = ad (1 \<cdot> x)"
using a2_eq by blast
finally show ?thesis
by simp
qed
text \<open>The following counterexamples show that some of the antidomain monoid axioms do not need to hold.\<close>
lemma "x \<cdot> ad 1 = ad 1"
(*nitpick [expect=genuine]*)
oops
lemma "ad (x \<cdot> y) \<cdot> ad (x \<cdot> ad y) = ad x"
(*nitpick [expect=genuine]*)
oops
lemma "ad (x \<cdot> y) \<cdot> ad (x \<cdot> ad y) = ad x"
(*nitpick [expect=genuine]*)
oops
lemma phl_seq_inv: "d v \<cdot> x \<cdot> y \<cdot> ad w = 0 \<Longrightarrow> \<exists>z. d v \<cdot> x \<cdot> d z = 0 \<and> ad z \<cdot> y \<cdot> ad w = 0"
proof -
assume "d v \<cdot> x \<cdot> y \<cdot> ad w = 0"
hence "d v \<cdot> x \<cdot> d (y \<cdot> ad w) = 0 \<and> ad (y \<cdot> ad w) \<cdot> y \<cdot> ad w = 0"
by (metis dnsz.dom_weakly_local local.ans1 mult_assoc)
thus "\<exists>z. d v \<cdot> x \<cdot> d z = 0 \<and> ad z \<cdot> y \<cdot> ad w = 0"
by blast
qed
lemma a_fixpoint: "ad x = x \<Longrightarrow> (\<forall>y. y = 0)"
proof -
assume a1: "ad x = x"
{ fix aa :: 'a
have "aa = 0"
using a1 by (metis (no_types) a_mul_d ans_d_def local.annil local.ans3 local.join.sup.idem local.mult_1_left)
}
then show ?thesis
by blast
qed
no_notation ans_d ("d")
end
subsection \<open>Antidomain Pre-Dioids\<close>
text \<open>Antidomain pre-diods are based on a different set of axioms, which are again taken from~\cite{DesharnaisStruthAMAST}.\<close>
class antidomain_pre_dioid = pre_dioid_one_zerol + antidomain_op +
assumes apd1 [simp]: "ad x \<cdot> x = 0"
and apd2 [simp]: "ad (x \<cdot> y) \<le> ad (x \<cdot> ad (ad y))"
and apd3 [simp]: "ad (ad x) + ad x = 1"
begin
definition apd_d :: "'a \<Rightarrow> 'a" ("d") where
"d x = ad (ad x)"
lemma a_very_costrict'': "ad x = 1 \<longleftrightarrow> x = 0"
- by (metis add_commute local.add_zerol local.antisym local.apd1 local.apd3 local.join.bot_least local.mult_1_right local.phl_skip)
+ by (metis add_commute local.add_zerol order.antisym local.apd1 local.apd3 local.join.bot_least local.mult_1_right local.phl_skip)
lemma a_subid': "ad x \<le> 1"
using local.apd3 local.join.sup_ge2 by fastforce
lemma d1_a' [simp]: "d x \<cdot> x = x"
proof -
have "x = (d x + ad x) \<cdot> x"
by (simp add: apd_d_def)
also have "... = d x \<cdot> x + ad x \<cdot> x"
using distrib_right' by blast
also have "... = d x \<cdot> x + 0"
by simp
finally show ?thesis
by auto
qed
lemma a_subid_aux1': "ad x \<cdot> y \<le> y"
using a_subid' mult_isor by fastforce
lemma a_mul_d' [simp]: "ad x \<cdot> d x = 0"
proof -
have "1 = ad (ad x \<cdot> x)"
using a_very_costrict'' by force
thus ?thesis
- by (metis a_subid' a_very_costrict'' apd_d_def local.antisym local.apd2)
+ by (metis a_subid' a_very_costrict'' apd_d_def order.antisym local.apd2)
qed
lemma a_d_closed [simp]: "d (ad x) = ad x"
-proof (rule antisym)
+proof (rule order.antisym)
have "d (ad x) = (d x + ad x) \<cdot> d (ad x)"
by (simp add: apd_d_def)
also have "... = ad (ad x) \<cdot> ad (d x) + ad x \<cdot> d (ad x)"
using apd_d_def local.distrib_right' by presburger
also have "... = ad x \<cdot> d (ad x)"
using a_mul_d' apd_d_def by auto
finally show "d (ad x) \<le> ad x"
by (metis a_subid' mult_1_right mult_isol apd_d_def)
next
have "ad x = ad (1 \<cdot> x)"
by simp
also have "... \<le> ad (1 \<cdot> d x)"
using apd_d_def local.apd2 by presburger
also have "... = ad (d x)"
by simp
finally show "ad x \<le> d (ad x)"
by (simp add: apd_d_def)
qed
lemma meet_ord_def: "ad x \<le> ad y \<longleftrightarrow> ad x \<cdot> ad y = ad x"
- by (metis a_d_closed a_subid_aux1' d1_a' eq_iff mult_1_right mult_isol)
+ by (metis a_d_closed a_subid_aux1' d1_a' order.eq_iff mult_1_right mult_isol)
lemma d_weak_loc: "x \<cdot> y = 0 \<longleftrightarrow> x \<cdot> d y = 0"
proof -
have "x \<cdot> y = 0 \<longleftrightarrow> ad (x \<cdot> y) = 1"
by (simp add: a_very_costrict'')
also have "... \<longleftrightarrow> ad (x \<cdot> d y) = 1"
- by (metis apd1 apd2 a_subid' apd_d_def d1_a' eq_iff mult_1_left mult_assoc)
+ by (metis apd1 apd2 a_subid' apd_d_def d1_a' order.eq_iff mult_1_left mult_assoc)
finally show ?thesis
by (simp add: a_very_costrict'')
qed
lemma gla_1: "ad x \<cdot> y = 0 \<Longrightarrow> ad x \<le> ad y"
proof -
assume "ad x \<cdot> y = 0"
hence a: "ad x \<cdot> d y = 0"
using d_weak_loc by force
hence "d y = ad x \<cdot> d y + d y"
by simp
also have "... = (1 + ad x) \<cdot> d y"
using join.sup_commute by auto
also have "... = (d x + ad x) \<cdot> d y"
using apd_d_def calculation by auto
also have "... = d x \<cdot> d y"
by (simp add: a join.sup_commute)
finally have "d y \<le> d x"
by (metis apd_d_def a_subid' mult_1_right mult_isol)
hence "d y \<cdot> ad x = 0"
by (metis apd_d_def a_d_closed a_mul_d' distrib_right' less_eq_def no_trivial_inverse)
hence "ad x = ad y \<cdot> ad x"
by (metis apd_d_def apd3 add_0_left distrib_right' mult_1_left)
thus "ad x \<le> ad y"
by (metis add_commute apd3 mult_oner subdistl)
qed
lemma a2_eq' [simp]: "ad (x \<cdot> d y) = ad (x \<cdot> y)"
-proof (rule antisym)
+proof (rule order.antisym)
show "ad (x \<cdot> y) \<le> ad (x \<cdot> d y)"
by (simp add: apd_d_def)
next
show "ad (x \<cdot> d y) \<le> ad (x \<cdot> y)"
by (metis gla_1 apd1 a_mul_d' d1_a' mult_assoc)
qed
lemma a_supdist_var: "ad (x + y) \<le> ad x"
by (metis gla_1 apd1 join.le_bot subdistl)
lemma a_antitone': "x \<le> y \<Longrightarrow> ad y \<le> ad x"
using a_supdist_var local.order_prop by auto
lemma a_comm_var: "ad x \<cdot> ad y \<le> ad y \<cdot> ad x"
proof -
have "ad x \<cdot> ad y = d (ad x \<cdot> ad y) \<cdot> ad x \<cdot> ad y"
by (simp add: mult_assoc)
also have "... \<le> d (ad x \<cdot> ad y) \<cdot> ad x"
using a_subid' mult_isol by fastforce
also have "... \<le> d (ad y) \<cdot> ad x"
by (simp add: a_antitone' a_subid_aux1' apd_d_def local.mult_isor)
finally show ?thesis
by simp
qed
lemma a_comm': "ad x \<cdot> ad y = ad y \<cdot> ad x"
- by (simp add: a_comm_var eq_iff)
+ by (simp add: a_comm_var order.eq_iff)
lemma a_closed [simp]: "d (ad x \<cdot> ad y) = ad x \<cdot> ad y"
proof -
have f1: "\<And>x y. ad x \<le> ad (ad y \<cdot> x)"
by (simp add: a_antitone' a_subid_aux1')
have "\<And>x y. d (ad x \<cdot> y) \<le> ad x"
by (metis a2_eq' a_antitone' a_comm' a_d_closed apd_d_def f1)
hence "\<And>x y. d (ad x \<cdot> y) \<cdot> y = ad x \<cdot> y"
by (metis d1_a' meet_ord_def mult_assoc apd_d_def)
thus ?thesis
by (metis f1 a_comm' apd_d_def meet_ord_def)
qed
lemma a_export'' [simp]: "ad (ad x \<cdot> y) = d x + ad y"
-proof (rule antisym)
+proof (rule order.antisym)
have "ad (ad x \<cdot> y) \<cdot> ad x \<cdot> d y = 0"
using d_weak_loc mult_assoc by fastforce
hence a: "ad (ad x \<cdot> y) \<cdot> d y \<le> d x"
by (metis a_closed a_comm' apd_d_def gla_1 mult_assoc)
have "ad (ad x \<cdot> y) = ad (ad x \<cdot> y) \<cdot> d y + ad (ad x \<cdot> y) \<cdot> ad y"
by (metis apd3 a_comm' d1_a' distrib_right' mult_1_right apd_d_def)
thus "ad (ad x \<cdot> y) \<le> d x + ad y"
by (metis a_subid_aux1' a join.sup_mono)
next
have "ad y \<le> ad (ad x \<cdot> y)"
by (simp add: a_antitone' a_subid_aux1')
thus "d x + ad y \<le> ad (ad x \<cdot> y)"
by (metis apd_d_def a_mul_d' d1_a' gla_1 apd1 join.sup_least mult_assoc)
qed
lemma d1_sum_var: "x + y \<le> (d x + d y) \<cdot> (x + y)"
proof -
have "x + y = d x \<cdot> x + d y \<cdot> y"
by simp
also have "... \<le> (d x + d y) \<cdot> x + (d x + d y) \<cdot> y"
using local.distrib_right' local.join.sup_ge1 local.join.sup_ge2 local.join.sup_mono by presburger
finally show ?thesis
using order_trans subdistl_var by blast
qed
lemma a4': "ad (x + y) = ad x \<cdot> ad y"
-proof (rule antisym)
+proof (rule order.antisym)
show "ad (x + y) \<le> ad x \<cdot> ad y"
by (metis a_d_closed a_supdist_var add_commute d1_a' local.mult_isol_var)
hence "ad x \<cdot> ad y = ad x \<cdot> ad y + ad (x + y)"
using less_eq_def add_commute by simp
also have "... = ad (ad (ad x \<cdot> ad y) \<cdot> (x + y))"
by (metis a_closed a_export'')
finally show "ad x \<cdot> ad y \<le> ad (x + y)"
using a_antitone' apd_d_def d1_sum_var by auto
qed
text \<open>Antidomain pre-dioids are domain pre-dioids and antidomain near-semirings, but still not antidomain monoids.\<close>
sublocale dpdz: domain_pre_dioid_one_zerol "(+)" "(\<cdot>)" 1 0 "(\<le>)" "(<)" "\<lambda>x. ad (ad x)"
apply (unfold_locales)
using apd_d_def d1_a' apply auto[1]
using a2_eq' apd_d_def apply auto[1]
apply (simp add: a_subid')
apply (simp add: a4' apd_d_def)
by (metis a_mul_d' a_very_costrict'' apd_d_def local.mult_onel)
subclass antidomain_near_semiring
apply (unfold_locales)
apply simp
using local.apd2 local.less_eq_def apply blast
apply simp
by (simp add: a4')
lemma a_supdist: "ad (x + y) \<le> ad x + ad y"
using a_supdist_var local.join.le_supI1 by auto
lemma a_gla: "ad x \<cdot> y = 0 \<longleftrightarrow> ad x \<le> ad y"
using gla_1 a_gla2 by blast
lemma a_subid_aux2: "x \<cdot> ad y \<le> x"
using a_subid' mult_isol by fastforce
lemma a42_var: "d x \<cdot> d y \<le> ad (ad x + ad y)"
by (simp add: apd_d_def)
lemma d1_weak [simp]: "(d x + d y) \<cdot> x = x"
proof -
have "(d x + d y) \<cdot> x = (1 + d y) \<cdot> x"
by simp
thus ?thesis
by (metis add_commute apd_d_def dpdz.dnso3 local.mult_1_left)
qed
lemma "x \<cdot> ad 1 = ad 1"
(*nitpick [expect=genuine]*)
oops
lemma "ad x \<cdot> (y + z) = ad x \<cdot> y + ad x \<cdot> z"
(*nitpick [expect=genuine]*)
oops
lemma "ad (x \<cdot> y) \<cdot> ad (x \<cdot> ad y) = ad x"
(*nitpick [expect=genuine]*)
oops
lemma "ad (x \<cdot> y) \<cdot> ad (x \<cdot> ad y) = ad x"
(*nitpick [expect=genuine]*)
oops
no_notation apd_d ("d")
end
subsection \<open>Antidomain Semirings\<close>
text \<open>Antidomain semirings are direct expansions of antidomain pre-dioids, but do not require idempotency of addition. Hence we give a slightly different axiomatisation, following~\cite{DesharnaisStruthSCP}.\<close>
class antidomain_semiringl = semiring_one_zerol + plus_ord + antidomain_op +
assumes as1 [simp]: "ad x \<cdot> x = 0"
and as2 [simp]: "ad (x \<cdot> y) + ad (x \<cdot> ad (ad y)) = ad (x \<cdot> ad (ad y))"
and as3 [simp]: "ad (ad x) + ad x = 1"
begin
definition ads_d :: "'a \<Rightarrow> 'a" ("d") where
"d x = ad (ad x)"
lemma one_idem': "1 + 1 = 1"
by (metis as1 as2 as3 add_zeror mult.right_neutral)
text \<open>Every antidomain semiring is a dioid and an antidomain pre-dioid.\<close>
subclass dioid
by (standard, metis distrib_left mult.right_neutral one_idem')
subclass antidomain_pre_dioid
by (unfold_locales, auto simp: local.less_eq_def)
lemma am5_lem [simp]: "ad (x \<cdot> y) \<cdot> ad (x \<cdot> ad y) = ad x"
proof -
have "ad (x \<cdot> y ) \<cdot> ad (x \<cdot> ad y) = ad (x \<cdot> d y) \<cdot> ad (x \<cdot> ad y)"
using ads_d_def local.a2_eq' local.apd_d_def by auto
also have "... = ad (x \<cdot> d y + x \<cdot> ad y)"
using ans4 by presburger
also have "... = ad (x \<cdot> (d y + ad y))"
using distrib_left by presburger
finally show ?thesis
by (simp add: ads_d_def)
qed
lemma am6_lem [simp]: "ad (x \<cdot> y) \<cdot> x \<cdot> ad y = ad (x \<cdot> y) \<cdot> x"
proof -
fix x y
have "ad (x \<cdot> y) \<cdot> x \<cdot> ad y = ad (x \<cdot> y) \<cdot> x \<cdot> ad y + 0"
by simp
also have "... = ad (x \<cdot> y) \<cdot> x \<cdot> ad y + ad (x \<cdot> d y) \<cdot> x \<cdot> d y"
using ans1 mult_assoc by presburger
also have "... = ad (x \<cdot> y) \<cdot> x \<cdot> (ad y + d y)"
using ads_d_def local.a2_eq' local.apd_d_def local.distrib_left by auto
finally show "ad (x \<cdot> y) \<cdot> x \<cdot> ad y = ad (x \<cdot> y) \<cdot> x"
using add_commute ads_d_def local.as3 by auto
qed
lemma a_zero [simp]: "ad 0 = 1"
by (simp add: local.a_very_costrict'')
lemma a_one [simp]: "ad 1 = 0"
using a_zero local.dpdz.dpd5 by blast
subclass antidomain_left_monoid
by (unfold_locales, auto simp: local.a_comm')
text \<open>Every antidomain left semiring is a domain left semiring.\<close>
no_notation domain_semiringl_class.fd ("( |_\<rangle> _)" [61,81] 82)
definition fdia :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("( |_\<rangle> _)" [61,81] 82) where
"|x\<rangle> y = ad (ad (x \<cdot> y))"
sublocale ds: domain_semiringl "(+)" "(\<cdot>)" 1 0 "\<lambda>x. ad (ad x)" "(\<le>)" "(<)"
rewrites "ds.fd x y \<equiv> fdia x y"
proof -
show "class.domain_semiringl (+) (\<cdot>) 1 0 (\<lambda>x. ad (ad x)) (\<le>) (<) "
by (unfold_locales, auto simp: local.dpdz.dpd4 ans_d_def)
then interpret ds: domain_semiringl "(+)" "(\<cdot>)" 1 0 "\<lambda>x. ad (ad x)" "(\<le>)" "(<)" .
show "ds.fd x y \<equiv> fdia x y"
by (auto simp: fdia_def ds.fd_def)
qed
lemma fd_eq_fdia [simp]: "domain_semiringl.fd (\<cdot>) d x y \<equiv> fdia x y"
proof -
have "class.domain_semiringl (+) (\<cdot>) 1 0 d (\<le>) (<)"
by (unfold_locales, auto simp: ads_d_def local.ans_d_def)
hence "domain_semiringl.fd (\<cdot>) d x y = d ((\<cdot>) x y)"
by (rule domain_semiringl.fd_def)
also have "... = ds.fd x y"
by (simp add: ds.fd_def ads_d_def)
finally show "domain_semiringl.fd (\<cdot>) d x y \<equiv> |x\<rangle> y"
by auto
qed
end
class antidomain_semiring = antidomain_semiringl + semiring_one_zero
begin
text \<open>Every antidomain semiring is an antidomain monoid.\<close>
subclass antidomain_monoid
by (standard, metis ans1 mult_1_right annir)
lemma "a_zero = 0"
by (simp add: local.a_zero_def)
sublocale ds: domain_semiring "(+)" "(\<cdot>)" 1 0 "\<lambda>x. ad (ad x)" "(\<le>)" "(<)"
rewrites "ds.fd x y \<equiv> fdia x y"
by unfold_locales
end
subsection \<open>The Boolean Algebra of Domain Elements\<close>
typedef (overloaded) 'a a2_element = "{x :: 'a :: antidomain_semiring. x = d x}"
by (rule_tac x=1 in exI, auto simp: ads_d_def)
setup_lifting type_definition_a2_element
instantiation a2_element :: (antidomain_semiring) boolean_algebra
begin
lift_definition less_eq_a2_element :: "'a a2_element \<Rightarrow> 'a a2_element \<Rightarrow> bool" is "(\<le>)" .
lift_definition less_a2_element :: "'a a2_element \<Rightarrow> 'a a2_element \<Rightarrow> bool" is "(<)" .
lift_definition bot_a2_element :: "'a a2_element" is 0
by (simp add: ads_d_def)
lift_definition top_a2_element :: "'a a2_element" is 1
by (simp add: ads_d_def)
lift_definition inf_a2_element :: "'a a2_element \<Rightarrow> 'a a2_element \<Rightarrow> 'a a2_element" is "(\<cdot>)"
by (metis (no_types, lifting) ads_d_def dpdz.dom_mult_closed)
lift_definition sup_a2_element :: "'a a2_element \<Rightarrow> 'a a2_element \<Rightarrow> 'a a2_element" is "(+)"
by (metis ads_d_def ds.dsr5)
lift_definition minus_a2_element :: "'a a2_element \<Rightarrow> 'a a2_element \<Rightarrow> 'a a2_element" is "\<lambda>x y. x \<cdot> ad y"
by (metis (no_types, lifting) ads_d_def dpdz.domain_export'')
lift_definition uminus_a2_element :: "'a a2_element \<Rightarrow> 'a a2_element" is antidomain_op
by (simp add: ads_d_def)
instance
apply (standard; transfer)
apply (simp add: less_le_not_le)
apply simp
apply auto[1]
apply simp
apply (metis a_subid_aux2 ads_d_def)
apply (metis a_subid_aux1' ads_d_def)
apply (metis (no_types, lifting) ads_d_def dpdz.dom_glb)
apply simp
apply simp
apply simp
apply simp
apply (metis a_subid' ads_d_def)
apply (metis (no_types, lifting) ads_d_def dpdz.dom_distrib)
apply (metis ads_d_def ans1)
apply (metis ads_d_def ans3)
by simp
end
subsection \<open>Further Properties\<close>
context antidomain_semiringl
begin
lemma a_2_var: "ad x \<cdot> d y = 0 \<longleftrightarrow> ad x \<le> ad y"
using local.a_gla local.ads_d_def local.dpdz.dom_weakly_local by auto
text \<open>The following two lemmas give the Galois connection of Heyting algebras.\<close>
lemma da_shunt1: "x \<le> d y + z \<Longrightarrow> x \<cdot> ad y \<le> z"
proof -
assume "x \<le> d y + z"
hence "x \<cdot> ad y \<le> (d y + z) \<cdot> ad y"
using mult_isor by blast
also have "... = d y \<cdot> ad y + z \<cdot> ad y"
by simp
also have "... \<le> z"
by (simp add: a_subid_aux2 ads_d_def)
finally show "x \<cdot> ad y \<le> z"
by simp
qed
lemma da_shunt2: "x \<le> ad y + z \<Longrightarrow> x \<cdot> d y \<le> z"
using da_shunt1 local.a_add_idem local.ads_d_def am_add_op_def by auto
lemma d_a_galois1: "d x \<cdot> ad y \<le> d z \<longleftrightarrow> d x \<le> d z + d y"
by (metis add_assoc local.a_gla local.ads_d_def local.am2 local.ans4 local.ans_d_def local.dnsz.dnso4)
lemma d_a_galois2: "d x \<cdot> d y \<le> d z \<longleftrightarrow> d x \<le> d z + ad y"
proof -
have "\<And>a aa. ad ((a::'a) \<cdot> ad (ad aa)) = ad (a \<cdot> aa)"
using local.a2_eq' local.apd_d_def by force
then show ?thesis
by (metis d_a_galois1 local.a_export' local.ads_d_def local.ans_d_def)
qed
lemma d_cancellation_1: "d x \<le> d y + d x \<cdot> ad y"
proof -
have a: "d (d x \<cdot> ad y) = ad y \<cdot> d x"
using local.a_closure' local.ads_d_def local.am2 local.ans_d_def by auto
hence "d x \<le> d (d x \<cdot> ad y) + d y"
using d_a_galois1 local.a_comm_var local.ads_d_def by fastforce
thus ?thesis
using a add_commute local.ads_d_def local.am2 by auto
qed
lemma d_cancellation_2: "(d z + d y) \<cdot> ad y \<le> d z"
by (simp add: da_shunt1)
lemma a_de_morgan: "ad (ad x \<cdot> ad y) = d (x + y)"
by (simp add: local.ads_d_def)
lemma a_de_morgan_var_3: "ad (d x + d y) = ad x \<cdot> ad y"
using local.a_add_idem local.ads_d_def am_add_op_def by auto
lemma a_de_morgan_var_4: "ad (d x \<cdot> d y) = ad x + ad y"
using local.a_add_idem local.ads_d_def am_add_op_def by auto
lemma a_4: "ad x \<le> ad (x \<cdot> y)"
using local.a_add_idem local.a_antitone' local.dpdz.domain_1'' am_add_op_def by fastforce
lemma a_6: "ad (d x \<cdot> y) = ad x + ad y"
using a_de_morgan_var_4 local.ads_d_def by auto
lemma a_7: "d x \<cdot> ad (d y + d z) = d x \<cdot> ad y \<cdot> ad z"
using a_de_morgan_var_3 local.mult.semigroup_axioms semigroup.assoc by fastforce
lemma a_d_add_closure [simp]: "d (ad x + ad y) = ad x + ad y"
using local.a_add_idem local.ads_d_def am_add_op_def by auto
lemma d_6 [simp]: "d x + ad x \<cdot> d y = d x + d y"
proof -
have "ad (ad x \<cdot> (x + ad y)) = d (x + y)"
by (simp add: distrib_left ads_d_def)
thus ?thesis
by (simp add: local.ads_d_def local.ans_d_def)
qed
lemma d_7 [simp]: "ad x + d x \<cdot> ad y = ad x + ad y"
by (metis a_d_add_closure local.ads_d_def local.ans4 local.s4)
lemma a_mult_add: "ad x \<cdot> (y + x) = ad x \<cdot> y"
by (simp add: distrib_left)
lemma kat_2: "y \<cdot> ad z \<le> ad x \<cdot> y \<Longrightarrow> d x \<cdot> y \<cdot> ad z = 0"
proof -
assume a: "y \<cdot> ad z \<le> ad x \<cdot> y"
hence "d x \<cdot> y \<cdot> ad z \<le> d x \<cdot> ad x \<cdot> y"
using local.mult_isol mult_assoc by presburger
thus ?thesis
using local.join.le_bot ads_d_def by auto
qed
lemma kat_3: "d x \<cdot> y \<cdot> ad z = 0 \<Longrightarrow> d x \<cdot> y = d x \<cdot> y \<cdot> d z"
using local.a_zero_def local.ads_d_def local.am_d_def local.kat_3' by auto
lemma kat_4: "d x \<cdot> y = d x \<cdot> y \<cdot> d z \<Longrightarrow> d x \<cdot> y \<le> y \<cdot> d z"
using a_subid_aux1 mult_assoc ads_d_def by auto
lemma kat_2_equiv: "y \<cdot> ad z \<le> ad x \<cdot> y \<longleftrightarrow> d x \<cdot> y \<cdot> ad z = 0"
proof
assume "y \<cdot> ad z \<le> ad x \<cdot> y"
thus "d x \<cdot> y \<cdot> ad z = 0"
by (simp add: kat_2)
next
assume 1: "d x \<cdot> y \<cdot> ad z = 0"
have "y \<cdot> ad z = (d x + ad x) \<cdot> y \<cdot> ad z"
by (simp add: local.ads_d_def)
also have "... = d x \<cdot> y \<cdot> ad z + ad x \<cdot> y \<cdot> ad z"
using local.distrib_right by presburger
also have "... = ad x \<cdot> y \<cdot> ad z"
using "1" by auto
also have "... \<le> ad x \<cdot> y"
by (simp add: local.a_subid_aux2)
finally show "y \<cdot> ad z \<le> ad x \<cdot> y" .
qed
lemma kat_4_equiv: "d x \<cdot> y = d x \<cdot> y \<cdot> d z \<longleftrightarrow> d x \<cdot> y \<le> y \<cdot> d z"
using local.ads_d_def local.dpdz.d_preserves_equation by auto
lemma kat_3_equiv_opp: "ad z \<cdot> y \<cdot> d x = 0 \<longleftrightarrow> y \<cdot> d x = d z \<cdot> y \<cdot> d x"
proof -
have "ad z \<cdot> (y \<cdot> d x) = 0 \<longrightarrow> (ad z \<cdot> y \<cdot> d x = 0) = (y \<cdot> d x = d z \<cdot> y \<cdot> d x)"
by (metis (no_types, hide_lams) add_commute local.add_zerol local.ads_d_def local.as3 local.distrib_right' local.mult_1_left mult_assoc)
thus ?thesis
by (metis a_4 local.a_add_idem local.a_gla2 local.ads_d_def mult_assoc am_add_op_def)
qed
lemma kat_4_equiv_opp: "y \<cdot> d x = d z \<cdot> y \<cdot> d x \<longleftrightarrow> y \<cdot> d x \<le> d z \<cdot> y"
using kat_2_equiv kat_3_equiv_opp local.ads_d_def by auto
subsection \<open>Forward Box and Diamond Operators\<close>
lemma fdemodalisation22: "|x\<rangle> y \<le> d z \<longleftrightarrow> ad z \<cdot> x \<cdot> d y = 0"
proof -
have "|x\<rangle> y \<le> d z \<longleftrightarrow> d (x \<cdot> y) \<le> d z"
by (simp add: fdia_def ads_d_def)
also have "... \<longleftrightarrow> ad z \<cdot> d (x \<cdot> y) = 0"
by (metis add_commute local.a_gla local.ads_d_def local.ans4)
also have "... \<longleftrightarrow> ad z \<cdot> x \<cdot> y = 0"
using dpdz.dom_weakly_local mult_assoc ads_d_def by auto
finally show ?thesis
using dpdz.dom_weakly_local ads_d_def by auto
qed
lemma dia_diff_var: "|x\<rangle> y \<le> |x\<rangle> (d y \<cdot> ad z) + |x\<rangle> z"
proof -
have 1: "|x\<rangle> (d y \<cdot> d z) \<le> |x\<rangle> (1 \<cdot> d z)"
using dpdz.dom_glb_eq ds.fd_subdist fdia_def ads_d_def by force
have "|x\<rangle> y = |x\<rangle> (d y \<cdot> (ad z + d z))"
by (metis as3 add_comm ds.fdia_d_simp mult_1_right ads_d_def)
also have "... = |x\<rangle> (d y \<cdot> ad z) + |x\<rangle> (d y \<cdot> d z)"
by (simp add: local.distrib_left local.ds.fdia_add1)
also have "... \<le> |x\<rangle> (d y \<cdot> ad z) + |x\<rangle> (1 \<cdot> d z)"
using "1" local.join.sup.mono by blast
finally show ?thesis
by (simp add: fdia_def ads_d_def)
qed
lemma dia_diff: "|x\<rangle> y \<cdot> ad ( |x\<rangle> z ) \<le> |x\<rangle> (d y \<cdot> ad z)"
using fdia_def dia_diff_var d_a_galois2 ads_d_def by metis
lemma fdia_export_2: "ad y \<cdot> |x\<rangle> z = |ad y \<cdot> x\<rangle> z"
using local.am_d_def local.d_a_export local.fdia_def mult_assoc by auto
lemma fdia_split: "|x\<rangle> y = d z \<cdot> |x\<rangle> y + ad z \<cdot> |x\<rangle> y"
by (metis mult_onel ans3 distrib_right ads_d_def)
definition fbox :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("( |_] _)" [61,81] 82) where
"|x] y = ad (x \<cdot> ad y)"
text \<open>The next lemmas establish the De Morgan duality between boxes and diamonds.\<close>
lemma fdia_fbox_de_morgan_2: "ad ( |x\<rangle> y) = |x] ad y"
using fbox_def local.a_closure local.a_loc local.am_d_def local.fdia_def by auto
lemma fbox_simp: "|x] y = |x] d y"
using fbox_def local.a_add_idem local.ads_d_def am_add_op_def by auto
lemma fbox_dom [simp]: "|x] 0 = ad x"
by (simp add: fbox_def)
lemma fbox_add1: "|x] (d y \<cdot> d z) = |x] y \<cdot> |x] z"
using a_de_morgan_var_4 fbox_def local.distrib_left by auto
lemma fbox_add2: "|x + y] z = |x] z \<cdot> |y] z"
by (simp add: fbox_def)
lemma fbox_mult: "|x \<cdot> y] z = |x] |y] z"
using fbox_def local.a2_eq' local.apd_d_def mult_assoc by auto
lemma fbox_zero [simp]: "|0] x = 1"
by (simp add: fbox_def)
lemma fbox_one [simp]: "|1] x = d x"
by (simp add: fbox_def ads_d_def)
lemma fbox_iso: "d x \<le> d y \<Longrightarrow> |z] x \<le> |z] y"
proof -
assume "d x \<le> d y"
hence "ad y \<le> ad x"
using local.a_add_idem local.a_antitone' local.ads_d_def am_add_op_def by fastforce
hence "z \<cdot> ad y \<le> z \<cdot> ad x"
by (simp add: mult_isol)
thus "|z] x \<le> |z] y"
by (simp add: fbox_def a_antitone')
qed
lemma fbox_antitone_var: "x \<le> y \<Longrightarrow> |y] z \<le> |x] z"
by (simp add: fbox_def a_antitone mult_isor)
lemma fbox_subdist_1: "|x] (d y \<cdot> d z) \<le> |x] y"
using a_de_morgan_var_4 fbox_def local.a_supdist_var local.distrib_left by force
lemma fbox_subdist_2: "|x] y \<le>|x] (d y + d z)"
by (simp add: fbox_iso ads_d_def)
lemma fbox_subdist: "|x] d y + |x] d z \<le> |x] (d y + d z)"
by (simp add: fbox_iso sup_least ads_d_def)
lemma fbox_diff_var: "|x] (d y + ad z) \<cdot> |x] z \<le> |x] y"
proof -
have "ad (ad y) \<cdot> ad (ad z) = ad (ad z + ad y)"
using local.dpdz.dsg4 by auto
then have "d (d (d y + ad z) \<cdot> d z) \<le> d y"
by (simp add: local.a_subid_aux1' local.ads_d_def)
then show ?thesis
by (metis fbox_add1 fbox_iso)
qed
lemma fbox_diff: "|x] (d y + ad z) \<le> |x] y + ad ( |x] z )"
proof -
have f1: "\<And>a. ad (ad (ad (a::'a))) = ad a"
using local.a_closure' local.ans_d_def by force
have f2: "\<And>a aa. ad (ad (a::'a)) + ad aa = ad (ad a \<cdot> aa)"
using local.ans_d_def by auto
have f3: "\<And>a aa. ad ((a::'a) + aa) = ad (aa + a)"
by (simp add: local.am2)
then have f4: "\<And>a aa. ad (ad (ad (a::'a) \<cdot> aa)) = ad (ad aa + a)"
using f2 f1 by (metis (no_types) local.ans4)
have f5: "\<And>a aa ab. ad ((a::'a) \<cdot> (aa + ab)) = ad (a \<cdot> (ab + aa))"
using f3 local.distrib_left by presburger
have f6: "\<And>a aa. ad (ad (ad (a::'a) + aa)) = ad (ad aa \<cdot> a)"
using f3 f1 by fastforce
have "ad (x \<cdot> ad (y + ad z)) \<le> ad (ad (x \<cdot> ad z) \<cdot> (x \<cdot> ad y))"
using f5 f2 f1 by (metis (no_types) a_mult_add fbox_def fbox_subdist_1 local.a_gla2 local.ads_d_def local.ans4 local.distrib_left local.gla_1 mult_assoc)
then show ?thesis
using f6 f4 f3 f1 by (simp add: fbox_def local.ads_d_def)
qed
end
context antidomain_semiring
begin
lemma kat_1: "d x \<cdot> y \<le> y \<cdot> d z \<Longrightarrow> y \<cdot> ad z \<le> ad x \<cdot> y"
proof -
assume a: "d x \<cdot> y \<le> y \<cdot> d z"
have "y \<cdot> ad z = d x \<cdot> y \<cdot> ad z + ad x \<cdot> y \<cdot> ad z"
by (metis local.ads_d_def local.as3 local.distrib_right local.mult_1_left)
also have "... \<le> y \<cdot> (d z \<cdot> ad z) + ad x \<cdot> y \<cdot> ad z"
by (metis a add_iso mult_isor mult_assoc)
also have "... = ad x \<cdot> y \<cdot> ad z"
by (simp add: ads_d_def)
finally show "y \<cdot> ad z \<le> ad x \<cdot> y"
using local.a_subid_aux2 local.dual_order.trans by blast
qed
lemma kat_1_equiv: "d x \<cdot> y \<le> y \<cdot> d z \<longleftrightarrow> y \<cdot> ad z \<le> ad x \<cdot> y"
using kat_1 kat_2 kat_3 kat_4 by blast
lemma kat_3_equiv': "d x \<cdot> y \<cdot> ad z = 0 \<longleftrightarrow> d x \<cdot> y = d x \<cdot> y \<cdot> d z"
by (simp add: kat_1_equiv local.kat_2_equiv local.kat_4_equiv)
lemma kat_1_equiv_opp: "y \<cdot> d x \<le> d z \<cdot> y \<longleftrightarrow> ad z \<cdot> y \<le> y \<cdot> ad x"
by (metis kat_1_equiv local.a_closure' local.ads_d_def local.ans_d_def)
lemma kat_2_equiv_opp: "ad z \<cdot> y \<le> y \<cdot> ad x \<longleftrightarrow> ad z \<cdot> y \<cdot> d x = 0"
by (simp add: kat_1_equiv_opp local.kat_3_equiv_opp local.kat_4_equiv_opp)
lemma fbox_one_1 [simp]: "|x] 1 = 1"
by (simp add: fbox_def)
lemma fbox_demodalisation3: "d y \<le> |x] d z \<longleftrightarrow> d y \<cdot> x \<le> x \<cdot> d z"
by (simp add: fbox_def a_gla kat_2_equiv_opp mult_assoc ads_d_def)
end
subsection \<open>Antidomain Kleene Algebras\<close>
class antidomain_left_kleene_algebra = antidomain_semiringl + left_kleene_algebra_zerol
begin
sublocale dka: domain_left_kleene_algebra "(+)" "(\<cdot>)" 1 0 d "(\<le>)" "(<)" star
rewrites "domain_semiringl.fd (\<cdot>) d x y \<equiv> |x\<rangle> y"
by (unfold_locales, auto simp add: local.ads_d_def ans_d_def)
lemma a_star [simp]: "ad (x\<^sup>\<star>) = 0"
using dka.dom_star local.a_very_costrict'' local.ads_d_def by force
lemma fbox_star_unfold [simp]: "|1] z \<cdot> |x] |x\<^sup>\<star>] z = |x\<^sup>\<star>] z"
proof -
have "ad (ad z + x \<cdot> (x\<^sup>\<star> \<cdot> ad z)) = ad (x\<^sup>\<star> \<cdot> ad z)"
using local.conway.dagger_unfoldl_distr mult_assoc by auto
then show ?thesis
using local.a_closure' local.ans_d_def local.fbox_def local.fdia_def local.fdia_fbox_de_morgan_2 by fastforce
qed
lemma fbox_star_unfold_var [simp]: "d z \<cdot> |x] |x\<^sup>\<star>] z = |x\<^sup>\<star>] z"
using fbox_star_unfold by auto
lemma fbox_star_unfoldr [simp]: "|1] z \<cdot> |x\<^sup>\<star>] |x] z = |x\<^sup>\<star>] z"
by (metis fbox_star_unfold fbox_mult star_slide_var)
lemma fbox_star_unfoldr_var [simp]: "d z \<cdot> |x\<^sup>\<star>] |x] z = |x\<^sup>\<star>] z"
using fbox_star_unfoldr by auto
lemma fbox_star_induct_var: "d y \<le> |x] y \<Longrightarrow> d y \<le> |x\<^sup>\<star>] y"
proof -
assume a1: "d y \<le> |x] y"
have "\<And>a. ad (ad (ad (a::'a))) = ad a"
using local.a_closure' local.ans_d_def by auto
then have "ad (ad (x\<^sup>\<star> \<cdot> ad y)) \<le> ad y"
using a1 by (metis dka.fdia_star_induct local.a_export' local.ads_d_def local.ans4 local.ans_d_def local.eq_refl local.fbox_def local.fdia_def local.meet_ord_def)
then have "ad (ad y + ad (x\<^sup>\<star> \<cdot> ad y)) = zero_class.zero"
by (metis (no_types) add_commute local.a_2_var local.ads_d_def local.ans4)
then show ?thesis
using local.a_2_var local.ads_d_def local.fbox_def by auto
qed
lemma fbox_star_induct: "d y \<le> d z \<cdot> |x] y \<Longrightarrow> d y \<le> |x\<^sup>\<star>] z"
proof -
assume a1: "d y \<le> d z \<cdot> |x] y"
hence a: "d y \<le> d z" and "d y \<le> |x] y"
apply (metis local.a_subid_aux2 local.dual_order.trans local.fbox_def)
using a1 dka.dom_subid_aux2 local.dual_order.trans by blast
hence "d y \<le> |x\<^sup>\<star>] y"
using fbox_star_induct_var by blast
thus ?thesis
using a local.fbox_iso local.order.trans by blast
qed
lemma fbox_star_induct_eq: "d z \<cdot> |x] y = d y \<Longrightarrow> d y \<le> |x\<^sup>\<star>] z"
by (simp add: fbox_star_induct)
lemma fbox_export_1: "ad y + |x] y = |d y \<cdot> x] y"
by (simp add: local.a_6 local.fbox_def mult_assoc)
lemma fbox_export_2: "d y + |x] y = |ad y \<cdot> x] y"
by (simp add: local.ads_d_def local.ans_d_def local.fbox_def mult_assoc)
end
class antidomain_kleene_algebra = antidomain_semiring + kleene_algebra
begin
subclass antidomain_left_kleene_algebra ..
lemma "d p \<le> |(d t \<cdot> x)\<^sup>\<star> \<cdot> ad t] (d q \<cdot> ad t) \<Longrightarrow> d p \<le> |d t \<cdot> x] d q"
(*nitpick [expect=genuine]*)
oops
end
end
diff --git a/thys/KAD/Domain_Semiring.thy b/thys/KAD/Domain_Semiring.thy
--- a/thys/KAD/Domain_Semiring.thy
+++ b/thys/KAD/Domain_Semiring.thy
@@ -1,854 +1,860 @@
(* Title: Domain Semirings
Author: Victor B. F. Gomes, Walter Guttmann, Peter Höfner, Georg Struth, Tjark Weber
Maintainer: Walter Guttmann <walter.guttman at canterbury.ac.nz>
Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Domain Semirings\<close>
theory Domain_Semiring
imports Kleene_Algebra.Kleene_Algebra
begin
subsection \<open>Domain Semigroups and Domain Monoids\<close>
class domain_op =
fixes domain_op :: "'a \<Rightarrow> 'a" ("d")
text \<open>First we define the class of domain semigroups. Axioms are taken from~\cite{DesharnaisJipsenStruth}.\<close>
class domain_semigroup = semigroup_mult + domain_op +
assumes dsg1 [simp]: "d x \<cdot> x = x"
and dsg2 [simp]: "d (x \<cdot> d y) = d (x \<cdot> y)"
and dsg3 [simp]: "d (d x \<cdot> y) = d x \<cdot> d y"
and dsg4: "d x \<cdot> d y = d y \<cdot> d x"
begin
lemma domain_invol [simp]: "d (d x) = d x"
proof -
have "d (d x) = d (d (d x \<cdot> x))"
by simp
also have "... = d (d x \<cdot> d x)"
using dsg3 by presburger
also have "... = d (d x \<cdot> x)"
by simp
finally show ?thesis
by simp
qed
text \<open>The next lemmas show that domain elements form semilattices.\<close>
lemma dom_el_idem [simp]: "d x \<cdot> d x = d x"
proof -
have "d x \<cdot> d x = d (d x \<cdot> x)"
using dsg3 by presburger
thus ?thesis
by simp
qed
lemma dom_mult_closed [simp]: "d (d x \<cdot> d y) = d x \<cdot> d y"
by simp
lemma dom_lc3 [simp]: "d x \<cdot> d (x \<cdot> y) = d (x \<cdot> y)"
proof -
have "d x \<cdot> d (x \<cdot> y) = d (d x \<cdot> x \<cdot> y)"
using dsg3 mult_assoc by presburger
thus ?thesis
by simp
qed
lemma d_fixpoint: "(\<exists>y. x = d y) \<longleftrightarrow> x = d x"
by auto
lemma d_type: "\<forall>P. (\<forall>x. x = d x \<longrightarrow> P x) \<longleftrightarrow> (\<forall>x. P (d x))"
by (metis domain_invol)
text \<open>We define the semilattice ordering on domain semigroups and explore the semilattice of domain elements from the order point of view.\<close>
definition ds_ord :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50) where
"x \<sqsubseteq> y \<longleftrightarrow> x = d x \<cdot> y"
lemma ds_ord_refl: "x \<sqsubseteq> x"
by (simp add: ds_ord_def)
lemma ds_ord_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
proof -
assume "x \<sqsubseteq> y" and a: "y \<sqsubseteq> z"
hence b: "x = d x \<cdot> y"
using ds_ord_def by blast
hence "x = d x \<cdot> d y \<cdot> z"
using a ds_ord_def mult_assoc by force
also have "... = d (d x \<cdot> y) \<cdot> z"
by simp
also have "... = d x \<cdot> z"
using b by auto
finally show ?thesis
using ds_ord_def by blast
qed
lemma ds_ord_antisym: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y"
proof -
assume a: "x \<sqsubseteq> y" and "y \<sqsubseteq> x"
hence b: "y = d y \<cdot> x"
using ds_ord_def by auto
have "x = d x \<cdot> d y \<cdot> x"
using a b ds_ord_def mult_assoc by force
also have "... = d y \<cdot> x"
by (metis (full_types) b dsg3 dsg4)
thus ?thesis
using b calculation by presburger
qed
text \<open>This relation is indeed an order.\<close>
-sublocale ds: order "(\<sqsubseteq>)" "\<lambda>x y. (x \<sqsubseteq> y \<and> x \<noteq> y)"
+sublocale ds: ordering \<open>(\<sqsubseteq>)\<close> \<open>\<lambda>x y. x \<sqsubseteq> y \<and> x \<noteq> y\<close>
proof
- show "\<And>x y. (x \<sqsubseteq> y \<and> x \<noteq> y) = (x \<sqsubseteq> y \<and> \<not> y \<sqsubseteq> x)"
- using ds_ord_antisym by blast
- show "\<And>x. x \<sqsubseteq> x"
+ show \<open>x \<sqsubseteq> y \<and> x \<noteq> y \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<noteq> y\<close> for x y
+ by (rule refl)
+ show "x \<sqsubseteq> x" for x
by (rule ds_ord_refl)
- show "\<And>x y z. x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
+ show "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" for x y z
by (rule ds_ord_trans)
- show "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y"
+ show "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y" for x y
by (rule ds_ord_antisym)
qed
+declare ds.refl [simp]
+
lemma ds_ord_eq: "x \<sqsubseteq> d x \<longleftrightarrow> x = d x"
by (simp add: ds_ord_def)
lemma "x \<sqsubseteq> y \<Longrightarrow> z \<cdot> x \<sqsubseteq> z \<cdot> y"
(*nitpick [expect=genuine]*)
oops
lemma ds_ord_iso_right: "x \<sqsubseteq> y \<Longrightarrow> x \<cdot> z \<sqsubseteq> y \<cdot> z"
proof -
assume "x \<sqsubseteq> y"
hence a: "x = d x \<cdot> y"
by (simp add: ds_ord_def)
hence "x \<cdot> z = d x \<cdot> y \<cdot> z"
by auto
also have "... = d (d x \<cdot> y \<cdot> z) \<cdot> d x \<cdot> y \<cdot> z"
using dsg1 mult_assoc by presburger
also have "... = d (x \<cdot> z) \<cdot> d x \<cdot> y \<cdot> z"
using a by presburger
finally show ?thesis
using ds_ord_def dsg4 mult_assoc by auto
qed
text \<open>The order on domain elements could as well be defined based on multiplication/meet.\<close>
lemma ds_ord_sl_ord: "d x \<sqsubseteq> d y \<longleftrightarrow> d x \<cdot> d y = d x"
using ds_ord_def by auto
lemma ds_ord_1: "d (x \<cdot> y) \<sqsubseteq> d x"
by (simp add: ds_ord_sl_ord dsg4)
lemma ds_subid_aux: "d x \<cdot> y \<sqsubseteq> y"
by (simp add: ds_ord_def mult_assoc)
lemma "y \<cdot> d x \<sqsubseteq> y"
(*nitpick [expect=genuine]*)
oops
lemma ds_dom_iso: "x \<sqsubseteq> y \<Longrightarrow> d x \<sqsubseteq> d y"
proof -
assume "x \<sqsubseteq> y"
hence "x = d x \<cdot> y"
by (simp add: ds_ord_def)
hence "d x = d (d x \<cdot> y)"
by presburger
also have "... = d x \<cdot> d y"
by simp
finally show ?thesis
using ds_ord_sl_ord by auto
qed
lemma ds_dom_llp: "x \<sqsubseteq> d y \<cdot> x \<longleftrightarrow> d x \<sqsubseteq> d y"
proof
assume "x \<sqsubseteq> d y \<cdot> x"
hence "x = d y \<cdot> x"
- by (simp add: ds_subid_aux ds.order.antisym)
+ by (simp add: ds_subid_aux ds.antisym)
hence "d x = d (d y \<cdot> x)"
by presburger
thus "d x \<sqsubseteq> d y"
using ds_ord_sl_ord dsg4 by force
next
assume "d x \<sqsubseteq> d y"
thus "x \<sqsubseteq> d y \<cdot> x"
by (metis (no_types) ds_ord_iso_right dsg1)
qed
lemma ds_dom_llp_strong: "x = d y \<cdot> x \<longleftrightarrow> d x \<sqsubseteq> d y"
+ using ds.eq_iff
by (simp add: ds_dom_llp ds.eq_iff ds_subid_aux)
definition refines :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
where "refines x y \<equiv> d y \<sqsubseteq> d x \<and> (d y) \<cdot> x \<sqsubseteq> y"
lemma refines_refl: "refines x x"
using refines_def by simp
lemma refines_trans: "refines x y \<Longrightarrow> refines y z \<Longrightarrow> refines x z"
unfolding refines_def
- by (metis domain_invol ds.dual_order.trans dsg1 dsg3 ds_ord_def)
+ by (metis domain_invol ds.trans dsg1 dsg3 ds_ord_def)
lemma refines_antisym: "refines x y \<Longrightarrow> refines y x \<Longrightarrow> x = y"
- unfolding refines_def
- using ds_dom_llp ds_ord_antisym by fastforce
+ apply (rule ds.antisym)
+ apply (simp_all add: refines_def)
+ apply (metis ds_dom_llp_strong)
+ apply (metis ds_dom_llp_strong)
+ done
-sublocale ref: order "refines" "\<lambda>x y. (refines x y \<and> x \<noteq> y)"
+sublocale ref: ordering "refines" "\<lambda>x y. (refines x y \<and> x \<noteq> y)"
proof
- show "\<And>x y. (refines x y \<and> x \<noteq> y) = (refines x y \<and> \<not> refines y x)"
- using refines_antisym by blast
+ show "\<And>x y. refines x y \<and> x \<noteq> y \<longleftrightarrow> refines x y \<and> x \<noteq> y"
+ ..
show "\<And>x. refines x x"
by (rule refines_refl)
show "\<And>x y z. refines x y \<Longrightarrow> refines y z \<Longrightarrow> refines x z"
by (rule refines_trans)
show "\<And>x y. refines x y \<Longrightarrow> refines y x \<Longrightarrow> x = y"
by (rule refines_antisym)
qed
end
text \<open>We expand domain semigroups to domain monoids.\<close>
class domain_monoid = monoid_mult + domain_semigroup
begin
lemma dom_one [simp]: "d 1 = 1"
proof -
have "1 = d 1 \<cdot> 1"
using dsg1 by presburger
thus ?thesis
by simp
qed
lemma ds_subid_eq: "x \<sqsubseteq> 1 \<longleftrightarrow> x = d x"
by (simp add: ds_ord_def)
end
subsection \<open>Domain Near-Semirings\<close>
text \<open>The axioms for domain near-semirings are taken from~\cite{DesharnaisStruthAMAST}.\<close>
class domain_near_semiring = ab_near_semiring + plus_ord + domain_op +
assumes dns1 [simp]: "d x \<cdot> x = x"
and dns2 [simp]: "d (x \<cdot> d y) = d(x \<cdot> y)"
and dns3 [simp]: "d (x + y) = d x + d y"
and dns4: "d x \<cdot> d y = d y \<cdot> d x"
and dns5 [simp]: "d x \<cdot> (d x + d y) = d x"
begin
text \<open>Domain near-semirings are automatically dioids; addition is idempotent.\<close>
subclass near_dioid
proof
show "\<And>x. x + x = x"
proof -
fix x
have a: "d x = d x \<cdot> d (x + x)"
using dns3 dns5 by presburger
have "d (x + x) = d (x + x + (x + x)) \<cdot> d (x + x)"
by (metis (no_types) dns3 dns4 dns5)
hence "d (x + x) = d (x + x) + d (x + x)"
by simp
thus "x + x = x"
by (metis a dns1 dns4 distrib_right')
qed
qed
text \<open>Next we prepare to show that domain near-semirings are domain semigroups.\<close>
lemma dom_iso: "x \<le> y \<Longrightarrow> d x \<le> d y"
using order_prop by auto
lemma dom_add_closed [simp]: "d (d x + d y) = d x + d y"
proof -
have "d (d x + d y) = d (d x) + d (d y)"
by simp
thus ?thesis
by (metis dns1 dns2 dns3 dns4)
qed
lemma dom_absorp_2 [simp]: "d x + d x \<cdot> d y = d x"
proof -
have "d x + d x \<cdot> d y = d x \<cdot> d x + d x \<cdot> d y"
by (metis add_idem' dns5)
also have "... = (d x + d y) \<cdot> d x"
by (simp add: dns4)
also have "... = d x \<cdot> (d x + d y)"
by (metis dom_add_closed dns4)
finally show ?thesis
by simp
qed
lemma dom_1: "d (x \<cdot> y) \<le> d x"
proof -
have "d (x \<cdot> y) = d (d x \<cdot> d (x \<cdot> y))"
by (metis dns1 dns2 mult_assoc)
also have "... \<le> d (d x) + d (d x \<cdot> d (x \<cdot> y))"
by simp
also have "... = d (d x + d x \<cdot> d (x \<cdot> y))"
using dns3 by presburger
also have "... = d (d x)"
by simp
finally show ?thesis
by (metis dom_add_closed add_idem')
qed
lemma dom_subid_aux2: "d x \<cdot> y \<le> y"
proof -
have "d x \<cdot> y \<le> d (x + d y) \<cdot> y"
by (simp add: mult_isor)
also have "... = (d x + d (d y)) \<cdot> d y \<cdot> y"
using dns1 dns3 mult_assoc by presburger
also have "... = (d y + d y \<cdot> d x) \<cdot> y"
by (simp add: dns4 add_commute)
finally show ?thesis
by simp
qed
lemma dom_glb: "d x \<le> d y \<Longrightarrow> d x \<le> d z \<Longrightarrow> d x \<le> d y \<cdot> d z"
by (metis dns5 less_eq_def mult_isor)
lemma dom_glb_eq: "d x \<le> d y \<cdot> d z \<longleftrightarrow> d x \<le> d y \<and> d x \<le> d z"
proof -
have "d x \<le> d z \<longrightarrow> d x \<le> d z"
by meson
then show ?thesis
by (metis (no_types) dom_absorp_2 dom_glb dom_subid_aux2 local.dual_order.trans local.join.sup.coboundedI2)
qed
lemma dom_ord: "d x \<le> d y \<longleftrightarrow> d x \<cdot> d y = d x"
proof
assume "d x \<le> d y"
hence "d x + d y = d y"
by (simp add: less_eq_def)
thus "d x \<cdot> d y = d x"
by (metis dns5)
next
assume "d x \<cdot> d y = d x"
thus "d x \<le> d y"
by (metis dom_subid_aux2)
qed
lemma dom_export [simp]: "d (d x \<cdot> y) = d x \<cdot> d y"
-proof (rule antisym)
+proof (rule order.antisym)
have "d (d x \<cdot> y) = d (d (d x \<cdot> y)) \<cdot> d (d x \<cdot> y)"
using dns1 by presburger
also have "... = d (d x \<cdot> d y) \<cdot> d (d x \<cdot> y)"
by (metis dns1 dns2 mult_assoc)
finally show a: "d (d x \<cdot> y) \<le> d x \<cdot> d y"
by (metis (no_types) dom_add_closed dom_glb dom_1 add_idem' dns2 dns4)
have "d (d x \<cdot> y) = d (d x \<cdot> y) \<cdot> d x"
using a dom_glb_eq dom_ord by force
hence "d x \<cdot> d y = d (d x \<cdot> y) \<cdot> d y"
by (metis dns1 dns2 mult_assoc)
thus "d x \<cdot> d y \<le> d (d x \<cdot> y)"
using a dom_glb_eq dom_ord by auto
qed
subclass domain_semigroup
by (unfold_locales, auto simp: dns4)
text \<open>We compare the domain semigroup ordering with that of the dioid.\<close>
lemma d_two_orders: "d x \<sqsubseteq> d y \<longleftrightarrow> d x \<le> d y"
by (simp add: dom_ord ds_ord_sl_ord)
lemma two_orders: "x \<sqsubseteq> y \<Longrightarrow> x \<le> y"
by (metis dom_subid_aux2 ds_ord_def)
lemma "x \<le> y \<Longrightarrow> x \<sqsubseteq> y"
(*nitpick [expect=genuine]*)
oops
text \<open>Next we prove additional properties.\<close>
lemma dom_subdist: "d x \<le> d (x + y)"
by simp
lemma dom_distrib: "d x + d y \<cdot> d z = (d x + d y) \<cdot> (d x + d z)"
proof -
have "(d x + d y) \<cdot> (d x + d z) = d x \<cdot> (d x + d z) + d y \<cdot> (d x + d z)"
using distrib_right' by blast
also have "... = d x + (d x + d z) \<cdot> d y"
by (metis (no_types) dns3 dns5 dsg4)
also have "... = d x + d x \<cdot> d y + d z \<cdot> d y"
using add_assoc' distrib_right' by presburger
finally show ?thesis
by (simp add: dsg4)
qed
lemma dom_llp1: "x \<le> d y \<cdot> x \<Longrightarrow> d x \<le> d y"
proof -
assume "x \<le> d y \<cdot> x"
hence "d x \<le> d (d y \<cdot> x)"
using dom_iso by blast
also have "... = d y \<cdot> d x"
by simp
finally show "d x \<le> d y"
by (simp add: dom_glb_eq)
qed
lemma dom_llp2: "d x \<le> d y \<Longrightarrow> x \<le> d y \<cdot> x"
using d_two_orders local.ds_dom_llp two_orders by blast
lemma dom_llp: "x \<le> d y \<cdot> x \<longleftrightarrow> d x \<le> d y"
using dom_llp1 dom_llp2 by blast
end
text \<open>We expand domain near-semirings by an additive unit, using slightly different axioms.\<close>
class domain_near_semiring_one = ab_near_semiring_one + plus_ord + domain_op +
assumes dnso1 [simp]: "x + d x \<cdot> x = d x \<cdot> x"
and dnso2 [simp]: "d (x \<cdot> d y) = d (x \<cdot> y)"
and dnso3 [simp]: "d x + 1 = 1"
and dnso4 [simp]: "d (x + y) = d x + d y"
and dnso5: "d x \<cdot> d y = d y \<cdot> d x"
begin
text \<open>The previous axioms are derivable.\<close>
subclass domain_near_semiring
proof
show a: "\<And>x. d x \<cdot> x = x"
by (metis add_commute local.dnso3 local.distrib_right' local.dnso1 local.mult_onel)
show "\<And>x y. d (x \<cdot> d y) = d (x \<cdot> y)"
by simp
show "\<And>x y. d (x + y) = d x + d y"
by simp
show "\<And>x y. d x \<cdot> d y = d y \<cdot> d x"
by (simp add: dnso5)
show "\<And>x y. d x \<cdot> (d x + d y) = d x"
proof -
fix x y
have "\<And>x. 1 + d x = 1"
using add_commute dnso3 by presburger
thus "d x \<cdot> (d x + d y) = d x"
by (metis (no_types) a dnso2 dnso4 dnso5 distrib_right' mult_onel)
qed
qed
subclass domain_monoid ..
lemma dom_subid: "d x \<le> 1"
by (simp add: less_eq_def)
end
text \<open>We add a left unit of multiplication.\<close>
class domain_near_semiring_one_zerol = ab_near_semiring_one_zerol + domain_near_semiring_one +
assumes dnso6 [simp]: "d 0 = 0"
begin
lemma domain_very_strict: "d x = 0 \<longleftrightarrow> x = 0"
by (metis annil dns1 dnso6)
lemma dom_weakly_local: "x \<cdot> y = 0 \<longleftrightarrow> x \<cdot> d y = 0"
proof -
have "x \<cdot> y = 0 \<longleftrightarrow> d (x \<cdot> y) = 0"
by (simp add: domain_very_strict)
also have "... \<longleftrightarrow> d (x \<cdot> d y) = 0"
by simp
finally show ?thesis
using domain_very_strict by blast
qed
end
subsection \<open>Domain Pre-Dioids\<close>
text \<open>
Pre-semirings with one and a left zero are automatically dioids.
Hence there is no point defining domain pre-semirings separately from domain dioids. The axioms
are once again from~\cite{DesharnaisStruthAMAST}.
\<close>
class domain_pre_dioid_one = pre_dioid_one + domain_op +
assumes dpd1 : "x \<le> d x \<cdot> x"
and dpd2 [simp]: "d (x \<cdot> d y) = d (x \<cdot> y)"
and dpd3 [simp]: "d x \<le> 1"
and dpd4 [simp]: "d (x + y) = d x + d y"
begin
text \<open>We prepare to show that every domain pre-dioid with one is a domain near-dioid with one.\<close>
lemma dns1'' [simp]: "d x \<cdot> x = x"
-proof (rule antisym)
+proof (rule order.antisym)
show "d x \<cdot> x \<le> x"
using dpd3 mult_isor by fastforce
show "x \<le> d x \<cdot> x "
by (simp add: dpd1)
qed
lemma d_iso: "x \<le> y \<Longrightarrow> d x \<le> d y"
by (metis dpd4 less_eq_def)
lemma domain_1'': "d (x \<cdot> y) \<le> d x"
proof -
have "d (x \<cdot> y) = d (x \<cdot> d y)"
by simp
also have "... \<le> d (x \<cdot> 1)"
by (meson d_iso dpd3 mult_isol)
finally show ?thesis
by simp
qed
lemma domain_export'' [simp]: "d (d x \<cdot> y) = d x \<cdot> d y"
-proof (rule antisym)
+proof (rule order.antisym)
have one: "d (d x \<cdot> y) \<le> d x"
by (metis dpd2 domain_1'' mult_onel)
have two: "d (d x \<cdot> y) \<le> d y"
using d_iso dpd3 mult_isor by fastforce
have "d (d x \<cdot> y) = d (d (d x \<cdot> y)) \<cdot> d (d x \<cdot> y)"
by simp
also have "... = d (d x \<cdot> y) \<cdot> d (d x \<cdot> y)"
by (metis dns1'' dpd2 mult_assoc)
thus "d (d x \<cdot> y) \<le> d x \<cdot> d y"
using mult_isol_var one two by force
next
have "d x \<cdot> d y \<le> 1"
by (metis dpd3 mult_1_right mult_isol order.trans)
thus "d x \<cdot> d y \<le> d (d x \<cdot> y)"
by (metis dns1'' dpd2 mult_isol mult_oner)
qed
lemma dom_subid_aux1'': "d x \<cdot> y \<le> y"
proof -
have "d x \<cdot> y \<le> 1 \<cdot> y"
using dpd3 mult_isor by blast
thus ?thesis
by simp
qed
lemma dom_subid_aux2'': "x \<cdot> d y \<le> x"
using dpd3 mult_isol by fastforce
lemma d_comm: "d x \<cdot> d y = d y \<cdot> d x"
-proof (rule antisym)
+proof (rule order.antisym)
have "d x \<cdot> d y = (d x \<cdot> d y) \<cdot> (d x \<cdot> d y)"
by (metis dns1'' domain_export'')
thus "d x \<cdot> d y \<le> d y \<cdot> d x"
by (metis dom_subid_aux1'' dom_subid_aux2'' mult_isol_var)
next
have "d y \<cdot> d x = (d y \<cdot> d x) \<cdot> (d y \<cdot> d x)"
by (metis dns1'' domain_export'')
thus "d y \<cdot> d x \<le> d x \<cdot> d y"
by (metis dom_subid_aux1'' dom_subid_aux2'' mult_isol_var)
qed
subclass domain_near_semiring_one
by (unfold_locales, auto simp: d_comm local.join.sup.absorb2)
lemma domain_subid: "x \<le> 1 \<Longrightarrow> x \<le> d x"
by (metis dns1 mult_isol mult_oner)
lemma d_preserves_equation: "d y \<cdot> x \<le> x \<cdot> d z \<longleftrightarrow> d y \<cdot> x = d y \<cdot> x \<cdot> d z"
- by (metis dom_subid_aux2'' local.antisym local.dom_el_idem local.dom_subid_aux2 local.order_prop local.subdistl mult_assoc)
+ by (metis dom_subid_aux2'' order.antisym local.dom_el_idem local.dom_subid_aux2 local.order_prop local.subdistl mult_assoc)
lemma d_restrict_iff: "(x \<le> y) \<longleftrightarrow> (x \<le> d x \<cdot> y)"
by (metis dom_subid_aux2 dsg1 less_eq_def order_trans subdistl)
lemma d_restrict_iff_1: "(d x \<cdot> y \<le> z) \<longleftrightarrow> (d x \<cdot> y \<le> d x \<cdot> z)"
by (metis dom_subid_aux2 domain_1'' domain_invol dsg1 mult_isol_var order_trans)
end
text \<open>We add once more a left unit of multiplication.\<close>
class domain_pre_dioid_one_zerol = domain_pre_dioid_one + pre_dioid_one_zerol +
assumes dpd5 [simp]: "d 0 = 0"
begin
subclass domain_near_semiring_one_zerol
by (unfold_locales, simp)
end
subsection \<open>Domain Semirings\<close>
text \<open>We do not consider domain semirings without units separately at the moment. The axioms are taken from from~\cite{DesharnaisStruthSCP}\<close>
class domain_semiringl = semiring_one_zerol + plus_ord + domain_op +
assumes dsr1 [simp]: "x + d x \<cdot> x = d x \<cdot> x"
and dsr2 [simp]: "d (x \<cdot> d y) = d (x \<cdot> y)"
and dsr3 [simp]: "d x + 1 = 1"
and dsr4 [simp]: "d 0 = 0"
and dsr5 [simp]: "d (x + y) = d x + d y"
begin
text \<open>Every domain semiring is automatically a domain pre-dioid with one and left zero.\<close>
subclass dioid_one_zerol
by (standard, metis add_commute dsr1 dsr3 distrib_left mult_oner)
subclass domain_pre_dioid_one_zerol
by (standard, auto simp: less_eq_def)
end
class domain_semiring = domain_semiringl + semiring_one_zero
subsection \<open>The Algebra of Domain Elements\<close>
text \<open>We show that the domain elements of a domain semiring form a distributive lattice. Unfortunately we cannot prove this within the type class of domain semirings.\<close>
typedef (overloaded) 'a d_element = "{x :: 'a :: domain_semiring. x = d x}"
- by (rule_tac x = 1 in exI, simp add: domain_subid order_class.eq_iff)
+ by (rule_tac x = 1 in exI, simp add: domain_subid ds.eq_iff)
setup_lifting type_definition_d_element
instantiation d_element :: (domain_semiring) bounded_lattice
begin
lift_definition less_eq_d_element :: "'a d_element \<Rightarrow> 'a d_element \<Rightarrow> bool" is "(\<le>)" .
lift_definition less_d_element :: "'a d_element \<Rightarrow> 'a d_element \<Rightarrow> bool" is "(<)" .
lift_definition bot_d_element :: "'a d_element" is 0
by simp
lift_definition top_d_element :: "'a d_element" is 1
by simp
lift_definition inf_d_element :: "'a d_element \<Rightarrow> 'a d_element \<Rightarrow> 'a d_element" is "(\<cdot>)"
by (metis dsg3)
lift_definition sup_d_element :: "'a d_element \<Rightarrow> 'a d_element \<Rightarrow> 'a d_element" is "(+)"
by simp
instance
apply (standard; transfer)
apply (simp add: less_le_not_le)+
apply (metis dom_subid_aux2'')
apply (metis dom_subid_aux2)
apply (metis dom_glb)
apply simp+
by (metis dom_subid)
end
instance d_element :: (domain_semiring) distrib_lattice
by (standard, transfer, metis dom_distrib)
subsection \<open>Domain Semirings with a Greatest Element\<close>
text \<open>If there is a greatest element in the semiring, then we have another equality.\<close>
class domain_semiring_top = domain_semiring + order_top
begin
notation top ("\<top>")
lemma kat_equivalence_greatest: "d x \<le> d y \<longleftrightarrow> x \<le> d y \<cdot> \<top>"
proof
assume "d x \<le> d y"
thus "x \<le> d y \<cdot> \<top>"
by (metis dsg1 mult_isol_var top_greatest)
next
assume "x \<le> d y \<cdot> \<top>"
thus "d x \<le> d y"
using dom_glb_eq dom_iso by fastforce
qed
end
subsection \<open>Forward Diamond Operators\<close>
context domain_semiringl
begin
text \<open>We define a forward diamond operator over a domain semiring. A more modular consideration is not given at the moment.\<close>
definition fd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("( |_\<rangle> _)" [61,81] 82) where
"|x\<rangle> y = d (x \<cdot> y)"
lemma fdia_d_simp [simp]: "|x\<rangle> d y = |x\<rangle> y"
by (simp add: fd_def)
lemma fdia_dom [simp]: "|x\<rangle> 1 = d x"
by (simp add: fd_def)
lemma fdia_add1: "|x\<rangle> (y + z) = |x\<rangle> y + |x\<rangle> z"
by (simp add: fd_def distrib_left)
lemma fdia_add2: "|x + y\<rangle> z = |x\<rangle> z + |y\<rangle> z"
by (simp add: fd_def distrib_right)
lemma fdia_mult: "|x \<cdot> y\<rangle> z = |x\<rangle> |y\<rangle> z"
by (simp add: fd_def mult_assoc)
lemma fdia_one [simp]: "|1\<rangle> x = d x"
by (simp add: fd_def)
lemma fdemodalisation1: "d z \<cdot> |x\<rangle> y = 0 \<longleftrightarrow> d z \<cdot> x \<cdot> d y = 0"
proof -
have "d z \<cdot> |x\<rangle> y = 0 \<longleftrightarrow> d z \<cdot> d (x \<cdot> y) = 0"
by (simp add: fd_def)
also have "... \<longleftrightarrow> d z \<cdot> x \<cdot> y = 0"
by (metis annil dnso6 dsg1 dsg3 mult_assoc)
finally show ?thesis
using dom_weakly_local by auto
qed
lemma fdemodalisation2: "|x\<rangle> y \<le> d z \<longleftrightarrow> x \<cdot> d y \<le> d z \<cdot> x"
proof
assume "|x\<rangle> y \<le> d z"
hence a: "d (x \<cdot> d y) \<le> d z"
by (simp add: fd_def)
have "x \<cdot> d y = d (x \<cdot> d y) \<cdot> x \<cdot> d y"
using dsg1 mult_assoc by presburger
also have "... \<le> d z \<cdot> x \<cdot> d y"
using a calculation dom_llp2 mult_assoc by auto
finally show "x \<cdot> d y \<le> d z \<cdot> x"
using dom_subid_aux2'' order_trans by blast
next
assume "x \<cdot> d y \<le> d z \<cdot> x"
hence "d (x \<cdot> d y) \<le> d (d z \<cdot> d x)"
using dom_iso by fastforce
also have "... \<le> d (d z)"
using domain_1'' by blast
finally show "|x\<rangle> y \<le> d z"
by (simp add: fd_def)
qed
lemma fd_iso1: "d x \<le> d y \<Longrightarrow> |z\<rangle> x \<le> |z\<rangle> y"
using fd_def local.dom_iso local.mult_isol by fastforce
lemma fd_iso2: "x \<le> y \<Longrightarrow> |x\<rangle> z \<le> |y\<rangle> z"
by (simp add: fd_def dom_iso mult_isor)
lemma fd_zero_var [simp]: "|0\<rangle> x = 0"
by (simp add: fd_def)
lemma fd_subdist_1: "|x\<rangle> y \<le> |x\<rangle> (y + z)"
by (simp add: fd_iso1)
lemma fd_subdist_2: "|x\<rangle> (d y \<cdot> d z) \<le> |x\<rangle> y"
by (simp add: fd_iso1 dom_subid_aux2'')
lemma fd_subdist: "|x\<rangle> (d y \<cdot> d z) \<le> |x\<rangle> y \<cdot> |x\<rangle> z"
using fd_def fd_iso1 fd_subdist_2 dom_glb dom_subid_aux2 by auto
lemma fdia_export_1: "d y \<cdot> |x\<rangle> z = |d y \<cdot> x\<rangle> z"
by (simp add: fd_def mult_assoc)
end
context domain_semiring
begin
lemma fdia_zero [simp]: "|x\<rangle> 0 = 0"
by (simp add: fd_def)
end
subsection \<open>Domain Kleene Algebras\<close>
text \<open>We add the Kleene star to our considerations. Special domain axioms are not needed.\<close>
class domain_left_kleene_algebra = left_kleene_algebra_zerol + domain_semiringl
begin
lemma dom_star [simp]: "d (x\<^sup>\<star>) = 1"
proof -
have "d (x\<^sup>\<star>) = d (1 + x \<cdot> x\<^sup>\<star>)"
by simp
also have "... = d 1 + d (x \<cdot> x\<^sup>\<star>)"
using dns3 by blast
finally show ?thesis
using add_commute local.dsr3 by auto
qed
lemma fdia_star_unfold [simp]: "|1\<rangle> y + |x\<rangle> |x\<^sup>\<star>\<rangle> y = |x\<^sup>\<star>\<rangle> y"
proof -
have "|1\<rangle> y + |x\<rangle> |x\<^sup>\<star>\<rangle> y = |1 + x \<cdot> x\<^sup>\<star>\<rangle> y"
using local.fdia_add2 local.fdia_mult by presburger
thus ?thesis
by simp
qed
lemma fdia_star_unfoldr [simp]: "|1\<rangle> y + |x\<^sup>\<star>\<rangle> |x\<rangle> y = |x\<^sup>\<star>\<rangle> y"
proof -
have "|1\<rangle> y + |x\<^sup>\<star>\<rangle> |x\<rangle> y = |1 + x\<^sup>\<star> \<cdot> x\<rangle> y"
using fdia_add2 fdia_mult by presburger
thus ?thesis
by simp
qed
lemma fdia_star_unfold_var [simp]: "d y + |x\<rangle> |x\<^sup>\<star>\<rangle> y = |x\<^sup>\<star>\<rangle> y"
proof -
have "d y + |x\<rangle> |x\<^sup>\<star>\<rangle> y = |1\<rangle> y + |x\<rangle> |x\<^sup>\<star>\<rangle> y"
by simp
also have "... = |1 + x \<cdot> x\<^sup>\<star>\<rangle> y"
using fdia_add2 fdia_mult by presburger
finally show ?thesis
by simp
qed
lemma fdia_star_unfoldr_var [simp]: "d y + |x\<^sup>\<star>\<rangle> |x\<rangle> y = |x\<^sup>\<star>\<rangle> y"
proof -
have "d y + |x\<^sup>\<star>\<rangle> |x\<rangle> y = |1\<rangle> y + |x\<^sup>\<star>\<rangle> |x\<rangle> y"
by simp
also have "... = |1 + x\<^sup>\<star> \<cdot> x\<rangle> y"
using fdia_add2 fdia_mult by presburger
finally show ?thesis
by simp
qed
lemma fdia_star_induct_var: "|x\<rangle> y \<le> d y \<Longrightarrow> |x\<^sup>\<star>\<rangle> y \<le> d y"
proof -
assume a1: "|x\<rangle> y \<le> d y"
hence "x \<cdot> d y \<le> d y \<cdot> x"
by (simp add: fdemodalisation2)
hence "x\<^sup>\<star> \<cdot> d y \<le> d y \<cdot> x\<^sup>\<star>"
by (simp add: star_sim1)
thus ?thesis
by (simp add: fdemodalisation2)
qed
lemma fdia_star_induct: "d z + |x\<rangle> y \<le> d y \<Longrightarrow> |x\<^sup>\<star>\<rangle> z \<le> d y"
proof -
assume a: "d z + |x\<rangle> y \<le> d y"
hence b: "d z \<le> d y" and c: "|x\<rangle> y \<le> d y"
apply (simp add: local.join.le_supE)
using a by auto
hence d: "|x\<^sup>\<star>\<rangle> z \<le> |x\<^sup>\<star>\<rangle> y"
using fd_def fd_iso1 by auto
have "|x\<^sup>\<star>\<rangle> y \<le> d y"
using c fdia_star_induct_var by blast
thus ?thesis
using d by fastforce
qed
lemma fdia_star_induct_eq: "d z + |x\<rangle> y = d y \<Longrightarrow> |x\<^sup>\<star>\<rangle> z \<le> d y"
by (simp add: fdia_star_induct)
end
class domain_kleene_algebra = kleene_algebra + domain_semiring
begin
subclass domain_left_kleene_algebra ..
end
end
diff --git a/thys/KAD/Modal_Kleene_Algebra.thy b/thys/KAD/Modal_Kleene_Algebra.thy
--- a/thys/KAD/Modal_Kleene_Algebra.thy
+++ b/thys/KAD/Modal_Kleene_Algebra.thy
@@ -1,113 +1,115 @@
(* Title: Modal Kleene Algebras
Author: Victor B. F. Gomes, Walter Guttmann, Peter Höfner, Georg Struth, Tjark Weber
Maintainer: Walter Guttmann <walter.guttman at canterbury.ac.nz>
Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Modal Kleene Algebras\<close>
text \<open>This section studies laws that relate antidomain and antirange semirings and Kleene algebra,
notably Galois connections and conjugations as those studied in~\cite{MoellerStruth,DesharnaisStruthSCP}.\<close>
theory Modal_Kleene_Algebra
imports Range_Semiring
begin
class modal_semiring = antidomain_semiring + antirange_semiring +
assumes domrange [simp]: "d (r x) = r x"
and rangedom [simp]: "r (d x) = d x"
begin
text \<open>These axioms force that the domain algebra and the range algebra coincide.\<close>
lemma domrangefix: "d x = x \<longleftrightarrow> r x = x"
by (metis domrange rangedom)
lemma box_diamond_galois_1:
assumes "d p = p" and "d q = q"
shows "\<langle>x| p \<le> q \<longleftrightarrow> p \<le> |x] q"
proof -
have "\<langle>x| p \<le> q \<longleftrightarrow> p \<cdot> x \<le> x \<cdot> q"
by (metis assms domrangefix local.ardual.ds.fdemodalisation2 local.ars_r_def)
thus ?thesis
by (metis assms fbox_demodalisation3)
qed
lemma box_diamond_galois_2:
assumes "d p = p" and "d q = q"
shows "|x\<rangle> p \<le> q \<longleftrightarrow> p \<le> [x| q"
proof -
have "|x\<rangle> p \<le> q \<longleftrightarrow> x \<cdot> p \<le> q \<cdot> x"
by (metis assms local.ads_d_def local.ds.fdemodalisation2)
thus ?thesis
by (metis assms domrangefix local.ardual.fbox_demodalisation3)
qed
lemma diamond_conjugation_var_1:
assumes "d p = p" and "d q = q"
shows "|x\<rangle> p \<le> ad q \<longleftrightarrow> \<langle>x| q \<le> ad p"
proof -
have "|x\<rangle> p \<le> ad q \<longleftrightarrow> x \<cdot> p \<le> ad q \<cdot> x"
by (metis assms local.ads_d_def local.ds.fdemodalisation2)
also have "... \<longleftrightarrow> q \<cdot> x \<le> x \<cdot> ad p"
by (metis assms local.ads_d_def local.kat_1_equiv_opp)
finally show ?thesis
by (metis assms box_diamond_galois_1 local.ads_d_def local.fbox_demodalisation3)
qed
lemma diamond_conjungation_aux:
assumes "d p = p" and "d q = q"
shows "\<langle>x| p \<le> ad q \<longleftrightarrow> q \<cdot> \<langle>x| p = 0"
apply standard
apply (metis assms(2) local.a_antitone' local.a_gla local.ads_d_def)
proof -
assume a1: "q \<cdot> \<langle>x| p = zero_class.zero"
have "ad (ad q) = q"
using assms(2) local.ads_d_def by fastforce
then show "\<langle>x| p \<le> ad q"
- using a1 by (metis (no_types) domrangefix local.a_gla local.ads_d_def local.antisym local.ardual.a_gla2 local.ardual.gla_1 local.ars_r_def local.bdia_def local.eq_refl)
+ using a1
+ by (metis a_closure' a_gla ardual.a_subid_aux1' bdia_def
+ dnsz.dsg4 dpdz.dsg1 dpdz.dsg3 ds.dsr4 order.trans)
qed
lemma diamond_conjugation:
assumes "d p = p" and "d q = q"
shows "p \<cdot> |x\<rangle> q = 0 \<longleftrightarrow> q \<cdot> \<langle>x| p = 0"
proof -
have "p \<cdot> |x\<rangle> q = 0 \<longleftrightarrow> |x\<rangle> q \<le> ad p"
by (metis assms(1) local.a_gla local.ads_d_def local.am2 local.fdia_def)
also have "... \<longleftrightarrow> \<langle>x| p \<le> ad q"
by (simp add: assms diamond_conjugation_var_1)
finally show ?thesis
by (simp add: assms diamond_conjungation_aux)
qed
lemma box_conjugation_var_1:
assumes "d p = p" and "d q = q"
shows "ad p \<le> [x| q \<longleftrightarrow> ad q \<le> |x] p"
by (metis assms box_diamond_galois_1 box_diamond_galois_2 diamond_conjugation_var_1 local.ads_d_def)
lemma box_diamond_cancellation_1: "d p = p \<Longrightarrow> p \<le> |x] \<langle>x| p"
using box_diamond_galois_1 domrangefix local.ars_r_def local.bdia_def by fastforce
lemma box_diamond_cancellation_2: "d p = p \<Longrightarrow> p \<le> [x| |x\<rangle> p"
by (metis box_diamond_galois_2 local.ads_d_def local.dpdz.domain_invol local.eq_refl local.fdia_def)
lemma box_diamond_cancellation_3: "d p = p \<Longrightarrow> |x\<rangle> [x| p \<le> p"
using box_diamond_galois_2 domrangefix local.ardual.fdia_fbox_de_morgan_2 local.ars_r_def local.bbox_def local.bdia_def by auto
lemma box_diamond_cancellation_4: "d p = p \<Longrightarrow> \<langle>x| |x] p \<le> p"
using box_diamond_galois_1 local.ads_d_def local.fbox_def local.fdia_def local.fdia_fbox_de_morgan_2 by auto
end
class modal_kleene_algebra = modal_semiring + kleene_algebra
begin
subclass antidomain_kleene_algebra ..
subclass antirange_kleene_algebra ..
end
end
diff --git a/thys/KAD/Modal_Kleene_Algebra_Applications.thy b/thys/KAD/Modal_Kleene_Algebra_Applications.thy
--- a/thys/KAD/Modal_Kleene_Algebra_Applications.thy
+++ b/thys/KAD/Modal_Kleene_Algebra_Applications.thy
@@ -1,567 +1,567 @@
(* Title: Applications of Modal Kleene Algebras
Author: Victor B. F. Gomes, Walter Guttmann, Peter Höfner, Georg Struth, Tjark Weber
Maintainer: Walter Guttmann <walter.guttman at canterbury.ac.nz>
Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Applications of Modal Kleene Algebras\<close>
theory Modal_Kleene_Algebra_Applications
imports Antidomain_Semiring
begin
text \<open>This file collects some applications of the theories developed so far. These are described
in~\cite{guttmannstruthweber11algmeth}.\<close>
context antidomain_kleene_algebra
begin
subsection \<open>A Reachability Result\<close>
text \<open>This example is taken from~\cite{desharnaismoellerstruth06kad}.\<close>
lemma opti_iterate_var [simp]: "|(ad y \<cdot> x)\<^sup>\<star>\<rangle> y = |x\<^sup>\<star>\<rangle> y"
-proof (rule antisym)
+proof (rule order.antisym)
show "|(ad y \<cdot> x)\<^sup>\<star>\<rangle> y \<le> |x\<^sup>\<star>\<rangle> y"
by (simp add: a_subid_aux1' ds.fd_iso2 star_iso)
have "d y + |x\<rangle> |(ad y \<cdot> x)\<^sup>\<star>\<rangle> y = d y + ad y \<cdot> |x\<rangle> |(ad y \<cdot> x)\<^sup>\<star>\<rangle> y"
using local.ads_d_def local.d_6 local.fdia_def by auto
also have "... = d y + |ad y \<cdot> x\<rangle> |(ad y \<cdot> x)\<^sup>\<star>\<rangle> y"
by (simp add: fdia_export_2)
finally have "d y + |x\<rangle> |(ad y \<cdot> x)\<^sup>\<star>\<rangle> y = |(ad y \<cdot> x)\<^sup>\<star>\<rangle> y"
by simp
thus "|x\<^sup>\<star>\<rangle> y \<le> |(ad y \<cdot> x)\<^sup>\<star>\<rangle> y"
using local.dka.fd_def local.dka.fdia_star_induct_eq by auto
qed
lemma opti_iterate [simp]: "d y + |(x \<cdot> ad y)\<^sup>\<star>\<rangle> |x\<rangle> y = |x\<^sup>\<star>\<rangle> y"
proof -
have "d y + |(x \<cdot> ad y)\<^sup>\<star>\<rangle> |x\<rangle> y = d y + |x\<rangle> |(ad y \<cdot> x)\<^sup>\<star>\<rangle> y"
by (metis local.conway.dagger_slide local.ds.fdia_mult)
also have "... = d y + |x\<rangle> |x\<^sup>\<star>\<rangle> y"
by simp
finally show ?thesis
by force
qed
lemma opti_iterate_var_2 [simp]: "d y + |ad y \<cdot> x\<rangle> |x\<^sup>\<star>\<rangle> y = |x\<^sup>\<star>\<rangle> y"
by (metis local.dka.fdia_star_unfold_var opti_iterate_var)
subsection \<open>Derivation of Segerberg's Formula\<close>
text \<open>This example is taken from~\cite{DesharnaisMoellerStruthLMCS}.\<close>
definition Alpha :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("A")
where "A x y = d (x \<cdot> y) \<cdot> ad y"
lemma A_dom [simp]: "d (A x y) = A x y"
using Alpha_def local.a_d_closed local.ads_d_def local.apd_d_def by auto
lemma A_fdia: "A x y = |x\<rangle> y \<cdot> ad y"
by (simp add: Alpha_def local.dka.fd_def)
lemma A_fdia_var: "A x y = |x\<rangle> d y \<cdot> ad y"
by (simp add: A_fdia)
lemma a_A: "ad (A x (ad y)) = |x] y + ad y"
using Alpha_def local.a_6 local.a_d_closed local.apd_d_def local.fbox_def by force
lemma fsegerberg [simp]: "d y + |x\<^sup>\<star>\<rangle> A x y = |x\<^sup>\<star>\<rangle> y"
-proof (rule antisym)
+proof (rule order.antisym)
have "d y + |x\<rangle> (d y + |x\<^sup>\<star>\<rangle> A x y) = d y + |x\<rangle> y + |x\<rangle> |x\<^sup>\<star>\<rangle> ( |x\<rangle> y \<cdot> ad y )"
by (simp add: A_fdia add_assoc local.ds.fdia_add1)
also have "... = d y + |x\<rangle> y \<cdot> ad y + |x\<rangle> |x\<^sup>\<star>\<rangle> ( |x\<rangle> y \<cdot> ad y )"
by (metis local.am2 local.d_6 local.dka.fd_def local.fdia_def)
finally have "d y + |x\<rangle> (d y + |x\<^sup>\<star>\<rangle> A x y) = d y + |x\<^sup>\<star>\<rangle> A x y"
by (metis A_dom A_fdia add_assoc local.dka.fdia_star_unfold_var)
thus "|x\<^sup>\<star>\<rangle> y \<le> d y + |x\<^sup>\<star>\<rangle> A x y"
by (metis local.a_d_add_closure local.ads_d_def local.dka.fdia_star_induct_eq local.fdia_def)
have "d y + |x\<^sup>\<star>\<rangle> A x y = d y + |x\<^sup>\<star>\<rangle> ( |x\<rangle> y \<cdot> ad y )"
by (simp add: A_fdia)
also have "... \<le> d y + |x\<^sup>\<star>\<rangle> |x\<rangle> y"
using local.dpdz.domain_1'' local.ds.fd_iso1 local.join.sup_mono by blast
finally show "d y + |x\<^sup>\<star>\<rangle> A x y \<le> |x\<^sup>\<star>\<rangle> y"
by simp
qed
lemma fbox_segerberg [simp]: "d y \<cdot> |x\<^sup>\<star>] ( |x] y + ad y ) = |x\<^sup>\<star>] y"
proof -
have "|x\<^sup>\<star>] ( |x] y + ad y ) = d ( |x\<^sup>\<star>] ( |x] y + ad y ) )"
using local.a_d_closed local.ads_d_def local.apd_d_def local.fbox_def by auto
hence f1: "|x\<^sup>\<star>] ( |x] y + ad y ) = ad ( |x\<^sup>\<star>\<rangle> (A x (ad y)) )"
by (simp add: a_A local.fdia_fbox_de_morgan_2)
have "ad y + |x\<^sup>\<star>\<rangle> (A x (ad y)) = |x\<^sup>\<star>\<rangle> ad y"
by (metis fsegerberg local.a_d_closed local.ads_d_def local.apd_d_def)
thus ?thesis
by (metis f1 local.ads_d_def local.ans4 local.fbox_simp local.fdia_fbox_de_morgan_2)
qed
subsection \<open>Wellfoundedness and Loeb's Formula\<close>
text \<open>This example is taken from~\cite{DesharnaisStruthSCP}.\<close>
definition Omega :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<Omega>")
where "\<Omega> x y = d y \<cdot> ad (x \<cdot> y)"
text \<open>If $y$ is a set, then $\Omega(x,y)$ describes those elements in $y$ from which no further $x$ transitions are possible.\<close>
lemma omega_fdia: "\<Omega> x y = d y \<cdot> ad ( |x\<rangle> y )"
using Omega_def local.a_d_closed local.ads_d_def local.apd_d_def local.dka.fd_def by auto
lemma omega_fbox: "\<Omega> x y = d y \<cdot> |x] (ad y)"
by (simp add: fdia_fbox_de_morgan_2 omega_fdia)
lemma omega_absorb1 [simp]: "\<Omega> x y \<cdot> ad ( |x\<rangle> y ) = \<Omega> x y"
by (simp add: mult_assoc omega_fdia)
lemma omega_absorb2 [simp]: "\<Omega> x y \<cdot> ad (x \<cdot> y) = \<Omega> x y"
by (simp add: Omega_def mult_assoc)
lemma omega_le_1: "\<Omega> x y \<le> d y"
by (simp add: Omega_def d_a_galois1)
lemma omega_subid: "\<Omega> x (d y) \<le> d y"
by (simp add: Omega_def local.a_subid_aux2)
lemma omega_le_2: "\<Omega> x y \<le> ad ( |x\<rangle> y )"
by (simp add: local.dka.dom_subid_aux2 omega_fdia)
lemma omega_dom [simp]: "d (\<Omega> x y) = \<Omega> x y"
using Omega_def local.a_d_closed local.ads_d_def local.apd_d_def by auto
lemma a_omega: "ad (\<Omega> x y) = ad y + |x\<rangle> y"
by (simp add: Omega_def local.a_6 local.ds.fd_def)
lemma omega_fdia_3 [simp]: "d y \<cdot> ad (\<Omega> x y) = d y \<cdot> |x\<rangle> y"
using Omega_def local.ads_d_def local.fdia_def local.s4 by presburger
lemma omega_zero_equiv_1: "\<Omega> x y = 0 \<longleftrightarrow> d y \<le> |x\<rangle> y"
by (simp add: Omega_def local.a_gla local.ads_d_def local.fdia_def)
definition Loebian :: "'a \<Rightarrow> bool"
where "Loebian x = (\<forall>y. |x\<rangle> y \<le> |x\<rangle> \<Omega> x y)"
definition PreLoebian :: "'a \<Rightarrow> bool"
where "PreLoebian x = (\<forall>y. d y \<le> |x\<^sup>\<star>\<rangle> \<Omega> x y)"
definition Noetherian :: "'a \<Rightarrow> bool"
where "Noetherian x = (\<forall>y. \<Omega> x y = 0 \<longrightarrow> d y = 0)"
lemma noetherian_alt: "Noetherian x \<longleftrightarrow> (\<forall>y. d y \<le> |x\<rangle> y \<longrightarrow> d y = 0)"
by (simp add: Noetherian_def omega_zero_equiv_1)
lemma Noetherian_iff_PreLoebian: "Noetherian x \<longleftrightarrow> PreLoebian x"
proof
assume hyp: "Noetherian x"
{
fix y
have "d y \<cdot> ad ( |x\<^sup>\<star>\<rangle> \<Omega> x y ) = d y \<cdot> ad (\<Omega> x y + |x\<rangle> |x\<^sup>\<star>\<rangle> \<Omega> x y)"
by (metis local.dka.fdia_star_unfold_var omega_dom)
also have "... = d y \<cdot> ad (\<Omega> x y) \<cdot> ad ( |x\<rangle> |x\<^sup>\<star>\<rangle> \<Omega> x y )"
using ans4 mult_assoc by presburger
also have "... \<le> |x\<rangle> d y \<cdot> ad ( |x\<rangle> |x\<^sup>\<star>\<rangle> \<Omega> x y )"
by (simp add: local.dka.dom_subid_aux2 local.mult_isor)
also have "... \<le> |x\<rangle> (d y \<cdot> ad ( |x\<^sup>\<star>\<rangle> \<Omega> x y ))"
by (simp add: local.dia_diff)
finally have "d y \<cdot> ad ( |x\<^sup>\<star>\<rangle> \<Omega> x y ) \<le> |x\<rangle> (d y \<cdot> ad ( |x\<^sup>\<star>\<rangle> \<Omega> x y ))"
by blast
hence "d y \<cdot> ad ( |x\<^sup>\<star>\<rangle> \<Omega> x y ) = 0"
by (metis hyp local.ads_d_def local.dpdz.dom_mult_closed local.fdia_def noetherian_alt)
hence "d y \<le> |x\<^sup>\<star>\<rangle> \<Omega> x y"
by (simp add: local.a_gla local.ads_d_def local.dka.fd_def)
}
thus "PreLoebian x"
using PreLoebian_def by blast
next
assume a: "PreLoebian x"
{
fix y
assume b: "\<Omega> x y = 0"
hence "d y \<le> |x\<rangle> d y"
using omega_zero_equiv_1 by auto
hence "d y = 0"
by (metis (no_types) PreLoebian_def a b a_one a_zero add_zeror annir fdia_def less_eq_def)
}
thus "Noetherian x"
by (simp add: Noetherian_def)
qed
lemma Loebian_imp_Noetherian: "Loebian x \<Longrightarrow> Noetherian x"
proof -
assume a: "Loebian x"
{
fix y
assume b: "\<Omega> x y = 0"
hence "d y \<le> |x\<rangle> d y"
using omega_zero_equiv_1 by auto
also have "... \<le> |x\<rangle> (\<Omega> x y)"
using Loebian_def a by auto
finally have "d y = 0"
- by (simp add: b local.antisym local.fdia_def)
+ by (simp add: b order.antisym fdia_def)
}
thus "Noetherian x"
by (simp add: Noetherian_def)
qed
lemma d_transitive: "(\<forall>y. |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y) \<Longrightarrow> (\<forall>y. |x\<rangle> y = |x\<^sup>\<star>\<rangle> |x\<rangle> y)"
proof -
assume a: "\<forall>y. |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y"
{
fix y
have "|x\<rangle> y + |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y"
by (simp add: a)
hence "|x\<^sup>\<star>\<rangle> |x\<rangle> y \<le> |x\<rangle> y"
using local.dka.fd_def local.dka.fdia_star_induct_var by auto
have "|x\<rangle> y \<le> |x\<^sup>\<star>\<rangle> |x\<rangle> y"
using local.dka.fd_def local.order_prop opti_iterate by force
}
thus ?thesis
- using a local.antisym local.dka.fd_def local.dka.fdia_star_induct_var by auto
+ using a order.antisym dka.fd_def dka.fdia_star_induct_var by auto
qed
lemma d_transitive_var: "(\<forall>y. |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y) \<Longrightarrow> (\<forall>y. |x\<rangle> y = |x\<rangle> |x\<^sup>\<star>\<rangle> y)"
proof -
assume "\<forall>y. |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y"
then have "\<forall>a. |x\<rangle> |x\<^sup>\<star>\<rangle> a = |x\<rangle> a"
by (metis (full_types) d_transitive local.dka.fd_def local.dka.fdia_d_simp local.star_slide_var mult_assoc)
then show ?thesis
by presburger
qed
lemma d_transitive_PreLoebian_imp_Loebian: "(\<forall>y. |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y) \<Longrightarrow> PreLoebian x \<Longrightarrow> Loebian x"
proof -
assume wt: "(\<forall>y. |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y)"
and "PreLoebian x"
hence "\<forall>y. |x\<rangle> y \<le> |x\<rangle> |x\<^sup>\<star>\<rangle> \<Omega> x y"
using PreLoebian_def local.ads_d_def local.dka.fd_def local.ds.fd_iso1 by auto
hence "\<forall>y. |x\<rangle> y \<le> |x\<rangle> \<Omega> x y"
by (metis wt d_transitive_var)
then show "Loebian x"
using Loebian_def fdia_def by auto
qed
lemma d_transitive_Noetherian_iff_Loebian: "\<forall>y. |x\<rangle> |x\<rangle> y \<le> |x\<rangle> y \<Longrightarrow> Noetherian x \<longleftrightarrow> Loebian x"
using Loebian_imp_Noetherian Noetherian_iff_PreLoebian d_transitive_PreLoebian_imp_Loebian by blast
lemma Loeb_iff_box_Loeb: "Loebian x \<longleftrightarrow> (\<forall>y. |x] (ad ( |x] y ) + d y) \<le> |x] y)"
proof -
have "Loebian x \<longleftrightarrow> (\<forall>y. |x\<rangle> y \<le> |x\<rangle> (d y \<cdot> |x] (ad y)))"
using Loebian_def omega_fbox by auto
also have "... \<longleftrightarrow> (\<forall>y. ad ( |x\<rangle> (d y \<cdot> |x] (ad y)) ) \<le> ad ( |x\<rangle> y ))"
using a_antitone' fdia_def by fastforce
also have "... \<longleftrightarrow> (\<forall>y. |x] ad (d y \<cdot> |x] (ad y)) \<le> |x] (ad y))"
by (simp add: fdia_fbox_de_morgan_2)
also have "... \<longleftrightarrow> (\<forall>y. |x] (d (ad y) + ad ( |x] (ad y) )) \<le> |x] (ad y))"
by (simp add: local.ads_d_def local.fbox_def)
finally show ?thesis
by (metis add_commute local.a_d_closed local.ads_d_def local.apd_d_def local.fbox_def)
qed
end
subsection \<open>Divergence Kleene Algebras and Separation of Termination\<close>
text \<open>The notion of divergence has been added to modal Kleene algebras in~\cite{DesharnaisMoellerStruthLMCS}.
More facts about divergence could be added in the future. Some could be adapted from omega algebras.\<close>
class nabla_op =
fixes nabla :: "'a \<Rightarrow> 'a" ("\<nabla>_" [999] 1000)
class fdivergence_kleene_algebra = antidomain_kleene_algebra + nabla_op +
assumes nabla_closure [simp]: "d \<nabla> x = \<nabla> x"
and nabla_unfold: "\<nabla> x \<le> |x\<rangle> \<nabla> x"
and nabla_coinduction: "d y \<le> |x\<rangle> y + d z \<Longrightarrow> d y \<le> \<nabla> x + |x\<^sup>\<star>\<rangle> z"
begin
lemma nabla_coinduction_var: "d y \<le> |x\<rangle> y \<Longrightarrow> d y \<le> \<nabla> x"
proof -
assume "d y \<le> |x\<rangle> y"
hence "d y \<le> |x\<rangle> y + d 0"
by simp
hence "d y \<le> \<nabla> x + |x\<^sup>\<star>\<rangle> 0"
using nabla_coinduction by blast
thus "d y \<le> \<nabla> x"
by (simp add: fdia_def)
qed
lemma nabla_unfold_eq [simp]: "|x\<rangle> \<nabla> x = \<nabla> x"
proof -
have f1: "ad (ad (x \<cdot> \<nabla>x)) = ad (ad (x \<cdot> \<nabla>x)) + \<nabla>x"
using local.ds.fd_def local.join.sup.order_iff local.nabla_unfold by presburger
then have "ad (ad (x \<cdot> \<nabla>x)) \<cdot> ad (ad \<nabla>x) = \<nabla>x"
by (metis local.ads_d_def local.dpdz.dns5 local.dpdz.dsg4 local.join.sup_commute local.nabla_closure)
then show ?thesis
using f1 by (metis local.ads_d_def local.ds.fd_def local.ds.fd_subdist_2 local.join.sup.order_iff local.join.sup_commute nabla_coinduction_var)
qed
lemma nabla_subdist: "\<nabla> x \<le> \<nabla> (x + y)"
proof -
have "d \<nabla> x \<le> \<nabla> (x + y)"
by (metis local.ds.fd_iso2 local.join.sup.cobounded1 local.nabla_closure nabla_coinduction_var nabla_unfold_eq)
thus ?thesis
by simp
qed
lemma nabla_iso: "x \<le> y \<Longrightarrow> \<nabla> x \<le> \<nabla> y"
by (metis less_eq_def nabla_subdist)
lemma nabla_omega: "\<Omega> x (d y) = 0 \<Longrightarrow> d y \<le> \<nabla> x"
using omega_zero_equiv_1 nabla_coinduction_var by fastforce
lemma nabla_noether: "\<nabla> x = 0 \<Longrightarrow> Noetherian x"
using local.join.le_bot local.noetherian_alt nabla_coinduction_var by blast
lemma nabla_preloeb: "\<nabla> x = 0 \<Longrightarrow> PreLoebian x"
using Noetherian_iff_PreLoebian nabla_noether by auto
lemma star_nabla_1 [simp]: "|x\<^sup>\<star>\<rangle> \<nabla> x = \<nabla> x"
-proof (rule antisym)
+proof (rule order.antisym)
show "|x\<^sup>\<star>\<rangle> \<nabla> x \<le> \<nabla> x"
- by (metis local.dka.fdia_star_induct_var local.eq_iff local.nabla_closure nabla_unfold_eq)
+ by (metis dka.fdia_star_induct_var order.eq_iff nabla_closure nabla_unfold_eq)
show "\<nabla> x \<le> |x\<^sup>\<star>\<rangle> \<nabla> x"
- by (metis local.ds.fd_iso2 local.star_ext nabla_unfold_eq)
+ by (metis ds.fd_iso2 star_ext nabla_unfold_eq)
qed
lemma nabla_sum_expand [simp]: "|x\<rangle> \<nabla> (x + y) + |y\<rangle> \<nabla> (x + y) = \<nabla> (x + y)"
proof -
have "d ((x + y) \<cdot> \<nabla>(x + y)) = \<nabla>(x + y)"
using local.dka.fd_def nabla_unfold_eq by presburger
thus ?thesis
by (simp add: local.dka.fd_def)
qed
lemma wagner_3:
assumes "d z + |x\<rangle> \<nabla> (x + y) = \<nabla> (x + y)"
shows "\<nabla> (x + y) = \<nabla> x + |x\<^sup>\<star>\<rangle> z"
-proof (rule antisym)
+proof (rule order.antisym)
have "d \<nabla>(x + y) \<le> d z + |x\<rangle> \<nabla>(x + y)"
by (simp add: assms)
thus "\<nabla> (x + y) \<le> \<nabla> x + |x\<^sup>\<star>\<rangle> z"
by (metis add_commute nabla_closure nabla_coinduction)
have "d z + |x\<rangle> \<nabla> (x + y) \<le> \<nabla> (x + y)"
using assms by auto
hence "|x\<^sup>\<star>\<rangle> z \<le> \<nabla> (x + y)"
by (metis local.dka.fdia_star_induct local.nabla_closure)
thus "\<nabla> x + |x\<^sup>\<star>\<rangle> z \<le> \<nabla> (x + y)"
by (simp add: sup_least nabla_subdist)
qed
lemma nabla_sum_unfold [simp]: "\<nabla> x + |x\<^sup>\<star>\<rangle> |y\<rangle> \<nabla> (x + y) = \<nabla> (x + y)"
proof -
have "\<nabla> (x + y) = |x\<rangle> \<nabla> (x + y) + |y\<rangle> \<nabla> (x + y)"
by simp
thus ?thesis
by (metis add_commute local.dka.fd_def local.ds.fd_def local.ds.fdia_d_simp wagner_3)
qed
lemma nabla_separation: "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star> \<Longrightarrow> (\<nabla> (x + y) = \<nabla> x + |x\<^sup>\<star>\<rangle> \<nabla> y)"
-proof (rule antisym)
+proof (rule order.antisym)
assume quasi_comm: "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
hence a: "y\<^sup>\<star> \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
using quasicomm_var by blast
have "\<nabla> (x + y) = \<nabla> y + |y\<^sup>\<star>\<cdot> x\<rangle> \<nabla> (x + y)"
by (metis local.ds.fdia_mult local.join.sup_commute nabla_sum_unfold)
also have "... \<le> \<nabla> y + |x \<cdot> (x + y)\<^sup>\<star>\<rangle> \<nabla> (x + y)"
using a local.ds.fd_iso2 local.join.sup.mono by blast
also have "... = \<nabla> y + |x\<rangle> |(x + y)\<^sup>\<star>\<rangle> \<nabla> (x + y)"
by (simp add: fdia_def mult_assoc)
finally have "\<nabla> (x + y) \<le> \<nabla> y + |x\<rangle> \<nabla> (x + y)"
by (metis star_nabla_1)
thus "\<nabla> (x + y) \<le> \<nabla> x + |x\<^sup>\<star>\<rangle> \<nabla> y"
by (metis add_commute nabla_closure nabla_coinduction)
next
have "\<nabla> x + |x\<^sup>\<star>\<rangle> \<nabla> y = \<nabla> x + |x\<^sup>\<star>\<rangle> |y\<rangle> \<nabla> y"
by simp
also have "... = \<nabla> x + |x\<^sup>\<star> \<cdot> y\<rangle> \<nabla> y"
by (simp add: fdia_def mult_assoc)
also have "... \<le> \<nabla> x + |x\<^sup>\<star> \<cdot> y\<rangle> \<nabla> (x + y)"
using dpdz.dom_iso eq_refl fdia_def join.sup_ge2 join.sup_mono mult_isol nabla_iso by presburger
also have "... = \<nabla> x + |x\<^sup>\<star>\<rangle> |y\<rangle> \<nabla> (x + y)"
by (simp add: fdia_def mult_assoc)
finally show "\<nabla> x + |x\<^sup>\<star>\<rangle> \<nabla> y \<le> \<nabla> (x + y)"
by (metis nabla_sum_unfold)
qed
text \<open>The next lemma is a separation of termination theorem by Bachmair and Dershowitz~\cite{bachmair86commutation}.\<close>
lemma bachmair_dershowitz: "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star> \<Longrightarrow> \<nabla> x + \<nabla> y = 0 \<longleftrightarrow> \<nabla> (x + y) = 0"
proof -
assume quasi_comm: "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
have "\<forall>x. |x\<rangle> 0 = 0"
by (simp add: fdia_def)
hence "\<nabla>y = 0 \<Longrightarrow> \<nabla>x + \<nabla>y = 0 \<longleftrightarrow> \<nabla>(x + y) = 0"
using quasi_comm nabla_separation by presburger
thus ?thesis
using add_commute local.join.le_bot nabla_subdist by fastforce
qed
text \<open>The next lemma is a more complex separation of termination theorem by Doornbos,
Backhouse and van der Woude~\cite{doornbos97calculational}.\<close>
lemma separation_of_termination:
assumes "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star> + y"
shows "\<nabla> x + \<nabla> y = 0 \<longleftrightarrow> \<nabla> (x + y) = 0"
proof
assume xy_wf: "\<nabla> x + \<nabla> y = 0"
hence x_preloeb: "\<nabla> (x + y) \<le> |x\<^sup>\<star>\<rangle> \<Omega> x (\<nabla> (x + y))"
by (metis PreLoebian_def no_trivial_inverse nabla_closure nabla_preloeb)
hence y_div: "\<nabla> y = 0"
using add_commute no_trivial_inverse xy_wf by blast
have "\<nabla> (x + y) \<le> |y\<rangle> \<nabla> (x + y) + |x\<rangle> \<nabla> (x + y)"
by (simp add: local.join.sup_commute)
hence "\<nabla> (x + y) \<cdot> ad ( |x\<rangle> \<nabla> (x + y) ) \<le> |y\<rangle> \<nabla>(x + y)"
by (simp add: local.da_shunt1 local.dka.fd_def local.join.sup_commute)
hence "\<Omega> x \<nabla>(x + y) \<le> |y\<rangle> \<nabla> (x + y)"
by (simp add: fdia_def omega_fdia)
also have "... \<le> |y\<rangle> |x\<^sup>\<star>\<rangle> (\<Omega> x \<nabla>(x + y))"
using local.dpdz.dom_iso local.ds.fd_iso1 x_preloeb by blast
also have "... = |y \<cdot> x\<^sup>\<star>\<rangle> (\<Omega> x \<nabla>(x + y))"
by (simp add: fdia_def mult_assoc)
also have "... \<le> |x \<cdot> (x + y)\<^sup>\<star> + y\<rangle> (\<Omega> x \<nabla>(x + y))"
using assms local.ds.fd_iso2 local.lazycomm_var by blast
also have "... = |x \<cdot> (x + y)\<^sup>\<star>\<rangle> (\<Omega> x \<nabla>(x + y)) + |y\<rangle> (\<Omega> x \<nabla>(x + y))"
by (simp add: local.dka.fd_def)
also have "... \<le> |(x \<cdot> (x + y)\<^sup>\<star>)\<rangle> \<nabla>(x + y) + |y\<rangle> (\<Omega> x \<nabla>(x + y))"
using local.add_iso local.dpdz.domain_1'' local.ds.fd_iso1 local.omega_fdia by auto
also have "... \<le> |x\<rangle> |(x + y)\<^sup>\<star>\<rangle> \<nabla>(x + y) + |y\<rangle> (\<Omega> x \<nabla>(x + y))"
by (simp add: fdia_def mult_assoc)
finally have "\<Omega> x \<nabla>(x + y) \<le> |x\<rangle> \<nabla>(x + y) + |y\<rangle> (\<Omega> x \<nabla>(x + y))"
by (metis star_nabla_1)
hence "\<Omega> x \<nabla>(x + y) \<cdot> ad ( |x\<rangle> \<nabla>(x + y) ) \<le> |y\<rangle> (\<Omega> x \<nabla>(x + y))"
by (simp add: local.da_shunt1 local.dka.fd_def)
hence "\<Omega> x \<nabla>(x + y) \<le> |y\<rangle> (\<Omega> x \<nabla>(x + y))"
by (simp add: omega_fdia mult_assoc)
hence "(\<Omega> x \<nabla>(x + y)) = 0"
by (metis noetherian_alt omega_dom nabla_noether y_div)
thus "\<nabla> (x + y) = 0"
using local.dka.fd_def local.join.le_bot x_preloeb by auto
next
assume "\<nabla> (x + y) = 0"
thus "(\<nabla> x) + (\<nabla> y) = 0"
by (metis local.join.le_bot local.join.sup.order_iff local.join.sup_commute nabla_subdist)
qed
text \<open>The final examples can be found in~\cite{guttmannstruthweber11algmeth}.\<close>
definition T :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<leadsto> _ \<leadsto> _" [61,61,61] 60)
where "p\<leadsto>x\<leadsto>q \<equiv> ad p + |x] d q"
lemma T_d [simp]: "d (p\<leadsto>x\<leadsto>q) = p\<leadsto>x\<leadsto>q"
using T_def local.a_d_add_closure local.fbox_def by auto
lemma T_p: "d p \<cdot> (p\<leadsto>x\<leadsto>q) = d p \<cdot> |x] d q"
proof -
have "d p \<cdot> (p\<leadsto>x\<leadsto>q) = ad (ad p + ad (p\<leadsto>x\<leadsto>q))"
using T_d local.ads_d_def by auto
thus ?thesis
using T_def add_commute local.a_mult_add local.ads_d_def by auto
qed
lemma T_a [simp]: "ad p \<cdot> (p\<leadsto>x\<leadsto>q) = ad p"
by (metis T_d T_def local.a_d_closed local.ads_d_def local.apd_d_def local.dpdz.dns5 local.join.sup.left_idem)
lemma T_seq: "(p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>y\<leadsto>s) \<le> p\<leadsto>x \<cdot> y\<leadsto>s"
proof -
have f1: "|x] q = |x] d q"
using local.fbox_simp by auto
have "ad p \<cdot> ad (x \<cdot> ad (q\<leadsto>y\<leadsto>s)) + |x] d q \<cdot> |x] (ad q + |y] d s) \<le> ad p + |x] d q \<cdot> |x] (ad q + |y] d s)"
using local.a_subid_aux2 local.add_iso by blast
hence "(p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>y\<leadsto>s) \<le> ad p + |x](d q \<cdot> (q\<leadsto>y\<leadsto>s))"
by (metis T_d T_def f1 local.distrib_right' local.fbox_add1 local.fbox_def)
also have "... = ad p + |x](d q \<cdot> (ad q + |y] d s))"
by (simp add: T_def)
also have "... = ad p + |x](d q \<cdot> |y] d s)"
using T_def T_p by auto
also have "... \<le> ad p + |x] |y] d s"
- by (metis (no_types, lifting) local.dka.dom_subid_aux2 local.dka.dsg3 local.eq_iff local.fbox_iso local.join.sup.mono)
+ by (metis (no_types, lifting) dka.dom_subid_aux2 dka.dsg3 order.eq_iff fbox_iso join.sup.mono)
finally show ?thesis
by (simp add: T_def fbox_mult)
qed
lemma T_square: "(p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>x\<leadsto>p) \<le> p\<leadsto>x\<cdot>x\<leadsto>p"
by (simp add: T_seq)
lemma T_segerberg [simp]: "d p \<cdot> |x\<^sup>\<star>](p\<leadsto>x\<leadsto>p) = |x\<^sup>\<star>] d p"
using T_def add_commute local.fbox_segerberg local.fbox_simp by force
lemma lookahead [simp]: "|x\<^sup>\<star>](d p \<cdot> |x] d p) = |x\<^sup>\<star>] d p"
by (metis (full_types) local.dka.dsg3 local.fbox_add1 local.fbox_mult local.fbox_simp local.fbox_star_unfold_var local.star_slide_var local.star_trans_eq)
lemma alternation: "d p \<cdot> |x\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p)) = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> (q\<leadsto>x\<leadsto>p)) \<cdot> |x\<cdot>(x\<cdot>x)\<^sup>\<star>](d q \<cdot> (p\<leadsto>x\<leadsto>q))"
proof -
have fbox_simp_2: "\<And>x p. |x]p = d( |x] p )"
using local.a_d_closed local.ads_d_def local.apd_d_def local.fbox_def by fastforce
have "|(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>x\<leadsto>p) \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](p\<leadsto>x\<leadsto>q)) \<le> |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>x\<leadsto>p))"
using local.dka.domain_1'' local.fbox_iso local.order_trans by blast
also have "... \<le> |(x\<cdot>x)\<^sup>\<star>](p\<leadsto>x\<cdot>x\<leadsto>p)"
using T_seq local.dka.dom_iso local.fbox_iso by blast
finally have 1: "|(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>x\<leadsto>p) \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](p\<leadsto>x\<leadsto>q)) \<le> |(x\<cdot>x)\<^sup>\<star>](p\<leadsto>x\<cdot>x\<leadsto>p)".
have "d p \<cdot> |x\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p)) = d p \<cdot> |1+x] |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p))"
by (metis (full_types) fbox_mult meyer_1)
also have "... = d p \<cdot> |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p)) \<cdot> |x\<cdot>(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p))"
using fbox_simp_2 fbox_mult fbox_add2 mult_assoc by auto
also have "... = d p \<cdot> |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p)) \<cdot> |(x\<cdot>x)\<^sup>\<star>\<cdot>x]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p))"
by (simp add: star_slide)
also have "... = d p \<cdot> |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p)) \<cdot> |(x\<cdot>x)\<^sup>\<star>] |x]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p))"
by (simp add: fbox_mult)
also have "... = d p \<cdot> |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x]((p\<leadsto>x\<leadsto>q) \<cdot> (q\<leadsto>x\<leadsto>p)))"
by (metis T_d fbox_simp_2 local.dka.dom_mult_closed local.fbox_add1 mult_assoc)
also have "... = d p \<cdot> |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>x\<leadsto>p) \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](p\<leadsto>x\<leadsto>q))"
proof -
have f1: "d ((q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q)) = (q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q)"
by (metis (full_types) T_d fbox_simp_2 local.dka.dsg3)
then have "|(x \<cdot> x)\<^sup>\<star>] (d ( |x] (q \<leadsto> x \<leadsto> p)) \<cdot> ((q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q))) = |(x \<cdot> x)\<^sup>\<star>] d ( |x] (q \<leadsto> x \<leadsto> p)) \<cdot> |(x \<cdot> x)\<^sup>\<star>] ((q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q))"
by (metis (full_types) fbox_simp_2 local.fbox_add1)
then have f2: "|(x \<cdot> x)\<^sup>\<star>] (d ( |x] (q \<leadsto> x \<leadsto> p)) \<cdot> ((q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q))) = ad ((x \<cdot> x)\<^sup>\<star> \<cdot> ad ((q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q)) + (x \<cdot> x)\<^sup>\<star> \<cdot> ad (d ( |x] (q \<leadsto> x \<leadsto> p))))"
by (simp add: add_commute local.fbox_def)
have "d ( |x] (p \<leadsto> x \<leadsto> q)) \<cdot> d ( |x] (q \<leadsto> x \<leadsto> p)) = |x] ((p \<leadsto> x \<leadsto> q) \<cdot> (q \<leadsto> x \<leadsto> p))"
by (metis (no_types) T_d fbox_simp_2 local.fbox_add1)
then have "d ((q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q)) \<cdot> d (d ( |x] (q \<leadsto> x \<leadsto> p))) = (q \<leadsto> x \<leadsto> p) \<cdot> |x] ((p \<leadsto> x \<leadsto> q) \<cdot> (q \<leadsto> x \<leadsto> p))"
using f1 fbox_simp_2 mult_assoc by force
then have "|(x \<cdot> x)\<^sup>\<star>] (d ( |x] (q \<leadsto> x \<leadsto> p)) \<cdot> ((q \<leadsto> x \<leadsto> p) \<cdot> |x] (p \<leadsto> x \<leadsto> q))) = |(x \<cdot> x)\<^sup>\<star>] ((q \<leadsto> x \<leadsto> p) \<cdot> |x] ((p \<leadsto> x \<leadsto> q) \<cdot> (q \<leadsto> x \<leadsto> p)))"
using f2 by (metis (no_types) local.ans4 local.fbox_add1 local.fbox_def)
then show ?thesis
by (metis (no_types) T_d fbox_simp_2 local.dka.dsg3 local.fbox_add1 mult_assoc)
qed
also have "... = d p \<cdot> |(x\<cdot>x)\<^sup>\<star>](p\<leadsto>x\<cdot>x\<leadsto>p) \<cdot> |(x\<cdot>x)\<^sup>\<star>]((p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>x\<leadsto>p) \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](p\<leadsto>x\<leadsto>q))" using "1"
- by (metis fbox_simp_2 local.dka.dns5 local.dka.dsg4 local.join.sup.absorb2 mult_assoc)
+ by (metis fbox_simp_2 dka.dns5 dka.dsg4 join.sup.absorb2 mult_assoc)
also have "... = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> (p\<leadsto>x\<leadsto>q) \<cdot> |x](q\<leadsto>x\<leadsto>p) \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](p\<leadsto>x\<leadsto>q))"
using T_segerberg local.a_d_closed local.ads_d_def local.apd_d_def local.distrib_left local.fbox_def mult_assoc by auto
also have "... = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> |x] d q \<cdot> |x](q\<leadsto>x\<leadsto>p) \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](p\<leadsto>x\<leadsto>q))"
by (simp add: T_p)
also have "... = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> |x] d q \<cdot> |x] |x] d p \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](p\<leadsto>x\<leadsto>q))"
by (metis T_d T_p fbox_simp_2 fbox_add1 fbox_simp mult_assoc)
also have "... = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> |x\<cdot>x] d p \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x] d q \<cdot> |x](p\<leadsto>x\<leadsto>q))"
proof -
have f1: "ad (x \<cdot> ad ( |x] d p)) = |x \<cdot> x] d p"
using local.fbox_def local.fbox_mult by presburger
have f2: "ad (d q \<cdot> d (x \<cdot> ad (d p))) = q \<leadsto> x \<leadsto> p"
by (simp add: T_def local.a_de_morgan_var_4 local.fbox_def)
have "ad q + |x] d p = ad (d q \<cdot> d (x \<cdot> ad (d p)))"
by (simp add: local.a_de_morgan_var_4 local.fbox_def)
then have "ad (x \<cdot> ad ( |x] d p)) \<cdot> ((q \<leadsto> x \<leadsto> p) \<cdot> |x] d q) = ad (x \<cdot> ad ( |x] d p)) \<cdot> ad (x \<cdot> ad (d q)) \<cdot> (ad q + |x] d p)"
using f2 by (metis (no_types) local.am2 local.fbox_def mult_assoc)
then show ?thesis
using f1 by (simp add: T_def local.am2 local.fbox_def mult_assoc)
qed
also have "... = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> |x\<cdot>x] d p \<cdot> (q\<leadsto>x\<leadsto>p) \<cdot> |x](d q \<cdot> (p\<leadsto>x\<leadsto>q)))"
using local.a_d_closed local.ads_d_def local.apd_d_def local.distrib_left local.fbox_def mult_assoc by auto
also have "... = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> |x\<cdot>x] d p) \<cdot> |(x\<cdot>x)\<^sup>\<star>](q\<leadsto>x\<leadsto>p) \<cdot> |(x\<cdot>x)\<^sup>\<star>] |x](d q \<cdot> (p\<leadsto>x\<leadsto>q))"
by (metis T_d fbox_simp_2 local.dka.dom_mult_closed local.fbox_add1)
also have "... = |(x\<cdot>x)\<^sup>\<star>](d p \<cdot> (q\<leadsto>x\<leadsto>p)) \<cdot> |(x\<cdot>x)\<^sup>\<star>] |x] (d q \<cdot> (p\<leadsto>x\<leadsto>q))"
by (metis T_d local.fbox_add1 local.fbox_simp lookahead)
finally show ?thesis
by (metis fbox_mult star_slide)
qed
lemma "|(x\<cdot>x)\<^sup>\<star>] d p \<cdot> |x\<cdot>(x\<cdot>x)\<^sup>\<star>] ad p = d p \<cdot> |x\<^sup>\<star>]((p\<leadsto>x\<leadsto>ad p) \<cdot> (ad p\<leadsto>x\<leadsto>p))"
using alternation local.a_d_closed local.ads_d_def local.apd_d_def by auto
lemma "|x\<^sup>\<star>] d p = d p \<cdot> |x\<^sup>\<star>](p\<leadsto>x\<leadsto>p)"
by (simp add: alternation)
end
end
diff --git a/thys/KAT_and_DRA/SingleSorted/Conway_Tests.thy b/thys/KAT_and_DRA/SingleSorted/Conway_Tests.thy
--- a/thys/KAT_and_DRA/SingleSorted/Conway_Tests.thy
+++ b/thys/KAT_and_DRA/SingleSorted/Conway_Tests.thy
@@ -1,67 +1,67 @@
(* Title: Kleene algebra with tests
Author: Alasdair Armstrong, Victor B. F. Gomes, Georg Struth
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
*)
section \<open>Pre-Conway Algebra with Tests\<close>
theory Conway_Tests
imports Kleene_Algebra.Conway Test_Dioid
begin
class near_conway_zerol_tests = near_conway_zerol + test_near_semiring_zerol_distrib
begin
lemma n_preserve1_var: "n x \<cdot> y \<le> n x \<cdot> y \<cdot> n x \<Longrightarrow> n x \<cdot> (n x \<cdot> y + t x \<cdot> z)\<^sup>\<dagger> \<le> (n x \<cdot> y)\<^sup>\<dagger> \<cdot> n x"
proof -
assume a: "n x \<cdot> y \<le> n x \<cdot> y \<cdot> n x"
have "n x \<cdot> (n x \<cdot> y + t x \<cdot> z) = n x \<cdot> y"
by (metis (no_types) local.add_zeror local.annil local.n_left_distrib local.n_mult_idem local.test_mult_comp mult_assoc)
hence "n x \<cdot> (n x \<cdot> y + t x \<cdot> z) \<le> n x \<cdot> y \<cdot> n x"
by (simp add: a)
thus " n x \<cdot> (n x \<cdot> y + t x \<cdot> z)\<^sup>\<dagger> \<le> (n x \<cdot> y)\<^sup>\<dagger> \<cdot> n x"
by (simp add: local.dagger_simr)
qed
lemma test_preserve1_var: "test p \<Longrightarrow> p \<cdot> x \<le> p \<cdot> x \<cdot> p \<Longrightarrow> p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger> \<le> (p \<cdot> x)\<^sup>\<dagger> \<cdot> p"
by (metis local.test_double_comp_var n_preserve1_var)
end
class test_pre_conway = pre_conway + test_pre_dioid_zerol
begin
subclass near_conway_zerol_tests
by (unfold_locales)
lemma test_preserve: "test p \<Longrightarrow> p \<cdot> x \<le> p \<cdot> x \<cdot> p \<Longrightarrow> p \<cdot> x\<^sup>\<dagger> = (p \<cdot> x)\<^sup>\<dagger> \<cdot> p"
using local.preservation1_eq local.test_restrictr by auto
lemma test_preserve1: "test p \<Longrightarrow> p \<cdot> x \<le> p \<cdot> x \<cdot> p \<Longrightarrow> p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger> = (p \<cdot> x)\<^sup>\<dagger> \<cdot> p"
-proof (rule antisym)
+proof (rule order.antisym)
assume a: "test p"
and b: "p \<cdot> x \<le> p \<cdot> x \<cdot> p"
hence "p \<cdot> (p \<cdot> x + !p \<cdot> y) \<le> (p \<cdot> x) \<cdot> p"
by (metis local.add_0_right local.annil local.n_left_distrib_var local.test_comp_mult2 local.test_mult_idem_var mult_assoc)
thus "p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger> \<le> (p \<cdot> x)\<^sup>\<dagger> \<cdot> p"
using local.dagger_simr by blast
next
assume a: "test p"
and b: "p \<cdot> x \<le> p \<cdot> x \<cdot> p"
hence "(p \<cdot> x)\<^sup>\<dagger> \<cdot> p = p \<cdot> (p \<cdot> x \<cdot> p)\<^sup>\<dagger>"
by (metis dagger_slide local.test_mult_idem_var mult_assoc)
also have "... = p \<cdot> (p \<cdot> x)\<^sup>\<dagger>"
by (metis a b local.order.antisym local.test_restrictr)
finally show "(p \<cdot> x)\<^sup>\<dagger> \<cdot> p \<le> p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger>"
by (simp add: a local.dagger_iso local.test_mult_isol)
qed
lemma test_preserve2: "test p \<Longrightarrow> p \<cdot> x \<cdot> p = p \<cdot> x \<Longrightarrow> p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger> \<le> x\<^sup>\<dagger>"
by (metis (no_types) local.eq_refl local.test_restrictl test_preserve test_preserve1)
end
end
diff --git a/thys/KAT_and_DRA/SingleSorted/DRAT.thy b/thys/KAT_and_DRA/SingleSorted/DRAT.thy
--- a/thys/KAT_and_DRA/SingleSorted/DRAT.thy
+++ b/thys/KAT_and_DRA/SingleSorted/DRAT.thy
@@ -1,255 +1,255 @@
(* Title: Demonic refinement algebra
Author: Alasdair Armstrong, Victor B. F. Gomes, Georg Struth
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
*)
section \<open>Demonic Refinement Algebra with Tests\<close>
theory DRAT
imports KAT Kleene_Algebra.DRA
begin
text \<open>
In this section, we define demonic refinement algebras with tests and prove the most important theorems from
the literature. In this context, tests are also known as guards.
\<close>
class drat = dra + test_semiring_zerol
begin
subclass kat_zerol ..
text \<open>
An assertion is a mapping from a guard to a subset similar to tests, but it aborts if the predicate does
not hold.
\<close>
definition assertion :: "'a \<Rightarrow> 'a" ("_\<^sup>o" [101] 100) where
"test p \<Longrightarrow> p\<^sup>o = !p\<cdot>\<top> + 1"
lemma asg: "\<lbrakk>test p; test q\<rbrakk> \<Longrightarrow> q \<le> 1 \<and> 1 \<le> p\<^sup>o"
by (simp add: assertion_def local.test_subid)
lemma assertion_isol: "test p \<Longrightarrow> y \<le> p\<^sup>o\<cdot>x \<longleftrightarrow> p\<cdot>y \<le> x"
proof
assume assms: "test p" "y \<le> p\<^sup>o\<cdot>x"
hence "p\<cdot>y \<le> p\<cdot>!p\<cdot>\<top>\<cdot>x + p\<cdot>x"
by (metis add_commute assertion_def local.distrib_left local.iteration_prod_unfold local.iteration_unfoldl_distr local.mult_isol local.top_mult_annil mult_assoc)
also have "... \<le> x"
by (simp add: assms(1) local.test_restrictl)
finally show "p\<cdot>y \<le> x"
by metis
next
assume assms: "test p" "p\<cdot>y \<le> x"
hence "p\<^sup>o\<cdot>p\<cdot>y = !p\<cdot>\<top>\<cdot>p\<cdot>y + p\<cdot>y"
by (metis assertion_def distrib_right' mult_1_left mult.assoc)
also have "... = !p\<cdot>\<top> + p\<cdot>y"
by (metis mult.assoc top_mult_annil)
moreover have "p\<^sup>o\<cdot>p\<cdot>y \<le> p\<^sup>o\<cdot>x"
by (metis assms(2) mult.assoc mult_isol)
moreover have "!p\<cdot>y + p\<cdot>y \<le> !p\<cdot>\<top> + p\<cdot>y"
using local.add_iso local.top_elim by blast
ultimately show "y \<le> p\<^sup>o\<cdot>x"
by (metis add.commute assms(1) distrib_right' mult_1_left order_trans test_add_comp)
qed
lemma assertion_isor: "test p \<Longrightarrow> y \<le> x\<cdot>p \<longleftrightarrow> y\<cdot>p\<^sup>o \<le> x"
proof
assume assms: "test p" "y \<le> x\<cdot>p"
hence "y\<cdot>p\<^sup>o \<le> x\<cdot>p\<cdot>!p\<cdot>\<top> + x\<cdot>p"
by (metis mult_isor assertion_def assms(1) distrib_left mult_1_right mult.assoc)
also have "... \<le> x"
by (metis assms(1) local.iteration_idep local.join.sup.absorb_iff1 local.join.sup_commute local.join.sup_ge2 local.mult_1_right local.mult_isol_var local.mult_isor local.mult_onel local.test_add_comp local.test_comp_mult2 mult_assoc)
finally show "y\<cdot>p\<^sup>o \<le> x"
by metis
next
assume assms: "test p" "y\<cdot>p\<^sup>o \<le> x"
have "y \<le> y\<cdot>(!p\<cdot>\<top> + p)"
by (metis join.sup_mono mult_isol order_refl order_refl top_elim add.commute assms(1) mult_1_right test_add_comp)
also have "... = y\<cdot>p\<^sup>o\<cdot>p"
by (metis assertion_def assms(1) distrib_right' mult_1_left mult.assoc top_mult_annil)
finally show "y \<le> x\<cdot>p"
by (metis assms(2) mult_isor order_trans)
qed
lemma assertion_iso: "\<lbrakk>test p; test q\<rbrakk> \<Longrightarrow> x\<cdot>q\<^sup>o \<le> p\<^sup>o\<cdot>x \<longleftrightarrow> p\<cdot>x \<le> x\<cdot>q"
by (metis assertion_isol assertion_isor mult.assoc)
lemma total_correctness: "\<lbrakk>test p; test q\<rbrakk> \<Longrightarrow> p\<cdot>x\<cdot>!q = 0 \<longleftrightarrow> x\<cdot>!q \<le> !p\<cdot>\<top>"
apply standard
apply (metis local.test_eq4 local.top_elim mult_assoc)
- by (metis annil antisym test_comp_mult2 join.bot_least mult_assoc mult_isol)
+ by (metis annil order.antisym test_comp_mult2 join.bot_least mult_assoc mult_isol)
lemma test_iteration_sim: "\<lbrakk>test p; p\<cdot>x \<le> x\<cdot>p\<rbrakk> \<Longrightarrow> p\<cdot>x\<^sup>\<infinity> \<le> x\<^sup>\<infinity>\<cdot>p"
by (metis iteration_sim)
lemma test_iteration_annir: "test p \<Longrightarrow> !p\<cdot>(p\<cdot>x)\<^sup>\<infinity> = !p"
by (metis annil test_comp_mult1 iteration_idep mult.assoc)
text \<open>Next we give an example of a program transformation from von Wright~\cite{Wright02}.\<close>
lemma loop_refinement: "\<lbrakk>test p; test q\<rbrakk> \<Longrightarrow> (p\<cdot>x)\<^sup>\<infinity>\<cdot>!p \<le> (p\<cdot>q\<cdot>x)\<^sup>\<infinity>\<cdot>!(p\<cdot>q)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p"
proof -
assume assms: "test p" "test q"
hence "(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p = ((p\<cdot>q) + !(p\<cdot>q))\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p"
by (simp add: local.test_mult_closed)
also have "... = (p\<cdot>q)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p + !(p\<cdot>q)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p"
by (metis distrib_right')
also have "... = (p\<cdot>q)\<cdot>!p + (p\<cdot>q)\<cdot>(p\<cdot>x)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p + !(p\<cdot>q)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p"
by (metis iteration_unfoldr_distr mult.assoc iteration_unfold_eq distrib_left mult.assoc)
also have "... = (p\<cdot>q)\<cdot>(p\<cdot>x)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p + !(p\<cdot>q)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p"
by (metis assms less_eq_def test_preserve2 join.bot_least)
finally have "(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p \<le> p\<cdot>q\<cdot>x\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p + !(p\<cdot>q)\<cdot>(p\<cdot>x)\<^sup>\<infinity>\<cdot>!p"
- by (metis assms(1) assms(2) local.eq_iff local.test_mult_comm_var local.test_preserve mult_assoc)
+ by (metis assms(1) assms(2) order.eq_iff local.test_mult_comm_var local.test_preserve mult_assoc)
thus ?thesis
by (metis coinduction add.commute mult.assoc)
qed
text \<open>Finally, we prove different versions of Back's atomicity refinement theorem for action systems.\<close>
lemma atom_step1: "r\<cdot>b \<le> b\<cdot>r \<Longrightarrow> (a + b + r)\<^sup>\<infinity> = b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
apply (subgoal_tac "(a + b + r)\<^sup>\<infinity> = (b + r)\<^sup>\<infinity>\<cdot>(a\<cdot>(b + r)\<^sup>\<infinity>)\<^sup>\<infinity>")
apply (metis iteration_sep mult.assoc)
by (metis add_assoc' add.commute iteration_denest)
lemma atom_step2:
assumes "s = s\<cdot>q" "q\<cdot>b = 0" "r\<cdot>q \<le> q\<cdot>r" "q\<cdot>l \<le> l\<cdot>q" "r\<^sup>\<infinity> = r\<^sup>\<star>" "test q"
shows "s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity> \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
proof -
have "s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity> \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(3) assms(5) star_sim1 mult.assoc mult_isol iteration_iso)
also have "... \<le> s\<cdot>q\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
using assms(1) assms(6) local.mult_isor local.test_restrictr by auto
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>q\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(4) iteration_sim mult.assoc mult_double_iso mult_double_iso)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(2) join.bot_least iteration_sim mult.assoc mult_double_iso)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(6) mult.assoc mult_isol test_restrictl iteration_idem mult.assoc)
finally show "s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity> \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by metis
qed
lemma atom_step3:
assumes "r\<cdot>l \<le> l\<cdot>r" "a\<cdot>l \<le> l\<cdot>a" "b\<cdot>l \<le> l\<cdot>b" "q\<cdot>l \<le> l\<cdot>q" "b\<^sup>\<infinity> = b\<^sup>\<star>"
shows "l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity> = (a\<cdot>b\<^sup>\<infinity>\<cdot>q + l + r)\<^sup>\<infinity>"
proof -
have "(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r)\<cdot>l \<le> a\<cdot>b\<^sup>\<infinity>\<cdot>l\<cdot>q + l\<cdot>r"
by (metis distrib_right join.sup_mono assms(1,4) mult.assoc mult_isol)
also have "... \<le> a\<cdot>l\<cdot>b\<^sup>\<infinity>\<cdot>q + l\<cdot>r"
by (metis assms(3) assms(5) star_sim1 add_iso mult.assoc mult_double_iso)
also have "... \<le> l\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r)"
by (metis add_iso assms(2) mult_isor distrib_left mult.assoc)
finally have "(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r)\<cdot>l \<le> l\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r)"
by metis
moreover have "l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity> = l\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r)\<^sup>\<infinity>"
by (metis add.commute mult.assoc iteration_denest)
ultimately show ?thesis
by (metis add.commute add.left_commute iteration_sep)
qed
text \<open>
This is Back's atomicity refinement theorem, as specified by von Wright~\cite {Wright02}.
\<close>
theorem atom_ref_back:
assumes "s = s\<cdot>q" "a = q\<cdot>a" "q\<cdot>b = 0"
"r\<cdot>b \<le> b\<cdot>r" "r\<cdot>l \<le> l\<cdot>r" "r\<cdot>q \<le> q\<cdot>r"
"a\<cdot>l \<le> l\<cdot>a" "b\<cdot>l \<le> l\<cdot>b" "q\<cdot>l \<le> l\<cdot>q"
"r\<^sup>\<infinity> = r\<^sup>\<star>" "b\<^sup>\<infinity> = b\<^sup>\<star>" "test q"
shows "s\<cdot>(a + b + r + l)\<^sup>\<infinity>\<cdot>q \<le> s\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r + l)\<^sup>\<infinity>"
proof -
have "(a + b + r)\<cdot>l \<le> l\<cdot>(a + b + r)"
by (metis join.sup_mono distrib_right' assms(5) assms(7) assms(8) distrib_left)
hence "s\<cdot>(l + a + b + r)\<^sup>\<infinity>\<cdot>q = s\<cdot>l\<^sup>\<infinity>\<cdot>(a + b + r)\<^sup>\<infinity>\<cdot>q"
by (metis add.commute add.left_commute mult.assoc iteration_sep)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity>"
by (metis assms(2,4,10,11) atom_step1 iteration_slide eq_refl mult.assoc)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(1) assms(10) assms(12) assms(3) assms(6) assms(9) atom_step2)
also have "... \<le> s\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + l + r)\<^sup>\<infinity>"
by (metis assms(11) assms(5) assms(7) assms(8) assms(9) atom_step3 eq_refl mult.assoc)
finally show ?thesis
by (metis add.commute add.left_commute)
qed
text \<open>
This variant is due to H\"ofner, Struth and Sutcliffe~\cite{HofnerSS09}.
\<close>
theorem atom_ref_back_struth:
assumes "s \<le> s\<cdot>q" "a \<le> q\<cdot>a" "q\<cdot>b = 0"
"r\<cdot>b \<le> b\<cdot>r" "r\<cdot>q \<le> q\<cdot>r"
"(a + r + b)\<cdot>l \<le> l\<cdot>(a + r + b)" "q\<cdot>l \<le> l\<cdot>q"
"r\<^sup>\<infinity> = r\<^sup>\<star>" "q \<le> 1"
shows "s\<cdot>(a + b + r + l)\<^sup>\<infinity>\<cdot>q \<le> s\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r + l)\<^sup>\<infinity>"
proof -
have "s\<cdot>(a + b + r + l)\<^sup>\<infinity>\<cdot>q = s\<cdot>l\<^sup>\<infinity>\<cdot>(a + b + r)\<^sup>\<infinity>\<cdot>q"
by (metis add.commute add.left_commute assms(6) iteration_sep mult.assoc)
also have "... = s\<cdot>l\<^sup>\<infinity>\<cdot>(b + r)\<^sup>\<infinity>\<cdot>(a\<cdot>(b + r)\<^sup>\<infinity>)\<^sup>\<infinity>\<cdot>q"
by (metis add_assoc' add.commute iteration_denest add.commute mult.assoc)
also have "... = s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>\<cdot>q"
by (metis assms(4) iteration_sep mult.assoc)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(q\<cdot>a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>\<cdot>q"
by (metis assms(2) iteration_iso mult_isol_var eq_refl order_refl)
also have "... = s\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity>"
by (metis iteration_slide mult.assoc)
also have "... \<le> s\<cdot>q\<cdot>l\<^sup>\<infinity>\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity>"
by (metis assms(1) mult_isor)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>q\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity>"
by (metis assms(7) iteration_sim mult.assoc mult_double_iso)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q)\<^sup>\<infinity>"
by (metis assms(3) iteration_idep mult.assoc order_refl)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>r\<^sup>\<star>\<cdot>q)\<^sup>\<infinity>"
by (metis assms(8) eq_refl)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<star>)\<^sup>\<infinity>"
by (metis assms(5) iteration_iso mult.assoc mult_isol star_sim1)
also have "... = s\<cdot>l\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(8))
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>q\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(9) mult_1_right mult_double_iso mult_isor)
also have "... \<le> s\<cdot>l\<^sup>\<infinity>\<cdot>r\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q\<cdot>r\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis assms(9) mult_1_right mult_double_iso)
also have "... = s\<cdot>l\<^sup>\<infinity>\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r)\<^sup>\<infinity>"
by (metis add.commute mult.assoc iteration_denest)
also have "... \<le> s\<cdot>(a\<cdot>b\<^sup>\<infinity>\<cdot>q + r + l)\<^sup>\<infinity>"
by (metis add.commute iteration_subdenest mult.assoc mult_isol)
finally show ?thesis .
qed
text \<open>Finally, we prove Cohen's~\cite{Cohen00} variation of the atomicity refinement theorem.\<close>
lemma atom_ref_cohen:
assumes "r\<cdot>p\<cdot>y \<le> y\<cdot>r" "y\<cdot>p\<cdot>l \<le> l\<cdot>y" "r\<cdot>p\<cdot>l \<le> l\<cdot>r"
"p\<cdot>r\<cdot>!p = 0" "p\<cdot>l\<cdot>!p = 0" "!p\<cdot>l\<cdot>p = 0"
"y\<cdot>0 = 0" "r\<cdot>0 = 0" "test p"
shows "(y + r + l)\<^sup>\<infinity> = (p\<cdot>l)\<^sup>\<infinity>\<cdot>(y + !p\<cdot>l + r\<cdot>!p)\<^sup>\<infinity>\<cdot>(r\<cdot>p)\<^sup>\<infinity>"
proof -
have "(y + r)\<cdot>p\<cdot>l \<le> l\<cdot>y + l\<cdot>r"
by (metis distrib_right' join.sup_mono assms(2) assms(3))
hence stepA: "(y + r)\<cdot>p\<cdot>l \<le> (p\<cdot>l + !p\<cdot>l)\<cdot>(y + r)\<^sup>\<star>"
by (metis assms(9) distrib_left distrib_right' mult_1_left mult_isol order_trans star_ext test_add_comp)
have subStepB: "(!p\<cdot>l + y + p\<cdot>r + !p\<cdot>r)\<^sup>\<infinity> = (!p\<cdot>l + y + r\<cdot>p + r\<cdot>!p)\<^sup>\<infinity>"
by (metis add_assoc' annil assms(8) assms(9) distrib_left distrib_right' star_slide star_subid test_add_comp join.bot_least)
have "r\<cdot>p\<cdot>(y + r\<cdot>!p + !p\<cdot>l) \<le> y\<cdot>(r\<cdot>p + r\<cdot>!p)"
by (metis assms(1,4,9) distrib_left add_0_left add.commute annil mult.assoc test_comp_mult2 distrib_left mult_oner test_add_comp)
also have "... \<le> (y + r\<cdot>!p + !p\<cdot>l)\<cdot>(r\<cdot>p + (y + r\<cdot>!p + !p\<cdot>l))"
by (meson local.eq_refl local.join.sup_ge1 local.join.sup_ge2 local.join.sup_mono local.mult_isol_var local.order_trans)
finally have "r\<cdot>p\<cdot>(y + r\<cdot>!p + !p\<cdot>l) \<le> (y + r\<cdot>!p + !p\<cdot>l)\<cdot>(y + r\<cdot>!p + !p\<cdot>l + r\<cdot>p)"
by (metis add.commute)
hence stepB: "(!p\<cdot>l + y + p\<cdot>r + !p\<cdot>r)\<^sup>\<infinity> = (y + !p\<cdot>l + r\<cdot>!p)\<^sup>\<infinity>\<cdot>(r\<cdot>p)\<^sup>\<infinity>"
by (metis subStepB iteration_sep3[of "r\<cdot>p" "y + r\<cdot>!p + !p\<cdot>l"] add_assoc' add.commute)
have "(y + r + l)\<^sup>\<infinity> = (p\<cdot>l + !p\<cdot>l + y + r)\<^sup>\<infinity>"
by (metis add_comm add.left_commute assms(9) distrib_right' mult_onel test_add_comp)
also have "... = (p\<cdot>l)\<^sup>\<infinity>\<cdot>(!p\<cdot>l + y + r)\<^sup>\<infinity>" using stepA
by (metis assms(6-8) annil add.assoc add_0_left distrib_right' add.commute mult.assoc iteration_sep4[of "y+r" "!p\<cdot>l" "p\<cdot>l"])
also have "... = (p\<cdot>l)\<^sup>\<infinity>\<cdot>(!p\<cdot>l + y + p\<cdot>r + !p\<cdot>r)\<^sup>\<infinity>"
by (metis add.commute assms(9) combine_common_factor mult_1_left test_add_comp)
finally show ?thesis using stepB
by (metis mult.assoc)
qed
end
end
diff --git a/thys/KAT_and_DRA/SingleSorted/FolkTheorem.thy b/thys/KAT_and_DRA/SingleSorted/FolkTheorem.thy
--- a/thys/KAT_and_DRA/SingleSorted/FolkTheorem.thy
+++ b/thys/KAT_and_DRA/SingleSorted/FolkTheorem.thy
@@ -1,165 +1,165 @@
(* Title: Kleene algebra with tests
Author: Alasdair Armstrong, Victor B. F. Gomes, Georg Struth
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
*)
section \<open>Transformation Theorem for while Loops\<close>
theory FolkTheorem
imports Conway_Tests KAT DRAT
begin
text \<open>
We prove Kozen's transformation theorem for while loops \cite{Kozen97} in a weak setting that unifies
previous proofs in Kleene algebra with tests, demonic refinement algebras and a variant of probabilistic
Kleene algebra.
\<close>
context test_pre_conway
begin
abbreviation pres :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
"pres x y \<equiv> y \<cdot> x = y \<cdot> x \<cdot> y"
lemma pres_comp: "pres y z \<Longrightarrow> pres x z \<Longrightarrow> pres (x \<cdot> y) z"
by (metis mult_assoc)
lemma test_pres1: "test p \<Longrightarrow> test q \<Longrightarrow> pres p q"
by (simp add: local.test_mult_comm_var mult_assoc)
lemma test_pres2: "test p \<Longrightarrow> test q \<Longrightarrow> pres x q \<Longrightarrow> pres (p \<cdot> x) q"
using pres_comp test_pres1 by blast
lemma cond_helper1:
assumes "test p" and "pres x p"
shows "p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger> \<cdot> (p \<cdot> z + !p \<cdot> w) = p \<cdot> x\<^sup>\<dagger> \<cdot> z"
proof -
have "p \<cdot> (p \<cdot> z + !p \<cdot> w) = p \<cdot> z"
by (metis assms(1) local.add_zerol local.annil local.join.sup_commute local.n_left_distrib_var local.test_comp_mult2 local.test_mult_idem_var mult_assoc)
hence "p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger> \<cdot> (p \<cdot> z + !p \<cdot> w) = (p \<cdot> x)\<^sup>\<dagger> \<cdot> p \<cdot> z"
using assms(1) assms(2) local.test_preserve1 mult_assoc by auto
thus ?thesis
using assms(1) assms(2) local.test_preserve mult_assoc by auto
qed
lemma cond_helper2:
assumes "test p" and "pres y (!p)"
shows "!p \<cdot> (p \<cdot> x + !p \<cdot> y)\<^sup>\<dagger> \<cdot> (p \<cdot> z + !p \<cdot> w) = !p \<cdot> y\<^sup>\<dagger> \<cdot> w"
proof -
have "!p \<cdot> (!p \<cdot> y + !(!p) \<cdot> x)\<^sup>\<dagger> \<cdot> (!p \<cdot> w + !(!p) \<cdot> z) = !p \<cdot> y\<^sup>\<dagger> \<cdot> w"
using assms(1) local.test_comp_closed assms(2) cond_helper1 by blast
thus ?thesis
using add_commute assms(1) by auto
qed
lemma cond_distr_var:
assumes "test p" and "test q" and "test r"
shows "(q \<cdot> p + r \<cdot> !p) \<cdot> (p \<cdot> x + !p \<cdot> y) = q \<cdot> p \<cdot> x + r \<cdot> !p \<cdot> y"
proof -
have "(q \<cdot> p + r \<cdot> !p) \<cdot> (p \<cdot> x + !p \<cdot> y) = q \<cdot> p \<cdot> p \<cdot> x + q \<cdot> p \<cdot> !p \<cdot> y + r \<cdot> !p \<cdot> p \<cdot> x + r \<cdot> !p \<cdot> !p \<cdot> y"
using assms(1) assms(2) assms(3) local.distrib_right' local.join.sup_assoc local.n_left_distrib_var local.test_comp_closed mult_assoc by presburger
also have "... = q \<cdot> p \<cdot> x + q \<cdot> 0 \<cdot> y + r \<cdot> 0 \<cdot> x + r \<cdot> !p \<cdot> y"
by (simp add: assms(1) mult_assoc)
finally show ?thesis
by (metis assms(2) assms(3) local.add_zerol local.annil local.join.sup_commute local.test_mult_comm_var local.test_zero_var)
qed
lemma cond_distr:
assumes "test p" and "test q" and "test r"
shows "(p \<cdot> q + !p \<cdot> r) \<cdot> (p \<cdot> x + !p \<cdot> y) = p \<cdot> q \<cdot> x + !p \<cdot> r \<cdot> y"
using assms(1) assms(2) assms(3) local.test_mult_comm_var assms(1) assms(2) assms(3) cond_distr_var local.test_mult_comm_var by auto
theorem conditional:
assumes "test p" and "test q" and "test r1" and "test r2"
and "pres x1 q" and "pres y1 q" and "pres x2 (!q)" and "pres y2 (!q)"
shows "(p \<cdot> q + !p \<cdot> !q) \<cdot> (q \<cdot> x1 + !q \<cdot> x2) \<cdot> ((q \<cdot> r1 + !q \<cdot> r2) \<cdot> (q \<cdot> y1 + !q \<cdot> y2))\<^sup>\<dagger> \<cdot> !(q \<cdot> r1 + !q \<cdot> r2) =
(p \<cdot> q + !p \<cdot> !q) \<cdot> (p \<cdot> x1 \<cdot> (r1 \<cdot> y1)\<^sup>\<dagger> \<cdot> !r1 + !p \<cdot> x2 \<cdot> (r2 \<cdot> y2)\<^sup>\<dagger> \<cdot> !r2)"
proof -
have a: "p \<cdot> q \<cdot> x1 \<cdot> q \<cdot> (q \<cdot> r1 \<cdot> y1 + !q \<cdot> r2 \<cdot> y2)\<^sup>\<dagger> \<cdot> (q \<cdot> !r1 + !q \<cdot> !r2) = p \<cdot> q \<cdot> x1 \<cdot> q \<cdot> (r1 \<cdot> y1)\<^sup>\<dagger> \<cdot> !r1"
by (metis assms(2) assms(3) assms(6) cond_helper1 mult_assoc test_pres2)
have b: "!q \<cdot> (q \<cdot> r1 \<cdot> y1 + !q \<cdot> r2 \<cdot> y2)\<^sup>\<dagger> \<cdot> (q \<cdot> !r1 + !q \<cdot> !r2) = !q \<cdot> (r2 \<cdot> y2)\<^sup>\<dagger> \<cdot> !r2"
by (metis assms(2) assms(4) assms(8) local.test_comp_closed cond_helper2 mult_assoc test_pres2)
have "(p \<cdot> q + !p \<cdot> !q) \<cdot> (q \<cdot> x1 + !q \<cdot> x2) \<cdot> ((q \<cdot> r1 + !q \<cdot> r2) \<cdot> (q \<cdot> y1 + !q \<cdot> y2))\<^sup>\<dagger> \<cdot> !(q \<cdot> r1 + !q \<cdot> r2) =
p \<cdot> q \<cdot> x1 \<cdot> q \<cdot> (q \<cdot> r1 \<cdot> y1 + !q \<cdot> r2 \<cdot> y2)\<^sup>\<dagger> \<cdot> (q \<cdot> !r1 + !q \<cdot> !r2) + !p \<cdot> !q \<cdot> x2 \<cdot> !q \<cdot> (q \<cdot> r1 \<cdot> y1 + !q \<cdot> r2 \<cdot> y2)\<^sup>\<dagger> \<cdot> (q \<cdot> !r1 + !q \<cdot> !r2)"
using assms(1) assms(2) assms(3) assms(4) assms(5) assms(7) cond_distr cond_distr_var local.test_dist_var2 mult_assoc by auto
also have "... = p \<cdot> q \<cdot> x1 \<cdot> (r1 \<cdot> y1)\<^sup>\<dagger> \<cdot> !r1 + !p \<cdot> !q \<cdot> x2 \<cdot> (r2 \<cdot> y2)\<^sup>\<dagger> \<cdot> !r2"
by (metis a assms(5) assms(7) b mult_assoc)
finally show ?thesis
using assms(1) assms(2) cond_distr mult_assoc by auto
qed
theorem nested_loops:
assumes "test p" and "test q"
shows "p \<cdot> x \<cdot> ((p + q) \<cdot> (q \<cdot> y + !q \<cdot> x))\<^sup>\<dagger> \<cdot> !(p + q) + !p = (p \<cdot> x \<cdot> (q \<cdot> y)\<^sup>\<dagger> \<cdot> !q)\<^sup>\<dagger> \<cdot> !p"
proof -
have "p \<cdot> x \<cdot> ((p + q) \<cdot> (q \<cdot> y + !q \<cdot> x))\<^sup>\<dagger> \<cdot> !(p + q) + !p = p \<cdot> x \<cdot> (q \<cdot> y)\<^sup>\<dagger> \<cdot> (!q \<cdot> p \<cdot> x \<cdot> (q \<cdot> y)\<^sup>\<dagger>)\<^sup>\<dagger> \<cdot> !q \<cdot> !p + !p"
using assms(1) assms(2) local.dagger_denest2 local.test_distrib mult_assoc test_mult_comm_var by auto
also have "... = p \<cdot> x \<cdot> (q \<cdot> y)\<^sup>\<dagger> \<cdot> !q \<cdot> (p \<cdot> x \<cdot> (q \<cdot> y)\<^sup>\<dagger> \<cdot> !q)\<^sup>\<dagger> \<cdot> !p + !p"
by (metis local.dagger_slide mult_assoc)
finally show ?thesis
using add_commute by force
qed
lemma postcomputation:
assumes "test p" and "pres y (!p)"
shows "!p \<cdot> y + p \<cdot> (p \<cdot> x \<cdot> (!p \<cdot> y + p))\<^sup>\<dagger> \<cdot> !p = (p \<cdot> x)\<^sup>\<dagger> \<cdot> !p \<cdot> y"
proof -
have "p \<cdot> (p \<cdot> x \<cdot> (!p \<cdot> y + p))\<^sup>\<dagger> \<cdot> !p = p \<cdot> (1 + p \<cdot> x \<cdot> ((!p \<cdot> y + p) \<cdot> p \<cdot> x)\<^sup>\<dagger> \<cdot> (!p \<cdot> y + p)) \<cdot> !p"
by (metis dagger_prod_unfold mult.assoc)
also have "... = (p + p \<cdot> p \<cdot> x \<cdot> ((!p \<cdot> y + p) \<cdot> p \<cdot> x)\<^sup>\<dagger> \<cdot> (!p \<cdot> y + p)) \<cdot> !p"
using assms(1) local.mult_oner local.n_left_distrib_var mult_assoc by presburger
also have "... = p \<cdot> x \<cdot> (!p \<cdot> y \<cdot> p \<cdot> x + p \<cdot> x)\<^sup>\<dagger> \<cdot> !p \<cdot> y \<cdot> !p"
by (simp add: assms(1) mult_assoc)
also have "... = p \<cdot> x \<cdot> (!p \<cdot> y \<cdot> 0 + p \<cdot> x)\<^sup>\<dagger> \<cdot> !p \<cdot> y"
by (metis assms(1) assms(2) local.annil local.test_comp_mult1 mult_assoc)
also have "... = p \<cdot> x \<cdot> (p \<cdot> x)\<^sup>\<dagger> \<cdot> (!p \<cdot> y \<cdot> 0 \<cdot> (p \<cdot> x)\<^sup>\<dagger>)\<^sup>\<dagger> \<cdot> !p \<cdot> y"
by (metis mult.assoc add.commute dagger_denest2)
finally have "p \<cdot> (p \<cdot> x \<cdot> (!p \<cdot> y + p))\<^sup>\<dagger> \<cdot> !p = p \<cdot> x \<cdot> (p \<cdot> x)\<^sup>\<dagger> \<cdot> !p \<cdot> y"
by (metis local.add_zeror local.annil local.dagger_prod_unfold local.dagger_slide local.mult_oner mult_assoc)
thus ?thesis
by (metis dagger_unfoldl_distr mult.assoc)
qed
lemma composition_helper:
assumes "test p" and "test q"
and "pres x p"
shows "p \<cdot> (q \<cdot> x)\<^sup>\<dagger> \<cdot> !q \<cdot> p = p \<cdot> (q \<cdot> x)\<^sup>\<dagger> \<cdot> !q"
-proof (rule antisym)
+proof (rule order.antisym)
show "p \<cdot> (q \<cdot> x)\<^sup>\<dagger> \<cdot> !q \<cdot> p \<le> p \<cdot> (q \<cdot> x)\<^sup>\<dagger> \<cdot> !q"
by (simp add: assms(1) local.test_restrictr)
next
have "p \<cdot> q \<cdot> x \<le> q \<cdot> x \<cdot> p"
by (metis assms(1) assms(2) assms(3) local.test_kat_2 mult_assoc test_pres2)
hence "p \<cdot> (q \<cdot> x)\<^sup>\<dagger> \<le> (q \<cdot> x)\<^sup>\<dagger> \<cdot> p"
by (simp add: local.dagger_simr mult_assoc)
thus "p \<cdot> (q \<cdot> x)\<^sup>\<dagger> \<cdot> !q \<le> p \<cdot> (q \<cdot> x)\<^sup>\<dagger> \<cdot> !q \<cdot> p"
- by (metis assms(1) assms(2) local.eq_iff local.test_comp_closed local.test_kat_2 local.test_mult_comm_var mult_assoc)
+ by (metis assms(1) assms(2) order.eq_iff local.test_comp_closed local.test_kat_2 local.test_mult_comm_var mult_assoc)
qed
theorem composition:
assumes "test p" and "test q"
and "pres y p" and "pres y (!p)"
shows "(p \<cdot> x)\<^sup>\<dagger> \<cdot> !p \<cdot> (q \<cdot> y)\<^sup>\<dagger> \<cdot> !q = !p \<cdot> (q \<cdot> y)\<^sup>\<dagger> \<cdot> !q + p \<cdot> (p \<cdot> x \<cdot> (!p \<cdot> (q \<cdot> y)\<^sup>\<dagger> \<cdot> !q + p))\<^sup>\<dagger> \<cdot> !p"
by (metis assms(1) assms(2) assms(4) composition_helper local.test_comp_closed mult_assoc postcomputation)
end
text \<open>
Kleene algebras with tests form pre-Conway algebras, therefore the transformation theorem is valid for KAT as well.
\<close>
sublocale kat \<subseteq> pre_conway star ..
text \<open>
Demonic refinement algebras form pre-Conway algebras, therefore the transformation theorem is valid for DRA as well.
\<close>
sublocale drat \<subseteq> pre_conway strong_iteration
apply standard
apply (simp add: local.iteration_denest local.iteration_slide)
apply (metis local.iteration_prod_unfold)
by (simp add: local.iteration_sim)
text \<open>
We do not currently consider an expansion of probabilistic Kleene algebra.
\<close>
end
diff --git a/thys/KAT_and_DRA/SingleSorted/KAT.thy b/thys/KAT_and_DRA/SingleSorted/KAT.thy
--- a/thys/KAT_and_DRA/SingleSorted/KAT.thy
+++ b/thys/KAT_and_DRA/SingleSorted/KAT.thy
@@ -1,117 +1,117 @@
(* Title: Kleene algebra with tests
Author: Alasdair Armstrong, Victor B. F. Gomes, Georg Struth
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
*)
section \<open>Kleene Algebra with Tests\<close>
theory KAT
imports Kleene_Algebra.Kleene_Algebra Conway_Tests
begin
text \<open>
First, we study left Kleene algebras with tests which also have only a left zero.
These structures can be expanded to demonic refinement algebras.
\<close>
class left_kat_zerol = left_kleene_algebra_zerol + test_semiring_zerol
begin
lemma star_n_export1: "(n x \<cdot> y)\<^sup>\<star> \<cdot> n x \<le> n x \<cdot> y\<^sup>\<star>"
by (simp add: local.n_restrictr local.star_sim1)
lemma star_test_export1: "test p \<Longrightarrow> (p \<cdot> x)\<^sup>\<star> \<cdot> p \<le> p \<cdot> x\<^sup>\<star>"
using star_n_export1 by auto
lemma star_n_export2: "(n x \<cdot> y)\<^sup>\<star> \<cdot> n x \<le> y\<^sup>\<star> \<cdot> n x"
by (simp add: local.mult_isor local.n_restrictl local.star_iso)
lemma star_test_export2: "test p \<Longrightarrow> (p \<cdot> x)\<^sup>\<star> \<cdot> p \<le> x\<^sup>\<star> \<cdot> p"
using star_n_export2 by auto
lemma star_n_export_left: "x \<cdot> n y \<le> n y \<cdot> x \<Longrightarrow> x\<^sup>\<star> \<cdot> n y = n y \<cdot> (x \<cdot> n y)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
assume a1: "x \<cdot> n y \<le> n y \<cdot> x"
hence "x \<cdot> n y = n y \<cdot> x \<cdot> n y"
by (simp add: local.n_kat_2_opp)
thus "x\<^sup>\<star> \<cdot> n y \<le> n y \<cdot> (x \<cdot> n y)\<^sup>\<star>"
by (simp add: local.star_sim1 mult_assoc)
next
assume a1: "x \<cdot> n y \<le> n y \<cdot> x"
thus "n y \<cdot> (x \<cdot> n y)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> n y"
using local.star_slide star_n_export2 by force
qed
lemma star_test_export_left: "test p \<Longrightarrow> x \<cdot> p \<le> p \<cdot> x \<longrightarrow> x\<^sup>\<star> \<cdot> p = p \<cdot> (x \<cdot> p)\<^sup>\<star>"
using star_n_export_left by auto
lemma star_n_export_right: "x \<cdot> n y \<le> n y \<cdot> x \<Longrightarrow> x\<^sup>\<star> \<cdot> n y = (n y \<cdot> x)\<^sup>\<star> \<cdot> n y"
by (simp add: local.star_slide star_n_export_left)
lemma star_test_export_right: "test p \<Longrightarrow> x \<cdot> p \<le> p \<cdot> x \<longrightarrow> x\<^sup>\<star> \<cdot> p = (p \<cdot> x)\<^sup>\<star> \<cdot> p"
using star_n_export_right by auto
lemma star_n_folk: "n z \<cdot> x = x \<cdot> n z \<Longrightarrow> n z \<cdot> y = y \<cdot> n z \<Longrightarrow> (n z \<cdot> x + t z \<cdot> y)\<^sup>\<star> \<cdot> n z = n z \<cdot> (n z \<cdot> x)\<^sup>\<star>"
proof -
assume a: "n z \<cdot> x = x \<cdot> n z" and b: "n z \<cdot> y = y \<cdot> n z"
hence "n z \<cdot> (n z \<cdot> x + t z \<cdot> y) = (n z \<cdot> x + t z \<cdot> y) \<cdot> n z"
using local.comm_add_var local.t_n_closed local.test_def by blast
hence "(n z \<cdot> x + t z \<cdot> y)\<^sup>\<star> \<cdot> n z = n z \<cdot> ((n z \<cdot> x + t z \<cdot> y) \<cdot> n z)\<^sup>\<star>"
using local.order_refl star_n_export_left by presburger
also have "... = n z \<cdot> (n z \<cdot> x \<cdot> n z + t z \<cdot> y \<cdot> n z)\<^sup>\<star>"
by simp
also have "... = n z \<cdot> (n z \<cdot> n z \<cdot> x + t z \<cdot> n z \<cdot> y)\<^sup>\<star>"
by (simp add: a b mult_assoc)
also have "... = n z \<cdot> (n z \<cdot> x + 0 \<cdot> y)\<^sup>\<star>"
by (simp add: local.n_mult_comm)
finally show "(n z \<cdot> x + t z \<cdot> y)\<^sup>\<star> \<cdot> n z = n z \<cdot> (n z \<cdot> x)\<^sup>\<star>"
by simp
qed
lemma star_test_folk: "test p \<Longrightarrow> p \<cdot> x = x \<cdot> p \<longrightarrow> p \<cdot> y = y \<cdot> p \<longrightarrow> (p \<cdot> x + !p \<cdot> y)\<^sup>\<star> \<cdot> p = p \<cdot> (p \<cdot> x)\<^sup>\<star>"
using star_n_folk by auto
end
class kat_zerol = kleene_algebra_zerol + test_semiring_zerol
begin
sublocale conway: near_conway_zerol_tests star ..
lemma n_star_sim_right: "n y \<cdot> x = x \<cdot> n y \<Longrightarrow> n y \<cdot> x\<^sup>\<star> = (n y \<cdot> x)\<^sup>\<star> \<cdot> n y"
by (metis local.n_mult_idem local.star_sim3 mult_assoc)
lemma star_sim_right: "test p \<Longrightarrow> p \<cdot> x = x \<cdot> p \<longrightarrow> p \<cdot> x\<^sup>\<star> = (p \<cdot> x)\<^sup>\<star> \<cdot> p"
using n_star_sim_right by auto
lemma n_star_sim_left: "n y \<cdot> x = x \<cdot> n y \<Longrightarrow> n y \<cdot> x\<^sup>\<star> = n y \<cdot> (x \<cdot> n y)\<^sup>\<star>"
by (metis local.star_slide n_star_sim_right)
lemma star_sim_left: "test p \<Longrightarrow> p \<cdot> x = x \<cdot> p \<longrightarrow> p \<cdot> x\<^sup>\<star> = p \<cdot> (x \<cdot> p)\<^sup>\<star>"
using n_star_sim_left by auto
lemma n_comm_star: "n z \<cdot> x = x \<cdot> n z \<Longrightarrow> n z \<cdot> y = y \<cdot> n z \<Longrightarrow> n z \<cdot> x \<cdot> (n z \<cdot> y)\<^sup>\<star> = n z \<cdot> x \<cdot> y\<^sup>\<star>"
using mult_assoc n_star_sim_left by presburger
lemma comm_star: "test p \<Longrightarrow> p \<cdot> x = x \<cdot> p \<longrightarrow> p \<cdot> y = y \<cdot> p \<longrightarrow> p \<cdot> x \<cdot> (p \<cdot> y)\<^sup>\<star> = p \<cdot> x \<cdot> y\<^sup>\<star>"
using n_comm_star by auto
lemma n_star_sim_right_var: "n y \<cdot> x = x \<cdot> n y \<Longrightarrow> x\<^sup>\<star> \<cdot> n y = n y \<cdot> (x \<cdot> n y)\<^sup>\<star>"
using local.star_sim3 n_star_sim_left by force
lemma star_sim_right_var: "test p \<Longrightarrow> p \<cdot> x = x \<cdot> p \<longrightarrow> x\<^sup>\<star> \<cdot> p = p \<cdot> (x \<cdot> p)\<^sup>\<star>"
using n_star_sim_right_var by auto
end
text \<open>Finally, we define Kleene algebra with tests.\<close>
class kat = kleene_algebra + test_semiring
begin
sublocale conway: test_pre_conway star ..
end
end
diff --git a/thys/Kleene_Algebra/Conway.thy b/thys/Kleene_Algebra/Conway.thy
--- a/thys/Kleene_Algebra/Conway.thy
+++ b/thys/Kleene_Algebra/Conway.thy
@@ -1,278 +1,279 @@
(* Title: Conway Algebra
Author: Alasdair Armstrong, Victor B. F. Gomes, Georg Struth
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Conway Algebras\<close>
theory Conway
imports Dioid
begin
text \<open>
We define a weak regular algebra which can serve as a common basis for Kleene algebra and demonic reginement algebra.
It is closely related to an axiomatisation given by Conway~\cite{conway71regular}.
\<close>
class dagger_op =
fixes dagger :: "'a \<Rightarrow> 'a" ("_\<^sup>\<dagger>" [101] 100)
subsection\<open>Near Conway Algebras\<close>
class near_conway_base = near_dioid_one + dagger_op +
assumes dagger_denest: "(x + y)\<^sup>\<dagger> = (x\<^sup>\<dagger> \<cdot> y)\<^sup>\<dagger> \<cdot> x\<^sup>\<dagger>"
and dagger_prod_unfold [simp]: "1 + x \<cdot> (y \<cdot> x)\<^sup>\<dagger> \<cdot> y = (x \<cdot> y)\<^sup>\<dagger>"
begin
lemma dagger_unfoldl_eq [simp]: "1 + x \<cdot> x\<^sup>\<dagger> = x\<^sup>\<dagger>"
by (metis dagger_prod_unfold mult_1_left mult_1_right)
lemma dagger_unfoldl: "1 + x \<cdot> x\<^sup>\<dagger> \<le> x\<^sup>\<dagger>"
by auto
lemma dagger_unfoldr_eq [simp]: "1 + x\<^sup>\<dagger> \<cdot> x = x\<^sup>\<dagger>"
by (metis dagger_prod_unfold mult_1_right mult_1_left)
lemma dagger_unfoldr: "1 + x\<^sup>\<dagger> \<cdot> x \<le> x\<^sup>\<dagger>"
by auto
lemma dagger_unfoldl_distr [simp]: "y + x \<cdot> x\<^sup>\<dagger> \<cdot> y = x\<^sup>\<dagger> \<cdot> y"
by (metis distrib_right' mult_1_left dagger_unfoldl_eq)
lemma dagger_unfoldr_distr [simp]: "y + x\<^sup>\<dagger> \<cdot> x \<cdot> y = x\<^sup>\<dagger> \<cdot> y"
by (metis dagger_unfoldr_eq distrib_right' mult_1_left mult.assoc)
lemma dagger_refl: "1 \<le> x\<^sup>\<dagger>"
using dagger_unfoldl local.join.sup.bounded_iff by blast
lemma dagger_plus_one [simp]: "1 + x\<^sup>\<dagger> = x\<^sup>\<dagger>"
by (simp add: dagger_refl local.join.sup_absorb2)
lemma star_1l: "x \<cdot> x\<^sup>\<dagger> \<le> x\<^sup>\<dagger>"
using dagger_unfoldl local.join.sup.bounded_iff by blast
lemma star_1r: "x\<^sup>\<dagger> \<cdot> x \<le> x\<^sup>\<dagger>"
using dagger_unfoldr local.join.sup.bounded_iff by blast
lemma dagger_ext: "x \<le> x\<^sup>\<dagger>"
by (metis dagger_unfoldl_distr local.join.sup.boundedE star_1r)
lemma dagger_trans_eq [simp]: "x\<^sup>\<dagger> \<cdot> x\<^sup>\<dagger> = x\<^sup>\<dagger>"
by (metis dagger_unfoldr_eq local.dagger_denest local.join.sup.idem mult_assoc)
lemma dagger_subdist: "x\<^sup>\<dagger> \<le> (x + y)\<^sup>\<dagger>"
by (metis dagger_unfoldr_distr local.dagger_denest local.order_prop)
lemma dagger_subdist_var: "x\<^sup>\<dagger> + y\<^sup>\<dagger> \<le> (x + y)\<^sup>\<dagger>"
using dagger_subdist local.join.sup_commute by fastforce
lemma dagger_iso [intro]: "x \<le> y \<Longrightarrow> x\<^sup>\<dagger> \<le> y\<^sup>\<dagger>"
by (metis less_eq_def dagger_subdist)
lemma star_square: "(x \<cdot> x)\<^sup>\<dagger> \<le> x\<^sup>\<dagger>"
- by (metis dagger_plus_one dagger_subdist dagger_trans_eq dagger_unfoldr_distr local.dagger_denest local.distrib_right' local.eq_iff local.join.sup_commute local.less_eq_def local.mult_onel mult_assoc)
+ by (metis dagger_plus_one dagger_subdist dagger_trans_eq dagger_unfoldr_distr dagger_denest
+ distrib_right' order.eq_iff join.sup_commute less_eq_def mult_onel mult_assoc)
lemma dagger_rtc1_eq [simp]: "1 + x + x\<^sup>\<dagger> \<cdot> x\<^sup>\<dagger> = x\<^sup>\<dagger>"
by (simp add: local.dagger_ext local.dagger_refl local.join.sup_absorb2)
text \<open>Nitpick refutes the next lemmas.\<close>
lemma " y + y \<cdot> x\<^sup>\<dagger> \<cdot> x = y \<cdot> x\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
lemma "y \<cdot> x\<^sup>\<dagger> = y + y \<cdot> x \<cdot> x\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
lemma "(x + y)\<^sup>\<dagger> = x\<^sup>\<dagger> \<cdot> (y \<cdot> x\<^sup>\<dagger>)\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
lemma "(x\<^sup>\<dagger>)\<^sup>\<dagger> = x\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
lemma "(1 + x)\<^sup>\<star> = x\<^sup>\<star>"
(*nitpick [expect = genuine]*)
oops
lemma "x\<^sup>\<dagger> \<cdot> x = x \<cdot> x\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
end
subsection\<open>Pre-Conway Algebras\<close>
class pre_conway_base = near_conway_base + pre_dioid_one
begin
lemma dagger_subdist_var_3: "x\<^sup>\<dagger> \<cdot> y\<^sup>\<dagger> \<le> (x + y)\<^sup>\<dagger>"
using local.dagger_subdist_var local.mult_isol_var by fastforce
lemma dagger_subdist_var_2: "x \<cdot> y \<le> (x + y)\<^sup>\<dagger>"
by (meson dagger_subdist_var_3 local.dagger_ext local.mult_isol_var local.order.trans)
lemma dagger_sum_unfold [simp]: "x\<^sup>\<dagger> + x\<^sup>\<dagger> \<cdot> y \<cdot> (x + y)\<^sup>\<dagger> = (x + y)\<^sup>\<dagger>"
by (metis local.dagger_denest local.dagger_unfoldl_distr mult_assoc)
end
subsection \<open>Conway Algebras\<close>
class conway_base = pre_conway_base + dioid_one
begin
lemma troeger: "(x + y)\<^sup>\<dagger> \<cdot> z = x\<^sup>\<dagger> \<cdot> (y \<cdot> (x + y)\<^sup>\<dagger> \<cdot> z + z)"
proof -
have "(x + y)\<^sup>\<dagger> \<cdot> z = x\<^sup>\<dagger> \<cdot> z + x\<^sup>\<dagger> \<cdot> y \<cdot> (x + y)\<^sup>\<dagger> \<cdot> z"
by (metis dagger_sum_unfold local.distrib_right')
thus ?thesis
by (metis add_commute local.distrib_left mult_assoc)
qed
lemma dagger_slide_var1: "x\<^sup>\<dagger> \<cdot> x \<le> x \<cdot> x\<^sup>\<dagger>"
- by (metis local.dagger_unfoldl_distr local.dagger_unfoldr_eq local.distrib_left local.eq_iff local.mult_1_right mult_assoc)
+ by (metis local.dagger_unfoldl_distr local.dagger_unfoldr_eq local.distrib_left order.eq_iff local.mult_1_right mult_assoc)
lemma dagger_slide_var1_eq: "x\<^sup>\<dagger> \<cdot> x = x \<cdot> x\<^sup>\<dagger>"
by (metis local.dagger_unfoldl_distr local.dagger_unfoldr_eq local.distrib_left local.mult_1_right mult_assoc)
lemma dagger_slide_eq: "(x \<cdot> y)\<^sup>\<dagger> \<cdot> x = x \<cdot> (y \<cdot> x)\<^sup>\<dagger>"
proof -
have "(x \<cdot> y)\<^sup>\<dagger> \<cdot> x = x + x \<cdot> (y \<cdot> x)\<^sup>\<dagger> \<cdot> y \<cdot> x"
by (metis local.dagger_prod_unfold local.distrib_right' local.mult_onel)
also have "... = x \<cdot> (1 + (y \<cdot> x)\<^sup>\<dagger> \<cdot> y \<cdot> x)"
using local.distrib_left local.mult_1_right mult_assoc by presburger
finally show ?thesis
by (simp add: mult_assoc)
qed
end
subsection \<open>Conway Algebras with Zero\<close>
class near_conway_base_zerol = near_conway_base + near_dioid_one_zerol
begin
lemma dagger_annil [simp]: "1 + x \<cdot> 0 = (x \<cdot> 0)\<^sup>\<dagger>"
by (metis annil dagger_unfoldl_eq mult.assoc)
lemma zero_dagger [simp]: "0\<^sup>\<dagger> = 1"
by (metis add_0_right annil dagger_annil)
end
class pre_conway_base_zerol = near_conway_base_zerol + pre_dioid
class conway_base_zerol = pre_conway_base_zerol + dioid
subclass (in pre_conway_base_zerol) pre_conway_base ..
subclass (in conway_base_zerol) conway_base ..
context conway_base_zerol
begin
lemma "z \<cdot> x \<le> y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<dagger> \<le> y\<^sup>\<dagger> \<cdot> z"
(*nitpick [expect=genuine]*)
oops
end
subsection \<open>Conway Algebras with Simulation\<close>
class near_conway = near_conway_base +
assumes dagger_simr: "z \<cdot> x \<le> y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<dagger> \<le> y\<^sup>\<dagger> \<cdot> z"
begin
lemma dagger_slide_var: "x \<cdot> (y \<cdot> x)\<^sup>\<dagger> \<le> (x \<cdot> y)\<^sup>\<dagger> \<cdot> x"
by (metis eq_refl dagger_simr mult.assoc)
text \<open>Nitpick refutes the next lemma.\<close>
lemma dagger_slide: "x \<cdot> (y \<cdot> x)\<^sup>\<dagger> = (x \<cdot> y)\<^sup>\<dagger> \<cdot> x"
(*nitpick [expect=genuine]*)
oops
text \<open>
We say that $y$ preserves $x$ if $x \cdot y \cdot x = x \cdot y$ and $!x \cdot y \cdot !x = !x \cdot y$. This definition is taken
from Solin~\cite{Solin11}. It is useful for program transformation.
\<close>
lemma preservation1: "x \<cdot> y \<le> x \<cdot> y \<cdot> x \<Longrightarrow> x \<cdot> y\<^sup>\<dagger> \<le> (x \<cdot> y + z)\<^sup>\<dagger> \<cdot> x"
proof -
assume "x \<cdot> y \<le> x \<cdot> y \<cdot> x"
hence "x \<cdot> y \<le> (x \<cdot> y + z) \<cdot> x"
by (simp add: local.join.le_supI1)
thus ?thesis
by (simp add: local.dagger_simr)
qed
end
class near_conway_zerol = near_conway + near_dioid_one_zerol
class pre_conway = near_conway + pre_dioid_one
begin
subclass pre_conway_base ..
lemma dagger_slide: "x \<cdot> (y \<cdot> x)\<^sup>\<dagger> = (x \<cdot> y)\<^sup>\<dagger> \<cdot> x"
- by (metis add.commute dagger_prod_unfold join.sup_least mult_1_right mult.assoc subdistl dagger_slide_var dagger_unfoldl_distr antisym)
+ by (metis add.commute dagger_prod_unfold join.sup_least mult_1_right mult.assoc subdistl dagger_slide_var dagger_unfoldl_distr order.antisym)
lemma dagger_denest2: "(x + y)\<^sup>\<dagger> = x\<^sup>\<dagger> \<cdot> (y \<cdot> x\<^sup>\<dagger>)\<^sup>\<dagger>"
by (metis dagger_denest dagger_slide)
lemma preservation2: "y \<cdot> x \<le> y \<Longrightarrow> (x \<cdot> y)\<^sup>\<dagger> \<cdot> x \<le> x \<cdot> y\<^sup>\<dagger>"
by (metis dagger_slide local.dagger_iso local.mult_isol)
lemma preservation1_eq: "x \<cdot> y \<le> x \<cdot> y \<cdot> x \<Longrightarrow> y \<cdot> x \<le> y \<Longrightarrow> (x \<cdot> y)\<^sup>\<dagger> \<cdot> x = x \<cdot> y\<^sup>\<dagger>"
- by (simp add: local.dagger_simr local.eq_iff preservation2)
+ by (simp add: local.dagger_simr order.eq_iff preservation2)
end
class pre_conway_zerol = near_conway_zerol + pre_dioid_one_zerol
begin
subclass pre_conway ..
end
class conway = pre_conway + dioid_one
class conway_zerol = pre_conway + dioid_one_zerol
begin
subclass conway_base ..
text \<open>Nitpick refutes the next lemmas.\<close>
lemma "1 = 1\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
lemma "(x\<^sup>\<dagger>)\<^sup>\<dagger> = x\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
lemma dagger_denest_var [simp]: "(x + y)\<^sup>\<dagger> = (x\<^sup>\<dagger> \<cdot> y\<^sup>\<dagger>)\<^sup>\<dagger>"
(* nitpick [expect=genuine]*)
oops
lemma star2 [simp]: "(1 + x)\<^sup>\<dagger> = x\<^sup>\<dagger>"
(*nitpick [expect=genuine]*)
oops
end
end
diff --git a/thys/Kleene_Algebra/DRA.thy b/thys/Kleene_Algebra/DRA.thy
--- a/thys/Kleene_Algebra/DRA.thy
+++ b/thys/Kleene_Algebra/DRA.thy
@@ -1,364 +1,364 @@
(* Title: Demonic refinement algebra
Author: Alasdair Armstrong, Victor B. F. Gomes, Georg Struth
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Demonic Refinement Algebras\<close>
theory DRA
imports Kleene_Algebra
begin
text \<open>
A demonic refinement algebra *DRA)~\cite{vonwright04refinement} is a Kleene algebra without right annihilation plus
an operation for possibly infinite iteration.
\<close>
class dra = kleene_algebra_zerol +
fixes strong_iteration :: "'a \<Rightarrow> 'a" ("_\<^sup>\<infinity>" [101] 100)
assumes iteration_unfoldl [simp] : "1 + x \<cdot> x\<^sup>\<infinity> = x\<^sup>\<infinity>"
and coinduction: "y \<le> z + x \<cdot> y \<longrightarrow> y \<le> x\<^sup>\<infinity> \<cdot> z"
and isolation [simp]: "x\<^sup>\<star> + x\<^sup>\<infinity> \<cdot> 0 = x\<^sup>\<infinity>"
begin
text \<open>$\top$ is an abort statement, defined as an infinite skip. It is the maximal element of any DRA.\<close>
abbreviation top_elem :: "'a" ("\<top>") where "\<top> \<equiv> 1\<^sup>\<infinity>"
text \<open>Simple/basic lemmas about the iteration operator\<close>
lemma iteration_refl: "1 \<le> x\<^sup>\<infinity>"
using local.iteration_unfoldl local.order_prop by blast
lemma iteration_1l: "x \<cdot> x\<^sup>\<infinity> \<le> x\<^sup>\<infinity>"
by (metis local.iteration_unfoldl local.join.sup.cobounded2)
lemma top_ref: "x \<le> \<top>"
proof -
have "x \<le> 1 + 1 \<cdot> x"
by simp
thus ?thesis
using local.coinduction by fastforce
qed
lemma it_ext: "x \<le> x\<^sup>\<infinity>"
proof -
have "x \<le> x \<cdot> x\<^sup>\<infinity>"
using iteration_refl local.mult_isol by fastforce
thus ?thesis
by (metis (full_types) local.isolation local.join.sup.coboundedI1 local.star_ext)
qed
lemma it_idem [simp]: "(x\<^sup>\<infinity>)\<^sup>\<infinity> = x\<^sup>\<infinity>"
(*nitpick [expect=genuine]*)
oops
lemma top_mult_annil [simp]: "\<top> \<cdot> x = \<top>"
by (simp add: local.coinduction local.order.antisym top_ref)
lemma top_add_annil [simp]: "\<top> + x = \<top>"
by (simp add: local.join.sup.absorb1 top_ref)
lemma top_elim: "x \<cdot> y \<le> x \<cdot> \<top>"
by (simp add: local.mult_isol top_ref)
lemma iteration_unfoldl_distl [simp]: " y + y \<cdot> x \<cdot> x\<^sup>\<infinity> = y \<cdot> x\<^sup>\<infinity>"
by (metis distrib_left mult.assoc mult_oner iteration_unfoldl)
lemma iteration_unfoldl_distr [simp]: " y + x \<cdot> x\<^sup>\<infinity> \<cdot> y = x\<^sup>\<infinity> \<cdot> y"
by (metis distrib_right' mult_1_left iteration_unfoldl)
lemma iteration_unfoldl' [simp]: "z \<cdot> y + z \<cdot> x \<cdot> x\<^sup>\<infinity> \<cdot> y = z \<cdot> x\<^sup>\<infinity> \<cdot> y"
by (metis iteration_unfoldl_distl local.distrib_right)
lemma iteration_idem [simp]: "x\<^sup>\<infinity> \<cdot> x\<^sup>\<infinity> = x\<^sup>\<infinity>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x\<^sup>\<infinity> \<cdot> x\<^sup>\<infinity> \<le> 1 + x \<cdot> x\<^sup>\<infinity> \<cdot> x\<^sup>\<infinity>"
by (metis add_assoc iteration_unfoldl_distr local.eq_refl local.iteration_unfoldl local.subdistl_eq mult_assoc)
thus "x\<^sup>\<infinity> \<cdot> x\<^sup>\<infinity> \<le> x\<^sup>\<infinity>"
using local.coinduction mult_assoc by fastforce
show "x\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> x\<^sup>\<infinity>"
using local.coinduction by auto
qed
lemma iteration_induct: "x \<cdot> x\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> x"
proof -
have "x + x \<cdot> (x \<cdot> x\<^sup>\<infinity>) = x \<cdot> x\<^sup>\<infinity>"
by (metis (no_types) local.distrib_left local.iteration_unfoldl local.mult_oner)
thus ?thesis
by (simp add: local.coinduction)
qed
lemma iteration_ref_star: "x\<^sup>\<star> \<le> x\<^sup>\<infinity>"
by (simp add: local.star_inductl_one)
lemma iteration_subdist: "x\<^sup>\<infinity> \<le> (x + y)\<^sup>\<infinity>"
by (metis add_assoc' distrib_right' mult_oner coinduction join.sup_ge1 iteration_unfoldl)
lemma iteration_iso: "x \<le> y \<Longrightarrow> x\<^sup>\<infinity> \<le> y\<^sup>\<infinity>"
using iteration_subdist local.order_prop by auto
lemma iteration_unfoldr [simp]: "1 + x\<^sup>\<infinity> \<cdot> x = x\<^sup>\<infinity>"
by (metis add_0_left annil eq_refl isolation mult.assoc iteration_idem iteration_unfoldl iteration_unfoldl_distr star_denest star_one star_prod_unfold star_slide tc)
lemma iteration_unfoldr_distl [simp]: " y + y \<cdot> x\<^sup>\<infinity> \<cdot> x = y \<cdot> x\<^sup>\<infinity>"
by (metis distrib_left mult.assoc mult_oner iteration_unfoldr)
lemma iteration_unfoldr_distr [simp]: " y + x\<^sup>\<infinity> \<cdot> x \<cdot> y = x\<^sup>\<infinity> \<cdot> y"
by (metis iteration_unfoldl_distr iteration_unfoldr_distl)
lemma iteration_unfold_eq: "x\<^sup>\<infinity> \<cdot> x = x \<cdot> x\<^sup>\<infinity>"
by (metis iteration_unfoldl_distr iteration_unfoldr_distl)
lemma iteration_unfoldr' [simp]: "z \<cdot> y + z \<cdot> x\<^sup>\<infinity> \<cdot> x \<cdot> y = z \<cdot> x\<^sup>\<infinity> \<cdot> y"
by (metis distrib_left mult.assoc iteration_unfoldr_distr)
lemma iteration_double [simp]: "(x\<^sup>\<infinity>)\<^sup>\<infinity> = \<top>"
- by (simp add: iteration_iso iteration_refl local.eq_iff top_ref)
+ by (simp add: iteration_iso iteration_refl order.eq_iff top_ref)
lemma star_iteration [simp]: "(x\<^sup>\<star>)\<^sup>\<infinity> = \<top>"
- by (simp add: iteration_iso local.eq_iff top_ref)
+ by (simp add: iteration_iso order.eq_iff top_ref)
lemma iteration_star [simp]: "(x\<^sup>\<infinity>)\<^sup>\<star> = x\<^sup>\<infinity>"
by (metis (no_types) iteration_idem iteration_refl local.star_inductr_var_eq2 local.sup_id_star1)
lemma iteration_star2 [simp]: "x\<^sup>\<star> \<cdot> x\<^sup>\<infinity> = x\<^sup>\<infinity>"
proof -
have f1: "(x\<^sup>\<infinity>)\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<infinity>"
by (metis (no_types) it_ext iteration_induct iteration_star local.bubble_sort local.join.sup.absorb1)
have "x\<^sup>\<infinity> = x\<^sup>\<infinity> \<cdot> x\<^sup>\<infinity>"
by simp
hence "x\<^sup>\<star> \<cdot> x\<^sup>\<infinity> = x\<^sup>\<star> \<cdot> (x\<^sup>\<infinity>)\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> (x\<^sup>\<infinity>)\<^sup>\<star>)\<^sup>\<star>"
using f1 by (metis (no_types) iteration_star local.star_denest_var_4 mult_assoc)
thus ?thesis
using f1 by (metis (no_types) iteration_star local.star_denest_var_4 local.star_denest_var_8)
qed
lemma iteration_zero [simp]: "0\<^sup>\<infinity> = 1"
by (metis add_zeror annil iteration_unfoldl)
lemma iteration_annil [simp]: "(x \<cdot> 0)\<^sup>\<infinity> = 1 + x \<cdot> 0"
by (metis annil iteration_unfoldl mult.assoc)
lemma iteration_subdenest: "x\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> \<le> (x + y)\<^sup>\<infinity>"
by (metis add_commute iteration_idem iteration_subdist local.mult_isol_var)
lemma sup_id_top: "1 \<le> y \<Longrightarrow> y \<cdot> \<top> = \<top>"
- using local.eq_iff local.mult_isol_var top_ref by fastforce
+ using order.eq_iff local.mult_isol_var top_ref by fastforce
lemma iteration_top [simp]: "x\<^sup>\<infinity> \<cdot> \<top> = \<top>"
by (simp add: iteration_refl sup_id_top)
text \<open>Next, we prove some simulation laws for data refinement.\<close>
lemma iteration_sim: "z \<cdot> y \<le> x \<cdot> z \<Longrightarrow> z \<cdot> y\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> z"
proof -
assume assms: "z \<cdot> y \<le> x \<cdot> z"
have "z \<cdot> y\<^sup>\<infinity> = z + z \<cdot> y \<cdot> y\<^sup>\<infinity>"
by simp
also have "... \<le> z + x \<cdot> z \<cdot> y\<^sup>\<infinity>"
by (metis assms add.commute add_iso mult_isor)
finally show "z \<cdot> y\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> z"
by (simp add: local.coinduction mult_assoc)
qed
text \<open>Nitpick gives a counterexample to the dual simulation law.\<close>
lemma "y \<cdot> z \<le> z \<cdot> x \<Longrightarrow> y\<^sup>\<infinity> \<cdot> z \<le> z \<cdot> x\<^sup>\<infinity>"
(*nitpick [expect=genuine]*)
oops
text \<open>Next, we prove some sliding laws.\<close>
lemma iteration_slide_var: "x \<cdot> (y \<cdot> x)\<^sup>\<infinity> \<le> (x \<cdot> y)\<^sup>\<infinity> \<cdot> x"
by (simp add: iteration_sim mult_assoc)
lemma iteration_prod_unfold [simp]: "1 + y \<cdot> (x \<cdot> y)\<^sup>\<infinity> \<cdot> x = (y \<cdot> x)\<^sup>\<infinity>"
-proof (rule antisym)
+proof (rule order.antisym)
have "1 + y \<cdot> (x \<cdot> y)\<^sup>\<infinity> \<cdot> x \<le> 1 + (y \<cdot> x)\<^sup>\<infinity> \<cdot> y \<cdot> x"
using iteration_slide_var local.join.sup_mono local.mult_isor by blast
thus "1 + y \<cdot> (x \<cdot> y)\<^sup>\<infinity> \<cdot> x \<le> (y \<cdot> x)\<^sup>\<infinity>"
by (simp add: mult_assoc)
have "(y \<cdot> x)\<^sup>\<infinity> = 1 + y \<cdot> x \<cdot> (y \<cdot> x)\<^sup>\<infinity>"
by simp
thus "(y \<cdot> x)\<^sup>\<infinity> \<le> 1 + y \<cdot> (x \<cdot> y)\<^sup>\<infinity> \<cdot> x"
by (metis iteration_sim local.eq_refl local.join.sup.mono local.mult_isol mult_assoc)
qed
lemma iteration_slide: "x \<cdot> (y \<cdot> x)\<^sup>\<infinity> = (x \<cdot> y)\<^sup>\<infinity> \<cdot> x"
by (metis iteration_prod_unfold iteration_unfoldl_distr distrib_left mult_1_right mult.assoc)
lemma star_iteration_slide [simp]: " y\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y)\<^sup>\<infinity> = (x\<^sup>\<star> \<cdot> y)\<^sup>\<infinity>"
by (metis iteration_star2 local.conway.dagger_unfoldl_distr local.join.sup.orderE local.mult_isor local.star_invol local.star_subdist local.star_trans_eq)
text \<open>The following laws are called denesting laws.\<close>
lemma iteration_sub_denest: "(x + y)\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> (y \<cdot> x\<^sup>\<infinity>)\<^sup>\<infinity>"
proof -
have "(x + y)\<^sup>\<infinity> = x \<cdot> (x + y)\<^sup>\<infinity> + y \<cdot> (x + y)\<^sup>\<infinity> + 1"
by (metis add.commute distrib_right' iteration_unfoldl)
hence "(x + y)\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> (y \<cdot> (x + y)\<^sup>\<infinity> + 1)"
by (metis add_assoc' join.sup_least join.sup_ge1 join.sup_ge2 coinduction)
moreover hence "x\<^sup>\<infinity> \<cdot> (y \<cdot> (x + y)\<^sup>\<infinity> + 1) \<le> x\<^sup>\<infinity> \<cdot> (y \<cdot> x\<^sup>\<infinity>)\<^sup>\<infinity>"
by (metis add_iso mult.assoc mult_isol add.commute coinduction mult_oner mult_isol)
ultimately show ?thesis
using local.order_trans by blast
qed
lemma iteration_denest: "(x + y)\<^sup>\<infinity> = x\<^sup>\<infinity> \<cdot> (y \<cdot> x\<^sup>\<infinity>)\<^sup>\<infinity>"
proof -
have "x\<^sup>\<infinity> \<cdot> (y \<cdot> x\<^sup>\<infinity>)\<^sup>\<infinity> \<le> x \<cdot> x\<^sup>\<infinity> \<cdot> (y \<cdot> x\<^sup>\<infinity>)\<^sup>\<infinity> + y \<cdot> x\<^sup>\<infinity> \<cdot> (y \<cdot> x\<^sup>\<infinity>)\<^sup>\<infinity> + 1"
by (metis add.commute iteration_unfoldl_distr add_assoc' add.commute iteration_unfoldl order_refl)
thus ?thesis
by (metis add.commute iteration_sub_denest order.antisym coinduction distrib_right' iteration_sub_denest mult.assoc mult_oner order.antisym)
qed
(*
end
sublocale dra \<subseteq> conway_zerol strong_iteration
apply (unfold_locales)
apply (simp add: iteration_denest iteration_slide)
apply simp
by (simp add: iteration_sim)
context dra
begin
*)
lemma iteration_denest2 [simp]: "y\<^sup>\<star> \<cdot> x \<cdot> (x + y)\<^sup>\<infinity> + y\<^sup>\<infinity> = (x + y)\<^sup>\<infinity>"
proof -
have "(x + y)\<^sup>\<infinity> = y\<^sup>\<infinity> \<cdot> x \<cdot> (y\<^sup>\<infinity> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> + y\<^sup>\<infinity>"
by (metis add.commute iteration_denest iteration_slide iteration_unfoldl_distr)
also have "... = y\<^sup>\<star> \<cdot> x \<cdot> (y\<^sup>\<infinity> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> + y\<^sup>\<infinity> \<cdot> 0 + y\<^sup>\<infinity>"
by (metis isolation mult.assoc distrib_right' annil mult.assoc)
also have "... = y\<^sup>\<star> \<cdot> x \<cdot> (y\<^sup>\<infinity> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> + y\<^sup>\<infinity>"
by (metis add.assoc distrib_left mult_1_right add_0_left mult_1_right)
finally show ?thesis
by (metis add.commute iteration_denest iteration_slide mult.assoc)
qed
lemma iteration_denest3: "(y\<^sup>\<star> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> = (x + y)\<^sup>\<infinity>"
-proof (rule antisym)
+proof (rule order.antisym)
have "(y\<^sup>\<star> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> \<le> (y\<^sup>\<infinity> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
by (simp add: iteration_iso iteration_ref_star local.mult_isor)
thus "(y\<^sup>\<star> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> \<le> (x + y)\<^sup>\<infinity>"
by (metis iteration_denest iteration_slide local.join.sup_commute)
have "(x + y)\<^sup>\<infinity> = y\<^sup>\<infinity> + y\<^sup>\<star> \<cdot> x \<cdot> (x + y)\<^sup>\<infinity>"
by (metis iteration_denest2 local.join.sup_commute)
thus "(x + y)\<^sup>\<infinity> \<le> (y\<^sup>\<star> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
by (simp add: local.coinduction)
qed
text \<open>Now we prove separation laws for reasoning about distributed systems in the context of action systems.\<close>
lemma iteration_sep: "y \<cdot> x \<le> x \<cdot> y \<Longrightarrow> (x + y)\<^sup>\<infinity> = x\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
proof -
assume "y \<cdot> x \<le> x \<cdot> y"
hence "y\<^sup>\<star> \<cdot> x \<le> x\<cdot>(x + y)\<^sup>\<star>"
by (metis star_sim1 add.commute mult_isol order_trans star_subdist)
hence "y\<^sup>\<star> \<cdot> x \<cdot> (x + y)\<^sup>\<infinity> + y\<^sup>\<infinity> \<le> x \<cdot> (x + y)\<^sup>\<infinity> + y\<^sup>\<infinity>"
by (metis mult_isor mult.assoc iteration_star2 join.sup.mono eq_refl)
thus ?thesis
by (metis iteration_denest2 add.commute coinduction add.commute less_eq_def iteration_subdenest)
qed
lemma iteration_sim2: "y \<cdot> x \<le> x \<cdot> y \<Longrightarrow> y\<^sup>\<infinity> \<cdot> x\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
by (metis add.commute iteration_sep iteration_subdenest)
lemma iteration_sep2: "y \<cdot> x \<le> x \<cdot> y\<^sup>\<star> \<Longrightarrow> (x + y)\<^sup>\<infinity> = x\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
proof -
assume "y \<cdot> x \<le> x \<cdot> y\<^sup>\<star>"
hence "y\<^sup>\<star> \<cdot> (y\<^sup>\<star> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> y\<^sup>\<star> \<cdot> y\<^sup>\<infinity>"
by (metis mult.assoc mult_isor iteration_sim star_denest_var_2 star_sim1 star_slide_var star_trans_eq tc_eq)
moreover have "x\<^sup>\<infinity> \<cdot> y\<^sup>\<star> \<cdot> y\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
by (metis eq_refl mult.assoc iteration_star2)
moreover have "(y\<^sup>\<star> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity> \<le> y\<^sup>\<star> \<cdot> (y\<^sup>\<star> \<cdot> x)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
by (metis mult_isor mult_onel star_ref)
ultimately show ?thesis
- by (metis antisym iteration_denest3 iteration_subdenest order_trans)
+ by (metis order.antisym iteration_denest3 iteration_subdenest order_trans)
qed
lemma iteration_sep3: "y \<cdot> x \<le> x \<cdot> (x + y) \<Longrightarrow> (x + y)\<^sup>\<infinity> = x\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
proof -
assume "y \<cdot> x \<le> x \<cdot> (x + y)"
hence "y\<^sup>\<star> \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
by (metis star_sim1)
hence "y\<^sup>\<star> \<cdot> x \<cdot> (x + y)\<^sup>\<infinity> + y\<^sup>\<infinity> \<le> x \<cdot> (x + y)\<^sup>\<star> \<cdot> (x + y)\<^sup>\<infinity> + y\<^sup>\<infinity>"
by (metis add_iso mult_isor)
hence "(x + y)\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
by (metis mult.assoc iteration_denest2 iteration_star2 add.commute coinduction)
thus ?thesis
by (metis add.commute less_eq_def iteration_subdenest)
qed
lemma iteration_sep4: "y \<cdot> 0 = 0 \<Longrightarrow> z \<cdot> x = 0 \<Longrightarrow> y \<cdot> x \<le> (x + z) \<cdot> y\<^sup>\<star> \<Longrightarrow> (x + y + z)\<^sup>\<infinity> = x\<^sup>\<infinity> \<cdot> (y + z)\<^sup>\<infinity>"
proof -
assume assms: "y \<cdot> 0 = 0" "z \<cdot> x = 0" "y \<cdot> x \<le> (x + z) \<cdot> y\<^sup>\<star>"
have "y \<cdot> y\<^sup>\<star> \<cdot> z \<le> y\<^sup>\<star> \<cdot> z \<cdot> y\<^sup>\<star>"
by (metis mult_isor star_1l mult_oner order_trans star_plus_one subdistl)
have "y\<^sup>\<star> \<cdot> z \<cdot> x \<le> x \<cdot> y\<^sup>\<star> \<cdot> z"
by (metis join.bot_least assms(1) assms(2) independence1 mult.assoc)
have "y \<cdot> (x + y\<^sup>\<star> \<cdot> z) \<le> (x + z) \<cdot> y\<^sup>\<star> + y \<cdot> y\<^sup>\<star> \<cdot> z"
by (metis assms(3) distrib_left mult.assoc add_iso)
also have "... \<le> (x + y\<^sup>\<star> \<cdot> z) \<cdot> y\<^sup>\<star> + y \<cdot> y\<^sup>\<star> \<cdot> z"
by (metis star_ref join.sup.mono eq_refl mult_1_left mult_isor)
also have "... \<le> (x + y\<^sup>\<star> \<cdot> z) \<cdot> y\<^sup>\<star> + y\<^sup>\<star> \<cdot> z \<cdot> y\<^sup>\<star>" using \<open>y \<cdot> y\<^sup>\<star> \<cdot> z \<le> y\<^sup>\<star> \<cdot> z \<cdot> y\<^sup>\<star>\<close>
by (metis add.commute add_iso)
finally have "y \<cdot> (x + y\<^sup>\<star> \<cdot> z) \<le> (x + y\<^sup>\<star> \<cdot> z) \<cdot> y\<^sup>\<star>"
by (metis add.commute add_idem' add.left_commute distrib_right)
moreover have "(x + y + z)\<^sup>\<infinity> \<le> (x + y + y\<^sup>\<star> \<cdot> z)\<^sup>\<infinity>"
by (metis star_ref join.sup.mono eq_refl mult_1_left mult_isor iteration_iso)
moreover have "... = (x + y\<^sup>\<star> \<cdot> z)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>"
by (metis add_commute calculation(1) iteration_sep2 local.add_left_comm)
moreover have "... = x\<^sup>\<infinity> \<cdot> (y\<^sup>\<star> \<cdot> z)\<^sup>\<infinity> \<cdot> y\<^sup>\<infinity>" using \<open>y\<^sup>\<star> \<cdot> z \<cdot> x \<le> x \<cdot> y\<^sup>\<star> \<cdot> z\<close>
by (metis iteration_sep mult.assoc)
ultimately have "(x + y + z)\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> (y + z)\<^sup>\<infinity>"
by (metis add.commute mult.assoc iteration_denest3)
thus ?thesis
by (metis add.commute add.left_commute less_eq_def iteration_subdenest)
qed
text \<open>Finally, we prove some blocking laws.\<close>
text \<open>Nitpick refutes the next lemma.\<close>
lemma "x \<cdot> y = 0 \<Longrightarrow> x\<^sup>\<infinity> \<cdot> y = y"
(*nitpick*)
oops
lemma iteration_idep: "x \<cdot> y = 0 \<Longrightarrow> x \<cdot> y\<^sup>\<infinity> = x"
by (metis add_zeror annil iteration_unfoldl_distl)
text \<open>Nitpick refutes the next lemma.\<close>
lemma "y \<cdot> w \<le> x \<cdot> y + z \<Longrightarrow> y \<cdot> w\<^sup>\<infinity> \<le> x\<^sup>\<infinity> \<cdot> z"
(*nitpick [expect=genuine]*)
oops
text \<open>At the end of this file, we consider a data refinement example from von Wright~\cite{Wright02}.\<close>
lemma data_refinement:
assumes "s' \<le> s \<cdot> z" and "z \<cdot> e' \<le> e" and "z \<cdot> a' \<le> a \<cdot> z" and "z \<cdot> b \<le> z" and "b\<^sup>\<infinity> = b\<^sup>\<star>"
shows "s' \<cdot> (a' + b)\<^sup>\<infinity> \<cdot> e' \<le> s \<cdot> a\<^sup>\<infinity> \<cdot> e"
proof -
have "z \<cdot> b\<^sup>\<star> \<le> z"
by (metis assms(4) star_inductr_var)
have "(z \<cdot> a') \<cdot> b\<^sup>\<star> \<le> (a \<cdot> z) \<cdot> b\<^sup>\<star>"
by (metis assms(3) mult.assoc mult_isor)
hence "z \<cdot> (a' \<cdot> b\<^sup>\<star>)\<^sup>\<infinity> \<le> a\<^sup>\<infinity> \<cdot> z" using \<open>z \<cdot> b\<^sup>\<star> \<le> z\<close>
by (metis mult.assoc mult_isol order_trans iteration_sim mult.assoc)
have "s' \<cdot> (a' + b)\<^sup>\<infinity> \<cdot> e' \<le> s' \<cdot> b\<^sup>\<star> \<cdot> (a' \<cdot> b\<^sup>\<star>)\<^sup>\<infinity> \<cdot> e'"
by (metis add.commute assms(5) eq_refl iteration_denest mult.assoc)
also have "... \<le> s \<cdot> z \<cdot> b\<^sup>\<star> \<cdot> (a' \<cdot> b\<^sup>\<star>)\<^sup>\<infinity> \<cdot> e'"
by (metis assms(1) mult_isor)
also have "... \<le> s \<cdot> z \<cdot> (a' \<cdot> b\<^sup>\<star>)\<^sup>\<infinity> \<cdot> e'" using \<open>z \<cdot> b\<^sup>\<star> \<le> z\<close>
by (metis mult.assoc mult_isol mult_isor)
also have "... \<le> s \<cdot> a\<^sup>\<infinity> \<cdot> z \<cdot> e'" using \<open>z \<cdot> (a' \<cdot> b\<^sup>\<star>)\<^sup>\<infinity> \<le> a\<^sup>\<infinity> \<cdot> z\<close>
by (metis mult.assoc mult_isol mult_isor)
finally show ?thesis
by (metis assms(2) mult.assoc mult_isol mult.assoc mult_isol order_trans)
qed
end
end
diff --git a/thys/Kleene_Algebra/Kleene_Algebra.thy b/thys/Kleene_Algebra/Kleene_Algebra.thy
--- a/thys/Kleene_Algebra/Kleene_Algebra.thy
+++ b/thys/Kleene_Algebra/Kleene_Algebra.thy
@@ -1,970 +1,970 @@
(* Title: Kleene Algebra
Author: Alasdair Armstrong, Georg Struth, Tjark Weber
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Kleene Algebras\<close>
theory Kleene_Algebra
imports Conway
begin
subsection \<open>Left Near Kleene Algebras\<close>
text \<open>Extending the hierarchy developed in @{theory Kleene_Algebra.Dioid} we now
add an operation of Kleene star, finite iteration, or reflexive
transitive closure to variants of Dioids. Since a multiplicative unit
is needed for defining the star we only consider variants with~$1$;
$0$ can be added separately. We consider the left star induction axiom
and the right star induction axiom independently since in some
applications, e.g., Salomaa's axioms, probabilistic Kleene algebras,
or completeness proofs with respect to the equational theoy of regular
expressions and regular languages, the right star induction axiom is
not needed or not valid.
We start with near dioids, then consider pre-dioids and finally
dioids. It turns out that many of the known laws of Kleene algebras
hold already in these more general settings. In fact, all our
equational theorems have been proved within left Kleene algebras, as
expected.
Although most of the proofs in this file could be fully automated by
Sledgehammer and Metis, we display step-wise proofs as they would
appear in a text book. First, this file may then be useful as a
reference manual on Kleene algebra. Second, it is better protected
against changes in the underlying theories and supports easy
translation of proofs into other settings.\<close>
class left_near_kleene_algebra = near_dioid_one + star_op +
assumes star_unfoldl: "1 + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
and star_inductl: "z + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> y"
begin
text \<open>First we prove two immediate consequences of the unfold
axiom. The first one states that starred elements are reflexive.\<close>
lemma star_ref [simp]: "1 \<le> x\<^sup>\<star>"
using star_unfoldl by auto
text \<open>Reflexivity of starred elements implies, by definition of the
order, that~$1$ is an additive unit for starred elements.\<close>
lemma star_plus_one [simp]: "1 + x\<^sup>\<star> = x\<^sup>\<star>"
using less_eq_def star_ref by blast
lemma star_1l [simp]: "x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
using star_unfoldl by auto
lemma "x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
text \<open>Next we show that starred elements are transitive.\<close>
lemma star_trans_eq [simp]: "x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
-proof (rule antisym) \<comment> \<open>this splits an equation into two inequalities\<close>
+proof (rule order.antisym) \<comment> \<open>this splits an equation into two inequalities\<close>
have "x\<^sup>\<star> + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by auto
thus "x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (simp add: star_inductl)
next show "x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> x\<^sup>\<star>"
using mult_isor star_ref by fastforce
qed
lemma star_trans: "x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by simp
text \<open>We now derive variants of the star induction axiom.\<close>
lemma star_inductl_var: "x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y"
proof -
assume "x \<cdot> y \<le> y"
hence "y + x \<cdot> y \<le> y"
by simp
thus "x\<^sup>\<star> \<cdot> y \<le> y"
by (simp add: star_inductl)
qed
lemma star_inductl_var_equiv [simp]: "x\<^sup>\<star> \<cdot> y \<le> y \<longleftrightarrow> x \<cdot> y \<le> y"
proof
assume "x \<cdot> y \<le> y"
thus "x\<^sup>\<star> \<cdot> y \<le> y"
by (simp add: star_inductl_var)
next
assume "x\<^sup>\<star> \<cdot> y \<le> y"
hence "x\<^sup>\<star> \<cdot> y = y"
- by (metis eq_iff mult_1_left mult_isor star_ref)
+ by (metis order.eq_iff mult_1_left mult_isor star_ref)
moreover hence "x \<cdot> y = x \<cdot> x\<^sup>\<star> \<cdot> y"
by (simp add: mult.assoc)
moreover have "... \<le> x\<^sup>\<star> \<cdot> y"
by (metis mult_isor star_1l)
ultimately show "x \<cdot> y \<le> y"
by auto
qed
lemma star_inductl_var_eq: "x \<cdot> y = y \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y"
- by (metis eq_iff star_inductl_var)
+ by (metis order.eq_iff star_inductl_var)
lemma star_inductl_var_eq2: "y = x \<cdot> y \<Longrightarrow> y = x\<^sup>\<star> \<cdot> y"
proof -
assume hyp: "y = x \<cdot> y"
hence "y \<le> x\<^sup>\<star> \<cdot> y"
using mult_isor star_ref by fastforce
thus "y = x\<^sup>\<star> \<cdot> y"
- using hyp eq_iff by auto
+ using hyp order.eq_iff by auto
qed
lemma "y = x \<cdot> y \<longleftrightarrow> y = x\<^sup>\<star> \<cdot> y"
(*nitpick [expect=genuine]*)
oops
lemma "x\<^sup>\<star> \<cdot> z \<le> y \<Longrightarrow> z + x \<cdot> y \<le> y"
(*nitpick [expect=genuine]*)
oops
lemma star_inductl_one: "1 + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y"
using star_inductl by force
lemma star_inductl_star: "x \<cdot> y\<^sup>\<star> \<le> y\<^sup>\<star> \<Longrightarrow> x\<^sup>\<star> \<le> y\<^sup>\<star>"
by (simp add: star_inductl_one)
lemma star_inductl_eq: "z + x \<cdot> y = y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> y"
by (simp add: star_inductl)
text \<open>We now prove two facts related to~$1$.\<close>
lemma star_subid: "x \<le> 1 \<Longrightarrow> x\<^sup>\<star> = 1"
proof -
assume "x \<le> 1"
hence "1 + x \<cdot> 1 \<le> 1"
by simp
hence "x\<^sup>\<star> \<le> 1"
by (metis mult_oner star_inductl)
thus "x\<^sup>\<star> = 1"
by (simp add: order.antisym)
qed
lemma star_one [simp]: "1\<^sup>\<star> = 1"
by (simp add: star_subid)
text \<open>We now prove a subdistributivity property for the star (which
is equivalent to isotonicity of star).\<close>
lemma star_subdist: "x\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
proof -
have "x \<cdot> (x + y)\<^sup>\<star> \<le> (x + y) \<cdot> (x + y)\<^sup>\<star>"
by simp
also have "... \<le> (x + y)\<^sup>\<star>"
by (metis star_1l)
thus ?thesis
using calculation order_trans star_inductl_star by blast
qed
lemma star_subdist_var: "x\<^sup>\<star> + y\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
using join.sup_commute star_subdist by force
lemma star_iso [intro]: "x \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y\<^sup>\<star>"
by (metis less_eq_def star_subdist)
text \<open>We now prove some more simple properties.\<close>
lemma star_invol [simp]: "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
by (fact star_trans_eq)
thus "(x\<^sup>\<star>)\<^sup>\<star> \<le> x\<^sup>\<star>"
by (simp add: star_inductl_star)
have"(x\<^sup>\<star>)\<^sup>\<star> \<cdot> (x\<^sup>\<star>)\<^sup>\<star> \<le> (x\<^sup>\<star>)\<^sup>\<star>"
by (fact star_trans)
hence "x \<cdot> (x\<^sup>\<star>)\<^sup>\<star> \<le> (x\<^sup>\<star>)\<^sup>\<star>"
by (meson star_inductl_var_equiv)
thus "x\<^sup>\<star> \<le> (x\<^sup>\<star>)\<^sup>\<star>"
by (simp add: star_inductl_star)
qed
lemma star2 [simp]: "(1 + x)\<^sup>\<star> = x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "x\<^sup>\<star> \<le> (1 + x)\<^sup>\<star>"
by auto
have "x\<^sup>\<star> + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by simp
thus "(1 + x)\<^sup>\<star> \<le> x\<^sup>\<star>"
by (simp add: star_inductl_star)
qed
lemma "1 + x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "x \<le> x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "x \<cdot> z \<le> z \<cdot> y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> z \<cdot> y\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
text \<open>The following facts express inductive conditions that are used
to show that @{term "(x + y)\<^sup>\<star>"} is the greatest term that can be built
from~@{term x} and~@{term y}.\<close>
lemma prod_star_closure: "x \<le> z\<^sup>\<star> \<Longrightarrow> y \<le> z\<^sup>\<star> \<Longrightarrow> x \<cdot> y \<le> z\<^sup>\<star>"
proof -
assume assm: "x \<le> z\<^sup>\<star>" "y \<le> z\<^sup>\<star>"
hence "y + z\<^sup>\<star> \<cdot> z\<^sup>\<star> \<le> z\<^sup>\<star>"
by simp
hence "z\<^sup>\<star> \<cdot> y \<le> z\<^sup>\<star>"
by (simp add: star_inductl)
also have "x \<cdot> y \<le> z\<^sup>\<star> \<cdot> y"
by (simp add: assm mult_isor)
thus "x \<cdot> y \<le> z\<^sup>\<star>"
using calculation order.trans by blast
qed
lemma star_star_closure: "x\<^sup>\<star> \<le> z\<^sup>\<star> \<Longrightarrow> (x\<^sup>\<star>)\<^sup>\<star> \<le> z\<^sup>\<star>"
by (metis star_invol)
lemma star_closed_unfold: "x\<^sup>\<star> = x \<Longrightarrow> x = 1 + x \<cdot> x"
by (metis star_plus_one star_trans_eq)
lemma "x\<^sup>\<star> = x \<longleftrightarrow> x = 1 + x \<cdot> x"
(*nitpick [expect=genuine]*)
oops
end (* left_near_kleene_algebra *)
subsection \<open>Left Pre-Kleene Algebras\<close>
class left_pre_kleene_algebra = left_near_kleene_algebra + pre_dioid_one
begin
text \<open>We first prove that the star operation is extensive.\<close>
lemma star_ext [simp]: "x \<le> x\<^sup>\<star>"
proof -
have "x \<le> x \<cdot> x\<^sup>\<star>"
by (metis mult_oner mult_isol star_ref)
thus ?thesis
by (metis order_trans star_1l)
qed
text \<open>We now prove a right star unfold law.\<close>
lemma star_1r [simp]: "x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
proof -
have "x + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by simp
thus ?thesis
by (fact star_inductl)
qed
lemma star_unfoldr: "1 + x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
by simp
lemma "1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
text \<open>Next we prove a simulation law for the star. It is
instrumental in proving further properties.\<close>
lemma star_sim1: "x \<cdot> z \<le> z \<cdot> y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> z \<cdot> y\<^sup>\<star>"
proof -
assume "x \<cdot> z \<le> z \<cdot> y"
hence "x \<cdot> z \<cdot> y\<^sup>\<star> \<le> z \<cdot> y \<cdot> y\<^sup>\<star>"
by (simp add: mult_isor)
also have "... \<le> z \<cdot> y\<^sup>\<star>"
by (simp add: mult_isol mult_assoc)
finally have "x \<cdot> z \<cdot> y\<^sup>\<star> \<le> z \<cdot> y\<^sup>\<star>"
by simp
hence "z + x \<cdot> z \<cdot> y\<^sup>\<star> \<le> z \<cdot> y\<^sup>\<star>"
by (metis join.sup_least mult_isol mult_oner star_ref)
thus "x\<^sup>\<star> \<cdot> z \<le> z \<cdot> y\<^sup>\<star>"
by (simp add: star_inductl mult_assoc)
qed
text \<open>The next lemma is used in omega algebras to prove, for
instance, Bachmair and Dershowitz's separation of termination
theorem~\cite{bachmair86commutation}. The property at the left-hand
side of the equivalence is known as \emph{quasicommutation}.\<close>
lemma quasicomm_var: "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star> \<longleftrightarrow> y\<^sup>\<star> \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
proof
assume "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
thus "y\<^sup>\<star> \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
using star_sim1 by force
next
assume "y\<^sup>\<star> \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
thus "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
by (meson mult_isor order_trans star_ext)
qed
lemma star_slide1: "(x \<cdot> y)\<^sup>\<star> \<cdot> x \<le> x \<cdot> (y \<cdot> x)\<^sup>\<star>"
by (simp add: mult_assoc star_sim1)
lemma "(x \<cdot> y)\<^sup>\<star> \<cdot> x = x \<cdot> (y \<cdot> x)\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma star_slide_var1: "x\<^sup>\<star> \<cdot> x \<le> x \<cdot> x\<^sup>\<star>"
by (simp add: star_sim1)
text \<open>We now show that the (left) star unfold axiom can be strengthened to an equality.\<close>
lemma star_unfoldl_eq [simp]: "1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "1 + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (fact star_unfoldl)
have "1 + x \<cdot> (1 + x \<cdot> x\<^sup>\<star>) \<le> 1 + x \<cdot> x\<^sup>\<star>"
by (meson join.sup_mono eq_refl mult_isol star_unfoldl)
thus "x\<^sup>\<star> \<le> 1 + x \<cdot> x\<^sup>\<star>"
by (simp add: star_inductl_one)
qed
lemma "1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
text \<open>Next we relate the star and the reflexive transitive closure
operation.\<close>
lemma star_rtc1_eq [simp]: "1 + x + x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
by (simp add: join.sup.absorb2)
lemma star_rtc1: "1 + x + x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by simp
lemma star_rtc2: "1 + x \<cdot> x \<le> x \<longleftrightarrow> x = x\<^sup>\<star>"
proof
assume "1 + x \<cdot> x \<le> x"
thus "x = x\<^sup>\<star>"
- by (simp add: local.eq_iff local.star_inductl_one)
+ by (simp add: order.eq_iff local.star_inductl_one)
next
assume "x = x\<^sup>\<star>"
thus "1 + x \<cdot> x \<le> x"
using local.star_closed_unfold by auto
qed
lemma star_rtc3: "1 + x \<cdot> x = x \<longleftrightarrow> x = x\<^sup>\<star>"
by (metis order_refl star_plus_one star_rtc2 star_trans_eq)
lemma star_rtc_least: "1 + x + y \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y"
proof -
assume "1 + x + y \<cdot> y \<le> y"
hence "1 + x \<cdot> y \<le> y"
by (metis join.le_sup_iff mult_isol_var star_trans_eq star_rtc2)
thus "x\<^sup>\<star> \<le> y"
by (simp add: star_inductl_one)
qed
lemma star_rtc_least_eq: "1 + x + y \<cdot> y = y \<Longrightarrow> x\<^sup>\<star> \<le> y"
by (simp add: star_rtc_least)
lemma "1 + x + y \<cdot> y \<le> y \<longleftrightarrow> x\<^sup>\<star> \<le> y"
(*nitpick [expect=genuine]*)
oops
text \<open>The next lemmas are again related to closure conditions\<close>
lemma star_subdist_var_1: "x \<le> (x + y)\<^sup>\<star>"
by (meson join.sup.boundedE star_ext)
lemma star_subdist_var_2: "x \<cdot> y \<le> (x + y)\<^sup>\<star>"
by (meson join.sup.boundedE prod_star_closure star_ext)
lemma star_subdist_var_3: "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (simp add: prod_star_closure star_iso)
text \<open>We now prove variants of sum-elimination laws under a star.
These are also known a denesting laws or as sum-star laws.\<close>
lemma star_denest [simp]: "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x + y \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis join.sup.bounded_iff mult_1_right mult_isol_var mult_onel star_ref star_ext)
thus "(x + y)\<^sup>\<star> \<le> (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (fact star_iso)
have "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (fact star_subdist_var_3)
thus "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (simp add: prod_star_closure star_inductl_star)
qed
lemma star_sum_var [simp]: "(x\<^sup>\<star> + y\<^sup>\<star>)\<^sup>\<star> = (x + y)\<^sup>\<star>"
by simp
lemma star_denest_var [simp]: "x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> = (x + y)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "1 \<le> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (metis mult_isol_var mult_oner star_ref)
moreover have "x \<cdot> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (simp add: mult_isor)
moreover have "y \<cdot> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (metis mult_isol_var mult_onel star_1l star_ref)
ultimately have "1 + (x + y) \<cdot> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by auto
thus "(x + y)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (metis mult.assoc mult_oner star_inductl)
have "(y \<cdot> x\<^sup>\<star>)\<^sup>\<star> \<le> (y\<^sup>\<star> \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (simp add: mult_isol_var star_iso)
hence "(y \<cdot> x\<^sup>\<star>)\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (metis add.commute star_denest)
moreover have "x\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (fact star_subdist)
ultimately show "x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
using prod_star_closure by blast
qed
lemma star_denest_var_2 [simp]: "x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by simp
lemma star_denest_var_3 [simp]: "x\<^sup>\<star> \<cdot> (y\<^sup>\<star> \<cdot> x\<^sup>\<star>)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by simp
lemma star_denest_var_4 [ac_simps]: "(y\<^sup>\<star> \<cdot> x\<^sup>\<star>)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (metis add_comm star_denest)
lemma star_denest_var_5 [ac_simps]: "x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> = y\<^sup>\<star> \<cdot> (x \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (simp add: star_denest_var_4)
lemma "x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma star_denest_var_6 [simp]: "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> (x + y)\<^sup>\<star> = (x + y)\<^sup>\<star>"
using mult_assoc by simp
lemma star_denest_var_7 [simp]: "(x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star> = (x + y)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "(x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (simp add: mult_assoc)
thus "(x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by simp
have "1 \<le> (x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis dual_order.trans mult_1_left mult_isor star_ref)
moreover have "(x + y) \<cdot> (x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
using mult_isor star_1l by presburger
ultimately have "1 + (x + y) \<cdot> (x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by simp
thus "(x + y)\<^sup>\<star> \<le> (x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis mult.assoc star_inductl_one)
qed
lemma star_denest_var_8 [simp]: "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (simp add: mult_assoc)
lemma star_denest_var_9 [simp]: "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
using star_denest_var_7 by simp
text \<open>The following statements are well known from term
rewriting. They are all variants of the Church-Rosser theorem in
Kleene algebra~\cite{struth06churchrosser}. But first we prove a law
relating two confluence properties.\<close>
lemma confluence_var [iff]: "y \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<longleftrightarrow> y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
proof
assume "y \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
thus "y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
using star_sim1 by fastforce
next
assume "y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
thus "y \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (meson mult_isor order_trans star_ext)
qed
lemma church_rosser [intro]: "y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<Longrightarrow> (x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
assume "y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
hence "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y\<^sup>\<star>) \<le> x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis mult_isol mult_isor mult.assoc)
hence "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y\<^sup>\<star>) \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (simp add: mult_assoc)
moreover have "1 \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis dual_order.trans mult_1_right mult_isol star_ref)
ultimately have "1 + x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y\<^sup>\<star>) \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by simp
hence "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (simp add: star_inductl_one)
thus "(x + y)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by simp
thus "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by simp
qed
lemma church_rosser_var: "y \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<Longrightarrow> (x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by fastforce
lemma church_rosser_to_confluence: "(x + y)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<Longrightarrow> y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
- by (metis add_comm eq_iff star_subdist_var_3)
+ by (metis add_comm order.eq_iff star_subdist_var_3)
lemma church_rosser_equiv: "y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<longleftrightarrow> (x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
- using church_rosser_to_confluence eq_iff by blast
+ using church_rosser_to_confluence order.eq_iff by blast
lemma confluence_to_local_confluence: "y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<Longrightarrow> y \<cdot> x \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (meson mult_isol_var order_trans star_ext)
lemma "y \<cdot> x \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<Longrightarrow> y\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "y \<cdot> x \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<Longrightarrow> (x + y)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
(* nitpick [expect=genuine] *) oops
text \<open>
More variations could easily be proved. The last counterexample shows
that Newman's lemma needs a wellfoundedness assumption. This is well
known.
\<close>
text \<open>The next lemmas relate the reflexive transitive closure and
the transitive closure.\<close>
lemma sup_id_star1: "1 \<le> x \<Longrightarrow> x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
proof -
assume "1 \<le> x"
hence " x\<^sup>\<star> \<le> x \<cdot> x\<^sup>\<star>"
using mult_isor by fastforce
thus "x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
- by (simp add: eq_iff)
+ by (simp add: order.eq_iff)
qed
lemma sup_id_star2: "1 \<le> x \<Longrightarrow> x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
by (metis order.antisym mult_isol mult_oner star_1r)
lemma "1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "(x \<cdot> y)\<^sup>\<star> \<cdot> x = x \<cdot> (y \<cdot> x)\<^sup>\<star>"
(*nitpick [expect=genuine]*)
oops
lemma "x \<cdot> x = x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
(* nitpick [expect=genuine] *)
oops
end (* left_pre_kleene_algebra *)
subsection \<open>Left Kleene Algebras\<close>
class left_kleene_algebra = left_pre_kleene_algebra + dioid_one
begin
text \<open>In left Kleene algebras the non-fact @{prop "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"}
is a good challenge for counterexample generators. A model of left
Kleene algebras in which the right star induction law does not hold
has been given by Kozen~\cite{kozen90kleene}.\<close>
text \<open>We now show that the right unfold law becomes an equality.\<close>
lemma star_unfoldr_eq [simp]: "1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "1 + x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
by (fact star_unfoldr)
have "1 + x \<cdot> (1 + x\<^sup>\<star> \<cdot> x) = 1 + (1 + x \<cdot> x\<^sup>\<star>) \<cdot> x"
using distrib_left distrib_right mult_1_left mult_1_right mult_assoc by presburger
also have "... = 1 + x\<^sup>\<star> \<cdot> x"
by simp
finally show "x\<^sup>\<star> \<le> 1 + x\<^sup>\<star> \<cdot> x"
by (simp add: star_inductl_one)
qed
text \<open>The following more complex unfold law has been used as an
axiom, called prodstar, by Conway~\cite{conway71regular}.\<close>
lemma star_prod_unfold [simp]: "1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y = (x \<cdot> y)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "(x \<cdot> y)\<^sup>\<star> = 1 + (x \<cdot> y)\<^sup>\<star> \<cdot> x \<cdot> y"
by (simp add: mult_assoc)
thus "(x \<cdot> y)\<^sup>\<star> \<le> 1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y"
by (metis join.sup_mono mult_isor order_refl star_slide1)
have "1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y \<le> 1 + x \<cdot> y \<cdot> (x \<cdot> y)\<^sup>\<star>"
by (metis join.sup_mono eq_refl mult.assoc mult_isol star_slide1)
thus "1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y \<le> (x \<cdot> y)\<^sup>\<star>"
by simp
qed
text \<open>The slide laws, which have previously been inequalities, now
become equations.\<close>
lemma star_slide [ac_simps]: "(x \<cdot> y)\<^sup>\<star> \<cdot> x = x \<cdot> (y \<cdot> x)\<^sup>\<star>"
proof -
have "x \<cdot> (y \<cdot> x)\<^sup>\<star> = x \<cdot> (1 + y \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> x)"
by simp
also have "... = (1 + x \<cdot> y \<cdot> (x \<cdot> y)\<^sup>\<star>) \<cdot> x"
by (simp add: distrib_left mult_assoc)
finally show ?thesis
by simp
qed
lemma star_slide_var [ac_simps]: "x\<^sup>\<star> \<cdot> x = x \<cdot> x\<^sup>\<star>"
by (metis mult_onel mult_oner star_slide)
lemma star_sum_unfold_var [simp]: "1 + x\<^sup>\<star> \<cdot> (x + y)\<^sup>\<star> \<cdot> y\<^sup>\<star> = (x + y)\<^sup>\<star>"
by (metis star_denest star_denest_var_3 star_denest_var_4 star_plus_one star_slide)
text \<open>The following law shows how starred sums can be unfolded.\<close>
lemma star_sum_unfold [simp]: "x\<^sup>\<star> + x\<^sup>\<star> \<cdot> y \<cdot> (x + y)\<^sup>\<star> = (x + y)\<^sup>\<star>"
proof -
have "(x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star> )\<^sup>\<star>"
by simp
also have "... = x\<^sup>\<star> \<cdot> (1 + y \<cdot> x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star> )\<^sup>\<star>)"
by simp
also have "... = x\<^sup>\<star> \<cdot> (1 + y \<cdot> (x + y)\<^sup>\<star>)"
by (simp add: mult.assoc)
finally show ?thesis
by (simp add: distrib_left mult_assoc)
qed
text \<open>The following property appears in process algebra.\<close>
lemma troeger: "(x + y)\<^sup>\<star> \<cdot> z = x\<^sup>\<star> \<cdot> (y \<cdot> (x + y)\<^sup>\<star> \<cdot> z + z)"
proof -
have "(x + y)\<^sup>\<star> \<cdot> z = x\<^sup>\<star> \<cdot> z + x\<^sup>\<star> \<cdot> y \<cdot> (x + y)\<^sup>\<star> \<cdot> z"
by (metis (full_types) distrib_right star_sum_unfold)
thus ?thesis
by (simp add: add_commute distrib_left mult_assoc)
qed
text \<open>The following properties are related to a property from
propositional dynamic logic which has been attributed to Albert
Meyer~\cite{harelkozentiuryn00dynamic}. Here we prove it as a theorem
of Kleene algebra.\<close>
lemma star_square: "(x \<cdot> x)\<^sup>\<star> \<le> x\<^sup>\<star>"
proof -
have "x \<cdot> x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (simp add: prod_star_closure)
thus ?thesis
by (simp add: star_inductl_star)
qed
lemma meyer_1 [simp]: "(1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> = x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x \<cdot> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> = x \<cdot> (x \<cdot> x)\<^sup>\<star> + x \<cdot> x \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (simp add: distrib_left)
also have "... \<le> x \<cdot> (x \<cdot> x)\<^sup>\<star> + (x \<cdot> x)\<^sup>\<star>"
using join.sup_mono star_1l by blast
finally have "x \<cdot> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> \<le> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (simp add: join.sup_commute)
moreover have "1 \<le> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
using join.sup.coboundedI1 by auto
ultimately have "1 + x \<cdot> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> \<le> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
by auto
thus "x\<^sup>\<star> \<le> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (simp add: star_inductl_one mult_assoc)
show "(1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> \<le> x\<^sup>\<star>"
by (simp add: prod_star_closure star_square)
qed
text \<open>The following lemma says that transitive elements are equal to
their transitive closure.\<close>
lemma tc: "x \<cdot> x \<le> x \<Longrightarrow> x\<^sup>\<star> \<cdot> x = x"
proof -
assume "x \<cdot> x \<le> x"
hence "x + x \<cdot> x \<le> x"
by simp
hence "x\<^sup>\<star> \<cdot> x \<le> x"
by (fact star_inductl)
thus "x\<^sup>\<star> \<cdot> x = x"
- by (metis mult_isol mult_oner star_ref star_slide_var eq_iff)
+ by (metis mult_isol mult_oner star_ref star_slide_var order.eq_iff)
qed
lemma tc_eq: "x \<cdot> x = x \<Longrightarrow> x\<^sup>\<star> \<cdot> x = x"
by (auto intro: tc)
text \<open>The next fact has been used by Boffa~\cite{boffa90remarque} to
axiomatise the equational theory of regular expressions.\<close>
lemma boffa_var: "x \<cdot> x \<le> x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
proof -
assume "x \<cdot> x \<le> x"
moreover have "x\<^sup>\<star> = 1 + x\<^sup>\<star> \<cdot> x"
by simp
ultimately show "x\<^sup>\<star> = 1 + x"
by (simp add: tc)
qed
lemma boffa: "x \<cdot> x = x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
by (auto intro: boffa_var)
(*
text {*
For the following two lemmas, Isabelle could neither find a
counterexample nor a proof automatically.
*}
lemma "z \<cdot> x \<le> y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y\<^sup>\<star> \<cdot> z"
-- "unfortunately, no counterexample found"
oops
lemma star_sim3: "z \<cdot> x = y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<star> = y\<^sup>\<star> \<cdot> z"
-- "we conjecture that this is not provable"
oops
*)
end (* left_kleene_algebra *)
subsection \<open>Left Kleene Algebras with Zero\<close>
text \<open>
There are applications where only a left zero is assumed, for instance
in the context of total correctness and for demonic refinement
algebras~\cite{vonwright04refinement}.
\<close>
class left_kleene_algebra_zerol = left_kleene_algebra + dioid_one_zerol
begin
sublocale conway: near_conway_base_zerol star
by standard (simp_all add: local.star_slide)
lemma star_zero [simp]: "0\<^sup>\<star> = 1"
by (rule local.conway.zero_dagger)
text \<open>
In principle,~$1$ could therefore be defined from~$0$ in this setting.
\<close>
end (* left_kleene_algebra_zerol *)
class left_kleene_algebra_zero = left_kleene_algebra_zerol + dioid_one_zero
subsection \<open>Pre-Kleene Algebras\<close>
text \<open>Pre-Kleene algebras are essentially probabilistic Kleene
algebras~\cite{mciverweber05pka}. They have a weaker right star
unfold axiom. We are still looking for theorems that could be proved
in this setting.\<close>
class pre_kleene_algebra = left_pre_kleene_algebra +
assumes weak_star_unfoldr: "z + y \<cdot> (x + 1) \<le> y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"
subsection \<open>Kleene Algebras\<close>
class kleene_algebra_zerol = left_kleene_algebra_zerol +
assumes star_inductr: "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"
begin
lemma star_sim2: "z \<cdot> x \<le> y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y\<^sup>\<star> \<cdot> z"
proof -
assume "z \<cdot> x \<le> y \<cdot> z"
hence "y\<^sup>\<star> \<cdot> z \<cdot> x \<le> y\<^sup>\<star> \<cdot> y \<cdot> z"
using mult_isol mult_assoc by auto
also have "... \<le> y\<^sup>\<star> \<cdot> z"
by (simp add: mult_isor)
finally have "y\<^sup>\<star> \<cdot> z \<cdot> x \<le> y\<^sup>\<star> \<cdot> z"
by simp
moreover have "z \<le> y\<^sup>\<star> \<cdot> z"
using mult_isor star_ref by fastforce
ultimately have "z + y\<^sup>\<star> \<cdot> z \<cdot> x \<le> y\<^sup>\<star> \<cdot> z"
by simp
thus "z \<cdot> x\<^sup>\<star> \<le> y\<^sup>\<star> \<cdot> z"
by (simp add: star_inductr)
qed
sublocale conway: pre_conway star
by standard (simp add: star_sim2)
lemma star_inductr_var: "y \<cdot> x \<le> y \<Longrightarrow> y \<cdot> x\<^sup>\<star> \<le> y"
by (simp add: star_inductr)
lemma star_inductr_var_equiv: "y \<cdot> x \<le> y \<longleftrightarrow> y \<cdot> x\<^sup>\<star> \<le> y"
by (meson order_trans mult_isol star_ext star_inductr_var)
lemma star_sim3: "z \<cdot> x = y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<star> = y\<^sup>\<star> \<cdot> z"
- by (simp add: eq_iff star_sim1 star_sim2)
+ by (simp add: order.eq_iff star_sim1 star_sim2)
lemma star_sim4: "x \<cdot> y \<le> y \<cdot> x \<Longrightarrow> x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> y\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (auto intro: star_sim1 star_sim2)
lemma star_inductr_eq: "z + y \<cdot> x = y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"
by (auto intro: star_inductr)
lemma star_inductr_var_eq: "y \<cdot> x = y \<Longrightarrow> y \<cdot> x\<^sup>\<star> \<le> y"
by (auto intro: star_inductr_var)
lemma star_inductr_var_eq2: "y \<cdot> x = y \<Longrightarrow> y \<cdot> x\<^sup>\<star> = y"
by (metis mult_onel star_one star_sim3)
lemma bubble_sort: "y \<cdot> x \<le> x \<cdot> y \<Longrightarrow> (x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (fastforce intro: star_sim4)
lemma independence1: "x \<cdot> y = 0 \<Longrightarrow> x\<^sup>\<star> \<cdot> y = y"
proof -
assume "x \<cdot> y = 0"
moreover have "x\<^sup>\<star> \<cdot> y = y + x\<^sup>\<star> \<cdot> x \<cdot> y"
by (metis distrib_right mult_onel star_unfoldr_eq)
ultimately show "x\<^sup>\<star> \<cdot> y = y"
- by (metis add_0_left add.commute join.sup_ge1 eq_iff star_inductl_eq)
+ by (metis add_0_left add.commute join.sup_ge1 order.eq_iff star_inductl_eq)
qed
lemma independence2: "x \<cdot> y = 0 \<Longrightarrow> x \<cdot> y\<^sup>\<star> = x"
by (metis annil mult_onel star_sim3 star_zero)
lemma lazycomm_var: "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star> + y \<longleftrightarrow> y \<cdot> x\<^sup>\<star> \<le> x \<cdot> (x + y)\<^sup>\<star> + y"
proof
let ?t = "x \<cdot> (x + y)\<^sup>\<star>"
assume hyp: "y \<cdot> x \<le> ?t + y"
have "(?t + y) \<cdot> x = ?t \<cdot> x + y \<cdot> x"
by (fact distrib_right)
also have "... \<le> ?t \<cdot> x + ?t + y"
using hyp join.sup.coboundedI2 join.sup_assoc by auto
also have "... \<le> ?t + y"
using eq_refl join.sup_least join.sup_mono mult_isol prod_star_closure star_subdist_var_1 mult_assoc by presburger
finally have "y + (?t + y) \<cdot> x \<le> ?t + y"
by simp
thus "y \<cdot> x\<^sup>\<star> \<le> x \<cdot> (x + y)\<^sup>\<star> + y"
by (fact star_inductr)
next
assume "y \<cdot> x\<^sup>\<star> \<le> x \<cdot> (x + y)\<^sup>\<star> + y"
thus "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star> + y"
using dual_order.trans mult_isol star_ext by blast
qed
lemma arden_var: "(\<forall>y v. y \<le> x \<cdot> y + v \<longrightarrow> y \<le> x\<^sup>\<star> \<cdot> v) \<Longrightarrow> z = x \<cdot> z + w \<Longrightarrow> z = x\<^sup>\<star> \<cdot> w"
- by (auto simp: add_comm eq_iff star_inductl_eq)
+ by (auto simp: add_comm order.eq_iff star_inductl_eq)
lemma "(\<forall>x y. y \<le> x \<cdot> y \<longrightarrow> y = 0) \<Longrightarrow> y \<le> x \<cdot> y + z \<Longrightarrow> y \<le> x\<^sup>\<star> \<cdot> z"
by (metis eq_refl mult_onel)
end
text \<open>Finally, here come the Kleene algebras \`a~la
Kozen~\cite{kozen94complete}. We only prove quasi-identities in this
section. Since left Kleene algebras are complete with respect to the
equational theory of regular expressions and regular languages, all
identities hold already without the right star induction axiom.\<close>
class kleene_algebra = left_kleene_algebra_zero +
assumes star_inductr': "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"
begin
subclass kleene_algebra_zerol
by standard (simp add: star_inductr')
sublocale conway_zerol: conway star ..
text \<open>
The next lemma shows that opposites of Kleene algebras (i.e., Kleene
algebras with the order of multiplication swapped) are again Kleene
algebras.
\<close>
lemma dual_kleene_algebra:
"class.kleene_algebra (+) (\<odot>) 1 0 (\<le>) (<) star"
proof
fix x y z :: 'a
show "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis opp_mult_def distrib_left)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (fact add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "x + x = x"
by (fact add_idem)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "z \<odot> x \<le> z \<odot> (x + y)"
by (metis mult_isor opp_mult_def order_prop)
show "1 + x \<odot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis opp_mult_def order_refl star_slide_var star_unfoldl_eq)
show "z + x \<odot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<odot> z \<le> y"
by (metis opp_mult_def star_inductr)
show "z + y \<odot> x \<le> y \<Longrightarrow> z \<odot> x\<^sup>\<star> \<le> y"
by (metis opp_mult_def star_inductl)
qed
end (* kleene_algebra *)
text \<open>We finish with some properties on (multiplicatively)
commutative Kleene algebras. A chapter in Conway's
book~\cite{conway71regular} is devoted to this topic.\<close>
class commutative_kleene_algebra = kleene_algebra +
assumes mult_comm [ac_simps]: "x \<cdot> y = y \<cdot> x"
begin
lemma conway_c3 [simp]: "(x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
using church_rosser mult_comm by auto
lemma conway_c4: "(x\<^sup>\<star> \<cdot> y)\<^sup>\<star> = 1 + x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> y"
by (metis conway_c3 star_denest_var star_prod_unfold)
lemma cka_1: "(x \<cdot> y)\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis conway_c3 star_invol star_iso star_subdist_var_2)
lemma cka_2 [simp]: "x\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis conway_c3 mult_comm star_denest_var)
lemma conway_c4_var [simp]: "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis conway_c3 star_invol)
lemma conway_c2_var: "(x \<cdot> y)\<^sup>\<star> \<cdot> x \<cdot> y \<cdot> y\<^sup>\<star> \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis mult_isor star_1r mult_assoc)
lemma conway_c2 [simp]: "(x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>) = x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "(x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>) \<le> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis cka_1 conway_c3 prod_star_closure star_ext star_sum_var)
have "x \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>) = x \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + 1 + y \<cdot> y\<^sup>\<star>)"
by (simp add: add_assoc)
also have "... = x \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y \<cdot> y\<^sup>\<star>)"
by (simp add: add_commute)
also have "... = (x \<cdot> y)\<^sup>\<star> \<cdot> (x \<cdot> x\<^sup>\<star>) + (x \<cdot> y)\<^sup>\<star> \<cdot> x \<cdot> y \<cdot> y\<^sup>\<star>"
using distrib_left mult_comm mult_assoc by force
also have "... \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star> + (x \<cdot> y)\<^sup>\<star> \<cdot> x \<cdot> y \<cdot> y\<^sup>\<star>"
using add_iso mult_isol by force
also have "... \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star> + (x \<cdot> y)\<^sup>\<star> \<cdot> y\<^sup>\<star>"
using conway_c2_var join.sup_mono by blast
also have "... = (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>)"
by (simp add: distrib_left)
finally have "x \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>) \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>)" .
moreover have "y\<^sup>\<star> \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>)"
by (metis dual_order.trans join.sup_ge2 mult_1_left mult_isor star_ref)
ultimately have "y\<^sup>\<star> + x \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>) \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>)"
by simp
thus "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> (x\<^sup>\<star> + y\<^sup>\<star>)"
by (simp add: mult.assoc star_inductl)
qed
end (* commutative_kleene_algebra *)
end
diff --git a/thys/Kleene_Algebra/Omega_Algebra.thy b/thys/Kleene_Algebra/Omega_Algebra.thy
--- a/thys/Kleene_Algebra/Omega_Algebra.thy
+++ b/thys/Kleene_Algebra/Omega_Algebra.thy
@@ -1,459 +1,459 @@
(* Title: Omega Algebra
Author: Alasdair Armstrong, Georg Struth, Tjark Weber
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Omega Algebras\<close>
theory Omega_Algebra
imports Kleene_Algebra
begin
text \<open>
\emph{Omega algebras}~\cite{cohen00omega} extend Kleene algebras by an
$\omega$-operation that axiomatizes infinite iteration (just like the
Kleene star axiomatizes finite iteration).
\<close>
subsection \<open>Left Omega Algebras\<close>
text \<open>
In this section we consider \emph{left omega algebras}, i.e., omega
algebras based on left Kleene algebras. Surprisingly, we are still
looking for statements mentioning~$\omega$ that are true in omega
algebras, but do not already hold in left omega algebras.
\<close>
class left_omega_algebra = left_kleene_algebra_zero + omega_op +
assumes omega_unfold: "x\<^sup>\<omega> \<le> x \<cdot> x\<^sup>\<omega>"
and omega_coinduct: "y \<le> z + x \<cdot> y \<Longrightarrow> y \<le> x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z"
begin
text \<open>First we prove some variants of the coinduction axiom.\<close>
lemma omega_coinduct_var1: "y \<le> 1 + x \<cdot> y \<Longrightarrow> y \<le> x\<^sup>\<omega> + x\<^sup>\<star>"
using local.omega_coinduct by fastforce
lemma omega_coinduct_var2: "y \<le> x \<cdot> y \<Longrightarrow> y \<le> x\<^sup>\<omega>"
by (metis add.commute add_zero_l annir omega_coinduct)
lemma omega_coinduct_eq: "y = z + x \<cdot> y \<Longrightarrow> y \<le> x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z"
by (simp add: local.omega_coinduct)
lemma omega_coinduct_eq_var1: "y = 1 + x \<cdot> y \<Longrightarrow> y \<le> x\<^sup>\<omega> + x\<^sup>\<star>"
by (simp add: omega_coinduct_var1)
lemma omega_coinduct_eq_var2: "y = x \<cdot> y \<Longrightarrow> y \<le> x\<^sup>\<omega>"
by (simp add: omega_coinduct_var2)
lemma "y = x \<cdot> y + z \<Longrightarrow> y = x\<^sup>\<star> \<cdot> z + x\<^sup>\<omega>"
(* nitpick [expect=genuine] -- "2-element counterexample" *)
oops
lemma "y = 1 + x \<cdot> y \<Longrightarrow> y = x\<^sup>\<omega> + x\<^sup>\<star>"
(* nitpick [expect=genuine] -- "3-element counterexample" *)
oops
lemma "y = x \<cdot> y \<Longrightarrow> y = x\<^sup>\<omega>"
(* nitpick [expect=genuine] -- "2-element counterexample" *)
oops
text \<open>Next we strengthen the unfold law to an equation.\<close>
lemma omega_unfold_eq [simp]: "x \<cdot> x\<^sup>\<omega> = x\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x \<cdot> x\<^sup>\<omega> \<le> x \<cdot> x \<cdot> x\<^sup>\<omega>"
by (simp add: local.mult_isol local.omega_unfold mult_assoc)
thus "x \<cdot> x\<^sup>\<omega> \<le> x\<^sup>\<omega>"
by (simp add: mult_assoc omega_coinduct_var2)
show "x\<^sup>\<omega> \<le> x \<cdot> x\<^sup>\<omega>"
by (fact omega_unfold)
qed
lemma omega_unfold_var: "z + x \<cdot> x\<^sup>\<omega> \<le> x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z"
by (simp add: local.omega_coinduct)
lemma "z + x \<cdot> x\<^sup>\<omega> = x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z"
(* nitpick [expect=genuine] -- "4-element counterexample" *)
oops
text \<open>We now prove subdistributivity and isotonicity of omega.\<close>
lemma omega_subdist: "x\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
proof -
have "x\<^sup>\<omega> \<le> (x + y) \<cdot> x\<^sup>\<omega>"
by simp
thus ?thesis
by (rule omega_coinduct_var2)
qed
lemma omega_iso: "x \<le> y \<Longrightarrow> x\<^sup>\<omega> \<le> y\<^sup>\<omega>"
by (metis less_eq_def omega_subdist)
lemma omega_subdist_var: "x\<^sup>\<omega> + y\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
by (simp add: omega_iso)
lemma zero_omega [simp]: "0\<^sup>\<omega> = 0"
by (metis annil omega_unfold_eq)
text \<open>The next lemma is another variant of omega unfold\<close>
lemma star_omega_1 [simp]: "x\<^sup>\<star> \<cdot> x\<^sup>\<omega> = x\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x \<cdot> x\<^sup>\<omega> \<le> x\<^sup>\<omega>"
by simp
thus "x\<^sup>\<star> \<cdot> x\<^sup>\<omega> \<le> x\<^sup>\<omega>"
by simp
show "x\<^sup>\<omega> \<le> x\<^sup>\<star> \<cdot> x\<^sup>\<omega>"
using local.star_inductl_var_eq2 by auto
qed
text \<open>The next lemma says that~@{term "1\<^sup>\<omega>"} is the maximal element
of omega algebra. We therefore baptise it~$\top$.\<close>
lemma max_element: "x \<le> 1\<^sup>\<omega>"
by (simp add: omega_coinduct_eq_var2)
definition top ("\<top>")
where "\<top> = 1\<^sup>\<omega>"
lemma star_omega_3 [simp]: "(x\<^sup>\<star>)\<^sup>\<omega> = \<top>"
proof -
have "1 \<le> x\<^sup>\<star>"
by (fact star_ref)
hence "\<top> \<le> (x\<^sup>\<star>)\<^sup>\<omega>"
by (simp add: omega_iso top_def)
thus ?thesis
by (simp add: local.order.antisym max_element top_def)
qed
text \<open>The following lemma is strange since it is counterintuitive
that one should be able to append something after an infinite
iteration.\<close>
lemma omega_1: "x\<^sup>\<omega> \<cdot> y \<le> x\<^sup>\<omega>"
proof -
have "x\<^sup>\<omega> \<cdot> y \<le> x \<cdot> x\<^sup>\<omega> \<cdot> y"
by simp
thus ?thesis
by (metis mult.assoc omega_coinduct_var2)
qed
lemma "x\<^sup>\<omega> \<cdot> y = x\<^sup>\<omega>"
(*nitpick [expect=genuine] -- "2-element counterexample"*)
oops
lemma omega_sup_id: "1 \<le> y \<Longrightarrow> x\<^sup>\<omega> \<cdot> y = x\<^sup>\<omega>"
- using local.eq_iff local.mult_isol omega_1 by fastforce
+ using order.eq_iff local.mult_isol omega_1 by fastforce
lemma omega_top [simp]: "x\<^sup>\<omega> \<cdot> \<top> = x\<^sup>\<omega>"
by (simp add: max_element omega_sup_id top_def)
lemma supid_omega: "1 \<le> x \<Longrightarrow> x\<^sup>\<omega> = \<top>"
by (simp add: local.order.antisym max_element omega_iso top_def)
lemma "x\<^sup>\<omega> = \<top> \<Longrightarrow> 1 \<le> x"
(* nitpick [expect=genuine] -- "4-element counterexample" *)
oops
text \<open>Next we prove a simulation law for the omega operation\<close>
lemma omega_simulation: "z \<cdot> x \<le> y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<omega> \<le> y\<^sup>\<omega>"
proof -
assume hyp: "z \<cdot> x \<le> y \<cdot> z"
have "z \<cdot> x\<^sup>\<omega> = z \<cdot> x \<cdot> x\<^sup>\<omega>"
by (simp add: mult_assoc)
also have "... \<le> y \<cdot> z \<cdot> x\<^sup>\<omega>"
by (simp add: hyp local.mult_isor)
finally show "z \<cdot> x\<^sup>\<omega> \<le> y\<^sup>\<omega>"
by (simp add: mult_assoc omega_coinduct_var2)
qed
lemma "z \<cdot> x \<le> y \<cdot> z \<Longrightarrow> z \<cdot> x\<^sup>\<omega> \<le> y\<^sup>\<omega> \<cdot> z"
(* nitpick [expect=genuine] -- "4-element counterexample" *)
oops
lemma "y \<cdot> z \<le> z \<cdot> x \<Longrightarrow> y\<^sup>\<omega> \<le> z \<cdot> x\<^sup>\<omega>"
(* nitpick [expect=genuine] -- "2-element counterexample" *)
oops
lemma "y \<cdot> z \<le> z \<cdot> x \<Longrightarrow> y\<^sup>\<omega> \<cdot> z \<le> x\<^sup>\<omega>"
(* nitpick [expect=genuine] -- "4-element counterexample" *)
oops
text \<open>Next we prove transitivity of omega elements.\<close>
lemma omega_omega: "(x\<^sup>\<omega>)\<^sup>\<omega> \<le> x\<^sup>\<omega>"
by (metis omega_1 omega_unfold_eq)
(*
lemma "x\<^sup>\<omega> \<cdot> x\<^sup>\<omega> = x\<^sup>\<omega>"
nitpick -- "no proof, no counterexample"
lemma "(x\<^sup>\<omega>)\<^sup>\<omega> = x\<^sup>\<omega>"
nitpick -- "no proof, no counterexample"
*)
text \<open>The next lemmas are axioms of Wagner's complete axiomatisation
for omega-regular languages~\cite{Wagner77omega}, but in a slightly
different setting.\<close>
lemma wagner_1 [simp]: "(x \<cdot> x\<^sup>\<star>)\<^sup>\<omega> = x\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
have "(x \<cdot> x\<^sup>\<star>)\<^sup>\<omega> = x \<cdot> x\<^sup>\<star> \<cdot> x \<cdot> x\<^sup>\<star> \<cdot> (x \<cdot> x\<^sup>\<star>)\<^sup>\<omega>"
by (metis mult.assoc omega_unfold_eq)
also have "... = x \<cdot> x \<cdot> x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> (x \<cdot> x\<^sup>\<star>)\<^sup>\<omega>"
by (simp add: local.star_slide_var mult_assoc)
also have "... = x \<cdot> x \<cdot> x\<^sup>\<star> \<cdot> (x \<cdot> x\<^sup>\<star>)\<^sup>\<omega>"
by (simp add: mult_assoc)
also have "... = x \<cdot> (x \<cdot> x\<^sup>\<star>)\<^sup>\<omega>"
by (simp add: mult_assoc)
thus "(x \<cdot> x\<^sup>\<star>)\<^sup>\<omega> \<le> x\<^sup>\<omega>"
using calculation omega_coinduct_eq_var2 by auto
show "x\<^sup>\<omega> \<le> (x \<cdot> x\<^sup>\<star>)\<^sup>\<omega>"
by (simp add: mult_assoc omega_coinduct_eq_var2)
qed
lemma wagner_2_var: "x \<cdot> (y \<cdot> x)\<^sup>\<omega> \<le> (x \<cdot> y)\<^sup>\<omega>"
proof -
have "x \<cdot> y \<cdot> x \<le> x \<cdot> y \<cdot> x"
by auto
thus "x \<cdot> (y \<cdot> x)\<^sup>\<omega> \<le> (x \<cdot> y)\<^sup>\<omega>"
by (simp add: mult_assoc omega_simulation)
qed
lemma wagner_2 [simp]: "x \<cdot> (y \<cdot> x)\<^sup>\<omega> = (x \<cdot> y)\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
show "x \<cdot> (y \<cdot> x)\<^sup>\<omega> \<le> (x \<cdot> y)\<^sup>\<omega>"
by (rule wagner_2_var)
have "(x \<cdot> y)\<^sup>\<omega> = x \<cdot> y \<cdot> (x \<cdot> y)\<^sup>\<omega>"
by simp
thus "(x \<cdot> y)\<^sup>\<omega> \<le> x \<cdot> (y \<cdot> x)\<^sup>\<omega>"
by (metis mult.assoc mult_isol wagner_2_var)
qed
text \<open>
This identity is called~(A8) in Wagner's paper.
\<close>
lemma wagner_3:
assumes "x \<cdot> (x + y)\<^sup>\<omega> + z = (x + y)\<^sup>\<omega>"
shows "(x + y)\<^sup>\<omega> = x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z"
-proof (rule antisym)
+proof (rule order.antisym)
show "(x + y)\<^sup>\<omega> \<le> x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z"
using assms local.join.sup_commute omega_coinduct_eq by auto
have "x\<^sup>\<star> \<cdot> z \<le> (x + y)\<^sup>\<omega>"
using assms local.join.sup_commute local.star_inductl_eq by auto
thus "x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z \<le> (x + y)\<^sup>\<omega>"
by (simp add: omega_subdist)
qed
text \<open>
This identity is called~(R4) in Wagner's paper.
\<close>
lemma wagner_1_var [simp]: "(x\<^sup>\<star> \<cdot> x)\<^sup>\<omega> = x\<^sup>\<omega>"
by (simp add: local.star_slide_var)
lemma star_omega_4 [simp]: "(x\<^sup>\<omega>)\<^sup>\<star> = 1 + x\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
have "(x\<^sup>\<omega>)\<^sup>\<star> = 1 + x\<^sup>\<omega> \<cdot> (x\<^sup>\<omega>)\<^sup>\<star>"
by simp
also have "... \<le> 1 + x\<^sup>\<omega> \<cdot> \<top>"
by (simp add: omega_sup_id)
finally show "(x\<^sup>\<omega>)\<^sup>\<star> \<le> 1 + x\<^sup>\<omega>"
by simp
show "1 + x\<^sup>\<omega> \<le> (x\<^sup>\<omega>)\<^sup>\<star>"
by simp
qed
lemma star_omega_5 [simp]: "x\<^sup>\<omega> \<cdot> (x\<^sup>\<omega>)\<^sup>\<star> = x\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
show "x\<^sup>\<omega> \<cdot> (x\<^sup>\<omega>)\<^sup>\<star> \<le> x\<^sup>\<omega>"
by (rule omega_1)
show "x\<^sup>\<omega> \<le> x\<^sup>\<omega> \<cdot> (x\<^sup>\<omega>)\<^sup>\<star>"
by (simp add: omega_sup_id)
qed
text \<open>The next law shows how omegas below a sum can be unfolded.\<close>
lemma omega_sum_unfold: "x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> y \<cdot> (x + y)\<^sup>\<omega> = (x + y)\<^sup>\<omega>"
proof -
have "(x + y)\<^sup>\<omega> = x \<cdot> (x + y)\<^sup>\<omega> + y \<cdot> (x+y)\<^sup>\<omega>"
by (metis distrib_right omega_unfold_eq)
thus ?thesis
by (metis mult.assoc wagner_3)
qed
text \<open>
The next two lemmas apply induction and coinduction to this law.
\<close>
lemma omega_sum_unfold_coind: "(x + y)\<^sup>\<omega> \<le> (x\<^sup>\<star> \<cdot> y)\<^sup>\<omega> + (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<omega>"
by (simp add: omega_coinduct_eq omega_sum_unfold)
lemma omega_sum_unfold_ind: "(x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
by (simp add: local.star_inductl_eq omega_sum_unfold)
lemma wagner_1_gen: "(x \<cdot> y\<^sup>\<star>)\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
proof -
have "(x \<cdot> y\<^sup>\<star>)\<^sup>\<omega> \<le> ((x + y) \<cdot> (x + y)\<^sup>\<star>)\<^sup>\<omega>"
using local.join.le_sup_iff local.join.sup.cobounded1 local.mult_isol_var local.star_subdist_var omega_iso by presburger
thus ?thesis
by (metis wagner_1)
qed
lemma wagner_1_var_gen: "(x\<^sup>\<star> \<cdot> y)\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
proof -
have "(x\<^sup>\<star> \<cdot> y)\<^sup>\<omega> = x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<omega>"
by simp
also have "... \<le> x\<^sup>\<star> \<cdot> (x + y)\<^sup>\<omega>"
by (metis add.commute mult_isol wagner_1_gen)
also have "... \<le> (x + y)\<^sup>\<star> \<cdot> (x + y)\<^sup>\<omega>"
using local.mult_isor local.star_subdist by auto
thus ?thesis
by (metis calculation order_trans star_omega_1)
qed
text \<open>The next lemma is a variant of the denest law for the star at
the level of omega.\<close>
lemma omega_denest [simp]: "(x + y)\<^sup>\<omega> = (x\<^sup>\<star> \<cdot> y)\<^sup>\<omega> + (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
show "(x + y)\<^sup>\<omega> \<le> (x\<^sup>\<star> \<cdot> y)\<^sup>\<omega> + (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<omega>"
by (rule omega_sum_unfold_coind)
have "(x\<^sup>\<star> \<cdot> y)\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
by (rule wagner_1_var_gen)
hence "(x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
by (simp add: omega_sum_unfold_ind)
thus "(x\<^sup>\<star> \<cdot> y)\<^sup>\<omega> + (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
by (simp add: wagner_1_var_gen)
qed
text \<open>The next lemma yields a separation theorem for infinite
iteration in the presence of a quasicommutation property. A
nondeterministic loop over~@{term x} and~@{term y} can be refined into
separate infinite loops over~@{term x} and~@{term y}.\<close>
lemma omega_sum_refine:
assumes "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
shows "(x + y)\<^sup>\<omega> = x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> y\<^sup>\<omega>"
-proof (rule antisym)
+proof (rule order.antisym)
have a: "y\<^sup>\<star> \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
using assms local.quasicomm_var by blast
have "(x + y)\<^sup>\<omega> = y\<^sup>\<omega> + y\<^sup>\<star> \<cdot> x \<cdot> (x + y)\<^sup>\<omega>"
by (metis add.commute omega_sum_unfold)
also have "... \<le> y\<^sup>\<omega> + x \<cdot> (x + y)\<^sup>\<star> \<cdot> (x + y)\<^sup>\<omega>"
using a local.join.sup_mono local.mult_isol_var by blast
also have "... \<le> x \<cdot> (x + y)\<^sup>\<omega> + y\<^sup>\<omega>"
using local.eq_refl local.join.sup_commute mult_assoc star_omega_1 by presburger
finally show "(x + y)\<^sup>\<omega> \<le> x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> y\<^sup>\<omega>"
by (metis add_commute local.omega_coinduct)
have "x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> y\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega> + (x + y)\<^sup>\<star> \<cdot> (x + y)\<^sup>\<omega>"
using local.join.sup.cobounded2 local.join.sup.mono local.mult_isol_var local.star_subdist omega_iso omega_subdist by presburger
thus "x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> y\<^sup>\<omega> \<le> (x + y)\<^sup>\<omega>"
by (metis local.join.sup_idem star_omega_1)
qed
text \<open>The following theorem by Bachmair and
Dershowitz~\cite{bachmair86commutation} is a corollary.\<close>
lemma bachmair_dershowitz:
assumes "y \<cdot> x \<le> x \<cdot> (x + y)\<^sup>\<star>"
shows "(x + y)\<^sup>\<omega> = 0 \<longleftrightarrow> x\<^sup>\<omega> + y\<^sup>\<omega> = 0"
by (metis add_commute assms local.annir local.join.le_bot local.no_trivial_inverse omega_subdist omega_sum_refine)
text \<open>
The next lemmas consider an abstract variant of the empty word
property from language theory and match it with the absence of
infinite iteration~\cite{struth12regeq}.
\<close>
definition (in dioid_one_zero) ewp
where "ewp x \<equiv> \<not>(\<forall>y. y \<le> x \<cdot> y \<longrightarrow> y = 0)"
lemma ewp_super_id1: "0 \<noteq> 1 \<Longrightarrow> 1 \<le> x \<Longrightarrow> ewp x"
by (metis ewp_def mult_oner)
lemma "0 \<noteq> 1 \<Longrightarrow> 1 \<le> x \<longleftrightarrow> ewp x"
(* nitpick [expect=genuine] -- "3-element counterexample" *)
oops
text \<open>The next facts relate the absence of the empty word property
with the absence of infinite iteration.\<close>
lemma ewp_neg_and_omega: "\<not> ewp x \<longleftrightarrow> x\<^sup>\<omega> = 0"
proof
assume "\<not> ewp x"
hence "\<forall> y. y \<le> x \<cdot> y \<longrightarrow> y = 0"
by (meson ewp_def)
thus "x\<^sup>\<omega> = 0"
by simp
next
assume "x\<^sup>\<omega> = 0"
hence "\<forall> y. y \<le> x \<cdot> y \<longrightarrow> y = 0"
using local.join.le_bot local.omega_coinduct_var2 by blast
thus "\<not> ewp x"
by (meson ewp_def)
qed
lemma ewp_alt1: "(\<forall>z. x\<^sup>\<omega> \<le> x\<^sup>\<star> \<cdot> z) \<longleftrightarrow> (\<forall>y z. y \<le> x \<cdot> y + z \<longrightarrow> y \<le> x\<^sup>\<star> \<cdot> z)"
by (metis add_comm less_eq_def omega_coinduct omega_unfold_eq order_prop)
lemma ewp_alt: "x\<^sup>\<omega> = 0 \<longleftrightarrow> (\<forall>y z. y \<le> x \<cdot> y + z \<longrightarrow> y \<le> x\<^sup>\<star> \<cdot> z)"
- by (metis annir antisym ewp_alt1 join.bot_least)
+ by (metis annir order.antisym ewp_alt1 join.bot_least)
text \<open>So we have obtained a condition for Arden's lemma in omega
algebra.\<close>
lemma omega_super_id1: "0 \<noteq> 1 \<Longrightarrow> 1 \<le> x \<Longrightarrow> x\<^sup>\<omega> \<noteq> 0"
using ewp_neg_and_omega ewp_super_id1 by blast
lemma omega_super_id2: "0 \<noteq> 1 \<Longrightarrow> x\<^sup>\<omega> = 0 \<Longrightarrow> \<not>(1 \<le> x)"
using omega_super_id1 by blast
text \<open>The next lemmas are abstract versions of Arden's lemma from
language theory.\<close>
lemma ardens_lemma_var:
assumes "x\<^sup>\<omega> = 0"
and "z + x \<cdot> y = y"
shows "x\<^sup>\<star> \<cdot> z = y"
proof -
have "y \<le> x\<^sup>\<omega> + x\<^sup>\<star> \<cdot> z"
by (simp add: assms(2) local.omega_coinduct_eq)
hence "y \<le> x\<^sup>\<star> \<cdot> z"
by (simp add: assms(1))
thus "x\<^sup>\<star> \<cdot> z = y"
- by (simp add: assms(2) local.eq_iff local.star_inductl_eq)
+ by (simp add: assms(2) order.eq_iff local.star_inductl_eq)
qed
lemma ardens_lemma: "\<not> ewp x \<Longrightarrow> z + x \<cdot> y = y \<Longrightarrow> x\<^sup>\<star> \<cdot> z = y"
by (simp add: ardens_lemma_var ewp_neg_and_omega)
lemma ardens_lemma_equiv:
assumes "\<not> ewp x"
shows "z + x \<cdot> y = y \<longleftrightarrow> x\<^sup>\<star> \<cdot> z = y"
by (metis ardens_lemma_var assms ewp_neg_and_omega local.conway.dagger_unfoldl_distr mult_assoc)
lemma ardens_lemma_var_equiv: "x\<^sup>\<omega> = 0 \<Longrightarrow> (z + x \<cdot> y = y \<longleftrightarrow> x\<^sup>\<star> \<cdot> z = y)"
by (simp add: ardens_lemma_equiv ewp_neg_and_omega)
lemma arden_conv1: "(\<forall>y z. z + x \<cdot> y = y \<longrightarrow> x\<^sup>\<star> \<cdot> z = y) \<Longrightarrow> \<not> ewp x"
by (metis add_zero_l annir ewp_neg_and_omega omega_unfold_eq)
lemma arden_conv2: "(\<forall>y z. z + x \<cdot> y = y \<longrightarrow> x\<^sup>\<star> \<cdot> z = y) \<Longrightarrow> x\<^sup>\<omega> = 0"
using arden_conv1 ewp_neg_and_omega by blast
lemma arden_var3: "(\<forall>y z. z + x \<cdot> y = y \<longrightarrow> x\<^sup>\<star> \<cdot> z = y) \<longleftrightarrow> x\<^sup>\<omega> = 0"
using arden_conv2 ardens_lemma_var by blast
end
subsection \<open>Omega Algebras\<close>
class omega_algebra = kleene_algebra + left_omega_algebra
end
diff --git a/thys/Knuth_Morris_Pratt/KMP.thy b/thys/Knuth_Morris_Pratt/KMP.thy
--- a/thys/Knuth_Morris_Pratt/KMP.thy
+++ b/thys/Knuth_Morris_Pratt/KMP.thy
@@ -1,939 +1,939 @@
theory KMP
imports Refine_Imperative_HOL.IICF
"HOL-Library.Sublist"
begin
declare len_greater_imp_nonempty[simp del] min_absorb2[simp]
no_notation Ref.update ("_ := _" 62)
section\<open>Specification\<close>text_raw\<open>\label{sec:spec}\<close>
subsection\<open>Sublist-predicate with a position check\<close>
subsubsection\<open>Definition\<close>
text \<open>One could define\<close>
definition "sublist_at' xs ys i \<equiv> take (length xs) (drop i ys) = xs"
text\<open>However, this doesn't handle out-of-bound indexes uniformly:\<close>
value[nbe] "sublist_at' [] [a] 5"
value[nbe] "sublist_at' [a] [a] 5"
value[nbe] "sublist_at' [] [] 5"
text\<open>Instead, we use a recursive definition:\<close>
fun sublist_at :: "'a list \<Rightarrow> 'a list \<Rightarrow> nat \<Rightarrow> bool" where
"sublist_at (x#xs) (y#ys) 0 \<longleftrightarrow> x=y \<and> sublist_at xs ys 0" |
"sublist_at xs (y#ys) (Suc i) \<longleftrightarrow> sublist_at xs ys i" |
"sublist_at [] ys 0 \<longleftrightarrow> True" |
"sublist_at _ [] _ \<longleftrightarrow> False"
text\<open>In the relevant cases, both definitions agree:\<close>
lemma "i \<le> length ys \<Longrightarrow> sublist_at xs ys i \<longleftrightarrow> sublist_at' xs ys i"
unfolding sublist_at'_def
by (induction xs ys i rule: sublist_at.induct) auto
text\<open>However, the new definition has some reasonable properties:\<close>
subsubsection\<open>Properties\<close>
lemma sublist_lengths: "sublist_at xs ys i \<Longrightarrow> i + length xs \<le> length ys"
by (induction xs ys i rule: sublist_at.induct) auto
lemma Nil_is_sublist: "sublist_at ([] :: 'x list) ys i \<longleftrightarrow> i \<le> length ys"
by (induction "[] :: 'x list" ys i rule: sublist_at.induct) auto
text\<open>Furthermore, we need:\<close>
lemma sublist_step[intro]:
"\<lbrakk>i + length xs < length ys; sublist_at xs ys i; ys!(i + length xs) = x\<rbrakk> \<Longrightarrow> sublist_at (xs@[x]) ys i"
apply (induction xs ys i rule: sublist_at.induct)
apply auto
using sublist_at.elims(3) by fastforce
lemma all_positions_sublist:
"\<lbrakk>i + length xs \<le> length ys; \<forall>jj<length xs. ys!(i+jj) = xs!jj\<rbrakk> \<Longrightarrow> sublist_at xs ys i"
proof (induction xs rule: rev_induct)
case Nil
then show ?case by (simp add: Nil_is_sublist)
next
case (snoc x xs)
from \<open>i + length (xs @ [x]) \<le> length ys\<close> have "i + length xs \<le> length ys" by simp
moreover have "\<forall>jj<length xs. ys!(i + jj) = xs!jj"
by (simp add: nth_append snoc.prems(2))
ultimately have "sublist_at xs ys i"
using snoc.IH by simp
then show ?case
using snoc.prems by auto
qed
lemma sublist_all_positions: "sublist_at xs ys i \<Longrightarrow> \<forall>jj<length xs. ys!(i+jj) = xs!jj"
by (induction xs ys i rule: sublist_at.induct) (auto simp: nth_Cons')
text\<open>It also connects well to theory @{theory "HOL-Library.Sublist"} (compare @{thm[source] sublist_def}):\<close>
lemma sublist_at_altdef:
"sublist_at xs ys i \<longleftrightarrow> (\<exists>ps ss. ys = ps@xs@ss \<and> i = length ps)"
proof (induction xs ys i rule: sublist_at.induct)
case (2 ss t ts i)
show "sublist_at ss (t#ts) (Suc i) \<longleftrightarrow> (\<exists>xs ys. t#ts = xs@ss@ys \<and> Suc i = length xs)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
then have "sublist_at ss ts i" by simp
with "2.IH" obtain xs where "\<exists>ys. ts = xs@ss@ys \<and> i = length xs" by auto
then have "\<exists>ys. t#ts = (t#xs)@ss@ys \<and> Suc i = length (t#xs)" by simp
then show ?rhs by blast
next
assume ?rhs
then obtain xs where "\<exists>ys. t#ts = xs@ss@ys \<and> length xs = Suc i"
by (blast dest: sym)
then have "\<exists>ys. ts = (tl xs)@ss@ys \<and> i = length (tl xs)"
by (auto simp add: length_Suc_conv)
then have "\<exists>xs ys. ts = xs@ss@ys \<and> i = length xs" by blast
with "2.IH" show ?lhs by simp
qed
qed auto
corollary sublist_iff_sublist_at: "Sublist.sublist xs ys \<longleftrightarrow> (\<exists>i. sublist_at xs ys i)"
by (simp add: sublist_at_altdef Sublist.sublist_def)
subsection\<open>Sublist-check algorithms\<close>
text\<open>
We use the Isabelle Refinement Framework (Theory @{theory Refine_Monadic.Refine_Monadic}) to
phrase the specification and the algorithm.
\<close>
text\<open>@{term s} for "searchword" / "searchlist", @{term t} for "text"\<close>
definition "kmp_SPEC s t = SPEC (\<lambda>
None \<Rightarrow> \<nexists>i. sublist_at s t i |
Some i \<Rightarrow> sublist_at s t i \<and> (\<forall>ii<i. \<not>sublist_at s t ii))"
lemma is_arg_min_id: "is_arg_min id P i \<longleftrightarrow> P i \<and> (\<forall>ii<i. \<not>P ii)"
unfolding is_arg_min_def by auto
lemma kmp_result: "kmp_SPEC s t =
RETURN (if sublist s t then Some (LEAST i. sublist_at s t i) else None)"
unfolding kmp_SPEC_def sublist_iff_sublist_at
apply (auto intro: LeastI dest: not_less_Least split: option.splits)
by (meson LeastI nat_neq_iff not_less_Least)
corollary weak_kmp_SPEC: "kmp_SPEC s t \<le> SPEC (\<lambda>pos. pos\<noteq>None \<longleftrightarrow> Sublist.sublist s t)"
by (simp add: kmp_result)
lemmas kmp_SPEC_altdefs =
kmp_SPEC_def[folded is_arg_min_id]
kmp_SPEC_def[folded sublist_iff_sublist_at]
kmp_result
section\<open>Naive algorithm\<close>
text\<open>Since KMP is a direct advancement of the naive "test-all-starting-positions" approach, we provide it here for comparison:\<close>
subsection\<open>Invariants\<close>
definition "I_out_na s t \<equiv> \<lambda>(i,j,pos).
(\<forall>ii<i. \<not>sublist_at s t ii) \<and>
(case pos of None \<Rightarrow> j = 0
| Some p \<Rightarrow> p=i \<and> sublist_at s t i)"
definition "I_in_na s t i \<equiv> \<lambda>(j,pos).
case pos of None \<Rightarrow> j < length s \<and> (\<forall>jj<j. t!(i+jj) = s!(jj))
| Some p \<Rightarrow> sublist_at s t i"
subsection\<open>Algorithm\<close>
(*Algorithm is common knowledge \<longrightarrow> remove citation here, move explanations to KMP below?*)
text\<open>The following definition is taken from Helmut Seidl's lecture on algorithms and data structures@{cite GAD} except that we
\<^item> output the identified position @{term \<open>pos :: nat option\<close>} instead of just @{const True}
\<^item> use @{term \<open>pos :: nat option\<close>} as break-flag to support the abort within the loops
\<^item> rewrite @{prop \<open>i \<le> length t - length s\<close>} in the first while-condition to @{prop \<open>i + length s \<le> length t\<close>} to avoid having to use @{typ int} for list indexes (or the additional precondition @{prop \<open>length s \<le> length t\<close>})
\<close>
definition "naive_algorithm s t \<equiv> do {
let i=0;
let j=0;
let pos=None;
(_,_,pos) \<leftarrow> WHILEIT (I_out_na s t) (\<lambda>(i,_,pos). i + length s \<le> length t \<and> pos=None) (\<lambda>(i,j,pos). do {
(_,pos) \<leftarrow> WHILEIT (I_in_na s t i) (\<lambda>(j,pos). t!(i+j) = s!j \<and> pos=None) (\<lambda>(j,_). do {
let j=j+1;
if j=length s then RETURN (j,Some i) else RETURN (j,None)
}) (j,pos);
if pos=None then do {
let i = i + 1;
let j = 0;
RETURN (i,j,None)
} else RETURN (i,j,Some i)
}) (i,j,pos);
RETURN pos
}"
subsection\<open>Correctness\<close>
text\<open>The basic lemmas on @{const sublist_at} from the previous chapter together with @{theory Refine_Monadic.Refine_Monadic}'s verification condition generator / solver suffice:\<close>
lemma "s \<noteq> [] \<Longrightarrow> naive_algorithm s t \<le> kmp_SPEC s t"
unfolding naive_algorithm_def kmp_SPEC_def I_out_na_def I_in_na_def
apply (refine_vcg
WHILEIT_rule[where R="measure (\<lambda>(i,_,pos). length t - i + (if pos = None then 1 else 0))"]
WHILEIT_rule[where R="measure (\<lambda>(j,_::nat option). length s - j)"]
)
apply (vc_solve solve: asm_rl)
subgoal by (metis add_Suc_right all_positions_sublist less_antisym)
subgoal using less_Suc_eq by blast
subgoal by (metis less_SucE sublist_all_positions)
subgoal by (auto split: option.splits) (metis sublist_lengths add_less_cancel_right leI le_less_trans)
done
text\<open>Note that the precondition cannot be removed without an extra branch: If @{prop \<open>s = []\<close>}, the inner while-condition accesses out-of-bound memory. This will apply to KMP, too.\<close>
section\<open>Knuth--Morris--Pratt algorithm\<close>
text\<open>Just like our templates@{cite KMP77}@{cite GAD}, we first verify the main routine and discuss the computation of the auxiliary values @{term \<open>\<ff> s\<close>} only in a later section.\<close>
subsection\<open>Preliminaries: Borders of lists\<close>
definition "border xs ys \<longleftrightarrow> prefix xs ys \<and> suffix xs ys"
definition "strict_border xs ys \<longleftrightarrow> border xs ys \<and> length xs < length ys"
definition "intrinsic_border ls \<equiv> ARG_MAX length b. strict_border b ls"
subsubsection\<open>Properties\<close>
interpretation border_order: order border strict_border
by standard (auto simp: border_def suffix_def strict_border_def)
interpretation border_bot: order_bot Nil border strict_border
by standard (simp add: border_def)
lemma borderE[elim]:
fixes xs ys :: "'a list"
assumes "border xs ys"
obtains "prefix xs ys" and "suffix xs ys"
using assms unfolding border_def by blast
lemma strict_borderE[elim]:
fixes xs ys :: "'a list"
assumes "strict_border xs ys"
obtains "border xs ys" and "length xs < length ys"
using assms unfolding strict_border_def by blast
lemma strict_border_simps[simp]:
"strict_border xs [] \<longleftrightarrow> False"
"strict_border [] (x # xs) \<longleftrightarrow> True"
by (simp_all add: strict_border_def)
lemma strict_border_prefix: "strict_border xs ys \<Longrightarrow> strict_prefix xs ys"
and strict_border_suffix: "strict_border xs ys \<Longrightarrow> strict_suffix xs ys"
and strict_border_imp_nonempty: "strict_border xs ys \<Longrightarrow> ys \<noteq> []"
and strict_border_prefix_suffix: "strict_border xs ys \<longleftrightarrow> strict_prefix xs ys \<and> strict_suffix xs ys"
by (auto simp: border_order.order.strict_iff_order border_def)
lemma border_length_le: "border xs ys \<Longrightarrow> length xs \<le> length ys"
unfolding border_def by (simp add: prefix_length_le)
lemma border_length_r_less (*rm*): "\<forall>xs. strict_border xs ys \<longrightarrow> length xs < length ys"
using strict_borderE by auto
lemma border_positions: "border xs ys \<Longrightarrow> \<forall>i<length xs. ys!i = ys!(length ys - length xs + i)"
unfolding border_def
by (metis diff_add_inverse diff_add_inverse2 length_append not_add_less1 nth_append prefixE suffixE)
lemma all_positions_drop_length_take: "\<lbrakk>i \<le> length w; i \<le> length x;
\<forall>j<i. x ! j = w ! (length w + j - i)\<rbrakk>
\<Longrightarrow> drop (length w - i) w = take i x"
by (cases "i = length x") (auto intro: nth_equalityI)
lemma all_positions_suffix_take: "\<lbrakk>i \<le> length w; i \<le> length x;
\<forall>j<i. x ! j = w ! (length w + j - i)\<rbrakk>
\<Longrightarrow> suffix (take i x) w"
by (metis all_positions_drop_length_take suffix_drop)
lemma suffix_butlast: "suffix xs ys \<Longrightarrow> suffix (butlast xs) (butlast ys)"
unfolding suffix_def by (metis append_Nil2 butlast.simps(1) butlast_append)
lemma positions_border: "\<forall>j<l. w!j = w!(length w - l + j) \<Longrightarrow> border (take l w) w"
by (cases "l < length w") (simp_all add: border_def all_positions_suffix_take take_is_prefix)
lemma positions_strict_border: "l < length w \<Longrightarrow> \<forall>j<l. w!j = w!(length w - l + j) \<Longrightarrow> strict_border (take l w) w"
by (simp add: positions_border strict_border_def)
lemmas intrinsic_borderI = arg_max_natI[OF _ border_length_r_less, folded intrinsic_border_def]
lemmas intrinsic_borderI' = border_bot.bot.not_eq_extremum[THEN iffD1, THEN intrinsic_borderI]
lemmas intrinsic_border_max = arg_max_nat_le[OF _ border_length_r_less, folded intrinsic_border_def]
lemma nonempty_is_arg_max_ib: "ys \<noteq> [] \<Longrightarrow> is_arg_max length (\<lambda>xs. strict_border xs ys) (intrinsic_border ys)"
by (simp add: intrinsic_borderI' intrinsic_border_max is_arg_max_linorder)
lemma intrinsic_border_less: "w \<noteq> [] \<Longrightarrow> length (intrinsic_border w) < length w"
using intrinsic_borderI[of w] border_length_r_less intrinsic_borderI' by blast
lemma intrinsic_border_take_less: "j > 0 \<Longrightarrow> w \<noteq> [] \<Longrightarrow> length (intrinsic_border (take j w)) < length w"
by (metis intrinsic_border_less length_take less_not_refl2 min_less_iff_conj take_eq_Nil)
subsubsection\<open>Examples\<close>
lemma border_example: "{b. border b ''aabaabaa''} = {'''', ''a'', ''aa'', ''aabaa'', ''aabaabaa''}"
(is "{b. border b ?l} = {?take0, ?take1, ?take2, ?take5, ?l}")
proof
show "{?take0, ?take1, ?take2, ?take5, ?l} \<subseteq> {b. border b ?l}"
by simp eval
have "\<not>border ''aab'' ?l" "\<not>border ''aaba'' ?l" "\<not>border ''aabaab'' ?l" "\<not>border ''aabaaba'' ?l"
by eval+
moreover have "{b. border b ?l} \<subseteq> set (prefixes ?l)"
using border_def in_set_prefixes by blast
ultimately show "{b. border b ?l} \<subseteq> {?take0, ?take1, ?take2, ?take5, ?l}"
by auto
qed
corollary strict_border_example: "{b. strict_border b ''aabaabaa''} = {'''', ''a'', ''aa'', ''aabaa''}"
(is "?l = ?r")
proof
have "?l \<subseteq> {b. border b ''aabaabaa''}"
by auto
also have "\<dots> = {'''', ''a'', ''aa'', ''aabaa'', ''aabaabaa''}"
by (fact border_example)
finally show "?l \<subseteq> ?r" by auto
show "?r \<subseteq> ?l" by simp eval
qed
corollary "intrinsic_border ''aabaabaa'' = ''aabaa''"
proof - \<comment> \<open>We later obtain a fast algorithm for that.\<close>
have exhaust: "strict_border b ''aabaabaa'' \<longleftrightarrow> b \<in> {'''', ''a'', ''aa'', ''aabaa''}" for b
using strict_border_example by auto
then have
"\<not>is_arg_max length (\<lambda>b. strict_border b ''aabaabaa'') ''''"
"\<not>is_arg_max length (\<lambda>b. strict_border b ''aabaabaa'') ''a''"
"\<not>is_arg_max length (\<lambda>b. strict_border b ''aabaabaa'') ''aa''"
"is_arg_max length (\<lambda>b. strict_border b ''aabaabaa'') ''aabaa''"
unfolding is_arg_max_linorder by auto
moreover have "strict_border (intrinsic_border ''aabaabaa'') ''aabaabaa''"
using intrinsic_borderI' by blast
note this[unfolded exhaust]
ultimately show ?thesis
by simp (metis list.discI nonempty_is_arg_max_ib)
qed
subsection\<open>Main routine\<close>
text\<open>The following is Seidl's "border"-table@{cite GAD} (values shifted by 1 so we don't need @{typ int}),
or equivalently, "f" from Knuth's, Morris' and Pratt's paper@{cite KMP77} (with indexes starting at 0).\<close>
fun \<ff> :: "'a list \<Rightarrow> nat \<Rightarrow> nat" where
"\<ff> s 0 = 0" \<comment> \<open>This increments the compare position while @{prop \<open>j=(0::nat)\<close>}\<close> |
"\<ff> s j = length (intrinsic_border (take j s)) + 1"
text\<open>Note that we use their "next" only implicitly.\<close>
subsubsection\<open>Invariants\<close>
definition "I_outer s t \<equiv> \<lambda>(i,j,pos).
(\<forall>ii<i. \<not>sublist_at s t ii) \<and>
(case pos of None \<Rightarrow> (\<forall>jj<j. t!(i+jj) = s!(jj)) \<and> j < length s
| Some p \<Rightarrow> p=i \<and> sublist_at s t i)"
text\<open>For the inner loop, we can reuse @{const I_in_na}.\<close>
subsubsection\<open>Algorithm\<close>
text\<open>First, we use the non-evaluable function @{const \<ff>} directly:\<close>
definition "kmp s t \<equiv> do {
ASSERT (s \<noteq> []);
let i=0;
let j=0;
let pos=None;
(_,_,pos) \<leftarrow> WHILEIT (I_outer s t) (\<lambda>(i,j,pos). i + length s \<le> length t \<and> pos=None) (\<lambda>(i,j,pos). do {
ASSERT (i + length s \<le> length t);
(j,pos) \<leftarrow> WHILEIT (I_in_na s t i) (\<lambda>(j,pos). t!(i+j) = s!j \<and> pos=None) (\<lambda>(j,pos). do {
let j=j+1;
if j=length s then RETURN (j,Some i) else RETURN (j,None)
}) (j,pos);
if pos=None then do {
ASSERT (j < length s);
let i = i + (j - \<ff> s j + 1);
let j = max 0 (\<ff> s j - 1); \<comment> \<open>\<open>max\<close> not necessary\<close>
RETURN (i,j,None)
} else RETURN (i,j,Some i)
}) (i,j,pos);
RETURN pos
}"
subsubsection\<open>Correctness\<close>
lemma \<ff>_eq_0_iff_j_eq_0[simp]: "\<ff> s j = 0 \<longleftrightarrow> j = 0"
by (cases j) simp_all
lemma j_le_\<ff>_le: "j \<le> length s \<Longrightarrow> \<ff> s j \<le> j"
apply (cases j)
apply simp_all
by (metis Suc_leI intrinsic_border_less length_take list.size(3) min.absorb2 nat.simps(3) not_less)
lemma j_le_\<ff>_le': "0 < j \<Longrightarrow> j \<le> length s \<Longrightarrow> \<ff> s j - 1 < j"
by (metis diff_less j_le_\<ff>_le le_eq_less_or_eq less_imp_diff_less less_one)
lemma \<ff>_le: "s \<noteq> [] \<Longrightarrow> \<ff> s j - 1 < length s"
by (cases j) (simp_all add: intrinsic_border_take_less)
(*
Only needed for run-time analysis
lemma "p576 et seq":
assumes
"j \<le> length s" and
assignments:
"i' = i + (j + 1 - \<ff> s j)"
"j' = max 0 (\<ff> s j - 1)"
shows
sum_no_decrease: "i' + j' \<ge> i + j" and
i_increase: "i' > i"
using assignments by (simp_all add: j_le_\<ff>_le[OF assms(1), THEN le_imp_less_Suc])
*)
lemma reuse_matches:
assumes j_le: "j \<le> length s"
and old_matches: "\<forall>jj<j. t ! (i + jj) = s ! jj"
shows "\<forall>jj<\<ff> s j - 1. t ! (i + (j - \<ff> s j + 1) + jj) = s ! jj"
(is "\<forall>jj<?j'. t ! (?i' + jj) = s ! jj")
proof (cases "j>0")
assume "j>0"
have \<ff>_le: "\<ff> s j \<le> j"
by (simp add: j_le j_le_\<ff>_le)
with old_matches have 1: "\<forall>jj<?j'. t ! (?i' + jj) = s ! (j - \<ff> s j + 1 + jj)"
by (metis ab_semigroup_add_class.add.commute add.assoc diff_diff_cancel less_diff_conv)
have [simp]: "length (take j s) = j" "length (intrinsic_border (take j s)) = ?j'"
by (simp add: j_le) (metis \<open>0 < j\<close> diff_add_inverse2 \<ff>.elims nat_neq_iff)
then have "\<forall>jj<?j'. take j s ! jj = take j s ! (j - (\<ff> s j - 1) + jj)"
by (metis intrinsic_borderI' \<open>0 < j\<close> border_positions length_greater_0_conv strict_border_def)
then have "\<forall>jj<?j'. take j s ! jj = take j s ! (j - \<ff> s j + 1 + jj)"
by (simp add: \<ff>_le)
then have 2: "\<forall>jj<?j'. s ! (j - \<ff> s j + 1 + jj) = s ! jj"
using \<ff>_le by simp
from 1 2 show ?thesis by simp
qed simp
theorem shift_safe:
assumes
"\<forall>ii<i. \<not>sublist_at s t ii"
"t!(i+j) \<noteq> s!j" and
[simp]: "j < length s" and
matches: "\<forall>jj<j. t!(i+jj) = s!jj"
defines
assignment: "i' \<equiv> i + (j - \<ff> s j + 1)"
shows
"\<forall>ii<i'. \<not>sublist_at s t ii"
proof (standard, standard)
fix ii
assume "ii < i'"
then consider \<comment> \<open>The position falls into one of three categories:\<close>
(old) "ii < i" |
(current) "ii = i" |
(skipped) "ii > i"
by linarith
then show "\<not>sublist_at s t ii"
proof cases
case old \<comment> \<open>Old position, use invariant.\<close>
with \<open>\<forall>ii<i. \<not>sublist_at s t ii\<close> show ?thesis by simp
next
case current \<comment> \<open>The mismatch occurred while testing this alignment.\<close>
with \<open>t!(i+j) \<noteq> s!j\<close> show ?thesis
using sublist_all_positions[of s t i] by auto
next
case skipped \<comment> \<open>The skipped positions.\<close>
then have "0<j"
using \<open>ii < i'\<close> assignment by linarith
then have less_j[simp]: "j + i - ii < j" and le_s: "j + i - ii \<le> length s"
using \<open>ii < i'\<close> assms(3) skipped by linarith+
note \<ff>_le[simp] = j_le_\<ff>_le[OF assms(3)[THEN less_imp_le]]
have "0 < \<ff> s j"
using \<open>0 < j\<close> \<ff>_eq_0_iff_j_eq_0 neq0_conv by blast
then have "j + i - ii > \<ff> s j - 1"
using \<open>ii < i'\<close> assignment \<ff>_le by linarith
then have contradiction_goal: "j + i - ii > length (intrinsic_border (take j s))"
by (metis \<ff>.elims \<open>0 < j\<close> add_diff_cancel_right' not_gr_zero)
show ?thesis
proof
assume "sublist_at s t ii"
note sublist_all_positions[OF this]
with le_s have a: "\<forall>jj < j+i-ii. t!(ii+jj) = s!jj"
by simp
have ff1: "\<not> ii < i"
by (metis not_less_iff_gr_or_eq skipped)
then have "i + (ii - i + jj) = ii + jj" for jj
by (metis add.assoc add_diff_inverse_nat)
then have "\<not> jj < j + i - ii \<or> t ! (ii + jj) = s ! (ii - i + jj)" if "ii - i + jj < j" for jj
using that ff1 by (metis matches)
then have "\<not> jj < j + i - ii \<or> t ! (ii + jj) = s ! (ii - i + jj)" for jj
using ff1 by auto
with matches have "\<forall>jj < j+i-ii. t!(ii+jj) = s!(ii-i+jj)" by metis
then have "\<forall>jj < j+i-ii. s!jj = s!(ii-i+jj)"
using a by auto
then have "\<forall>jj < j+i-ii. (take j s)!jj = (take j s)!(ii-i+jj)"
using \<open>i<ii\<close> by auto
with positions_strict_border[of "j+i-ii" "take j s", simplified]
have "strict_border (take (j+i-ii) s) (take j s)".
note intrinsic_border_max[OF this]
also note contradiction_goal
also have "j+i-ii \<le> length s" by (fact le_s)
ultimately
show False by simp
qed
qed
qed
lemma kmp_correct: "s \<noteq> []
\<Longrightarrow> kmp s t \<le> kmp_SPEC s t"
unfolding kmp_def kmp_SPEC_def I_outer_def I_in_na_def
apply (refine_vcg
WHILEIT_rule[where R="measure (\<lambda>(i,_,pos). length t - i + (if pos = None then 1 else 0))"]
WHILEIT_rule[where R="measure (\<lambda>(j,_::nat option). length s - j)"]
)
apply (vc_solve solve: asm_rl)
subgoal by (metis add_Suc_right all_positions_sublist less_antisym)
subgoal using less_antisym by blast
subgoal for i jout j using shift_safe[of i s t j] by fastforce
subgoal for i jout j using reuse_matches[of j s t i] \<ff>_le by simp
subgoal by (auto split: option.splits) (metis sublist_lengths add_less_cancel_right leI le_less_trans)
done
subsubsection\<open>Storing the @{const \<ff>}-values\<close>
text\<open>We refine the algorithm to compute the @{const \<ff>}-values only once at the start:\<close>
definition compute_\<ff>s_SPEC :: "'a list \<Rightarrow> nat list nres" where
"compute_\<ff>s_SPEC s \<equiv> SPEC (\<lambda>\<ff>s. length \<ff>s = length s + 1 \<and> (\<forall>j\<le>length s. \<ff>s!j = \<ff> s j))"
definition "kmp1 s t \<equiv> do {
ASSERT (s \<noteq> []);
let i=0;
let j=0;
let pos=None;
\<ff>s \<leftarrow> compute_\<ff>s_SPEC (butlast s); \<comment> \<open>At the last char, we abort instead.\<close>
(_,_,pos) \<leftarrow> WHILEIT (I_outer s t) (\<lambda>(i,j,pos). i + length s \<le> length t \<and> pos=None) (\<lambda>(i,j,pos). do {
ASSERT (i + length s \<le> length t);
(j,pos) \<leftarrow> WHILEIT (I_in_na s t i) (\<lambda>(j,pos). t!(i+j) = s!j \<and> pos=None) (\<lambda>(j,pos). do {
let j=j+1;
if j=length s then RETURN (j,Some i) else RETURN (j,None)
}) (j,pos);
if pos=None then do {
ASSERT (j < length \<ff>s);
let i = i + (j - \<ff>s!j + 1);
let j = max 0 (\<ff>s!j - 1); \<comment> \<open>\<open>max\<close> not necessary\<close>
RETURN (i,j,None)
} else RETURN (i,j,Some i)
}) (i,j,pos);
RETURN pos
}"
lemma \<ff>_butlast[simp]: "j < length s \<Longrightarrow> \<ff> (butlast s) j = \<ff> s j"
by (cases j) (simp_all add: take_butlast)
lemma kmp1_refine: "kmp1 s t \<le> kmp s t"
apply (rule refine_IdD)
unfolding kmp1_def kmp_def Let_def compute_\<ff>s_SPEC_def nres_monad_laws
apply (intro ASSERT_refine_right ASSERT_refine_left)
apply simp
apply (rule Refine_Basic.intro_spec_refine)
apply refine_rcg
apply refine_dref_type
apply vc_solve
done
text\<open>Next, an algorithm that satisfies @{const compute_\<ff>s_SPEC}:\<close>
subsection\<open>Computing @{const \<ff>}\<close>
subsubsection\<open>Invariants\<close>
definition "I_out_cb s \<equiv> \<lambda>(\<ff>s,i,j).
length s + 1 = length \<ff>s \<and>
(\<forall>jj<j. \<ff>s!jj = \<ff> s jj) \<and>
\<ff>s!(j-1) = i \<and>
0 < j"
definition "I_in_cb s j \<equiv> \<lambda>i.
if j=1 then i=0 \<comment> \<open>first iteration\<close>
else
strict_border (take (i-1) s) (take (j-1) s) \<and>
\<ff> s j \<le> i + 1"
subsubsection\<open>Algorithm\<close>
text\<open>Again, we follow Seidl@{cite GAD}, p.582. Apart from the +1-shift, we make another modification:
Instead of directly setting @{term \<open>\<ff>s!1\<close>}, we let the first loop-iteration (if there is one) do that for us.
This allows us to remove the precondition @{prop \<open>s \<noteq> []\<close>}, as the index bounds are respected even in that corner case.\<close>
definition compute_\<ff>s :: "'a list \<Rightarrow> nat list nres" where
"compute_\<ff>s s = do {
let \<ff>s=replicate (length s + 1) 0; \<comment> \<open>only the first 0 is needed\<close>
let i=0;
let j=1;
(\<ff>s,_,_) \<leftarrow> WHILEIT (I_out_cb s) (\<lambda>(\<ff>s,_,j). j < length \<ff>s) (\<lambda>(\<ff>s,i,j). do {
i \<leftarrow> WHILEIT (I_in_cb s j) (\<lambda>i. i>0 \<and> s!(i-1) \<noteq> s!(j-1)) (\<lambda>i. do {
ASSERT (i-1 < length \<ff>s);
let i=\<ff>s!(i-1);
RETURN i
}) i;
let i=i+1;
ASSERT (j < length \<ff>s);
let \<ff>s=\<ff>s[j:=i];
let j=j+1;
RETURN (\<ff>s,i,j)
}) (\<ff>s,i,j);
RETURN \<ff>s
}"
subsubsection\<open>Correctness\<close>
lemma take_length_ib[simp]:
assumes "0 < j" "j \<le> length s"
shows "take (length (intrinsic_border (take j s))) s = intrinsic_border (take j s)"
proof -
from assms have "prefix (intrinsic_border (take j s)) (take j s)"
by (metis intrinsic_borderI' border_def list.size(3) neq0_conv not_less strict_border_def take_eq_Nil)
also have "prefix (take j s) s"
by (simp add: \<open>j \<le> length s\<close> take_is_prefix)
finally show ?thesis
by (metis append_eq_conv_conj prefixE)
qed
lemma ib_singleton[simp]: "intrinsic_border [z] = []"
by (metis intrinsic_border_less length_Cons length_greater_0_conv less_Suc0 list.size(3))
lemma border_butlast: "border xs ys \<Longrightarrow> border (butlast xs) (butlast ys)"
apply (auto simp: border_def)
apply (metis butlast_append prefixE prefix_order.eq_refl prefix_prefix prefixeq_butlast)
apply (metis Sublist.suffix_def append.right_neutral butlast.simps(1) butlast_append)
done
corollary strict_border_butlast: "xs \<noteq> [] \<Longrightarrow> strict_border xs ys \<Longrightarrow> strict_border (butlast xs) (butlast ys)"
unfolding strict_border_def by (simp add: border_butlast less_diff_conv)
lemma border_take_lengths: "i \<le> length s \<Longrightarrow> border (take i s) (take j s) \<Longrightarrow> i \<le> j"
using border_length_le by fastforce
lemma border_step: "border xs ys \<longleftrightarrow> border (xs@[ys!length xs]) (ys@[ys!length xs])"
apply (auto simp: border_def suffix_def)
using append_one_prefix prefixE apply fastforce
using append_prefixD apply blast
done
corollary strict_border_step: "strict_border xs ys \<longleftrightarrow> strict_border (xs@[ys!length xs]) (ys@[ys!length xs])"
unfolding strict_border_def using border_step by auto
lemma ib_butlast: "length w \<ge> 2 \<Longrightarrow> length (intrinsic_border w) \<le> length (intrinsic_border (butlast w)) + 1"
proof -
assume "length w \<ge> 2"
then have "w \<noteq> []" by auto
then have "strict_border (intrinsic_border w) w"
by (fact intrinsic_borderI')
with \<open>2 \<le> length w\<close> have "strict_border (butlast (intrinsic_border w)) (butlast w)"
by (metis One_nat_def border_bot.bot.not_eq_extremum butlast.simps(1) len_greater_imp_nonempty length_butlast lessI less_le_trans numerals(2) strict_border_butlast zero_less_diff)
then have "length (butlast (intrinsic_border w)) \<le> length (intrinsic_border (butlast w))"
using intrinsic_border_max by blast
then show ?thesis
by simp
qed
corollary \<ff>_Suc(*rm*): "Suc i \<le> length w \<Longrightarrow> \<ff> w (Suc i) \<le> \<ff> w i + 1"
apply (cases i)
apply (simp_all add: take_Suc0)
by (metis One_nat_def Suc_eq_plus1 Suc_to_right butlast_take diff_is_0_eq ib_butlast length_take min.absorb2 nat.simps(3) not_less_eq_eq numerals(2))
lemma \<ff>_step_bound(*rm*):
assumes "j \<le> length w"
shows "\<ff> w j \<le> \<ff> w (j-1) + 1"
using assms[THEN j_le_\<ff>_le] \<ff>_Suc assms
by (metis One_nat_def Suc_pred le_SucI not_gr_zero trans_le_add2)
lemma border_take_\<ff>: "border (take (\<ff> s i - 1) s ) (take i s)"
apply (cases i, simp_all)
- by (metis intrinsic_borderI' border_order.eq_iff border_order.less_imp_le border_positions nat.simps(3) nat_le_linear positions_border take_all take_eq_Nil take_length_ib zero_less_Suc)
+ by (metis intrinsic_borderI' border_order.order.eq_iff border_order.less_imp_le border_positions nat.simps(3) nat_le_linear positions_border take_all take_eq_Nil take_length_ib zero_less_Suc)
corollary \<ff>_strict_borderI: "y = \<ff> s (i-1) \<Longrightarrow> strict_border (take (i-1) s) (take (j-1) s) \<Longrightarrow> strict_border (take (y-1) s) (take (j-1) s)"
using border_order.less_le_not_le border_order.order.trans border_take_\<ff> by blast
corollary strict_border_take_\<ff>: "0 < i \<Longrightarrow> i \<le> length s \<Longrightarrow> strict_border (take (\<ff> s i - 1) s ) (take i s)"
by (meson border_order.less_le_not_le border_take_\<ff> border_take_lengths j_le_\<ff>_le' leD)
lemma \<ff>_is_max: "j \<le> length s \<Longrightarrow> strict_border b (take j s) \<Longrightarrow> \<ff> s j \<ge> length b + 1"
by (metis \<ff>.elims add_le_cancel_right add_less_same_cancel2 border_length_r_less intrinsic_border_max length_take min_absorb2 not_add_less2)
theorem skipping_ok:
assumes j_bounds[simp]: "1 < j" "j \<le> length s"
and mismatch: "s!(i-1) \<noteq> s!(j-1)"
and greater_checked: "\<ff> s j \<le> i + 1"
and "strict_border (take (i-1) s) (take (j-1) s)"
shows "\<ff> s j \<le> \<ff> s (i-1) + 1"
proof (rule ccontr)
assume "\<not>\<ff> s j \<le> \<ff> s (i-1) + 1"
then have i_bounds: "0 < i" "i \<le> length s"
using greater_checked assms(5) take_Nil by fastforce+
then have i_less_j: "i < j"
using assms(5) border_length_r_less nz_le_conv_less by auto
from \<open>\<not>\<ff> s j \<le> \<ff> s (i-1) + 1\<close> greater_checked consider
(tested) "\<ff> s j = i + 1" \<comment> \<open>This contradicts @{thm mismatch}\<close> |
(skipped) "\<ff> s (i-1) + 1 < \<ff> s j" "\<ff> s j \<le> i"
\<comment> \<open>This contradicts @{thm \<ff>_is_max[of "i-1" s]}\<close>
by linarith
then show False
proof cases
case tested
then have "\<ff> s j - 1 = i" by simp
moreover note border_positions[OF border_take_\<ff>[of s j, unfolded this]]
ultimately have "take j s ! (i-1) = s!(j-1)" using i_bounds i_less_j by simp
with \<open>i < j\<close> have "s!(i-1) = s!(j-1)"
by (simp add: less_imp_diff_less)
with mismatch show False..
next
case skipped
let ?border = "take (i-1) s"
\<comment> \<open>This border of @{term \<open>take (j-1) s\<close>} could not be extended to a border of @{term \<open>take j s\<close>} due to the mismatch.\<close>
let ?impossible = "take (\<ff> s j - 2) s"
\<comment> \<open>A strict border longer than @{term \<open>intrinsic_border ?border\<close>}, a contradiction.\<close>
have "length (take j s) = j"
by simp
have "\<ff> s j - 2 < i - 1"
using skipped by linarith
then have less_s: "\<ff> s j - 2 < length s" "i - 1 < length s"
using \<open>i < j\<close> j_bounds(2) by linarith+
then have strict: "length ?impossible < length ?border"
using \<open>\<ff> s j - 2 < i - 1\<close> by auto
moreover {
have "prefix ?impossible (take j s)"
using prefix_length_prefix take_is_prefix
by (metis (no_types, lifting) \<open>length (take j s) = j\<close> j_bounds(2) diff_le_self j_le_\<ff>_le length_take less_s(1) min_simps(2) order_trans)
moreover have "prefix ?border (take j s)"
by (metis (no_types, lifting) \<open>length (take j s) = j\<close> diff_le_self i_less_j le_trans length_take less_or_eq_imp_le less_s(2) min_simps(2) prefix_length_prefix take_is_prefix)
ultimately have "prefix ?impossible ?border"
using strict less_imp_le_nat prefix_length_prefix by blast
} moreover {
have "suffix (take (\<ff> s j - 1) s) (take j s)" using border_take_\<ff>
by (auto simp: border_def)
note suffix_butlast[OF this]
then have "suffix ?impossible (take (j-1) s)"
by (metis One_nat_def j_bounds(2) butlast_take diff_diff_left \<ff>_le len_greater_imp_nonempty less_or_eq_imp_le less_s(2) one_add_one)
then have "suffix ?impossible (take (j-1) s)" "suffix ?border (take (j-1) s)"
using assms(5) by auto
from suffix_length_suffix[OF this strict[THEN less_imp_le]]
have "suffix ?impossible ?border".
}
ultimately have "strict_border ?impossible ?border"
unfolding strict_border_def[unfolded border_def] by blast
note \<ff>_is_max[of "i-1" s, OF _ this]
then have "length (take (\<ff> s j - 2) s) + 1 \<le> \<ff> s (i-1)"
using less_imp_le_nat less_s(2) by blast
then have "\<ff> s j - 1 \<le> \<ff> s (i-1)"
by (simp add: less_s(1))
then have "\<ff> s j \<le> \<ff> s (i-1) + 1"
using le_diff_conv by blast
with skipped(1) show False
by linarith
qed
qed
lemma extend_border:
assumes "j \<le> length s"
assumes "s!(i-1) = s!(j-1)"
assumes "strict_border (take (i-1) s) (take (j-1) s)"
assumes "\<ff> s j \<le> i + 1"
shows "\<ff> s j = i + 1"
proof -
from assms(3) have pos_in_range: "i - 1 < length s " "length (take (i-1) s) = i - 1"
using border_length_r_less min_less_iff_conj by auto
with strict_border_step[THEN iffD1, OF assms(3)] have "strict_border (take (i-1) s @ [s!(i-1)]) (take (j-1) s @ [s!(i-1)])"
by (metis assms(3) border_length_r_less length_take min_less_iff_conj nth_take)
with pos_in_range have "strict_border (take i s) (take (j-1) s @ [s!(i-1)])"
by (metis Suc_eq_plus1 Suc_pred add.left_neutral border_bot.bot.not_eq_extremum border_order.less_asym neq0_conv take_0 take_Suc_conv_app_nth)
then have "strict_border (take i s) (take (j-1) s @ [s!(j-1)])"
by (simp only: \<open>s!(i-1) = s!(j-1)\<close>)
then have "strict_border (take i s) (take j s)"
by (metis One_nat_def Suc_pred assms(1,3) diff_le_self less_le_trans neq0_conv nz_le_conv_less strict_border_imp_nonempty take_Suc_conv_app_nth take_eq_Nil)
with \<ff>_is_max[OF assms(1) this] have "\<ff> s j \<ge> i + 1"
using Suc_leI by fastforce
with \<open>\<ff> s j \<le> i + 1\<close> show ?thesis
using le_antisym by presburger
qed
lemma compute_\<ff>s_correct: "compute_\<ff>s s \<le> compute_\<ff>s_SPEC s"
unfolding compute_\<ff>s_SPEC_def compute_\<ff>s_def I_out_cb_def I_in_cb_def
apply (simp, refine_vcg
WHILEIT_rule[where R="measure (\<lambda>(\<ff>s,i,j). length s + 1 - j)"]
WHILEIT_rule[where R="measure id"] \<comment> \<open>@{term \<open>i::nat\<close>} decreases with every iteration.\<close>
)
apply (vc_solve, fold One_nat_def)
subgoal for b j by (rule strict_border_take_\<ff>, auto)
subgoal by (metis Suc_eq_plus1 \<ff>_step_bound less_Suc_eq_le)
subgoal by fastforce
subgoal
by (metis (no_types, lifting) One_nat_def Suc_lessD Suc_pred border_length_r_less \<ff>_strict_borderI length_take less_Suc_eq less_Suc_eq_le min.absorb2)
subgoal for b j i
by (metis (no_types, lifting) One_nat_def Suc_diff_1 Suc_eq_plus1 Suc_leI border_take_lengths less_Suc_eq_le less_antisym skipping_ok strict_border_def)
subgoal by (metis Suc_diff_1 border_take_lengths j_le_\<ff>_le less_Suc_eq_le strict_border_def)
subgoal for b j i jj
by (metis Suc_eq_plus1 Suc_eq_plus1_left add.right_neutral extend_border \<ff>_eq_0_iff_j_eq_0 j_le_\<ff>_le le_zero_eq less_Suc_eq less_Suc_eq_le nth_list_update_eq nth_list_update_neq)
subgoal by linarith
done
subsubsection\<open>Index shift\<close>
text\<open>To avoid inefficiencies, we refine @{const compute_\<ff>s} to take @{term s}
instead of @{term \<open>butlast s\<close>} (it still only uses @{term \<open>butlast s\<close>}).\<close>
definition compute_butlast_\<ff>s :: "'a list \<Rightarrow> nat list nres" where
"compute_butlast_\<ff>s s = do {
let \<ff>s=replicate (length s) 0;
let i=0;
let j=1;
(\<ff>s,_,_) \<leftarrow> WHILEIT (I_out_cb (butlast s)) (\<lambda>(b,i,j). j < length b) (\<lambda>(\<ff>s,i,j). do {
ASSERT (j < length \<ff>s);
i \<leftarrow> WHILEIT (I_in_cb (butlast s) j) (\<lambda>i. i>0 \<and> s!(i-1) \<noteq> s!(j-1)) (\<lambda>i. do {
ASSERT (i-1 < length \<ff>s);
let i=\<ff>s!(i-1);
RETURN i
}) i;
let i=i+1;
ASSERT (j < length \<ff>s);
let \<ff>s=\<ff>s[j:=i];
let j=j+1;
RETURN (\<ff>s,i,j)
}) (\<ff>s,i,j);
RETURN \<ff>s
}"
lemma compute_\<ff>s_inner_bounds:
assumes "I_out_cb s (\<ff>s,ix,j)"
assumes "j < length \<ff>s"
assumes "I_in_cb s j i"
shows "i-1 < length s" "j-1 < length s"
using assms
by (auto simp: I_out_cb_def I_in_cb_def split: if_splits)
lemma compute_butlast_\<ff>s_refine[refine]:
assumes "(s,s') \<in> br butlast ((\<noteq>) [])"
shows "compute_butlast_\<ff>s s \<le> \<Down> Id (compute_\<ff>s_SPEC s')"
proof -
have "compute_butlast_\<ff>s s \<le> \<Down> Id (compute_\<ff>s s')"
unfolding compute_butlast_\<ff>s_def compute_\<ff>s_def
apply (refine_rcg)
apply (refine_dref_type)
using assms apply (vc_solve simp: in_br_conv)
apply (metis Suc_pred length_greater_0_conv replicate_Suc)
by (metis One_nat_def compute_\<ff>s_inner_bounds nth_butlast)
also note compute_\<ff>s_correct
finally show ?thesis by simp
qed
subsection\<open>Conflation\<close>
text\<open>We replace @{const compute_\<ff>s_SPEC} with @{const compute_butlast_\<ff>s}\<close>
definition "kmp2 s t \<equiv> do {
ASSERT (s \<noteq> []);
let i=0;
let j=0;
let pos=None;
\<ff>s \<leftarrow> compute_butlast_\<ff>s s;
(_,_,pos) \<leftarrow> WHILEIT (I_outer s t) (\<lambda>(i,j,pos). i + length s \<le> length t \<and> pos=None) (\<lambda>(i,j,pos). do {
ASSERT (i + length s \<le> length t \<and> pos=None);
(j,pos) \<leftarrow> WHILEIT (I_in_na s t i) (\<lambda>(j,pos). t!(i+j) = s!j \<and> pos=None) (\<lambda>(j,pos). do {
let j=j+1;
if j=length s then RETURN (j,Some i) else RETURN (j,None)
}) (j,pos);
if pos=None then do {
ASSERT (j < length \<ff>s);
let i = i + (j - \<ff>s!j + 1);
let j = max 0 (\<ff>s!j - 1); \<comment> \<open>\<open>max\<close> not necessary\<close>
RETURN (i,j,None)
} else RETURN (i,j,Some i)
}) (i,j,pos);
RETURN pos
}"
text\<open>Using @{thm [source] compute_butlast_\<ff>s_refine} (it has attribute @{attribute refine}), the proof is trivial:\<close>
lemma kmp2_refine: "kmp2 s t \<le> kmp1 s t"
apply (rule refine_IdD)
unfolding kmp2_def kmp1_def
apply refine_rcg
apply refine_dref_type
apply (vc_solve simp: in_br_conv)
done
lemma kmp2_correct: "s \<noteq> []
\<Longrightarrow> kmp2 s t \<le> kmp_SPEC s t"
proof -
assume "s \<noteq> []"
have "kmp2 s t \<le> kmp1 s t" by (fact kmp2_refine)
also have "... \<le> kmp s t" by (fact kmp1_refine)
also have "... \<le> kmp_SPEC s t" by (fact kmp_correct[OF \<open>s \<noteq> []\<close>])
finally show ?thesis.
qed
text\<open>For convenience, we also remove the precondition:\<close>
definition "kmp3 s t \<equiv> do {
if s=[] then RETURN (Some 0) else kmp2 s t
}"
lemma kmp3_correct: "kmp3 s t \<le> kmp_SPEC s t"
unfolding kmp3_def by (simp add: kmp2_correct) (simp add: kmp_SPEC_def)
section \<open>Refinement to Imperative/HOL\<close>
lemma eq_id_param: "((=), (=)) \<in> Id \<rightarrow> Id \<rightarrow> Id" by simp
lemmas in_bounds_aux = compute_\<ff>s_inner_bounds[of "butlast s" for s, simplified]
sepref_definition compute_butlast_\<ff>s_impl is compute_butlast_\<ff>s :: "(arl_assn id_assn)\<^sup>k \<rightarrow>\<^sub>a array_assn nat_assn"
unfolding compute_butlast_\<ff>s_def
supply in_bounds_aux[dest]
supply eq_id_param[where 'a='a, sepref_import_param]
apply (rewrite array_fold_custom_replicate)
by sepref
declare compute_butlast_\<ff>s_impl.refine[sepref_fr_rules]
sepref_register compute_\<ff>s
lemma kmp_inner_in_bound:
assumes "i + length s \<le> length t"
assumes "I_in_na s t i (j,None)"
shows "i + j < length t" "j < length s"
using assms
by (auto simp: I_in_na_def)
sepref_definition kmp_impl is "uncurry kmp3" :: "(arl_assn id_assn)\<^sup>k *\<^sub>a (arl_assn id_assn)\<^sup>k \<rightarrow>\<^sub>a option_assn nat_assn"
unfolding kmp3_def kmp2_def
apply (simp only: max_0L) \<comment> \<open>Avoid the unneeded @{const max}\<close>
apply (rewrite in "WHILEIT (I_in_na _ _ _) \<hole>" conj_commute)
apply (rewrite in "WHILEIT (I_in_na _ _ _) \<hole>" short_circuit_conv)
supply kmp_inner_in_bound[dest]
supply option.splits[split]
supply eq_id_param[where 'a='a, sepref_import_param]
by sepref
export_code kmp_impl in SML_imp module_name KMP
lemma kmp3_correct':
"(uncurry kmp3, uncurry kmp_SPEC) \<in> Id \<times>\<^sub>r Id \<rightarrow>\<^sub>f \<langle>Id\<rangle>nres_rel"
apply (intro frefI nres_relI; clarsimp)
apply (fact kmp3_correct)
done
lemmas kmp_impl_correct' = kmp_impl.refine[FCOMP kmp3_correct']
subsection \<open>Overall Correctness Theorem\<close>
text \<open>The following theorem relates the final Imperative HOL algorithm to its specification,
using, beyond basic HOL concepts
\<^item> Hoare triples for Imperative/HOL, provided by the Separation Logic Framework for Imperative/HOL (Theory @{theory Separation_Logic_Imperative_HOL.Sep_Main});
\<^item> The assertion @{const arl_assn} to specify array-lists, which we use to represent the input strings of the algorithm;
\<^item> The @{const sublist_at} function that we defined in section \ref{sec:spec}.
\<close>
theorem kmp_impl_correct:
"< arl_assn id_assn s si * arl_assn id_assn t ti >
kmp_impl si ti
<\<lambda>r. arl_assn id_assn s si * arl_assn id_assn t ti * \<up>(
case r of None \<Rightarrow> \<nexists>i. sublist_at s t i
| Some i \<Rightarrow> sublist_at s t i \<and> (\<forall>ii<i. \<not> sublist_at s t ii)
)>\<^sub>t"
by (sep_auto
simp: pure_def kmp_SPEC_def
split: option.split
heap: kmp_impl_correct'[THEN hfrefD, THEN hn_refineD, of "(s,t)" "(si,ti)", simplified])
definition "kmp_string_impl \<equiv> kmp_impl :: (char array \<times> nat) \<Rightarrow> _"
section \<open>Tests of Generated ML-Code\<close>
ML_val \<open>
fun str2arl s = (Array.fromList (@{code String.explode} s), @{code nat_of_integer} (String.size s))
fun kmp s t = map_option @{code integer_of_nat} (@{code kmp_string_impl} (str2arl s) (str2arl t) ())
val test1 = kmp "anas" "bananas"
val test2 = kmp "" "bananas"
val test3 = kmp "hide_fact" (File.read @{file \<open>~~/src/HOL/Main.thy\<close>})
val test4 = kmp "sorry" (File.read @{file \<open>~~/src/HOL/HOL.thy\<close>})
\<close>
end
diff --git a/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy b/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy
--- a/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy
+++ b/thys/Lambda_Free_KBOs/Lambda_Free_KBO_Std.thy
@@ -1,1269 +1,1269 @@
(* Title: The Graceful Standard Knuth-Bendix Order for Lambda-Free Higher-Order Terms
Author: Heiko Becker <heikobecker92@gmail.com>, 2016
Author: Jasmin Blanchette <jasmin.blanchette at inria.fr>, 2016
Author: Uwe Waldmann <waldmann at mpi-inf.mpg.de>, 2016
Author: Daniel Wand <dwand at mpi-inf.mpg.de>, 2016
Maintainer: Jasmin Blanchette <jasmin.blanchette at inria.fr>
*)
section \<open>The Graceful Standard Knuth--Bendix Order for Lambda-Free Higher-Order Terms\<close>
theory Lambda_Free_KBO_Std
imports Lambda_Free_KBO_Util Nested_Multisets_Ordinals.Multiset_More
abbrevs ">t" = ">\<^sub>t"
and "\<ge>t" = "\<ge>\<^sub>t"
begin
text \<open>
This theory defines the standard version of the graceful Knuth--Bendix order for \<open>\<lambda>\<close>-free
higher-order terms. Standard means that one symbol is allowed to have a weight of 0.
\<close>
subsection \<open>Setup\<close>
locale kbo_std = kbo_std_basis _ _ arity_sym arity_var wt_sym
for
arity_sym :: "'s \<Rightarrow> enat" and
arity_var :: "'v \<Rightarrow> enat" and
wt_sym :: "'s \<Rightarrow> nat"
begin
subsection \<open>Weights\<close>
primrec wt :: "('s, 'v) tm \<Rightarrow> nat" where
"wt (Hd \<zeta>) = (LEAST w. \<exists>f \<in> ground_heads \<zeta>. w = wt_sym f + the_enat (\<delta> * arity_sym f))"
| "wt (App s t) = (wt s - \<delta>) + wt t"
lemma wt_Hd_Sym: "wt (Hd (Sym f)) = wt_sym f + the_enat (\<delta> * arity_sym f)"
by simp
lemma exists_wt_sym: "\<exists>f \<in> ground_heads \<zeta>. wt (Hd \<zeta>) = wt_sym f + the_enat (\<delta> * arity_sym f)"
by (auto intro: Least_in_nonempty_set_imp_ex)
lemma wt_le_wt_sym: "f \<in> ground_heads \<zeta> \<Longrightarrow> wt (Hd \<zeta>) \<le> wt_sym f + the_enat (\<delta> * arity_sym f)"
using not_le_imp_less not_less_Least by fastforce
lemma enat_the_enat_\<delta>_times_arity_sym[simp]: "enat (the_enat (\<delta> * arity_sym f)) = \<delta> * arity_sym f"
using arity_sym_ne_infinity_if_\<delta>_gt_0 imult_is_infinity zero_enat_def by fastforce
lemma wt_arg_le: "wt (arg s) \<le> wt s"
by (cases s) auto
lemma wt_ge_\<epsilon>: "wt s \<ge> \<epsilon>"
by (induct s, metis exists_wt_sym of_nat_eq_enat le_diff_conv of_nat_id wt_sym_ge,
simp add: add_increasing)
lemma wt_ge_\<delta>: "wt s \<ge> \<delta>"
by (meson \<delta>_le_\<epsilon> order.trans enat_ord_simps(1) wt_ge_\<epsilon>)
lemma wt_gt_\<delta>_if_superunary: "arity_hd (head s) > 1 \<Longrightarrow> wt s > \<delta>"
proof (induct s)
case \<zeta>: (Hd \<zeta>)
obtain g where
g_in_grs: "g \<in> ground_heads \<zeta>" and
wt_\<zeta>: "wt (Hd \<zeta>) = wt_sym g + the_enat (\<delta> * arity_sym g)"
using exists_wt_sym by blast
have "arity_hd \<zeta> > 1"
using \<zeta> by auto
hence ary_g: "arity_sym g > 1"
using ground_heads_arity[OF g_in_grs] by simp
show ?case
proof (cases "\<delta> = 0")
case True
thus ?thesis
by (metis \<epsilon>_gt_0 gr0I leD wt_ge_\<epsilon>)
next
case \<delta>_ne_0: False
hence ary_g_ninf: "arity_sym g \<noteq> \<infinity>"
using arity_sym_ne_infinity_if_\<delta>_gt_0 by blast
hence "\<delta> < the_enat (enat \<delta> * arity_sym g)"
using \<delta>_ne_0 ary_g by (cases "arity_sym g") (auto simp: one_enat_def)
thus ?thesis
unfolding wt_\<zeta> by simp
qed
next
case (App s t)
thus ?case
using wt_ge_\<delta>[of t] by force
qed
lemma wt_App_\<delta>: "wt (App s t) = wt t \<Longrightarrow> wt s = \<delta>"
by (simp add: order.antisym wt_ge_\<delta>)
lemma wt_App_ge_fun: "wt (App s t) \<ge> wt s"
by (metis diff_le_mono2 wt_ge_\<delta> le_diff_conv wt.simps(2))
lemma wt_hd_le: "wt (Hd (head s)) \<le> wt s"
by (induct s, simp) (metis head_App leD le_less_trans not_le_imp_less wt_App_ge_fun)
lemma wt_\<delta>_imp_\<delta>_eq_\<epsilon>: "wt s = \<delta> \<Longrightarrow> \<delta> = \<epsilon>"
by (metis \<delta>_le_\<epsilon> le_antisym wt_ge_\<epsilon>)
lemma wt_ge_arity_head_if_\<delta>_gt_0:
assumes \<delta>_gt_0: "\<delta> > 0"
shows "wt s \<ge> arity_hd (head s)"
proof (induct s)
case (Hd \<zeta>)
obtain f where
f_in_\<zeta>: "f \<in> ground_heads \<zeta>" and
wt_\<zeta>: "wt (Hd \<zeta>) = wt_sym f + the_enat (\<delta> * arity_sym f)"
using exists_wt_sym by blast
have "arity_sym f \<ge> arity_hd \<zeta>"
by (rule ground_heads_arity[OF f_in_\<zeta>])
hence "the_enat (\<delta> * arity_sym f) \<ge> arity_hd \<zeta>"
using \<delta>_gt_0
by (metis One_nat_def Suc_ile_eq dual_order.trans enat_ord_simps(2)
enat_the_enat_\<delta>_times_arity_sym i0_lb mult.commute mult.right_neutral mult_left_mono
one_enat_def)
thus ?case
unfolding wt_\<zeta> by (metis add.left_neutral add_mono le_iff_add plus_enat_simps(1) tm.sel(1))
next
case App
thus ?case
by (metis dual_order.trans enat_ord_simps(1) head_App wt_App_ge_fun)
qed
lemma wt_ge_num_args_if_\<delta>_eq_0:
assumes \<delta>_eq_0: "\<delta> = 0"
shows "wt s \<ge> num_args s"
by (induct s, simp_all,
metis (no_types) \<delta>_eq_0 \<epsilon>_gt_0 wt_\<delta>_imp_\<delta>_eq_\<epsilon> add_le_same_cancel1 le_0_eq le_trans
minus_nat.diff_0 not_gr_zero not_less_eq_eq)
lemma wt_ge_num_args: "wary s \<Longrightarrow> wt s \<ge> num_args s"
using wt_ge_arity_head_if_\<delta>_gt_0 wt_ge_num_args_if_\<delta>_eq_0
by (meson order.trans enat_ord_simps(1) neq0_conv wary_num_args_le_arity_head)
subsection \<open>Inductive Definitions\<close>
inductive gt :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" (infix ">\<^sub>t" 50) where
gt_wt: "vars_mset t \<supseteq># vars_mset s \<Longrightarrow> wt t > wt s \<Longrightarrow> t >\<^sub>t s"
| gt_unary: "wt t = wt s \<Longrightarrow> \<not> head t \<le>\<ge>\<^sub>h\<^sub>d head s \<Longrightarrow> num_args t = 1 \<Longrightarrow>
(\<exists>f \<in> ground_heads (head t). arity_sym f = 1 \<and> wt_sym f = 0) \<Longrightarrow> arg t >\<^sub>t s \<or> arg t = s \<Longrightarrow>
t >\<^sub>t s"
| gt_diff: "vars_mset t \<supseteq># vars_mset s \<Longrightarrow> wt t = wt s \<Longrightarrow> head t >\<^sub>h\<^sub>d head s \<Longrightarrow> t >\<^sub>t s"
| gt_same: "vars_mset t \<supseteq># vars_mset s \<Longrightarrow> wt t = wt s \<Longrightarrow> head t = head s \<Longrightarrow>
(\<forall>f \<in> ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \<Longrightarrow> t >\<^sub>t s"
abbreviation ge :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" (infix "\<ge>\<^sub>t" 50) where
"t \<ge>\<^sub>t s \<equiv> t >\<^sub>t s \<or> t = s"
inductive gt_wt :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" where
gt_wtI: "vars_mset t \<supseteq># vars_mset s \<Longrightarrow> wt t > wt s \<Longrightarrow> gt_wt t s"
inductive gt_diff :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" where
gt_diffI: "vars_mset t \<supseteq># vars_mset s \<Longrightarrow> wt t = wt s \<Longrightarrow> head t >\<^sub>h\<^sub>d head s \<Longrightarrow> gt_diff t s"
inductive gt_unary :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" where
gt_unaryI: "wt t = wt s \<Longrightarrow> \<not> head t \<le>\<ge>\<^sub>h\<^sub>d head s \<Longrightarrow> num_args t = 1 \<Longrightarrow>
(\<exists>f \<in> ground_heads (head t). arity_sym f = 1 \<and> wt_sym f = 0) \<Longrightarrow> arg t \<ge>\<^sub>t s \<Longrightarrow> gt_unary t s"
inductive gt_same :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" where
gt_sameI: "vars_mset t \<supseteq># vars_mset s \<Longrightarrow> wt t = wt s \<Longrightarrow> head t = head s \<Longrightarrow>
(\<forall>f \<in> ground_heads (head t). extf f (>\<^sub>t) (args t) (args s)) \<Longrightarrow> gt_same t s"
lemma gt_iff_wt_unary_diff_same: "t >\<^sub>t s \<longleftrightarrow> gt_wt t s \<or> gt_unary t s \<or> gt_diff t s \<or> gt_same t s"
by (subst gt.simps) (auto simp: gt_wt.simps gt_unary.simps gt_diff.simps gt_same.simps)
lemma gt_imp_vars_mset: "t >\<^sub>t s \<Longrightarrow> vars_mset t \<supseteq># vars_mset s"
- by (induct rule: gt.induct) (auto intro: subset_mset.order.trans)
+ by (induct rule: gt.induct) (auto intro: subset_mset.trans)
lemma gt_imp_vars: "t >\<^sub>t s \<Longrightarrow> vars t \<supseteq> vars s"
using set_mset_mono[OF gt_imp_vars_mset] by simp
subsection \<open>Irreflexivity\<close>
theorem gt_irrefl: "wary s \<Longrightarrow> \<not> s >\<^sub>t s"
proof (induct "size s" arbitrary: s rule: less_induct)
case less
note ih = this(1) and wary_s = this(2)
show ?case
proof
assume s_gt_s: "s >\<^sub>t s"
show False
using s_gt_s
proof (cases rule: gt.cases)
case gt_same
then obtain f where f: "extf f (>\<^sub>t) (args s) (args s)"
by fastforce
thus False
using wary_s ih by (metis wary_args extf_irrefl size_in_args)
qed (auto simp: comp_hd_def gt_hd_irrefl)
qed
qed
subsection \<open>Transitivity\<close>
lemma gt_imp_wt_ge: "t >\<^sub>t s \<Longrightarrow> wt t \<ge> wt s"
by (induct rule: gt.induct) auto
lemma not_extf_gt_nil_singleton_if_\<delta>_eq_\<epsilon>:
assumes wary_s: "wary s" and \<delta>_eq_\<epsilon>: "\<delta> = \<epsilon>"
shows "\<not> extf f (>\<^sub>t) [] [s]"
proof
assume nil_gt_s: "extf f (>\<^sub>t) [] [s]"
note s_gt_nil = extf_singleton_nil_if_\<delta>_eq_\<epsilon>[OF \<delta>_eq_\<epsilon>, of f gt s]
have "\<not> extf f (>\<^sub>t) [] []"
by (rule extf_irrefl) simp
moreover have "extf f (>\<^sub>t) [] []"
using extf_trans_from_irrefl[of "{s}", OF _ _ _ _ _ _ nil_gt_s s_gt_nil] gt_irrefl[OF wary_s]
by fastforce
ultimately show False
by sat
qed
lemma gt_sub_arg: "wary (App s t) \<Longrightarrow> App s t >\<^sub>t t"
proof (induct t arbitrary: s rule: measure_induct_rule[of size])
case (less t)
note ih = this(1) and wary_st = this(2)
{
assume wt_st: "wt (App s t) = wt t"
hence \<delta>_eq_\<epsilon>: "\<delta> = \<epsilon>"
using wt_App_\<delta> wt_\<delta>_imp_\<delta>_eq_\<epsilon> by metis
hence \<delta>_gt_0: "\<delta> > 0"
using \<epsilon>_gt_0 by simp
have wt_s: "wt s = \<delta>"
by (rule wt_App_\<delta>[OF wt_st])
have
wary_t: "wary t" and
nargs_lt: "num_args s < arity_hd (head s)"
using wary_st wary.simps by blast+
have ary_hd_s: "arity_hd (head s) = 1"
by (metis One_nat_def arity.wary_AppE dual_order.order_iff_strict eSuc_enat enat_defs(1)
enat_defs(2) ileI1 linorder_not_le not_iless0 wary_st wt_gt_\<delta>_if_superunary wt_s)
hence nargs_s: "num_args s = 0"
by (metis enat_ord_simps(2) less_one nargs_lt one_enat_def)
have "s = Hd (head s)"
by (simp add: Hd_head_id nargs_s)
then obtain f where
f_in: "f \<in> ground_heads (head s)" and
wt_f_etc: "wt_sym f + the_enat (\<delta> * arity_sym f) = \<delta>"
using exists_wt_sym wt_s by fastforce
have ary_f_1: "arity_sym f = 1"
proof -
have ary_f_ge_1: "arity_sym f \<ge> 1"
using ary_hd_s f_in ground_heads_arity by fastforce
hence "enat \<delta> * arity_sym f = \<delta>"
using wt_f_etc by (metis enat_ord_simps(1) enat_the_enat_\<delta>_times_arity_sym le_add2
le_antisym mult.right_neutral mult_left_mono zero_le)
thus ?thesis
using \<delta>_gt_0 by (cases "arity_sym f") (auto simp: one_enat_def)
qed
hence wt_f_0: "wt_sym f = 0"
using wt_f_etc by simp
{
assume hd_s_ncmp_t: "\<not> head s \<le>\<ge>\<^sub>h\<^sub>d head t"
have ?case
by (rule gt_unary[OF wt_st]) (auto simp: hd_s_ncmp_t nargs_s intro: f_in ary_f_1 wt_f_0)
}
moreover
{
assume hd_s_gt_t: "head s >\<^sub>h\<^sub>d head t"
have ?case
by (rule gt_diff) (auto simp: hd_s_gt_t wt_s[folded \<delta>_eq_\<epsilon>])
}
moreover
{
assume "head t >\<^sub>h\<^sub>d head s"
hence False
using ary_f_1 exists_wt_sym f_in gt_hd_def gt_sym_antisym unary_wt_sym_0_gt wt_f_0 by blast
}
moreover
{
assume hd_t_eq_s: "head t = head s"
hence nargs_t_le: "num_args t \<le> 1"
using ary_hd_s wary_num_args_le_arity_head[OF wary_t] by (simp add: one_enat_def)
have extf: "extf f (>\<^sub>t) [t] (args t)" for f
proof (cases "args t")
case Nil
thus ?thesis
by (simp add: extf_singleton_nil_if_\<delta>_eq_\<epsilon>[OF \<delta>_eq_\<epsilon>])
next
case args_t: (Cons ta ts)
hence ts: "ts = []"
using ary_hd_s[folded hd_t_eq_s] wary_num_args_le_arity_head[OF wary_t]
nargs_t_le by simp
have ta: "ta = arg t"
by (metis apps.simps(1) apps.simps(2) args_t tm.sel(6) tm_collapse_apps ts)
hence t: "t = App (fun t) ta"
by (metis args.simps(1) args_t not_Cons_self2 tm.exhaust_sel ts)
have "t >\<^sub>t ta"
by (rule ih[of ta "fun t", folded t, OF _ wary_t]) (metis ta size_arg_lt t tm.disc(2))
thus ?thesis
unfolding args_t ts by (metis extf_singleton gt_irrefl wary_t)
qed
have ?case
by (rule gt_same)
(auto simp: hd_t_eq_s wt_s[folded \<delta>_eq_\<epsilon>] length_0_conv[THEN iffD1, OF nargs_s] extf)
}
ultimately have ?case
unfolding comp_hd_def by metis
}
thus ?case
using gt_wt by fastforce
qed
lemma gt_arg: "wary s \<Longrightarrow> is_App s \<Longrightarrow> s >\<^sub>t arg s"
by (cases s) (auto intro: gt_sub_arg)
theorem gt_trans: "wary u \<Longrightarrow> wary t \<Longrightarrow> wary s \<Longrightarrow> u >\<^sub>t t \<Longrightarrow> t >\<^sub>t s \<Longrightarrow> u >\<^sub>t s"
proof (simp only: atomize_imp,
rule measure_induct_rule[of "\<lambda>(u, t, s). {#size u, size t, size s#}"
"\<lambda>(u, t, s). wary u \<longrightarrow> wary t \<longrightarrow> wary s \<longrightarrow> u >\<^sub>t t \<longrightarrow> t >\<^sub>t s \<longrightarrow> u >\<^sub>t s" "(u, t, s)",
simplified prod.case],
simp only: split_paired_all prod.case atomize_imp[symmetric])
fix u t s
assume
ih: "\<And>ua ta sa. {#size ua, size ta, size sa#} < {#size u, size t, size s#} \<Longrightarrow>
wary ua \<Longrightarrow> wary ta \<Longrightarrow> wary sa \<Longrightarrow> ua >\<^sub>t ta \<Longrightarrow> ta >\<^sub>t sa \<Longrightarrow> ua >\<^sub>t sa" and
wary_u: "wary u" and wary_t: "wary t" and wary_s: "wary s" and
u_gt_t: "u >\<^sub>t t" and t_gt_s: "t >\<^sub>t s"
have "vars_mset u \<supseteq># vars_mset t" and "vars_mset t \<supseteq># vars_mset s"
using u_gt_t t_gt_s by (auto simp: gt_imp_vars_mset)
hence vars_u_s: "vars_mset u \<supseteq># vars_mset s"
by auto
have wt_u_ge_t: "wt u \<ge> wt t" and wt_t_ge_s: "wt t \<ge> wt s"
using gt_imp_wt_ge u_gt_t t_gt_s by auto
{
assume wt_t_s: "wt t = wt s" and wt_u_t: "wt u = wt t"
hence wt_u_s: "wt u = wt s"
by simp
have wary_arg_u: "wary (arg u)"
by (rule wary_arg[OF wary_u])
have wary_arg_t: "wary (arg t)"
by (rule wary_arg[OF wary_t])
have wary_arg_s: "wary (arg s)"
by (rule wary_arg[OF wary_s])
have "u >\<^sub>t s"
using t_gt_s
proof cases
case gt_unary_t_s: gt_unary
have t_app: "is_App t"
by (metis args_Nil_iff_is_Hd gt_unary_t_s(3) length_greater_0_conv less_numeral_extra(1))
have \<delta>_eq_\<epsilon>: "\<delta> = \<epsilon>"
using gt_unary_t_s(4) unary_wt_sym_0_imp_\<delta>_eq_\<epsilon> by blast
show ?thesis
using u_gt_t
proof cases
case gt_unary_u_t: gt_unary
have u_app: "is_App u"
by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1))
hence arg_u_gt_s: "arg u >\<^sub>t s"
using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t
by force
hence arg_u_ge_s: "arg u \<ge>\<^sub>t s"
by sat
{
assume "size (arg u) < size t"
hence ?thesis
using ih[of u "arg u" s] arg_u_gt_s gt_arg by (simp add: u_app wary_arg_u wary_s wary_u)
}
moreover
{
assume "size (arg t) < size s"
hence "u >\<^sub>t arg t"
using ih[of u t "arg t"] args_Nil_iff_is_Hd gt_arg gt_unary_t_s(3) u_gt_t wary_t wary_u
by force
hence ?thesis
using ih[of u "arg t" s] args_Nil_iff_is_Hd gt_unary_t_s(3,5) size_arg_lt wary_arg_t
wary_s wary_u by force
}
moreover
{
assume sz_u_gt_t: "size u > size t" and sz_t_gt_s: "size t > size s"
have wt_fun_u: "wt (fun u) = \<delta>"
by (metis antisym gt_imp_wt_ge gt_unary_u_t(5) tm.collapse(2) u_app wt_App_\<delta> wt_arg_le
wt_t_s wt_u_s)
have nargs_fun_u: "num_args (fun u) = 0"
by (metis args.simps(1) gt_unary_u_t(3) list.size(3) one_arg_imp_Hd tm.collapse(2)
u_app)
{
assume hd_u_eq_s: "head u = head s"
hence ary_hd_s: "arity_hd (head s) = 1"
using ground_heads_arity gt_unary_u_t(3,4) hd_u_eq_s one_enat_def
wary_num_args_le_arity_head wary_u by fastforce
have extf: "extf f (>\<^sub>t) (args u) (args s)" for f
proof (cases "args s")
case Nil
thus ?thesis
by (metis Hd_head_id \<delta>_eq_\<epsilon> append_Nil args.simps(2) extf_singleton_nil_if_\<delta>_eq_\<epsilon>
gt_unary_u_t(3) head_fun length_greater_0_conv less_irrefl_nat nargs_fun_u
tm.exhaust_sel zero_neq_one)
next
case args_s: (Cons sa ss)
hence ss: "ss = []"
by (cases s, simp, metis One_nat_def antisym_conv ary_hd_s diff_Suc_1
enat_ord_simps(1) le_add2 length_0_conv length_Cons list.size(4) one_enat_def
wary_num_args_le_arity_head wary_s)
have sa: "sa = arg s"
by (metis apps.simps(1) apps.simps(2) args_s tm.sel(6) tm_collapse_apps ss)
have s_app: "is_App s"
using args_Nil_iff_is_Hd args_s by force
have args_u: "args u = [arg u]"
by (metis append_Nil args.simps(2) args_Nil_iff_is_Hd gt_unary_u_t(3) length_0_conv
nargs_fun_u tm.collapse(2) zero_neq_one)
have max_sz_u_t_s: "Max {size s, size t, size u} = size u"
using sz_t_gt_s sz_u_gt_t by auto
have max_sz_arg_u_t_arg_t: "Max {size (arg t), size t, size (arg u)} < size u"
using size_arg_lt sz_u_gt_t t_app u_app by fastforce
have "{#size (arg u), size t, size (arg t)#} < {#size u, size t, size s#}"
using max_sz_arg_u_t_arg_t
by (simp add: Max_lt_imp_lt_mset insert_commute max_sz_u_t_s)
hence arg_u_gt_arg_t: "arg u >\<^sub>t arg t"
using ih[OF _ wary_arg_u wary_t wary_arg_t] args_Nil_iff_is_Hd gt_arg
gt_unary_t_s(3) gt_unary_u_t(5) wary_t by force
have max_sz_arg_s_s_arg_t: "Max {size (arg s), size s, size (arg t)} < size u"
using s_app t_app size_arg_lt sz_t_gt_s sz_u_gt_t by force
have "{#size (arg t), size s, size (arg s)#} < {#size u, size t, size s#}"
by (meson add_mset_lt_lt_lt less_trans mset_lt_single_iff s_app size_arg_lt
sz_t_gt_s sz_u_gt_t t_app)
hence arg_t_gt_arg_s: "arg t >\<^sub>t arg s"
using ih[OF _ wary_arg_t wary_s wary_arg_s]
gt_unary_t_s(5) gt_arg args_Nil_iff_is_Hd args_s wary_s by force
have "arg u >\<^sub>t arg s"
using ih[of "arg u" "arg t" "arg s"] arg_u_gt_arg_t arg_t_gt_arg_s
by (simp add: add_mset_lt_le_lt less_imp_le_nat s_app size_arg_lt t_app u_app
wary_arg_s wary_arg_t wary_arg_u)
thus ?thesis
unfolding args_u args_s ss sa by (metis extf_singleton gt_irrefl wary_arg_u)
qed
have ?thesis
by (rule gt_same[OF vars_u_s wt_u_s hd_u_eq_s]) (simp add: extf)
}
moreover
{
assume "head u >\<^sub>h\<^sub>d head s"
hence ?thesis
by (rule gt_diff[OF vars_u_s wt_u_s])
}
moreover
{
assume "head s >\<^sub>h\<^sub>d head u"
hence False
using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast
}
moreover
{
assume "\<not> head u \<le>\<ge>\<^sub>h\<^sub>d head s"
hence ?thesis
by (rule gt_unary[OF wt_u_s _ gt_unary_u_t(3,4) arg_u_ge_s])
}
ultimately have ?thesis
unfolding comp_hd_def by sat
}
ultimately show ?thesis
by (metis args_Nil_iff_is_Hd dual_order.strict_trans2 gt_unary_t_s(3) gt_unary_u_t(3)
length_0_conv not_le_imp_less size_arg_lt zero_neq_one)
next
case gt_diff_u_t: gt_diff
have False
using gt_diff_u_t(3) gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_t_s(4)
unary_wt_sym_0_gt by blast
thus ?thesis
by sat
next
case gt_same_u_t: gt_same
have hd_u_ncomp_s: "\<not> head u \<le>\<ge>\<^sub>h\<^sub>d head s"
by (rule gt_unary_t_s(2)[folded gt_same_u_t(3)])
have "num_args u \<le> 1"
by (metis enat_ord_simps(1) ground_heads_arity gt_same_u_t(3) gt_unary_t_s(4) one_enat_def
order_trans wary_num_args_le_arity_head wary_u)
hence nargs_u: "num_args u = 1"
by (cases "args u",
metis Hd_head_id \<delta>_eq_\<epsilon> append_Nil args.simps(2) gt_same_u_t(3,4) gt_unary_t_s(3,4)
head_fun list.size(3) not_extf_gt_nil_singleton_if_\<delta>_eq_\<epsilon> one_arg_imp_Hd
tm.collapse(2)[OF t_app] wary_arg_t,
simp)
have "arg u >\<^sub>t arg t"
by (metis extf_singleton[THEN iffD1] append_Nil args.simps args_Nil_iff_is_Hd
comp_hd_def gt_hd_def gt_irrefl gt_same_u_t(3,4) gt_unary_t_s(2,3) head_fun
length_0_conv nargs_u one_arg_imp_Hd t_app tm.collapse(2) u_gt_t wary_u)
hence "arg u >\<^sub>t s"
using ih[OF _ wary_arg_u wary_arg_t wary_s] gt_unary_t_s(5)
by (metis add_mset_lt_left_lt add_mset_lt_lt_lt args_Nil_iff_is_Hd list.size(3) nargs_u
size_arg_lt t_app zero_neq_one)
hence arg_u_ge_s: "arg u \<ge>\<^sub>t s"
by sat
show ?thesis
by (rule gt_unary[OF wt_u_s hd_u_ncomp_s nargs_u _ arg_u_ge_s])
(simp add: gt_same_u_t(3) gt_unary_t_s(4))
qed (simp add: wt_u_t)
next
case gt_diff_t_s: gt_diff
show ?thesis
using u_gt_t
proof cases
case gt_unary_u_t: gt_unary
have "is_App u"
by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1))
hence "arg u >\<^sub>t s"
using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t
by force
hence arg_u_ge_s: "arg u \<ge>\<^sub>t s"
by sat
{
assume "head u = head s"
hence False
using gt_diff_t_s(3) gt_unary_u_t(2) unfolding comp_hd_def by force
}
moreover
{
assume "head s >\<^sub>h\<^sub>d head u"
hence False
using gt_hd_def gt_hd_irrefl gt_sym_antisym gt_unary_u_t(4) unary_wt_sym_0_gt by blast
}
moreover
{
assume "head u >\<^sub>h\<^sub>d head s"
hence ?thesis
by (rule gt_diff[OF vars_u_s wt_u_s])
}
moreover
{
assume "\<not> head u \<le>\<ge>\<^sub>h\<^sub>d head s"
hence ?thesis
by (rule gt_unary[OF wt_u_s _ gt_unary_u_t(3,4) arg_u_ge_s])
}
ultimately show ?thesis
unfolding comp_hd_def by sat
next
case gt_diff_u_t: gt_diff
have "head u >\<^sub>h\<^sub>d head s"
using gt_diff_u_t(3) gt_diff_t_s(3) gt_hd_trans by blast
thus ?thesis
by (rule gt_diff[OF vars_u_s wt_u_s])
next
case gt_same_u_t: gt_same
have "head u >\<^sub>h\<^sub>d head s"
using gt_diff_t_s(3) gt_same_u_t(3) by simp
thus ?thesis
by (rule gt_diff[OF vars_u_s wt_u_s])
qed (simp add: wt_u_t)
next
case gt_same_t_s: gt_same
show ?thesis
using u_gt_t
proof cases
case gt_unary_u_t: gt_unary
have "is_App u"
by (metis args_Nil_iff_is_Hd gt_unary_u_t(3) length_greater_0_conv less_numeral_extra(1))
hence "arg u >\<^sub>t s"
using ih[of "arg u" t s] gt_unary_u_t(5) t_gt_s size_arg_lt wary_arg_u wary_s wary_t
by force
hence arg_u_ge_s: "arg u \<ge>\<^sub>t s"
by sat
have "\<not> head u \<le>\<ge>\<^sub>h\<^sub>d head s"
using gt_same_t_s(3) gt_unary_u_t(2) by simp
thus ?thesis
by (rule gt_unary[OF wt_u_s _ gt_unary_u_t(3,4) arg_u_ge_s])
next
case gt_diff_u_t: gt_diff
have "head u >\<^sub>h\<^sub>d head s"
using gt_diff_u_t(3) gt_same_t_s(3) by simp
thus ?thesis
by (rule gt_diff[OF vars_u_s wt_u_s])
next
case gt_same_u_t: gt_same
have hd_u_s: "head u = head s"
by (simp only: gt_same_t_s(3) gt_same_u_t(3))
let ?S = "set (args u) \<union> set (args t) \<union> set (args s)"
have gt_trans_args: "\<forall>ua \<in> ?S. \<forall>ta \<in> ?S. \<forall>sa \<in> ?S. ua >\<^sub>t ta \<longrightarrow> ta >\<^sub>t sa \<longrightarrow> ua >\<^sub>t sa"
proof clarify
fix sa ta ua
assume
ua_in: "ua \<in> ?S" and ta_in: "ta \<in> ?S" and sa_in: "sa \<in> ?S" and
ua_gt_ta: "ua >\<^sub>t ta" and ta_gt_sa: "ta >\<^sub>t sa"
have wary_sa: "wary sa" and wary_ta: "wary ta" and wary_ua: "wary ua"
using wary_args ua_in ta_in sa_in wary_u wary_t wary_s by blast+
show "ua >\<^sub>t sa"
by (auto intro!: ih[OF Max_lt_imp_lt_mset wary_ua wary_ta wary_sa ua_gt_ta ta_gt_sa])
(meson ua_in ta_in sa_in Un_iff max.strict_coboundedI1 max.strict_coboundedI2
size_in_args)+
qed
have "\<forall>f \<in> ground_heads (head u). extf f (>\<^sub>t) (args u) (args s)"
by (clarify, rule extf_trans_from_irrefl[of ?S _ "args t", OF _ _ _ _ _ gt_trans_args])
(auto simp: gt_same_u_t(3,4) gt_same_t_s(4) wary_args wary_u wary_t wary_s gt_irrefl)
thus ?thesis
by (rule gt_same[OF vars_u_s wt_u_s hd_u_s])
qed (simp add: wt_u_t)
qed (simp add: wt_t_s)
}
thus "u >\<^sub>t s"
using vars_u_s wt_t_ge_s wt_u_ge_t by (force intro: gt_wt)
qed
lemma gt_antisym: "wary s \<Longrightarrow> wary t \<Longrightarrow> t >\<^sub>t s \<Longrightarrow> \<not> s >\<^sub>t t"
using gt_irrefl gt_trans by blast
subsection \<open>Subterm Property\<close>
lemma gt_sub_fun: "App s t >\<^sub>t s"
proof (cases "wt (App s t) > wt s")
case True
thus ?thesis
using gt_wt by simp
next
case False
hence wt_st: "wt (App s t) = wt s"
by (meson order.antisym not_le_imp_less wt_App_ge_fun)
hence \<delta>_eq_\<epsilon>: "\<delta> = \<epsilon>"
by (metis add_diff_cancel_left' diff_diff_cancel wt_\<delta>_imp_\<delta>_eq_\<epsilon> wt_ge_\<delta> wt.simps(2))
have vars_st: "vars_mset (App s t) \<supseteq># vars_mset s"
by auto
have hd_st: "head (App s t) = head s"
by auto
have extf: "\<forall>f \<in> ground_heads (head (App s t)). extf f (>\<^sub>t) (args (App s t)) (args s)"
by (simp add: \<delta>_eq_\<epsilon> extf_snoc_if_\<delta>_eq_\<epsilon>)
show ?thesis
by (rule gt_same[OF vars_st wt_st hd_st extf])
qed
theorem gt_proper_sub: "wary t \<Longrightarrow> proper_sub s t \<Longrightarrow> t >\<^sub>t s"
by (induct t) (auto intro: gt_sub_fun gt_sub_arg gt_trans sub.intros wary_sub)
subsection \<open>Compatibility with Functions\<close>
theorem gt_compat_fun:
assumes
wary_t: "wary t" and
t'_gt_t: "t' >\<^sub>t t"
shows "App s t' >\<^sub>t App s t"
proof -
have vars_st': "vars_mset (App s t') \<supseteq># vars_mset (App s t)"
by (simp add: t'_gt_t gt_imp_vars_mset)
show ?thesis
proof (cases "wt t' > wt t")
case True
hence wt_st': "wt (App s t') > wt (App s t)"
by (simp only: wt.simps)
show ?thesis
by (rule gt_wt[OF vars_st' wt_st'])
next
case False
hence "wt t' = wt t"
using t'_gt_t gt_imp_wt_ge order.not_eq_order_implies_strict by fastforce
hence wt_st': "wt (App s t') = wt (App s t)"
by (simp only: wt.simps)
have head_st': "head (App s t') = head (App s t)"
by simp
have extf: "\<And>f. extf f (>\<^sub>t) (args s @ [t']) (args s @ [t])"
using t'_gt_t by (metis extf_compat_list gt_irrefl[OF wary_t])
show ?thesis
by (rule gt_same[OF vars_st' wt_st' head_st']) (simp add: extf)
qed
qed
subsection \<open>Compatibility with Arguments\<close>
theorem gt_compat_arg:
assumes wary_s't: "wary (App s' t)" and s'_gt_s: "s' >\<^sub>t s"
shows "App s' t >\<^sub>t App s t"
proof -
have vars_s't: "vars_mset (App s' t) \<supseteq># vars_mset (App s t)"
by (simp add: s'_gt_s gt_imp_vars_mset)
show ?thesis
using s'_gt_s
proof cases
case gt_wt_s'_s: gt_wt
have "wt (App s' t) > wt (App s t)"
by (simp add: wt_ge_\<delta>) (metis add_diff_assoc add_less_cancel_right gt_wt_s'_s(2) wt_ge_\<delta>)
thus ?thesis
by (rule gt_wt[OF vars_s't])
next
case gt_unary_s'_s: gt_unary
have False
by (metis ground_heads_arity gt_unary_s'_s(3) gt_unary_s'_s(4) leD one_enat_def wary_AppE
wary_s't)
thus ?thesis
by sat
next
case _: gt_diff
thus ?thesis
by (simp add: gt_diff)
next
case gt_same_s'_s: gt_same
have wt_s't: "wt (App s' t) = wt (App s t)"
by (simp add: gt_same_s'_s(2))
have hd_s't: "head (App s' t) = head (App s t)"
by (simp add: gt_same_s'_s(3))
have "\<forall>f \<in> ground_heads (head (App s' t)). extf f (>\<^sub>t) (args (App s' t)) (args (App s t))"
using gt_same_s'_s(4) by (auto intro: extf_compat_append_right)
thus ?thesis
by (rule gt_same[OF vars_s't wt_s't hd_s't])
qed
qed
subsection \<open>Stability under Substitution\<close>
definition extra_wt :: "('v \<Rightarrow> ('s, 'v) tm) \<Rightarrow> ('s, 'v) tm \<Rightarrow> nat" where
"extra_wt \<rho> s = (\<Sum>x \<in># vars_mset s. wt (\<rho> x) - wt (Hd (Var x)))"
lemma
extra_wt_Var[simp]: "extra_wt \<rho> (Hd (Var x)) = wt (\<rho> x) - wt (Hd (Var x))" and
extra_wt_Sym[simp]: "extra_wt \<rho> (Hd (Sym f)) = 0" and
extra_wt_App[simp]: "extra_wt \<rho> (App s t) = extra_wt \<rho> s + extra_wt \<rho> t"
unfolding extra_wt_def by simp+
lemma extra_wt_subseteq:
assumes vars_s: "vars_mset t \<supseteq># vars_mset s"
shows "extra_wt \<rho> t \<ge> extra_wt \<rho> s"
proof (unfold extra_wt_def)
let ?diff = "\<lambda>v. wt (\<rho> v) - wt (Hd (Var v))"
have "vars_mset s + (vars_mset t - vars_mset s) = vars_mset t"
using vars_s by (meson subset_mset.add_diff_inverse)
hence "{#?diff v. v \<in># vars_mset t#} =
{#?diff v. v \<in># vars_mset s#} + {#?diff v. v \<in># vars_mset t - vars_mset s#}"
by (metis image_mset_union)
thus "(\<Sum>v \<in># vars_mset t. ?diff v) \<ge> (\<Sum>v \<in># vars_mset s. ?diff v)"
by simp
qed
lemma wt_subst:
assumes wary_\<rho>: "wary_subst \<rho>" and wary_s: "wary s"
shows "wt (subst \<rho> s) = wt s + extra_wt \<rho> s"
using wary_s
proof (induct s rule: tm.induct)
case \<zeta>: (Hd \<zeta>)
show ?case
proof (cases \<zeta>)
case x: (Var x)
let ?\<xi> = "head (\<rho> x)"
obtain g where
g_in_grs_\<xi>: "g \<in> ground_heads ?\<xi>" and
wt_\<xi>: "wt (Hd ?\<xi>) = wt_sym g + the_enat (\<delta> * arity_sym g)"
using exists_wt_sym by blast
have "g \<in> ground_heads \<zeta>"
using x g_in_grs_\<xi> wary_\<rho> wary_subst_def by auto
hence wt_\<rho>x_ge: "wt (\<rho> x) \<ge> wt (Hd \<zeta>)"
by (metis (full_types) dual_order.trans wt_le_wt_sym wt_\<xi> wt_hd_le)
thus ?thesis
using x by (simp add: extra_wt_def)
qed auto
next
case (App s t)
note ih_s = this(1) and ih_t = this(2) and wary_st = this(3)
have "wary s"
using wary_st by (meson wary_AppE)
hence "\<And>n. extra_wt \<rho> s + (wt s - \<delta> + n) = wt (subst \<rho> s) - \<delta> + n"
using ih_s by (metis (full_types) add_diff_assoc2 ab_semigroup_add_class.add_ac(1)
add.left_commute wt_ge_\<delta>)
hence "extra_wt \<rho> s + (wt s + wt t - \<delta> + extra_wt \<rho> t) = wt (subst \<rho> s) + wt (subst \<rho> t) - \<delta>"
using ih_t wary_st
by (metis (no_types) add_diff_assoc2 ab_semigroup_add_class.add_ac(1) wary_AppE wt_ge_\<delta>)
thus ?case
by (simp add: wt_ge_\<delta>)
qed
theorem gt_subst:
assumes wary_\<rho>: "wary_subst \<rho>"
shows "wary t \<Longrightarrow> wary s \<Longrightarrow> t >\<^sub>t s \<Longrightarrow> subst \<rho> t >\<^sub>t subst \<rho> s"
proof (simp only: atomize_imp,
rule measure_induct_rule[of "\<lambda>(t, s). {#size t, size s#}"
"\<lambda>(t, s). wary t \<longrightarrow> wary s \<longrightarrow> t >\<^sub>t s \<longrightarrow> subst \<rho> t >\<^sub>t subst \<rho> s" "(t, s)",
simplified prod.case],
simp only: split_paired_all prod.case atomize_imp[symmetric])
fix t s
assume
ih: "\<And>ta sa. {#size ta, size sa#} < {#size t, size s#} \<Longrightarrow> wary ta \<Longrightarrow> wary sa \<Longrightarrow> ta >\<^sub>t sa \<Longrightarrow>
subst \<rho> ta >\<^sub>t subst \<rho> sa" and
wary_t: "wary t" and wary_s: "wary s" and t_gt_s: "t >\<^sub>t s"
show "subst \<rho> t >\<^sub>t subst \<rho> s"
proof (cases "wt (subst \<rho> t) = wt (subst \<rho> s)")
case wt_\<rho>t_ne_\<rho>s: False
have vars_s: "vars_mset t \<supseteq># vars_mset s"
by (simp add: t_gt_s gt_imp_vars_mset)
hence vars_\<rho>s: "vars_mset (subst \<rho> t) \<supseteq># vars_mset (subst \<rho> s)"
by (rule vars_mset_subst_subseteq)
have wt_t_ge_s: "wt t \<ge> wt s"
by (simp add: gt_imp_wt_ge t_gt_s)
have "wt (subst \<rho> t) > wt (subst \<rho> s)"
using wt_\<rho>t_ne_\<rho>s unfolding wt_subst[OF wary_\<rho> wary_s] wt_subst[OF wary_\<rho> wary_t]
by (metis add_le_cancel_left add_less_le_mono extra_wt_subseteq
order.not_eq_order_implies_strict vars_s wt_t_ge_s)
thus ?thesis
by (rule gt_wt[OF vars_\<rho>s])
next
case wt_\<rho>t_eq_\<rho>s: True
show ?thesis
using t_gt_s
proof cases
case gt_wt
hence False
using wt_\<rho>t_eq_\<rho>s wary_s wary_t
by (metis add_diff_cancel_right' diff_le_mono2 extra_wt_subseteq wt_subst leD wary_\<rho>)
thus ?thesis
by sat
next
case gt_unary
have wary_\<rho>t: "wary (subst \<rho> t)"
by (simp add: wary_subst_wary wary_t wary_\<rho>)
show ?thesis
proof (cases t)
case Hd
hence False
using gt_unary(3) by simp
thus ?thesis
by sat
next
case t: (App t1 t2)
hence t2: "t2 = arg t"
by simp
hence wary_t2: "wary t2"
using wary_t by blast
show ?thesis
proof (cases "t2 = s")
case True
moreover have "subst \<rho> t >\<^sub>t subst \<rho> t2"
using gt_sub_arg wary_\<rho>t unfolding t by simp
ultimately show ?thesis
by simp
next
case t2_ne_s: False
hence t2_gt_s: "t2 >\<^sub>t s"
using gt_unary(5) t2 by blast
have "subst \<rho> t2 >\<^sub>t subst \<rho> s"
by (rule ih[OF _ wary_t2 wary_s t2_gt_s]) (simp add: t)
thus ?thesis
by (metis gt_sub_arg gt_trans subst.simps(2) t wary_\<rho> wary_\<rho>t wary_s wary_subst_wary
wary_t2)
qed
qed
next
case _: gt_diff
note vars_s = this(1) and hd_t_gt_hd_s = this(3)
have vars_\<rho>s: "vars_mset (subst \<rho> t) \<supseteq># vars_mset (subst \<rho> s)"
by (rule vars_mset_subst_subseteq[OF vars_s])
have "head (subst \<rho> t) >\<^sub>h\<^sub>d head (subst \<rho> s)"
by (meson hd_t_gt_hd_s wary_subst_ground_heads gt_hd_def rev_subsetD wary_\<rho>)
thus ?thesis
by (rule gt_diff[OF vars_\<rho>s wt_\<rho>t_eq_\<rho>s])
next
case _: gt_same
note vars_s = this(1) and hd_s_eq_hd_t = this(3) and extf = this(4)
have vars_\<rho>s: "vars_mset (subst \<rho> t) \<supseteq># vars_mset (subst \<rho> s)"
by (rule vars_mset_subst_subseteq[OF vars_s])
have hd_\<rho>t: "head (subst \<rho> t) = head (subst \<rho> s)"
by (simp add: hd_s_eq_hd_t)
{
fix f
assume f_in_grs: "f \<in> ground_heads (head (subst \<rho> t))"
let ?S = "set (args t) \<union> set (args s)"
have extf_args_s_t: "extf f (>\<^sub>t) (args t) (args s)"
using extf f_in_grs wary_subst_ground_heads wary_\<rho> by blast
have "extf f (>\<^sub>t) (map (subst \<rho>) (args t)) (map (subst \<rho>) (args s))"
proof (rule extf_map[of ?S, OF _ _ _ _ _ _ extf_args_s_t])
show "\<forall>x \<in> ?S. \<not> subst \<rho> x >\<^sub>t subst \<rho> x"
using gt_irrefl wary_t wary_s wary_args wary_\<rho> wary_subst_wary by fastforce
next
show "\<forall>z \<in> ?S. \<forall>y \<in> ?S. \<forall>x \<in> ?S. subst \<rho> z >\<^sub>t subst \<rho> y \<longrightarrow> subst \<rho> y >\<^sub>t subst \<rho> x \<longrightarrow>
subst \<rho> z >\<^sub>t subst \<rho> x"
using gt_trans wary_t wary_s wary_args wary_\<rho> wary_subst_wary by (metis Un_iff)
next
have sz_a: "\<forall>ta \<in> ?S. \<forall>sa \<in> ?S. {#size ta, size sa#} < {#size t, size s#}"
by (fastforce intro: Max_lt_imp_lt_mset dest: size_in_args)
show "\<forall>y \<in> ?S. \<forall>x \<in> ?S. y >\<^sub>t x \<longrightarrow> subst \<rho> y >\<^sub>t subst \<rho> x"
using ih sz_a size_in_args wary_t wary_s wary_args wary_\<rho> wary_subst_wary by fastforce
qed auto
hence "extf f (>\<^sub>t) (args (subst \<rho> t)) (args (subst \<rho> s))"
by (auto simp: hd_s_eq_hd_t intro: extf_compat_append_left)
}
hence "\<forall>f \<in> ground_heads (head (subst \<rho> t)).
extf f (>\<^sub>t) (args (subst \<rho> t)) (args (subst \<rho> s))"
by blast
thus ?thesis
by (rule gt_same[OF vars_\<rho>s wt_\<rho>t_eq_\<rho>s hd_\<rho>t])
qed
qed
qed
subsection \<open>Totality on Ground Terms\<close>
theorem gt_total_ground:
assumes extf_total: "\<And>f. ext_total (extf f)"
shows "ground t \<Longrightarrow> ground s \<Longrightarrow> t >\<^sub>t s \<or> s >\<^sub>t t \<or> t = s"
proof (simp only: atomize_imp,
rule measure_induct_rule[of "\<lambda>(t, s). {# size t, size s #}"
"\<lambda>(t, s). ground t \<longrightarrow> ground s \<longrightarrow> t >\<^sub>t s \<or> s >\<^sub>t t \<or> t = s" "(t, s)", simplified prod.case],
simp only: split_paired_all prod.case atomize_imp[symmetric])
fix t s :: "('s, 'v) tm"
assume
ih: "\<And>ta sa. {# size ta, size sa #} < {# size t, size s #} \<Longrightarrow> ground ta \<Longrightarrow> ground sa \<Longrightarrow>
ta >\<^sub>t sa \<or> sa >\<^sub>t ta \<or> ta = sa" and
gr_t: "ground t" and gr_s: "ground s"
let ?case = "t >\<^sub>t s \<or> s >\<^sub>t t \<or> t = s"
have
vars_t: "vars_mset t \<supseteq># vars_mset s" and
vars_s: "vars_mset s \<supseteq># vars_mset t"
by (simp only: vars_mset_empty_iff[THEN iffD2, OF gr_s]
vars_mset_empty_iff[THEN iffD2, OF gr_t])+
show ?case
proof (cases "wt t = wt s")
case False
moreover
{
assume "wt t > wt s"
hence "t >\<^sub>t s"
by (rule gt_wt[OF vars_t])
}
moreover
{
assume "wt s > wt t"
hence "s >\<^sub>t t"
by (rule gt_wt[OF vars_s])
}
ultimately show ?thesis
by linarith
next
case wt_t: True
note wt_s = wt_t[symmetric]
obtain g where \<xi>: "head t = Sym g"
by (metis ground_head[OF gr_t] hd.collapse(2))
obtain f where \<zeta>: "head s = Sym f"
by (metis ground_head[OF gr_s] hd.collapse(2))
{
assume g_gt_f: "g >\<^sub>s f"
have "t >\<^sub>t s"
by (rule gt_diff[OF vars_t wt_t]) (simp add: \<xi> \<zeta> g_gt_f gt_hd_def)
}
moreover
{
assume f_gt_g: "f >\<^sub>s g"
have "s >\<^sub>t t"
by (rule gt_diff[OF vars_s wt_s]) (simp add: \<xi> \<zeta> f_gt_g gt_hd_def)
}
moreover
{
assume g_eq_f: "g = f"
hence hd_t: "head t = head s"
using \<xi> \<zeta> by force
note hd_s = hd_t[symmetric]
have gr_ts: "\<forall>t \<in> set (args t). ground t"
using gr_t ground_args by auto
have gr_ss: "\<forall>s \<in> set (args s). ground s"
using gr_s ground_args by auto
let ?ts = "args t"
let ?ss = "args s"
have ?thesis
proof (cases "?ts = ?ss")
case ts_eq_ss: True
show ?thesis
using \<xi> \<zeta> g_eq_f ts_eq_ss by (simp add: tm_expand_apps)
next
case False
hence "extf g (>\<^sub>t) (args t) ?ss \<or> extf g (>\<^sub>t) ?ss ?ts"
using ih gr_ss gr_ts
ext_total.total[OF extf_total, rule_format, of "set ?ts \<union> set ?ss" "(>\<^sub>t)" ?ts ?ss g]
by (metis Un_commute Un_iff in_lists_iff_set less_multiset_doubletons size_in_args sup_ge2)
moreover
{
assume extf: "extf g (>\<^sub>t) ?ts ?ss"
have "t >\<^sub>t s"
by (rule gt_same[OF vars_t wt_t hd_t]) (simp add: extf \<xi>)
}
moreover
{
assume extf: "extf g (>\<^sub>t) ?ss ?ts"
have "s >\<^sub>t t"
by (rule gt_same[OF vars_s wt_s hd_s]) (simp add: extf[unfolded g_eq_f] \<zeta>)
}
ultimately show ?thesis
by sat
qed
}
ultimately show ?thesis
using gt_sym_total by blast
qed
qed
subsection \<open>Well-foundedness\<close>
abbreviation gtw :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" (infix ">\<^sub>t\<^sub>w" 50) where
"(>\<^sub>t\<^sub>w) \<equiv> \<lambda>t s. wary t \<and> wary s \<and> t >\<^sub>t s"
abbreviation gtwg :: "('s, 'v) tm \<Rightarrow> ('s, 'v) tm \<Rightarrow> bool" (infix ">\<^sub>t\<^sub>w\<^sub>g" 50) where
"(>\<^sub>t\<^sub>w\<^sub>g) \<equiv> \<lambda>t s. ground t \<and> t >\<^sub>t\<^sub>w s"
lemma ground_gt_unary:
assumes gr_t: "ground t"
shows "\<not> gt_unary t s"
proof
assume gt_unary_t_s: "gt_unary t s"
hence "t >\<^sub>t s"
using gt_iff_wt_unary_diff_same by blast
hence gr_s: "ground s"
using gr_t gt_imp_vars by blast
have ngr_t_or_s: "\<not> ground t \<or> \<not> ground s"
using gt_unary_t_s by cases (blast dest: ground_head not_comp_hd_imp_Var)
show False
using gr_t gr_s ngr_t_or_s by sat
qed
theorem gt_wf: "wfP (\<lambda>s t. t >\<^sub>t\<^sub>w s)"
proof -
have ground_wfP: "wfP (\<lambda>s t. t >\<^sub>t\<^sub>w\<^sub>g s)"
unfolding wfP_iff_no_inf_chain
proof
assume "\<exists>f. inf_chain (>\<^sub>t\<^sub>w\<^sub>g) f"
then obtain t where t_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) t"
unfolding inf_chain_def bad_def by blast
let ?ff = "worst_chain (>\<^sub>t\<^sub>w\<^sub>g) (\<lambda>t s. size t > size s)"
note wf_sz = wf_app[OF wellorder_class.wf, of size, simplified]
have ffi_ground: "\<And>i. ground (?ff i)" and ffi_wary: "\<And>i. wary (?ff i)"
using worst_chain_bad[OF wf_sz t_bad, unfolded inf_chain_def] by fast+
have "inf_chain (>\<^sub>t\<^sub>w\<^sub>g) ?ff"
by (rule worst_chain_bad[OF wf_sz t_bad])
hence bad_wt_diff_same:
"inf_chain (\<lambda>t s. ground t \<and> (gt_wt t s \<or> gt_diff t s \<or> gt_same t s)) ?ff"
unfolding inf_chain_def using gt_iff_wt_unary_diff_same ground_gt_unary by blast
have wf_wt: "wf {(s, t). ground t \<and> gt_wt t s}"
by (rule wf_subset[OF wf_app[of _ wt, OF wf_less]]) (auto simp: gt_wt.simps)
have wt_O_diff_same: "{(s, t). ground t \<and> gt_wt t s}
O {(s, t). ground t \<and> (gt_diff t s \<or> gt_same t s)} \<subseteq> {(s, t). ground t \<and> gt_wt t s}"
unfolding gt_wt.simps gt_diff.simps gt_same.simps by auto
have wt_diff_same_as_union: "{(s, t). ground t \<and> (gt_wt t s \<or> gt_diff t s \<or> gt_same t s)} =
{(s, t). ground t \<and> gt_wt t s} \<union> {(s, t). ground t \<and> (gt_diff t s \<or> gt_same t s)}"
by auto
obtain k1 where bad_diff_same:
"inf_chain (\<lambda>t s. ground t \<and> (gt_diff t s \<or> gt_same t s)) (\<lambda>i. ?ff (i + k1))"
using wf_infinite_down_chain_compatible[OF wf_wt _ wt_O_diff_same, of ?ff] bad_wt_diff_same
unfolding inf_chain_def wt_diff_same_as_union[symmetric] by auto
have "wf {(s, t). ground s \<and> ground t \<and> sym (head t) >\<^sub>s sym (head s)}"
using gt_sym_wf unfolding wfP_def wf_iff_no_infinite_down_chain by fast
moreover have "{(s, t). ground t \<and> gt_diff t s}
\<subseteq> {(s, t). ground s \<and> ground t \<and> sym (head t) >\<^sub>s sym (head s)}"
proof (clarsimp, intro conjI)
fix s t
assume gr_t: "ground t" and gt_diff_t_s: "gt_diff t s"
thus gr_s: "ground s"
using gt_iff_wt_unary_diff_same gt_imp_vars by fastforce
show "sym (head t) >\<^sub>s sym (head s)"
using gt_diff_t_s by cases (simp add: gt_hd_def gr_s gr_t ground_hd_in_ground_heads)
qed
ultimately have wf_diff: "wf {(s, t). ground t \<and> gt_diff t s}"
by (rule wf_subset)
have diff_O_same: "{(s, t). ground t \<and> gt_diff t s} O {(s, t). ground t \<and> gt_same t s}
\<subseteq> {(s, t). ground t \<and> gt_diff t s}"
unfolding gt_diff.simps gt_same.simps by auto
have diff_same_as_union: "{(s, t). ground t \<and> (gt_diff t s \<or> gt_same t s)} =
{(s, t). ground t \<and> gt_diff t s} \<union> {(s, t). ground t \<and> gt_same t s}"
by auto
obtain k2 where bad_same: "inf_chain (\<lambda>t s. ground t \<and> gt_same t s) (\<lambda>i. ?ff (i + k2))"
using wf_infinite_down_chain_compatible[OF wf_diff _ diff_O_same, of "\<lambda>i. ?ff (i + k1)"]
bad_diff_same
unfolding inf_chain_def diff_same_as_union[symmetric] by (auto simp: add.assoc)
hence hd_sym: "\<And>i. is_Sym (head (?ff (i + k2)))"
unfolding inf_chain_def by (simp add: ground_head)
define f where "f = sym (head (?ff k2))"
have hd_eq_f: "head (?ff (i + k2)) = Sym f" for i
unfolding f_def
proof (induct i)
case 0
thus ?case
by (auto simp: hd.collapse(2)[OF hd_sym, of 0, simplified])
next
case (Suc ia)
thus ?case
using bad_same unfolding inf_chain_def gt_same.simps by simp
qed
define max_args where "max_args = wt (?ff k2)"
have wt_eq_max_args: "wt (?ff (i + k2)) = max_args" for i
unfolding max_args_def
proof (induct i)
case (Suc ia)
thus ?case
using bad_same unfolding inf_chain_def gt_same.simps by simp
qed auto
have nargs_le_max_args: "num_args (?ff (i + k2)) \<le> max_args" for i
unfolding wt_eq_max_args[of i, symmetric] by (rule wt_ge_num_args[OF ffi_wary])
let ?U_of = "\<lambda>i. set (args (?ff (i + k2)))"
define U where "U = (\<Union>i. ?U_of i)"
have gr_u: "\<And>u. u \<in> U \<Longrightarrow> ground u"
unfolding U_def by (blast dest: ground_args[OF _ ffi_ground])
have wary_u: "\<And>u. u \<in> U \<Longrightarrow> wary u"
unfolding U_def by (blast dest: wary_args[OF _ ffi_wary])
have "\<not> bad (>\<^sub>t\<^sub>w\<^sub>g) u" if u_in: "u \<in> ?U_of i" for u i
proof
assume u_bad: "bad (>\<^sub>t\<^sub>w\<^sub>g) u"
have sz_u: "size u < size (?ff (i + k2))"
by (rule size_in_args[OF u_in])
show False
proof (cases "i + k2")
case 0
thus False
using sz_u min_worst_chain_0[OF wf_sz u_bad] by simp
next
case Suc
hence gt: "?ff (i + k2 - 1) >\<^sub>t\<^sub>w ?ff (i + k2)"
using worst_chain_pred[OF wf_sz t_bad] by auto
moreover have "?ff (i + k2) >\<^sub>t\<^sub>w u"
using gt gt_proper_sub sub_args sz_u u_in wary_args by auto
ultimately have "?ff (i + k2 - 1) >\<^sub>t\<^sub>w u"
using gt_trans by blast
thus False
using Suc sz_u min_worst_chain_Suc[OF wf_sz u_bad] ffi_ground by fastforce
qed
qed
hence u_good: "\<And>u. u \<in> U \<Longrightarrow> \<not> bad (>\<^sub>t\<^sub>w\<^sub>g) u"
unfolding U_def by blast
let ?gtwu = "\<lambda>t s. t \<in> U \<and> t >\<^sub>t\<^sub>w s"
have gtwu_irrefl: "\<And>x. \<not> ?gtwu x x"
using gt_irrefl by auto
have "\<And>i j. \<forall>t \<in> set (args (?ff (i + k2))). \<forall>s \<in> set (args (?ff (j + k2))). t >\<^sub>t s \<longrightarrow>
t \<in> U \<and> t >\<^sub>t\<^sub>w s"
using wary_u unfolding U_def by blast
moreover have "\<And>i. extf f (>\<^sub>t) (args (?ff (i + k2))) (args (?ff (Suc i + k2)))"
using bad_same hd_eq_f unfolding inf_chain_def gt_same.simps by auto
ultimately have "\<And>i. extf f ?gtwu (args (?ff (i + k2))) (args (?ff (Suc i + k2)))"
by (rule extf_mono_strong)
hence "inf_chain (extf f ?gtwu) (\<lambda>i. args (?ff (i + k2)))"
unfolding inf_chain_def by blast
hence nwf_ext:
"\<not> wfP (\<lambda>xs ys. length ys \<le> max_args \<and> length xs \<le> max_args \<and> extf f ?gtwu ys xs)"
unfolding inf_chain_def wfP_def wf_iff_no_infinite_down_chain using nargs_le_max_args by fast
have gtwu_le_gtwg: "?gtwu \<le> (>\<^sub>t\<^sub>w\<^sub>g)"
by (auto intro!: gr_u)
have "wfP (\<lambda>s t. ?gtwu t s)"
unfolding wfP_iff_no_inf_chain
proof (intro notI, elim exE)
fix f
assume bad_f: "inf_chain ?gtwu f"
hence bad_f0: "bad ?gtwu (f 0)"
by (rule inf_chain_bad)
hence "f 0 \<in> U"
using bad_f unfolding inf_chain_def by blast
hence "\<not> bad (>\<^sub>t\<^sub>w\<^sub>g) (f 0)"
using u_good by blast
hence "\<not> bad ?gtwu (f 0)"
using bad_f inf_chain_bad inf_chain_subset[OF _ gtwu_le_gtwg] by blast
thus False
using bad_f0 by sat
qed
hence wf_ext: "wfP (\<lambda>xs ys. length ys \<le> max_args \<and> length xs \<le> max_args \<and> extf f ?gtwu ys xs)"
using extf_wf_bounded[of ?gtwu] gtwu_irrefl by blast
show False
using nwf_ext wf_ext by blast
qed
let ?subst = "subst grounding_\<rho>"
have "wfP (\<lambda>s t. ?subst t >\<^sub>t\<^sub>w\<^sub>g ?subst s)"
by (rule wfP_app[OF ground_wfP])
hence "wfP (\<lambda>s t. ?subst t >\<^sub>t\<^sub>w ?subst s)"
by (simp add: ground_grounding_\<rho>)
thus ?thesis
by (auto intro: wfP_subset wary_subst_wary[OF wary_grounding_\<rho>] gt_subst[OF wary_grounding_\<rho>])
qed
end
end
diff --git a/thys/LatticeProperties/Complete_Lattice_Prop.thy b/thys/LatticeProperties/Complete_Lattice_Prop.thy
--- a/thys/LatticeProperties/Complete_Lattice_Prop.thy
+++ b/thys/LatticeProperties/Complete_Lattice_Prop.thy
@@ -1,124 +1,109 @@
section \<open>Fixpoints and Complete Lattices\<close>
(*
Author: Viorel Preoteasa
*)
theory Complete_Lattice_Prop
imports WellFoundedTransitive
begin
text\<open>
This theory introduces some results about fixpoints of functions on
complete lattices. The main result is that a monotonic function
mapping momotonic functions to monotonic functions has the least
fixpoint monotonic.
\<close>
context complete_lattice begin
lemma inf_Inf: assumes nonempty: "A \<noteq> {}"
- shows "inf x (Inf A) = Inf ((inf x) ` A)"
- apply (rule antisym, simp_all)
- apply (rule INF_greatest)
- apply (safe, simp)
- apply (rule_tac y = "Inf A" in order_trans, simp_all)
- apply (rule Inf_lower, simp)
- apply (cut_tac nonempty)
- apply safe
- apply (erule notE)
- apply (rule_tac y = "inf x xa" in order_trans)
- apply (rule INF_lower, blast)
- apply simp
- apply (rule Inf_greatest)
- apply (rule_tac y = "inf x xa" in order_trans)
- apply (rule INF_lower)
- apply (simp add: image_def)
- by auto
+ shows "inf x (Inf A) = Inf ((inf x) ` A)"
+ using assms by (auto simp add: INF_inf_const1 nonempty)
end
(*
Monotonic applications which map monotonic to monotonic have monotonic fixpoints
*)
definition
"mono_mono F = (mono F \<and> (\<forall> f . mono f \<longrightarrow> mono (F f)))"
theorem lfp_mono [simp]:
"mono_mono F \<Longrightarrow> mono (lfp F)"
apply (simp add: mono_mono_def)
apply (rule_tac f="F" and P = "mono" in lfp_ordinal_induct)
apply (simp_all add: mono_def)
apply (intro allI impI SUP_least)
apply (rule_tac y = "f y" in order_trans)
apply (auto intro: SUP_upper)
done
lemma gfp_ordinal_induct:
fixes f :: "'a::complete_lattice => 'a"
assumes mono: "mono f"
and P_f: "!!S. P S ==> P (f S)"
and P_Union: "!!M. \<forall>S\<in>M. P S ==> P (Inf M)"
shows "P (gfp f)"
proof -
let ?M = "{S. gfp f \<le> S \<and> P S}"
have "P (Inf ?M)" using P_Union by simp
also have "Inf ?M = gfp f"
proof (rule antisym)
show "gfp f \<le> Inf ?M" by (blast intro: Inf_greatest)
hence "f (gfp f) \<le> f (Inf ?M)" by (rule mono [THEN monoD])
hence "gfp f \<le> f (Inf ?M)" using mono [THEN gfp_unfold] by simp
hence "f (Inf ?M) \<in> ?M" using P_f P_Union by simp
hence "Inf ?M \<le> f (Inf ?M)" by (rule Inf_lower)
thus "Inf ?M \<le> gfp f" by (rule gfp_upperbound)
qed
finally show ?thesis .
qed
theorem gfp_mono [simp]:
"mono_mono F \<Longrightarrow> mono (gfp F)"
apply (simp add: mono_mono_def)
apply (rule_tac f="F" and P = "mono" in gfp_ordinal_induct)
apply (simp_all, safe)
apply (simp_all add: mono_def)
apply (intro allI impI INF_greatest)
apply (rule_tac y = "f x" in order_trans)
apply (auto intro: INF_lower)
done
context complete_lattice begin
definition
"Sup_less x (w::'b::well_founded) = Sup {y ::'a . \<exists> v < w . y = x v}"
lemma Sup_less_upper:
"v < w \<Longrightarrow> P v \<le> Sup_less P w"
by (simp add: Sup_less_def, rule Sup_upper, blast)
lemma Sup_less_least:
"(!! v . v < w \<Longrightarrow> P v \<le> Q) \<Longrightarrow> Sup_less P w \<le> Q"
by (simp add: Sup_less_def, rule Sup_least, blast)
end
lemma Sup_less_fun_eq:
"((Sup_less P w) i) = (Sup_less (\<lambda> v . P v i)) w"
apply (simp add: Sup_less_def fun_eq_iff)
apply (rule arg_cong [of _ _ Sup])
apply auto
done
theorem fp_wf_induction:
"f x = x \<Longrightarrow> mono f \<Longrightarrow> (\<forall> w . (y w) \<le> f (Sup_less y w)) \<Longrightarrow> Sup (range y) \<le> x"
apply (rule Sup_least)
apply (simp add: image_def, safe, simp)
apply (rule less_induct1, simp_all)
apply (rule_tac y = "f (Sup_less y xa)" in order_trans, simp)
apply (drule_tac x = "Sup_less y xa" and y = "x" in monoD)
by (simp add: Sup_less_least, auto)
end
diff --git a/thys/LatticeProperties/Lattice_Ordered_Group.thy b/thys/LatticeProperties/Lattice_Ordered_Group.thy
--- a/thys/LatticeProperties/Lattice_Ordered_Group.thy
+++ b/thys/LatticeProperties/Lattice_Ordered_Group.thy
@@ -1,192 +1,192 @@
section\<open>Lattice Orderd Groups\<close>
(*
Author: Viorel Preoteasa
*)
theory Lattice_Ordered_Group
imports Modular_Distrib_Lattice
begin
text\<open>
This theory introduces lattice ordered groups \cite{birkhoff:1942}
and proves some results about them. The most important result is
that a lattice ordered group is also a distributive lattice.
\<close>
class lgroup = group_add + lattice +
assumes add_order_preserving: "a \<le> b \<Longrightarrow> u + a + v \<le> u + b + v"
begin
lemma add_order_preserving_left: "a \<le> b \<Longrightarrow> u + a \<le> u + b"
apply (cut_tac a = a and b = b and u = u and v = 0 in add_order_preserving)
by simp_all
lemma add_order_preserving_right: "a \<le> b \<Longrightarrow> a + v \<le> b + v"
apply (cut_tac a = a and b = b and u = 0 and v = v in add_order_preserving)
by simp_all
lemma minus_order: "-a \<le> -b \<Longrightarrow> b \<le> a"
apply (cut_tac a = "-a" and b = "-b" and u = a and v = b in add_order_preserving)
by simp_all
lemma right_move_to_left: "a + - c \<le> b \<Longrightarrow> a \<le> b + c"
apply (drule_tac v = "c" in add_order_preserving_right)
by (simp add: add.assoc)
lemma right_move_to_right: "a \<le> b + -c \<Longrightarrow> a + c \<le> b"
apply (drule_tac v = "c" in add_order_preserving_right)
by (simp add: add.assoc)
lemma [simp]: "(a \<sqinter> b) + c = (a + c) \<sqinter> (b + c)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply safe
apply (rule add_order_preserving_right)
apply simp
apply (rule add_order_preserving_right)
apply simp
apply (rule right_move_to_left)
apply simp
apply safe
apply (simp_all only: diff_conv_add_uminus)
apply (rule right_move_to_right)
apply simp
apply (rule right_move_to_right)
by simp
lemma [simp]: "(a \<sqinter> b) - c = (a - c) \<sqinter> (b - c)"
by (simp add: diff_conv_add_uminus del: add_uminus_conv_diff)
lemma left_move_to_left: "-c + a \<le> b \<Longrightarrow> a \<le> c + b"
apply (drule_tac u = "c" in add_order_preserving_left)
by (simp add: add.assoc [THEN sym])
lemma left_move_to_right: "a \<le> - c + b \<Longrightarrow> c + a \<le> b"
apply (drule_tac u = "c" in add_order_preserving_left)
by (simp add: add.assoc [THEN sym])
lemma [simp]: "c + (a \<sqinter> b) = (c + a) \<sqinter> (c + b)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply safe
apply (rule add_order_preserving_left)
apply simp
apply (rule add_order_preserving_left)
apply simp
apply (rule left_move_to_left)
apply simp
apply safe
apply (rule left_move_to_right)
apply simp
apply (rule left_move_to_right)
by simp
lemma [simp]: "- (a \<sqinter> b) = (- a) \<squnion> (- b)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule minus_order)
apply simp
apply safe
apply (rule minus_order)
apply simp
apply (rule minus_order)
apply simp
apply simp
apply safe
apply (rule minus_order)
apply simp
apply (rule minus_order)
by simp
lemma [simp]: "(a \<squnion> b) + c = (a + c) \<squnion> (b + c)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule right_move_to_right)
apply simp
apply safe
apply (simp_all only: diff_conv_add_uminus)
apply (rule right_move_to_left)
apply simp
apply (rule right_move_to_left)
apply simp
apply simp
apply safe
apply (rule add_order_preserving_right)
apply simp
apply (rule add_order_preserving_right)
by simp
lemma [simp]: "c + (a \<squnion> b) = (c + a) \<squnion> (c + b)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule left_move_to_right)
apply simp
apply safe
apply (rule left_move_to_left)
apply simp
apply (rule left_move_to_left)
apply simp
apply simp
apply safe
apply (rule add_order_preserving_left)
apply simp
apply (rule add_order_preserving_left)
by simp
lemma [simp]: "c - (a \<sqinter> b) = (c - a) \<squnion> (c - b)"
by (simp add: diff_conv_add_uminus del: add_uminus_conv_diff)
lemma [simp]: "(a \<squnion> b) - c = (a - c) \<squnion> (b - c)"
by (simp add: diff_conv_add_uminus del: add_uminus_conv_diff)
lemma [simp]: "- (a \<squnion> b) = (- a) \<sqinter> (- b)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply safe
apply (rule minus_order)
apply simp
apply (rule minus_order)
apply simp
apply (rule minus_order)
by simp
lemma [simp]: "c - (a \<squnion> b) = (c - a) \<sqinter> (c - b)"
by (simp add: diff_conv_add_uminus del: add_uminus_conv_diff)
lemma add_pos: "0 \<le> a \<Longrightarrow> b \<le> b + a"
apply (cut_tac a = 0 and b = a and u = b and v = 0 in add_order_preserving)
by simp_all
lemma add_pos_left: "0 \<le> a \<Longrightarrow> b \<le> a + b"
apply (rule right_move_to_left)
by simp
lemma inf_sup: "a - (a \<sqinter> b) + b = a \<squnion> b"
by (simp add: add.assoc sup_commute)
lemma inf_sup_2: "b = (a \<sqinter> b) - a + (a \<squnion> b)"
apply (unfold inf_sup [THEN sym])
proof -
fix a b:: 'a
have "b = (a \<sqinter> b) + (- a + a) + - (a \<sqinter> b) + b" by (simp only: right_minus left_minus add_0_right add_0_left)
also have "\<dots> = (a \<sqinter> b) + - a + (a + - (a \<sqinter> b) + b)" by (unfold add.assoc, simp)
also have "\<dots> = (a \<sqinter> b) - a + (a - (a \<sqinter> b) + b)" by simp
finally show "b = (a \<sqinter> b) - a + (a - (a \<sqinter> b) + b)" .
qed
subclass inf_sup_eq_lattice
proof
fix x y z:: 'a
assume A: "x \<sqinter> z = y \<sqinter> z"
assume B: "x \<squnion> z = y \<squnion> z"
have "x = (z \<sqinter> x) - z + (z \<squnion> x)" by (rule inf_sup_2)
also have "\<dots> = (z \<sqinter> y) - z + (z \<squnion> y)" by (simp add: sup_commute inf_commute A B)
also have "\<dots> = y" by (simp only: inf_sup_2 [THEN sym])
finally show "x = y" .
qed
end
end
diff --git a/thys/LatticeProperties/Modular_Distrib_Lattice.thy b/thys/LatticeProperties/Modular_Distrib_Lattice.thy
--- a/thys/LatticeProperties/Modular_Distrib_Lattice.thy
+++ b/thys/LatticeProperties/Modular_Distrib_Lattice.thy
@@ -1,455 +1,430 @@
section\<open>Modular and Distributive Lattices\<close>
(*
Author: Viorel Preoteasa
*)
theory Modular_Distrib_Lattice
imports Lattice_Prop
begin
text \<open>
The main result of this theory is the fact that a lattice is distributive
if and only if it satisfies the following property:
\<close>
term "(\<forall> x y z . x \<sqinter> z = y \<sqinter> z \<and> x \<squnion> z = y \<squnion> z \<Longrightarrow> x = y)"
text\<open>
This result was proved by Bergmann in \cite{bergmann:1929}. The formalization
presented here is based on \cite{birkhoff:1967,burris:sankappanavar:1981}.
\<close>
class modular_lattice = lattice +
assumes modular: "x \<le> y \<Longrightarrow> x \<squnion> (y \<sqinter> z) = y \<sqinter> (x \<squnion> z)"
context distrib_lattice begin
subclass modular_lattice
apply unfold_locales
by (simp add: inf_sup_distrib inf_absorb2)
end
context lattice begin
definition
"d_aux a b c = (a \<sqinter> b) \<squnion> (b \<sqinter> c) \<squnion> (c \<sqinter> a)"
lemma d_b_c_a: "d_aux b c a = d_aux a b c"
by (metis d_aux_def sup.assoc sup_commute)
lemma d_c_a_b: "d_aux c a b = d_aux a b c"
by (metis d_aux_def sup.assoc sup_commute)
definition
"e_aux a b c = (a \<squnion> b) \<sqinter> (b \<squnion> c) \<sqinter> (c \<squnion> a)"
lemma e_b_c_a: "e_aux b c a = e_aux a b c"
- apply (simp add: e_aux_def)
- apply (rule antisym)
- by (simp_all add: sup_commute)
+ by (simp add: e_aux_def ac_simps)
lemma e_c_a_b: "e_aux c a b = e_aux a b c"
- apply (simp add: e_aux_def)
- apply (rule antisym)
- by (simp_all add: sup_commute)
+ by (simp add: e_aux_def ac_simps)
definition
"a_aux a b c = (a \<sqinter> (e_aux a b c)) \<squnion> (d_aux a b c)"
definition
"b_aux a b c = (b \<sqinter> (e_aux a b c)) \<squnion> (d_aux a b c)"
definition
"c_aux a b c = (c \<sqinter> (e_aux a b c)) \<squnion> (d_aux a b c)"
lemma b_a: "b_aux a b c = a_aux b c a"
by (simp add: a_aux_def b_aux_def e_b_c_a d_b_c_a)
lemma c_a: "c_aux a b c = a_aux c a b"
by (simp add: a_aux_def c_aux_def e_c_a_b d_c_a_b)
lemma [simp]: "a_aux a b c \<le> e_aux a b c"
apply (simp add: a_aux_def e_aux_def d_aux_def)
apply (rule_tac y = "(a \<squnion> b) \<sqinter> (b \<squnion> c) \<sqinter> (c \<squnion> a)" in order_trans)
apply (rule inf_le2)
by simp
lemma [simp]: "b_aux a b c \<le> e_aux a b c"
apply (unfold b_a)
apply (subst e_b_c_a [THEN sym])
by simp
lemma [simp]: "c_aux a b c \<le> e_aux a b c"
apply (unfold c_a)
apply (subst e_c_a_b [THEN sym])
by simp
lemma [simp]: "d_aux a b c \<le> a_aux a b c"
by (simp add: a_aux_def e_aux_def d_aux_def)
lemma [simp]: "d_aux a b c \<le> b_aux a b c"
apply (unfold b_a)
apply (subst d_b_c_a [THEN sym])
by simp
lemma [simp]: "d_aux a b c \<le> c_aux a b c"
apply (unfold c_a)
apply (subst d_c_a_b [THEN sym])
by simp
lemma a_meet_e: "a \<sqinter> (e_aux a b c) = a \<sqinter> (b \<squnion> c)"
- apply (simp add: e_aux_def)
- apply (rule antisym)
- apply simp_all
- apply (rule_tac y = "(a \<squnion> b) \<sqinter> (b \<squnion> c) \<sqinter> (c \<squnion> a)" in order_trans)
- apply (rule inf_le2)
- by simp
+ by (rule order.antisym) (simp_all add: e_aux_def le_infI2)
lemma b_meet_e: "b \<sqinter> (e_aux a b c) = b \<sqinter> (c \<squnion> a)"
by (simp add: a_meet_e [THEN sym] e_b_c_a)
lemma c_meet_e: "c \<sqinter> (e_aux a b c) = c \<sqinter> (a \<squnion> b)"
by (simp add: a_meet_e [THEN sym] e_c_a_b)
lemma a_join_d: "a \<squnion> d_aux a b c = a \<squnion> (b \<sqinter> c)"
- apply (simp add: d_aux_def)
- apply (rule antisym)
- apply simp_all
- apply (rule_tac y = "a \<sqinter> b \<squnion> b \<sqinter> c \<squnion> c \<sqinter> a" in order_trans)
- by simp_all
+ by (rule order.antisym) (simp_all add: d_aux_def le_supI2)
lemma b_join_d: "b \<squnion> d_aux a b c = b \<squnion> (c \<sqinter> a)"
by (simp add: a_join_d [THEN sym] d_b_c_a)
end
context lattice begin
definition
"no_distrib a b c = (a \<sqinter> b \<squnion> c \<sqinter> a < a \<sqinter> (b \<squnion> c))"
definition
"incomp x y = (\<not> x \<le> y \<and> \<not> y \<le> x)"
definition
"N5_lattice a b c = (a \<sqinter> c = b \<sqinter> c \<and> a < b \<and> a \<squnion> c = b \<squnion> c)"
definition
"M5_lattice a b c = (a \<sqinter> b = b \<sqinter> c \<and> c \<sqinter> a = b \<sqinter> c \<and> a \<squnion> b = b \<squnion> c \<and> c \<squnion> a = b \<squnion> c \<and> a \<sqinter> b < a \<squnion> b)"
lemma M5_lattice_incomp: "M5_lattice a b c \<Longrightarrow> incomp a b"
apply (simp add: M5_lattice_def incomp_def)
apply safe
apply (simp_all add: inf_absorb1 inf_absorb2 )
apply (simp_all add: sup_absorb1 sup_absorb2 )
apply (subgoal_tac "c \<sqinter> (b \<squnion> c) = c")
apply simp
apply (subst sup_commute)
by simp
end
context modular_lattice begin
lemma a_meet_d: "a \<sqinter> (d_aux a b c) = (a \<sqinter> b) \<squnion> (c \<sqinter> a)"
proof -
have "a \<sqinter> (d_aux a b c) = a \<sqinter> ((a \<sqinter> b) \<squnion> (b \<sqinter> c) \<squnion> (c \<sqinter> a))" by (simp add: d_aux_def)
also have "... = a \<sqinter> (a \<sqinter> b \<squnion> c \<sqinter> a \<squnion> b \<sqinter> c)" by (simp add: sup_assoc, simp add: sup_commute)
also have "... = (a \<sqinter> b \<squnion> c \<sqinter> a) \<squnion> (a \<sqinter> (b \<sqinter> c))" by (simp add: modular)
- also have "... = (a \<sqinter> b) \<squnion> (c \<sqinter> a)" by (rule antisym, simp_all, rule_tac y = "a \<sqinter> b" in order_trans, simp_all)
+ also have "... = (a \<sqinter> b) \<squnion> (c \<sqinter> a)" by (rule order.antisym, simp_all, rule_tac y = "a \<sqinter> b" in order_trans, simp_all)
finally show ?thesis by simp
qed
lemma b_meet_d: "b \<sqinter> (d_aux a b c) = (b \<sqinter> c) \<squnion> (a \<sqinter> b)"
by (simp add: a_meet_d [THEN sym] d_b_c_a)
lemma c_meet_d: "c \<sqinter> (d_aux a b c) = (c \<sqinter> a) \<squnion> (b \<sqinter> c)"
by (simp add: a_meet_d [THEN sym] d_c_a_b)
lemma d_less_e: "no_distrib a b c \<Longrightarrow> d_aux a b c < e_aux a b c"
apply (subst less_le)
apply(case_tac "d_aux a b c = e_aux a b c")
apply simp_all
apply (simp add: no_distrib_def a_meet_e [THEN sym] a_meet_d [THEN sym])
apply (rule_tac y = "a_aux a b c" in order_trans)
by simp_all
lemma a_meet_b_eq_d: " d_aux a b c \<le> e_aux a b c \<Longrightarrow> a_aux a b c \<sqinter> b_aux a b c = d_aux a b c"
proof -
assume d_less_e: " d_aux a b c \<le> e_aux a b c"
have "(a \<sqinter> e_aux a b c \<squnion> d_aux a b c) \<sqinter> (b \<sqinter> e_aux a b c \<squnion> d_aux a b c) = (b \<sqinter> e_aux a b c \<squnion> d_aux a b c) \<sqinter> (d_aux a b c \<squnion> a \<sqinter> e_aux a b c)"
by (simp add: inf_commute sup_commute)
also have "\<dots> = d_aux a b c \<squnion> ((b \<sqinter> e_aux a b c \<squnion> d_aux a b c) \<sqinter> (a \<sqinter> e_aux a b c))"
by (simp add: modular)
also have "\<dots> = d_aux a b c \<squnion> (d_aux a b c \<squnion> e_aux a b c \<sqinter> b) \<sqinter> (a \<sqinter> e_aux a b c)"
by (simp add: inf_commute sup_commute)
also have "\<dots> = d_aux a b c \<squnion> (e_aux a b c \<sqinter> (d_aux a b c \<squnion> b)) \<sqinter> (a \<sqinter> e_aux a b c)"
by (cut_tac d_less_e, simp add: modular [THEN sym] less_le)
also have "\<dots> = d_aux a b c \<squnion> ((a \<sqinter> e_aux a b c) \<sqinter> (e_aux a b c \<sqinter> (b \<squnion> d_aux a b c)))"
by (simp add: inf_commute sup_commute)
also have "\<dots> = d_aux a b c \<squnion> (a \<sqinter> e_aux a b c \<sqinter> (b \<squnion> d_aux a b c))" by (simp add: inf_assoc)
also have "\<dots> = d_aux a b c \<squnion> (a \<sqinter> e_aux a b c \<sqinter> (b \<squnion> (c \<sqinter> a)))" by (simp add: b_join_d)
also have "\<dots> = d_aux a b c \<squnion> (a \<sqinter> (b \<squnion> c) \<sqinter> (b \<squnion> (c \<sqinter> a)))" by (simp add: a_meet_e)
also have "\<dots> = d_aux a b c \<squnion> (a \<sqinter> ((b \<squnion> c) \<sqinter> (b \<squnion> (c \<sqinter> a))))" by (simp add: inf_assoc)
also have "\<dots> = d_aux a b c \<squnion> (a \<sqinter> (b \<squnion> ((b \<squnion> c) \<sqinter> (c \<sqinter> a))))" by (simp add: modular)
also have "\<dots> = d_aux a b c \<squnion> (a \<sqinter> (b \<squnion> (c \<sqinter> a)))" by (simp add: inf_absorb2)
also have "\<dots> = d_aux a b c \<squnion> (a \<sqinter> ((c \<sqinter> a) \<squnion> b))" by (simp add: sup_commute inf_commute)
also have "\<dots> = d_aux a b c \<squnion> ((c \<sqinter> a) \<squnion> (a \<sqinter> b))" by (simp add: modular)
also have "\<dots> = d_aux a b c"
- by (rule antisym, simp_all add: d_aux_def)
+ by (rule order.antisym, simp_all add: d_aux_def)
finally show ?thesis by (simp add: a_aux_def b_aux_def)
qed
lemma b_meet_c_eq_d: " d_aux a b c \<le> e_aux a b c \<Longrightarrow> b_aux a b c \<sqinter> c_aux a b c = d_aux a b c"
apply (subst b_a)
apply (subgoal_tac "c_aux a b c = b_aux b c a")
apply simp
apply (subst a_meet_b_eq_d)
by (simp_all add: c_aux_def b_aux_def d_b_c_a e_b_c_a)
lemma c_meet_a_eq_d: "d_aux a b c \<le> e_aux a b c \<Longrightarrow> c_aux a b c \<sqinter> a_aux a b c = d_aux a b c"
apply (subst c_a)
apply (subgoal_tac "a_aux a b c = b_aux c a b")
apply simp
apply (subst a_meet_b_eq_d)
by (simp_all add: a_aux_def b_aux_def d_b_c_a e_b_c_a)
lemma a_def_equiv: "d_aux a b c \<le> e_aux a b c \<Longrightarrow> a_aux a b c = (a \<squnion> d_aux a b c) \<sqinter> e_aux a b c"
apply (simp add: a_aux_def)
apply (subst inf_commute)
apply (subst sup_commute)
apply (simp add: modular)
by (simp add: inf_commute sup_commute)
lemma b_def_equiv: "d_aux a b c \<le> e_aux a b c \<Longrightarrow> b_aux a b c = (b \<squnion> d_aux a b c) \<sqinter> e_aux a b c"
apply (cut_tac a = b and b = c and c = a in a_def_equiv)
by (simp_all add: d_b_c_a e_b_c_a b_a)
lemma c_def_equiv: "d_aux a b c \<le> e_aux a b c \<Longrightarrow> c_aux a b c = (c \<squnion> d_aux a b c) \<sqinter> e_aux a b c"
apply (cut_tac a = c and b = a and c = b in a_def_equiv)
by (simp_all add: d_c_a_b e_c_a_b c_a)
lemma a_join_b_eq_e: "d_aux a b c \<le> e_aux a b c \<Longrightarrow> a_aux a b c \<squnion> b_aux a b c = e_aux a b c"
proof -
assume d_less_e: " d_aux a b c \<le> e_aux a b c"
have "((a \<squnion> d_aux a b c) \<sqinter> e_aux a b c) \<squnion> ((b \<squnion> d_aux a b c) \<sqinter> e_aux a b c) = ((b \<squnion> d_aux a b c) \<sqinter> e_aux a b c) \<squnion> (e_aux a b c \<sqinter> (a \<squnion> d_aux a b c))"
by (simp add: inf_commute sup_commute)
also have "\<dots> = e_aux a b c \<sqinter> (((b \<squnion> d_aux a b c) \<sqinter> e_aux a b c) \<squnion> (a \<squnion> d_aux a b c))"
by (simp add: modular)
also have "\<dots> = e_aux a b c \<sqinter> ((e_aux a b c \<sqinter> (d_aux a b c \<squnion> b)) \<squnion> (a \<squnion> d_aux a b c))"
by (simp add: inf_commute sup_commute)
also have "\<dots> = e_aux a b c \<sqinter> ((d_aux a b c \<squnion> (e_aux a b c \<sqinter> b)) \<squnion> (a \<squnion> d_aux a b c))"
by (cut_tac d_less_e, simp add: modular)
also have "\<dots> = e_aux a b c \<sqinter> ((a \<squnion> d_aux a b c) \<squnion> (d_aux a b c \<squnion> (b \<sqinter> e_aux a b c)))"
by (simp add: inf_commute sup_commute)
also have "\<dots> = e_aux a b c \<sqinter> (a \<squnion> d_aux a b c \<squnion> (b \<sqinter> e_aux a b c))" by (simp add: sup_assoc)
also have "\<dots> = e_aux a b c \<sqinter> (a \<squnion> d_aux a b c \<squnion> (b \<sqinter> (c \<squnion> a)))" by (simp add: b_meet_e)
also have "\<dots> = e_aux a b c \<sqinter> (a \<squnion> (b \<sqinter> c) \<squnion> (b \<sqinter> (c \<squnion> a)))" by (simp add: a_join_d)
also have "\<dots> = e_aux a b c \<sqinter> (a \<squnion> ((b \<sqinter> c) \<squnion> (b \<sqinter> (c \<squnion> a))))" by (simp add: sup_assoc)
also have "\<dots> = e_aux a b c \<sqinter> (a \<squnion> (b \<sqinter> ((b \<sqinter> c) \<squnion> (c \<squnion> a))))" by (simp add: modular)
also have "\<dots> = e_aux a b c \<sqinter> (a \<squnion> (b \<sqinter> (c \<squnion> a)))" by (simp add: sup_absorb2)
also have "\<dots> = e_aux a b c \<sqinter> (a \<squnion> ((c \<squnion> a) \<sqinter> b))" by (simp add: sup_commute inf_commute)
also have "\<dots> = e_aux a b c \<sqinter> ((c \<squnion> a) \<sqinter> (a \<squnion> b))" by (simp add: modular)
also have "\<dots> = e_aux a b c"
- by (rule antisym, simp_all, simp_all add: e_aux_def)
+ by (rule order.antisym, simp_all, simp_all add: e_aux_def)
finally show ?thesis by (cut_tac d_less_e, simp add: a_def_equiv b_def_equiv)
qed
lemma b_join_c_eq_e: " d_aux a b c <= e_aux a b c \<Longrightarrow> b_aux a b c \<squnion> c_aux a b c = e_aux a b c"
apply (subst b_a)
apply (subgoal_tac "c_aux a b c = b_aux b c a")
apply simp
apply (subst a_join_b_eq_e)
by (simp_all add: c_aux_def b_aux_def d_b_c_a e_b_c_a)
lemma c_join_a_eq_e: "d_aux a b c <= e_aux a b c \<Longrightarrow> c_aux a b c \<squnion> a_aux a b c = e_aux a b c"
apply (subst c_a)
apply (subgoal_tac "a_aux a b c = b_aux c a b")
apply simp
apply (subst a_join_b_eq_e)
by (simp_all add: a_aux_def b_aux_def d_b_c_a e_b_c_a)
lemma "no_distrib a b c \<Longrightarrow> incomp a b"
- apply (simp add: no_distrib_def incomp_def)
- apply safe
- apply (simp add: inf_absorb1)
- apply (subgoal_tac "a \<squnion> c \<sqinter> a = a \<and> a \<sqinter> (b \<squnion> c) = a")
- apply simp
- apply safe
- apply (rule antisym)
- apply simp
- apply simp
- apply (rule antisym)
- apply simp_all
- apply (rule_tac y = b in order_trans)
- apply simp_all
- apply (simp add: inf_absorb2)
- apply (unfold modular [THEN sym])
- by (simp add: inf_commute)
+ apply (simp add: no_distrib_def incomp_def ac_simps)
+ using order.strict_iff_not inf.absorb_iff2 inf.commute modular
+ apply fastforce
+ done
lemma M5_modular: "no_distrib a b c \<Longrightarrow> M5_lattice (a_aux a b c) (b_aux a b c) (c_aux a b c)"
apply (frule d_less_e)
by (simp add: M5_lattice_def a_meet_b_eq_d b_meet_c_eq_d c_meet_a_eq_d a_join_b_eq_e b_join_c_eq_e c_join_a_eq_e)
lemma M5_modular_def: "M5_lattice a b c = (a \<sqinter> b = b \<sqinter> c \<and> c \<sqinter> a = b \<sqinter> c \<and> a \<squnion> b = b \<squnion> c \<and> c \<squnion> a = b \<squnion> c \<and> a \<sqinter> b < a \<squnion> b)"
by (simp add: M5_lattice_def)
end
context lattice begin
lemma not_modular_N5: "(\<not> class.modular_lattice inf ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) sup) =
(\<exists> a b c::'a . N5_lattice a b c)"
apply (subgoal_tac "class.lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) sup")
apply (unfold N5_lattice_def class.modular_lattice_def class.modular_lattice_axioms_def)
apply simp_all
apply safe
apply (subgoal_tac "x \<squnion> y \<sqinter> z < y \<sqinter> (x \<squnion> z)")
apply (rule_tac x = "x \<squnion> y \<sqinter> z" in exI)
apply (rule_tac x = "y \<sqinter> (x \<squnion> z)" in exI)
apply (rule_tac x = z in exI)
apply safe
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (rule_tac y = "x \<squnion> y \<sqinter> z" in order_trans)
apply simp_all
apply (rule_tac y = "y \<sqinter> z" in order_trans)
apply simp_all
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (rule_tac y = "y \<sqinter> (x \<squnion> z)" in order_trans)
apply simp_all
apply (rule_tac y = "x \<squnion> z" in order_trans)
apply simp_all
apply (rule neq_le_trans)
apply simp
apply simp
apply (rule_tac x = a in exI)
apply (rule_tac x = b in exI)
apply safe
apply (simp add: less_le)
apply (rule_tac x = c in exI)
apply simp
apply (simp add: less_le)
apply safe
apply (subgoal_tac "a \<squnion> a \<sqinter> c = b")
apply (unfold sup_inf_absorb) [1]
apply simp
apply simp
proof qed
lemma not_distrib_N5_M5: "(\<not> class.distrib_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)) =
((\<exists> a b c::'a . N5_lattice a b c) \<or> (\<exists> a b c::'a . M5_lattice a b c))"
apply (unfold not_modular_N5 [THEN sym])
proof
assume A: "\<not> class.distrib_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)"
have B: "\<exists> a b c:: 'a . (a \<sqinter> b) \<squnion> (a \<sqinter> c) < a \<sqinter> (b \<squnion> c)"
apply (cut_tac A)
apply (unfold class.distrib_lattice_def)
apply safe
apply simp_all
proof
fix x y z::'a
assume A: "\<forall>(a::'a) (b::'a) c::'a. \<not> a \<sqinter> b \<squnion> a \<sqinter> c < a \<sqinter> (b \<squnion> c)"
show "x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
apply (cut_tac A)
apply (rule distrib_imp1)
by (simp add: less_le)
qed
from B show "\<not> class.modular_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>) \<or> (\<exists>a b c::'a. M5_lattice a b c)"
proof (unfold disj_not1, safe)
fix a b c::'a
assume A: "a \<sqinter> b \<squnion> a \<sqinter> c < a \<sqinter> (b \<squnion> c)"
assume B: "class.modular_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)"
interpret modular: modular_lattice "(\<sqinter>)" "((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool)" "(<)" "(\<squnion>)"
by (fact B)
have H: "M5_lattice (a_aux a b c) (b_aux a b c) (c_aux a b c)"
apply (cut_tac a = a and b = b and c = c in modular.M5_modular)
apply (unfold no_distrib_def)
by (simp_all add: A inf_commute)
from H show "\<exists>a b c::'a. M5_lattice a b c" by blast
qed
next
assume A: "\<not> class.modular_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>) \<or> (\<exists>(a::'a) (b::'a) c::'a. M5_lattice a b c)"
show "\<not> class.distrib_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)"
apply (cut_tac A)
apply safe
apply (erule notE)
apply unfold_locales
apply (unfold class.distrib_lattice_def)
apply (unfold class.distrib_lattice_axioms_def)
apply safe
apply (simp add: sup_absorb2)
apply (frule M5_lattice_incomp)
apply (unfold M5_lattice_def)
apply (drule_tac x = a in spec)
apply (drule_tac x = b in spec)
apply (drule_tac x = c in spec)
apply safe
proof -
fix a b c:: 'a
assume A:"a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)"
assume B: "a \<sqinter> b = b \<sqinter> c"
assume D: "a \<squnion> b = b \<squnion> c"
assume E: "c \<squnion> a = b \<squnion> c"
assume G: "incomp a b"
have H: "a \<squnion> b \<sqinter> c = a" by (simp add: B [THEN sym] sup_absorb1)
have I: "(a \<squnion> b) \<sqinter> (a \<squnion> c) = a \<squnion> b" by (cut_tac E, simp add: sup_commute D)
have J: "a = a \<squnion> b" by (cut_tac A, simp add: H I)
show False
apply (cut_tac G J)
apply (subgoal_tac "b \<le> a")
apply (simp add: incomp_def)
apply (rule_tac y = "a \<squnion> b" in order_trans)
apply (rule sup_ge2)
by simp
qed
qed
lemma distrib_not_N5_M5: "(class.distrib_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)) =
((\<forall> a b c::'a . \<not> N5_lattice a b c) \<and> (\<forall> a b c::'a . \<not> M5_lattice a b c))"
apply (cut_tac not_distrib_N5_M5)
by auto
lemma distrib_inf_sup_eq:
"(class.distrib_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)) =
(\<forall> x y z::'a . x \<sqinter> z = y \<sqinter> z \<and> x \<squnion> z = y \<squnion> z \<longrightarrow> x = y)"
apply safe
proof -
fix x y z:: 'a
assume A: "class.distrib_lattice (\<sqinter>) ((\<le>) ::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)"
interpret distrib: distrib_lattice "(\<sqinter>)" "(\<le>) :: 'a \<Rightarrow> 'a \<Rightarrow> bool" "(<)" "(\<squnion>)"
by (fact A)
assume B: "x \<sqinter> z = y \<sqinter> z"
assume C: "x \<squnion> z = y \<squnion> z"
have "x = x \<sqinter> (x \<squnion> z)" by simp
also have "\<dots> = x \<sqinter> (y \<squnion> z)" by (simp add: C)
also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by (simp add: distrib.inf_sup_distrib)
also have "\<dots> = (y \<sqinter> x) \<squnion> (y \<sqinter> z)" by (simp add: B inf_commute)
also have "\<dots> = y \<sqinter> (x \<squnion> z)" by (simp add: distrib.inf_sup_distrib)
also have "\<dots> = y" by (simp add: C)
finally show "x = y" .
next
assume A: "(\<forall>x y z:: 'a. x \<sqinter> z = y \<sqinter> z \<and> x \<squnion> z = y \<squnion> z \<longrightarrow> x = y)"
have B: "!! x y z :: 'a. x \<sqinter> z = y \<sqinter> z \<and> x \<squnion> z = y \<squnion> z \<Longrightarrow> x = y"
by (cut_tac A, blast)
show "class.distrib_lattice (\<sqinter>) ((\<le>)::'a \<Rightarrow> 'a \<Rightarrow> bool) (<) (\<squnion>)"
apply (unfold distrib_not_N5_M5)
apply safe
apply (unfold N5_lattice_def)
apply (cut_tac x = a and y = b and z = c in B)
apply (simp_all)
apply (unfold M5_lattice_def)
apply (cut_tac x = a and y = b and z = c in B)
by (simp_all add: inf_commute sup_commute)
qed
end
class inf_sup_eq_lattice = lattice +
assumes inf_sup_eq: "x \<sqinter> z = y \<sqinter> z \<Longrightarrow> x \<squnion> z = y \<squnion> z \<Longrightarrow> x = y"
begin
subclass distrib_lattice
by (metis distrib_inf_sup_eq inf_sup_eq)
end
end
diff --git a/thys/Markov_Models/Markov_Models_Auxiliary.thy b/thys/Markov_Models/Markov_Models_Auxiliary.thy
--- a/thys/Markov_Models/Markov_Models_Auxiliary.thy
+++ b/thys/Markov_Models/Markov_Models_Auxiliary.thy
@@ -1,602 +1,602 @@
(* Author: Johannes Hölzl <hoelzl@in.tum.de> *)
section \<open>Auxiliary Theory\<close>
text \<open>Parts of it should be moved to the Isabelle repository\<close>
theory Markov_Models_Auxiliary
imports
"HOL-Probability.Probability"
"HOL-Library.Rewrite"
"HOL-Library.Linear_Temporal_Logic_on_Streams"
Coinductive.Coinductive_Stream
Coinductive.Coinductive_Nat
begin
lemma lfp_upperbound: "(\<And>y. x \<le> f y) \<Longrightarrow> x \<le> lfp f"
unfolding lfp_def by (intro Inf_greatest) (auto intro: order_trans)
(* ?? *)
lemma lfp_arg: "(\<lambda>t. lfp (F t)) = lfp (\<lambda>x t. F t (x t))"
apply (auto simp: lfp_def le_fun_def fun_eq_iff intro!: Inf_eqI Inf_greatest)
subgoal for x y
by (rule INF_lower2[of "top(x := y)"]) auto
done
lemma lfp_pair: "lfp (\<lambda>f (a, b). F (\<lambda>a b. f (a, b)) a b) (a, b) = lfp F a b"
unfolding lfp_def
by (auto intro!: INF_eq simp: le_fun_def)
(auto intro!: exI[of _ "\<lambda>(a, b). x a b" for x])
lemma all_Suc_split: "(\<forall>i. P i) \<longleftrightarrow> (P 0 \<and> (\<forall>i. P (Suc i)))"
using nat_induct by auto
definition "with P f d = (if \<exists>x. P x then f (SOME x. P x) else d)"
lemma withI[case_names default exists]:
"((\<And>x. \<not> P x) \<Longrightarrow> Q d) \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q (f x)) \<Longrightarrow> Q (with P f d)"
unfolding with_def by (auto intro: someI2)
context order
begin
definition
"maximal f S = {x\<in>S. \<forall>y\<in>S. f y \<le> f x}"
lemma maximalI: "x \<in> S \<Longrightarrow> (\<And>y. y \<in> S \<Longrightarrow> f y \<le> f x) \<Longrightarrow> x \<in> maximal f S"
by (simp add: maximal_def)
lemma maximalI_trans: "x \<in> maximal f S \<Longrightarrow> f x \<le> f y \<Longrightarrow> y \<in> S \<Longrightarrow> y \<in> maximal f S"
unfolding maximal_def by (blast intro: antisym order_trans)
lemma maximalD1: "x \<in> maximal f S \<Longrightarrow> x \<in> S"
by (simp add: maximal_def)
lemma maximalD2: "x \<in> maximal f S \<Longrightarrow> y \<in> S \<Longrightarrow> f y \<le> f x"
by (simp add: maximal_def)
lemma maximal_inject: "x \<in> maximal f S \<Longrightarrow> y \<in> maximal f S \<Longrightarrow> f x = f y"
- unfolding maximal_def by (blast intro: antisym)
+ by (rule order.antisym) (simp_all add: maximal_def)
lemma maximal_empty[simp]: "maximal f {} = {}"
by (simp add: maximal_def)
lemma maximal_singleton[simp]: "maximal f {x} = {x}"
by (auto simp add: maximal_def)
lemma maximal_in_S: "maximal f S \<subseteq> S"
by (auto simp: maximal_def)
end
context linorder
begin
lemma maximal_ne:
assumes "finite S" "S \<noteq> {}"
shows "maximal f S \<noteq> {}"
using assms
proof (induct rule: finite_ne_induct)
case (insert s S)
show ?case
proof cases
assume "\<forall>x\<in>S. f x \<le> f s"
with insert have "s \<in> maximal f (insert s S)"
by (auto intro!: maximalI)
then show ?thesis
by auto
next
assume "\<not> (\<forall>x\<in>S. f x \<le> f s)"
then have "maximal f (insert s S) = maximal f S"
by (auto simp: maximal_def)
with insert show ?thesis
by auto
qed
qed simp
end
lemma mono_les:
fixes s S N and l1 l2 :: "'a \<Rightarrow> real" and K :: "'a \<Rightarrow> 'a pmf"
defines "\<Delta> x \<equiv> l2 x - l1 x"
assumes s: "s \<in> S" and S: "(\<Union>s\<in>S. set_pmf (K s)) \<subseteq> S \<union> N"
assumes int_l1[simp]: "\<And>s. s \<in> S \<Longrightarrow> integrable (K s) l1"
assumes int_l2[simp]: "\<And>s. s \<in> S \<Longrightarrow> integrable (K s) l2"
assumes to_N: "\<And>s. s \<in> S \<Longrightarrow> \<exists>t\<in>N. (s, t) \<in> (SIGMA s:UNIV. K s)\<^sup>*"
assumes l1: "\<And>s. s \<in> S \<Longrightarrow> (\<integral>t. l1 t \<partial>K s) + c s \<le> l1 s"
assumes l2: "\<And>s. s \<in> S \<Longrightarrow> l2 s \<le> (\<integral>t. l2 t \<partial>K s) + c s"
assumes eq: "\<And>s. s \<in> N \<Longrightarrow> l2 s \<le> l1 s"
assumes finitary: "finite (\<Delta> ` (S\<union>N))"
shows "l2 s \<le> l1 s"
proof -
define M where "M = {s\<in>S\<union>N. \<forall>t\<in>S\<union>N. \<Delta> t \<le> \<Delta> s}"
have [simp]: "\<And>s. s\<in>S \<Longrightarrow> integrable (K s) \<Delta>"
by (simp add: \<Delta>_def[abs_def])
have M_unqiue: "\<And>s t. s \<in> M \<Longrightarrow> t \<in> M \<Longrightarrow> \<Delta> s = \<Delta> t"
by (auto intro!: antisym simp: M_def)
have M1: "\<And>s. s \<in> M \<Longrightarrow> s \<in> S \<union> N"
by (auto simp: M_def)
have M2: "\<And>s t. s \<in> M \<Longrightarrow> t \<in> S \<union> N \<Longrightarrow> \<Delta> t \<le> \<Delta> s"
by (auto simp: M_def)
have M3: "\<And>s t. s \<in> M \<Longrightarrow> t \<in> S \<union> N \<Longrightarrow> t \<notin> M \<Longrightarrow> \<Delta> t < \<Delta> s"
by (auto simp: M_def less_le)
have N: "\<forall>s\<in>N. \<Delta> s \<le> 0"
using eq by (simp add: \<Delta>_def)
{ fix s assume s: "s \<in> M" "M \<inter> N = {}"
then have "s \<in> S - N"
by (auto dest: M1)
with to_N[of s] obtain t where "(s, t) \<in> (SIGMA s:UNIV. K s)\<^sup>*" and "t \<in> N"
by (auto simp: M_def)
from this(1) \<open>s \<in> M\<close> have "\<Delta> s \<le> 0"
proof (induction rule: converse_rtrancl_induct)
case (step s s')
then have s: "s \<in> M" "s \<in> S" "s \<notin> N" and s': "s' \<in> S \<union> N" "s' \<in> K s"
using S \<open>M \<inter> N = {}\<close> by (auto dest: M1)
have "s' \<in> M"
proof (rule ccontr)
assume "s' \<notin> M"
with \<open>s \<in> S\<close> s' \<open>s \<in> M\<close>
have "0 < pmf (K s) s'" "\<Delta> s' < \<Delta> s"
by (auto intro: M2 M3 pmf_positive)
have "\<Delta> s \<le> ((\<integral>t. l2 t \<partial>K s) + c s) - ((\<integral>t. l1 t \<partial>K s) + c s)"
unfolding \<Delta>_def using \<open>s \<in> S\<close> \<open>s \<notin> N\<close> by (intro diff_mono l1 l2) auto
then have "\<Delta> s \<le> (\<integral>s'. \<Delta> s' \<partial>K s)"
using \<open>s \<in> S\<close> by (simp add: \<Delta>_def)
also have "\<dots> < (\<integral>s'. \<Delta> s \<partial>K s)"
using \<open>s' \<in> K s\<close> \<open>\<Delta> s' < \<Delta> s\<close> \<open>s\<in>S\<close> S \<open>s\<in>M\<close>
by (intro measure_pmf.integral_less_AE[where A="{s'}"])
(auto simp: emeasure_measure_pmf_finite AE_measure_pmf_iff set_pmf_iff[symmetric]
intro!: M2)
finally show False
using measure_pmf.prob_space[of "K s"] by simp
qed
with step.IH \<open>t\<in>N\<close> N have "\<Delta> s' \<le> 0" "s' \<in> M"
by auto
with \<open>s\<in>S\<close> show "\<Delta> s \<le> 0"
by (force simp: M_def)
qed (insert N \<open>t\<in>N\<close>, auto) }
show ?thesis
proof cases
assume "M \<inter> N = {}"
have "Max (\<Delta>`(S\<union>N)) \<in> \<Delta>`(S\<union>N)"
using \<open>s \<in> S\<close> by (intro Max_in finitary) auto
then obtain t where "t \<in> S \<union> N" "\<Delta> t = Max (\<Delta>`(S\<union>N))"
unfolding image_iff by metis
then have "t \<in> M"
by (auto simp: M_def finitary intro!: Max_ge)
have "\<Delta> s \<le> \<Delta> t"
using \<open>t\<in>M\<close> \<open>s\<in>S\<close> by (auto dest: M2)
also have "\<Delta> t \<le> 0"
using \<open>t\<in>M\<close> \<open>M \<inter> N = {}\<close> by fact
finally show ?thesis
by (simp add: \<Delta>_def)
next
assume "M \<inter> N \<noteq> {}"
then obtain t where "t \<in> M" "t \<in> N" by auto
with N \<open>s\<in>S\<close> have "\<Delta> s \<le> 0"
by (intro order_trans[of "\<Delta> s" "\<Delta> t" 0]) (auto simp: M_def)
then show ?thesis
by (simp add: \<Delta>_def)
qed
qed
lemma unique_les:
fixes s S N and l1 l2 :: "'a \<Rightarrow> real" and K :: "'a \<Rightarrow> 'a pmf"
defines "\<Delta> x \<equiv> l2 x - l1 x"
assumes s: "s \<in> S" and S: "(\<Union>s\<in>S. set_pmf (K s)) \<subseteq> S \<union> N"
assumes "\<And>s. s \<in> S \<Longrightarrow> integrable (K s) l1"
assumes "\<And>s. s \<in> S \<Longrightarrow> integrable (K s) l2"
assumes "\<And>s. s \<in> S \<Longrightarrow> \<exists>t\<in>N. (s, t) \<in> (SIGMA s:UNIV. K s)\<^sup>*"
assumes "\<And>s. s \<in> S \<Longrightarrow> l1 s = (\<integral>t. l1 t \<partial>K s) + c s"
assumes "\<And>s. s \<in> S \<Longrightarrow> l2 s = (\<integral>t. l2 t \<partial>K s) + c s"
assumes "\<And>s. s \<in> N \<Longrightarrow> l2 s = l1 s"
assumes 1: "finite (\<Delta> ` (S\<union>N))"
shows "l2 s = l1 s"
proof -
have "finite ((\<lambda>x. l2 x - l1 x) ` (S\<union>N))"
using 1 by (auto simp: \<Delta>_def[abs_def])
moreover then have "finite (uminus ` (\<lambda>x. l2 x - l1 x) ` (S\<union>N))"
by auto
ultimately show ?thesis
using assms
by (intro antisym mono_les[of s S K N l2 l1 c] mono_les[of s S K N l1 l2 c])
(auto simp: image_comp comp_def)
qed
lemma inf_continuous_suntil_disj[order_continuous_intros]:
assumes Q: "inf_continuous Q"
assumes disj: "\<And>x \<omega>. \<not> (P \<omega> \<and> Q x \<omega>)"
shows "inf_continuous (\<lambda>x. P suntil Q x)"
unfolding inf_continuous_def
proof (safe intro!: ext)
fix M \<omega> i assume "(P suntil Q (\<Sqinter>i. M i)) \<omega>" "decseq M" then show "(P suntil Q (M i)) \<omega>"
unfolding inf_continuousD[OF Q \<open>decseq M\<close>] by induction (auto intro: suntil.intros)
next
fix M \<omega> assume *: "(\<Sqinter>i. P suntil Q (M i)) \<omega>" "decseq M"
then have "(P suntil Q (M 0)) \<omega>"
by auto
from this * show "(P suntil Q (\<Sqinter>i. M i)) \<omega>"
unfolding inf_continuousD[OF Q \<open>decseq M\<close>]
proof induction
case (base \<omega>) with disj[of \<omega> "M _"] show ?case by (auto intro: suntil.intros elim: suntil.cases)
next
case (step \<omega>) with disj[of \<omega> "M _"] show ?case by (auto intro: suntil.intros elim: suntil.cases)
qed
qed
lemma inf_continuous_nxt[order_continuous_intros]: "inf_continuous P \<Longrightarrow> inf_continuous (\<lambda>x. nxt (P x) \<omega>)"
by (auto simp: inf_continuous_def image_comp)
lemma sup_continuous_nxt[order_continuous_intros]: "sup_continuous P \<Longrightarrow> sup_continuous (\<lambda>x. nxt (P x) \<omega>)"
by (auto simp: sup_continuous_def image_comp)
lemma mcont_ennreal_of_enat: "mcont Sup (\<le>) Sup (\<le>) ennreal_of_enat"
by (auto intro!: mcontI monotoneI contI ennreal_of_enat_Sup)
lemma mcont2mcont_ennreal_of_enat[cont_intro]:
"mcont lub ord Sup (\<le>) f \<Longrightarrow> mcont lub ord Sup (\<le>) (\<lambda>x. ennreal_of_enat (f x))"
by (auto intro: ccpo.mcont2mcont[OF complete_lattice_ccpo'] mcont_ennreal_of_enat)
declare stream.exhaust[cases type: stream]
lemma scount_eq_emeasure: "scount P \<omega> = emeasure (count_space UNIV) {i. P (sdrop i \<omega>)}"
proof cases
assume "alw (ev P) \<omega>"
moreover then have "infinite {i. P (sdrop i \<omega>)}"
using infinite_iff_alw_ev[of P \<omega>] by simp
ultimately show ?thesis
by (simp add: scount_infinite_iff[symmetric])
next
assume "\<not> alw (ev P) \<omega>"
moreover then have "finite {i. P (sdrop i \<omega>)}"
using infinite_iff_alw_ev[of P \<omega>] by simp
ultimately show ?thesis
by (simp add: not_alw_iff not_ev_iff scount_eq_card)
qed
lemma measurable_scount[measurable]:
assumes [measurable]: "Measurable.pred (stream_space M) P"
shows "scount P \<in> measurable (stream_space M) (count_space UNIV)"
unfolding scount_eq[abs_def] by measurable
lemma measurable_sfirst2:
assumes [measurable]: "Measurable.pred (N \<Otimes>\<^sub>M stream_space M) (\<lambda>(x, \<omega>). P x \<omega>)"
shows "(\<lambda>(x, \<omega>). sfirst (P x) \<omega>) \<in> measurable (N \<Otimes>\<^sub>M stream_space M) (count_space UNIV)"
apply (coinduction rule: measurable_enat_coinduct)
apply simp
apply (rule exI[of _ "\<lambda>x. 0"])
apply (rule exI[of _ "\<lambda>(x, \<omega>). (x, stl \<omega>)"])
apply (rule exI[of _ "\<lambda>(x, \<omega>). P x \<omega>"])
apply (subst sfirst.simps[abs_def])
apply (simp add: fun_eq_iff)
done
lemma measurable_sfirst2'[measurable (raw)]:
assumes [measurable (raw)]: "f \<in> N \<rightarrow>\<^sub>M stream_space M" "Measurable.pred (N \<Otimes>\<^sub>M stream_space M) (\<lambda>x. P (fst x) (snd x))"
shows "(\<lambda>x. sfirst (P x) (f x)) \<in> measurable N (count_space UNIV)"
using measurable_sfirst2[measurable] by measurable
lemma measurable_sfirst[measurable]:
assumes [measurable]: "Measurable.pred (stream_space M) P"
shows "sfirst P \<in> measurable (stream_space M) (count_space UNIV)"
by measurable
lemma measurable_epred[measurable]: "epred \<in> count_space UNIV \<rightarrow>\<^sub>M count_space UNIV"
by (rule measurable_count_space)
lemma nn_integral_stretch:
"f \<in> borel \<rightarrow>\<^sub>M borel \<Longrightarrow> c \<noteq> 0 \<Longrightarrow> (\<integral>\<^sup>+x. f (c * x) \<partial>lborel) = (1 / \<bar>c\<bar>::real) * (\<integral>\<^sup>+x. f x \<partial>lborel)"
using nn_integral_real_affine[of f c 0] by (simp add: mult.assoc[symmetric] ennreal_mult[symmetric])
lemma prod_sum_distrib:
fixes f g :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::comm_semiring_1"
assumes "finite I" shows "(\<And>i. i \<in> I \<Longrightarrow> finite (J i)) \<Longrightarrow> (\<Prod>i\<in>I. \<Sum>j\<in>J i. f i j) = (\<Sum>m\<in>Pi\<^sub>E I J. \<Prod>i\<in>I. f i (m i))"
using \<open>finite I\<close>
proof induction
case (insert i I) then show ?case
by (auto simp: PiE_insert_eq finite_PiE sum.reindex inj_combinator sum.swap[of _ "Pi\<^sub>E I J"]
sum_cartesian_product' sum_distrib_left sum_distrib_right
intro!: sum.cong prod.cong arg_cong[where f="(*) x" for x])
qed simp
lemma prod_add_distrib:
fixes f g :: "'a \<Rightarrow> 'b::comm_semiring_1"
assumes "finite I" shows "(\<Prod>i\<in>I. f i + g i) = (\<Sum>J\<in>Pow I. (\<Prod>i\<in>J. f i) * (\<Prod>i\<in>I - J. g i))"
proof -
have "(\<Prod>i\<in>I. f i + g i) = (\<Prod>i\<in>I. \<Sum>b\<in>{True, False}. if b then f i else g i)"
by simp
also have "\<dots> = (\<Sum>m\<in>I \<rightarrow>\<^sub>E {True, False}. \<Prod>i\<in>I. if m i then f i else g i)"
using \<open>finite I\<close> by (rule prod_sum_distrib) simp
also have "\<dots> = (\<Sum>J\<in>Pow I. (\<Prod>i\<in>J. f i) * (\<Prod>i\<in>I - J. g i))"
by (rule sum.reindex_bij_witness[where i="\<lambda>J. \<lambda>i\<in>I. i\<in>J" and j="\<lambda>m. {i\<in>I. m i}"])
(auto simp: fun_eq_iff prod.If_cases \<open>finite I\<close> intro!: arg_cong2[where f="(*)"] prod.cong)
finally show ?thesis .
qed
subclass (in linordered_nonzero_semiring) ordered_semiring_0
proof qed
lemma (in linordered_nonzero_semiring) prod_nonneg: "(\<forall>a\<in>A. 0 \<le> f a) \<Longrightarrow> 0 \<le> prod f A"
by (induct A rule: infinite_finite_induct) simp_all
lemma (in linordered_nonzero_semiring) prod_mono:
"\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i \<Longrightarrow> prod f A \<le> prod g A"
by (induct A rule: infinite_finite_induct) (auto intro!: prod_nonneg mult_mono)
lemma (in linordered_nonzero_semiring) prod_mono2:
assumes "finite J" "I \<subseteq> J" "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> g i \<and> g i \<le> f i" "(\<And>i. i \<in> J - I \<Longrightarrow> 1 \<le> f i)"
shows "prod g I \<le> prod f J"
proof -
have "prod g I = (\<Prod>i\<in>J. if i \<in> I then g i else 1)"
using \<open>finite J\<close> \<open>I \<subseteq> J\<close> by (simp add: prod.If_cases Int_absorb1)
also have "\<dots> \<le> prod f J"
using assms by (intro prod_mono) auto
finally show ?thesis .
qed
lemma (in linordered_nonzero_semiring) prod_mono3:
assumes "finite J" "I \<subseteq> J" "\<And>i. i \<in> J \<Longrightarrow> 0 \<le> g i" "\<And>i. i \<in> I \<Longrightarrow> g i \<le> f i" "(\<And>i. i \<in> J - I \<Longrightarrow> g i \<le> 1)"
shows "prod g J \<le> prod f I"
proof -
have "prod g J \<le> (\<Prod>i\<in>J. if i \<in> I then f i else 1)"
using assms by (intro prod_mono) auto
also have "\<dots> = prod f I"
using \<open>finite J\<close> \<open>I \<subseteq> J\<close> by (simp add: prod.If_cases Int_absorb1)
finally show ?thesis .
qed
lemma (in linordered_nonzero_semiring) one_le_prod: "(\<And>i. i \<in> I \<Longrightarrow> 1 \<le> f i) \<Longrightarrow> 1 \<le> prod f I"
proof (induction I rule: infinite_finite_induct)
case (insert i I) then show ?case
using mult_mono[of 1 "f i" 1 "prod f I"]
by (auto intro: order_trans[OF zero_le_one])
qed auto
lemma sum_plus_one_le_prod_plus_one:
fixes p :: "'a \<Rightarrow> 'b::linordered_nonzero_semiring"
assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> p i"
shows "(\<Sum>i\<in>I. p i) + 1 \<le> (\<Prod>i\<in>I. p i + 1)"
proof cases
assume [simp]: "finite I"
with assms have [simp]: "J \<subseteq> I \<Longrightarrow> 0 \<le> prod p J" for J
by (intro prod_nonneg) auto
have "1 + (\<Sum>i\<in>I. p i) = (\<Sum>J\<in>insert {} ((\<lambda>x. {x})`I). (\<Prod>i\<in>J. p i) * (\<Prod>i\<in>I - J. 1))"
by (subst sum.insert) (auto simp: sum.reindex)
also have "\<dots> \<le> (\<Sum>J\<in>Pow I. (\<Prod>i\<in>J. p i) * (\<Prod>i\<in>I - J. 1))"
using assms by (intro sum_mono2) auto
finally show ?thesis
by (subst prod_add_distrib) (auto simp: add.commute)
qed simp
lemma summable_iff_convergent_prod:
fixes p :: "nat \<Rightarrow> real" assumes p: "\<And>i. 0 \<le> p i"
shows "summable p \<longleftrightarrow> convergent (\<lambda>n. \<Prod>i<n. p i + 1)"
unfolding summable_iff_convergent
proof
assume "convergent (\<lambda>n. \<Prod>i<n. p i + 1)"
then obtain x where x: "(\<lambda>n. \<Prod>i<n. p i + 1) \<longlonglongrightarrow> x"
by (auto simp: convergent_def)
then have "1 \<le> x"
by (rule tendsto_lowerbound) (auto intro!: always_eventually one_le_prod p)
have "convergent (\<lambda>n. 1 + (\<Sum>i<n. p i))"
proof (intro Bseq_mono_convergent BseqI allI)
show "0 < x" using \<open>1 \<le> x\<close> by auto
next
fix n
have "norm ((\<Sum>i<n. p i) + 1) \<le> (\<Prod>i<n. p i + 1)"
using p by (simp add: sum_nonneg sum_plus_one_le_prod_plus_one p)
also have "\<dots> \<le> x"
using assms
by (intro tendsto_lowerbound[OF x])
(auto simp: eventually_sequentially intro!: exI[of _ n] prod_mono2)
finally show "norm (1 + sum p {..<n}) \<le> x"
by (simp add: add.commute)
qed (insert p, auto intro!: sum_mono2)
then show "convergent (\<lambda>n. \<Sum>i<n. p i)"
unfolding convergent_add_const_iff .
next
assume "convergent (\<lambda>n. \<Sum>i<n. p i)"
then obtain x where x: "(\<lambda>n. exp (\<Sum>i<n. p i)) \<longlonglongrightarrow> exp x"
by (force simp: convergent_def intro!: tendsto_exp)
show "convergent (\<lambda>n. \<Prod>i<n. p i + 1)"
proof (intro Bseq_mono_convergent BseqI allI)
show "0 < exp x" by simp
next
fix n
have "norm (\<Prod>i<n. p i + 1) \<le> exp (\<Sum>i<n. p i)"
using p exp_ge_add_one_self[of "p _"] by (auto simp add: prod_nonneg exp_sum add.commute intro!: prod_mono)
also have "\<dots> \<le> exp x"
using p
by (intro tendsto_lowerbound[OF x]) (auto simp: eventually_sequentially intro!: sum_mono2 )
finally show "norm (\<Prod>i<n. p i + 1) \<le> exp x" .
qed (insert p, auto intro!: prod_mono2)
qed
primrec eexp :: "ereal \<Rightarrow> ennreal"
where
"eexp MInfty = 0"
| "eexp (ereal r) = ennreal (exp r)"
| "eexp PInfty = top"
lemma
shows eexp_minus_infty[simp]: "eexp (-\<infinity>) = 0"
and eexp_infty[simp]: "eexp \<infinity> = top"
using eexp.simps by simp_all
lemma eexp_0[simp]: "eexp 0 = 1"
by (simp add: zero_ereal_def)
lemma eexp_inj[simp]: "eexp x = eexp y \<longleftrightarrow> x = y"
by (cases x; cases y; simp)
lemma eexp_mono[simp]: "eexp x \<le> eexp y \<longleftrightarrow> x \<le> y"
by (cases x; cases y; simp add: top_unique)
lemma eexp_strict_mono[simp]: "eexp x < eexp y \<longleftrightarrow> x < y"
by (simp add: less_le)
lemma exp_eq_0_iff[simp]: "eexp x = 0 \<longleftrightarrow> x = -\<infinity>"
using eexp_inj[of x "-\<infinity>"] unfolding eexp_minus_infty .
lemma eexp_surj: "range eexp = UNIV"
proof -
have part: "UNIV = {0} \<union> {0 <..< top} \<union> {top::ennreal}"
by (auto simp: less_top)
show ?thesis
unfolding part
by (force simp: image_iff less_top less_top_ennreal intro!: eexp.simps[symmetric] eexp.simps dest: exp_total)
qed
lemma continuous_on_eexp': "continuous_on UNIV eexp"
by (rule continuous_onI_mono) (auto simp: eexp_surj)
lemma continuous_on_eexp[continuous_intros]: "continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. eexp (f x))"
by (rule continuous_on_compose2[OF continuous_on_eexp']) auto
lemma tendsto_eexp[tendsto_intros]: "(f \<longlongrightarrow> x) F \<Longrightarrow> ((\<lambda>x. eexp (f x)) \<longlongrightarrow> eexp x) F"
by (rule continuous_on_tendsto_compose[OF continuous_on_eexp']) auto
lemma measurable_eexp[measurable]: "eexp \<in> borel \<rightarrow>\<^sub>M borel"
using continuous_on_eexp' by (rule borel_measurable_continuous_onI)
lemma eexp_add: "\<not> ((x = \<infinity> \<and> y = -\<infinity>) \<or> (x = -\<infinity> \<and> y = \<infinity>)) \<Longrightarrow> eexp (x + y) = eexp x * eexp y"
by (cases x; cases y; simp add: exp_add ennreal_mult ennreal_top_mult ennreal_mult_top)
lemma sum_Pinfty:
fixes f :: "'a \<Rightarrow> ereal"
shows "sum f I = \<infinity> \<longleftrightarrow> (finite I \<and> (\<exists>i\<in>I. f i = \<infinity>))"
by (induction I rule: infinite_finite_induct) auto
lemma sum_Minfty:
fixes f :: "'a \<Rightarrow> ereal"
shows "sum f I = -\<infinity> \<longleftrightarrow> (finite I \<and> \<not> (\<exists>i\<in>I. f i = \<infinity>) \<and> (\<exists>i\<in>I. f i = -\<infinity>))"
by (induction I rule: infinite_finite_induct)
(auto simp: sum_Pinfty)
lemma eexp_sum: "\<not> (\<exists>i\<in>I. \<exists>j\<in>I. f i = -\<infinity> \<and> f j = \<infinity>) \<Longrightarrow> eexp (\<Sum>i\<in>I. f i) = (\<Prod>i\<in>I. eexp (f i))"
proof (induction I rule: infinite_finite_induct)
case (insert i I)
have "eexp (sum f (insert i I)) = eexp (f i) * eexp (sum f I)"
using insert.prems insert.hyps by (auto simp: sum_Pinfty sum_Minfty intro!: eexp_add)
then show ?case
using insert by auto
qed simp_all
lemma eexp_suminf:
assumes wf_f: "\<not> {-\<infinity>, \<infinity>} \<subseteq> range f" and f: "summable f"
shows "(\<lambda>n. \<Prod>i<n. eexp (f i)) \<longlonglongrightarrow> eexp (\<Sum>i. f i)"
proof -
have "(\<lambda>n. eexp (\<Sum>i<n. f i)) \<longlonglongrightarrow> eexp (\<Sum>i. f i)"
by (intro tendsto_eexp summable_LIMSEQ f)
also have "(\<lambda>n. eexp (\<Sum>i<n. f i)) = (\<lambda>n. \<Prod>i<n. eexp (f i))"
using wf_f by (auto simp: fun_eq_iff image_iff eq_commute intro!: eexp_sum)
finally show ?thesis .
qed
lemma continuous_onI_antimono:
fixes f :: "'a::linorder_topology \<Rightarrow> 'b::{dense_order,linorder_topology}"
assumes "open (f`A)"
and mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f y \<le> f x"
shows "continuous_on A f"
proof (rule continuous_on_generate_topology[OF open_generated_order], safe)
have monoD: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> f y < f x \<Longrightarrow> x < y"
by (auto simp: not_le[symmetric] mono)
have "\<exists>x. x \<in> A \<and> f x < b \<and> x < a" if a: "a \<in> A" and fa: "f a < b" for a b
proof -
obtain y where "f a < y" "{f a ..< y} \<subseteq> f`A"
using open_right[OF \<open>open (f`A)\<close>, of "f a" b] a fa
by auto
obtain z where z: "f a < z" "z < min b y"
using dense[of "f a" "min b y"] \<open>f a < y\<close> \<open>f a < b\<close> by auto
then obtain c where "z = f c" "c \<in> A"
using \<open>{f a ..< y} \<subseteq> f`A\<close>[THEN subsetD, of z] by (auto simp: less_imp_le)
with a z show ?thesis
by (auto intro!: exI[of _ c] simp: monoD)
qed
then show "\<exists>C. open C \<and> C \<inter> A = f -` {..<b} \<inter> A" for b
by (intro exI[of _ "(\<Union>x\<in>{x\<in>A. f x < b}. {x <..})"])
(auto intro: le_less_trans[OF mono] less_imp_le)
have "\<exists>x. x \<in> A \<and> b < f x \<and> x > a" if a: "a \<in> A" and fa: "b < f a" for a b
proof -
note a fa
moreover
obtain y where "y < f a" "{y <.. f a} \<subseteq> f`A"
using open_left[OF \<open>open (f`A)\<close>, of "f a" b] a fa
by auto
then obtain z where z: "max b y < z" "z < f a"
using dense[of "max b y" "f a"] \<open>y < f a\<close> \<open>b < f a\<close> by auto
then obtain c where "z = f c" "c \<in> A"
using \<open>{y <.. f a} \<subseteq> f`A\<close>[THEN subsetD, of z] by (auto simp: less_imp_le)
with a z show ?thesis
by (auto intro!: exI[of _ c] simp: monoD)
qed
then show "\<exists>C. open C \<and> C \<inter> A = f -` {b <..} \<inter> A" for b
by (intro exI[of _ "(\<Union>x\<in>{x\<in>A. b < f x}. {..< x})"])
(auto intro: less_le_trans[OF _ mono] less_imp_le)
qed
lemma minus_add_eq_ereal: "\<not> ((a = \<infinity> \<and> b = -\<infinity>) \<or> (a = -\<infinity> \<and> b = \<infinity>)) \<Longrightarrow> - (a + b::ereal) = -a - b"
by (cases a; cases b; simp)
lemma setsum_negf_ereal: "\<not> {-\<infinity>, \<infinity>} \<subseteq> f`I \<Longrightarrow> (\<Sum>i\<in>I. - f i) = - (\<Sum>i\<in>I. f i::ereal)"
by (induction I rule: infinite_finite_induct)
(auto simp: minus_add_eq_ereal sum_Minfty sum_Pinfty,
(subst minus_add_eq_ereal; auto simp: sum_Pinfty sum_Minfty image_iff minus_ereal_def)+)
lemma convergent_minus_iff_ereal: "convergent (\<lambda>x. - f x::ereal) \<longleftrightarrow> convergent f"
unfolding convergent_def by (metis ereal_uminus_uminus ereal_Lim_uminus)
lemma summable_minus_ereal: "\<not> {-\<infinity>, \<infinity>} \<subseteq> range f \<Longrightarrow> summable (\<lambda>n. f n) \<Longrightarrow> summable (\<lambda>n. - f n::ereal)"
unfolding summable_iff_convergent
by (subst setsum_negf_ereal) (auto simp: convergent_minus_iff_ereal)
lemma (in product_prob_space) product_nn_integral_component:
assumes "f \<in> borel_measurable (M i)""i \<in> I"
shows "integral\<^sup>N (Pi\<^sub>M I M) (\<lambda>x. f (x i)) = integral\<^sup>N (M i) f"
proof -
from assms show ?thesis
apply (subst PiM_component[symmetric, OF \<open>i \<in> I\<close>])
apply (subst nn_integral_distr[OF measurable_component_singleton])
apply simp_all
done
qed
lemma ennreal_inverse_le[simp]: "inverse x \<le> inverse y \<longleftrightarrow> y \<le> (x::ennreal)"
by (cases "0 < x"; cases x; cases "0 < y"; cases y; auto simp: top_unique inverse_ennreal)
lemma inverse_inverse_ennreal[simp]: "inverse (inverse x::ennreal) = x"
by (cases "0 < x"; cases x; auto simp: inverse_ennreal)
lemma range_inverse_ennreal: "range inverse = (UNIV::ennreal set)"
proof -
have "\<exists>x. y = inverse x" for y :: ennreal
by (intro exI[of _ "inverse y"]) simp
then show ?thesis
unfolding surj_def by auto
qed
lemma continuous_on_inverse_ennreal': "continuous_on (UNIV :: ennreal set) inverse"
by (rule continuous_onI_antimono) (auto simp: range_inverse_ennreal)
lemma sums_minus_ereal: "\<not> {- \<infinity>, \<infinity>} \<subseteq> f ` UNIV \<Longrightarrow> (\<lambda>n. - f n::ereal) sums x \<Longrightarrow> f sums - x"
unfolding sums_def
apply (subst ereal_Lim_uminus)
apply (subst (asm) setsum_negf_ereal)
apply auto
done
lemma suminf_minus_ereal: "\<not> {- \<infinity>, \<infinity>} \<subseteq> f ` UNIV \<Longrightarrow> summable f \<Longrightarrow> (\<Sum>n. - f n :: ereal) = - suminf f"
apply (rule sums_unique[symmetric])
apply (rule sums_minus_ereal)
apply (auto simp: ereal_uminus_eq_reorder)
done
end
diff --git a/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy b/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy
--- a/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy
+++ b/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy
@@ -1,961 +1,961 @@
section \<open>Algebra of Monotonic Boolean Transformers\<close>
theory Mono_Bool_Tran_Algebra
imports Mono_Bool_Tran
begin
text\<open>
In this section we introduce the {\em algebra of monotonic boolean transformers}.
This is a bounded distributive lattice with a monoid operation, a
dual operator and an iteration operator. The standard model for this
algebra is the set of monotonic boolean transformers introduced
in the previous section.
\<close>
class dual =
fixes dual::"'a \<Rightarrow> 'a" ("_ ^ o" [81] 80)
class omega =
fixes omega::"'a \<Rightarrow> 'a" ("_ ^ \<omega>" [81] 80)
class star =
fixes star::"'a \<Rightarrow> 'a" ("(_ ^ *)" [81] 80)
class dual_star =
fixes dual_star::"'a \<Rightarrow> 'a" ("(_ ^ \<otimes>)" [81] 80)
class mbt_algebra = monoid_mult + dual + omega + distrib_lattice + order_top + order_bot + star + dual_star +
assumes
dual_le: "(x \<le> y) = (y ^ o \<le> x ^ o)"
and dual_dual [simp]: "(x ^ o) ^ o = x"
and dual_comp: "(x * y) ^ o = x ^ o * y ^ o"
and dual_one [simp]: "1 ^ o = 1"
and top_comp [simp]: "\<top> * x = \<top>"
and inf_comp: "(x \<sqinter> y) * z = (x * z) \<sqinter> (y * z)"
and le_comp: "x \<le> y \<Longrightarrow> z * x \<le> z * y"
and dual_neg: "(x * \<top>) \<sqinter> (x ^ o * \<bottom>) = \<bottom>"
and omega_fix: "x ^ \<omega> = (x * (x ^ \<omega>)) \<sqinter> 1"
and omega_least: "(x * z) \<sqinter> y \<le> z \<Longrightarrow> (x ^ \<omega>) * y \<le> z"
and star_fix: "x ^ * = (x * (x ^ *)) \<sqinter> 1"
and star_greatest: "z \<le> (x * z) \<sqinter> y \<Longrightarrow> z \<le> (x ^ *) * y"
and dual_star_def: "(x ^ \<otimes>) = (((x ^ o) ^ *) ^ o)"
begin
lemma le_comp_right: "x \<le> y \<Longrightarrow> x * z \<le> y * z"
apply (cut_tac x = x and y = y and z = z in inf_comp)
apply (simp add: inf_absorb1)
apply (subgoal_tac "x * z \<sqinter> (y * z) \<le> y * z")
apply simp
by (rule inf_le2)
subclass bounded_lattice
proof qed
end
instantiation MonoTran :: (complete_boolean_algebra) mbt_algebra
begin
lift_definition dual_MonoTran :: "'a MonoTran \<Rightarrow> 'a MonoTran"
is dual_fun
by (fact mono_dual_fun)
lift_definition omega_MonoTran :: "'a MonoTran \<Rightarrow> 'a MonoTran"
is omega_fun
by (fact mono_omega_fun)
lift_definition star_MonoTran :: "'a MonoTran \<Rightarrow> 'a MonoTran"
is star_fun
by (fact mono_star_fun)
definition dual_star_MonoTran :: "'a MonoTran \<Rightarrow> 'a MonoTran"
where
"(x::('a MonoTran)) ^ \<otimes> = ((x ^ o) ^ *) ^ o"
instance proof
fix x y :: "'a MonoTran" show "(x \<le> y) = (y ^ o \<le> x ^ o)"
apply transfer
apply (auto simp add: fun_eq_iff le_fun_def)
apply (drule_tac x = "-xa" in spec)
apply simp
done
next
fix x :: "'a MonoTran" show "(x ^ o) ^ o = x"
apply transfer
apply (simp add: fun_eq_iff)
done
next
fix x y :: "'a MonoTran" show "(x * y) ^ o = x ^ o * y ^ o"
apply transfer
apply (simp add: fun_eq_iff)
done
next
show "(1::'a MonoTran) ^ o = 1"
apply transfer
apply (simp add: fun_eq_iff)
done
next
fix x :: "'a MonoTran" show "\<top> * x = \<top>"
apply transfer
apply (simp add: fun_eq_iff)
done
next
fix x y z :: "'a MonoTran" show "(x \<sqinter> y) * z = (x * z) \<sqinter> (y * z)"
apply transfer
apply (simp add: fun_eq_iff)
done
next
fix x y z :: "'a MonoTran" assume A: "x \<le> y" from A show " z * x \<le> z * y"
apply transfer
apply (auto simp add: le_fun_def elim: monoE)
done
next
fix x :: "'a MonoTran" show "x * \<top> \<sqinter> (x ^ o * \<bottom>) = \<bottom>"
apply transfer
apply (simp add: fun_eq_iff)
done
next
fix x :: "'a MonoTran" show "x ^ \<omega> = x * x ^ \<omega> \<sqinter> 1"
apply transfer
apply (simp add: fun_eq_iff)
apply (simp add: omega_fun_def Omega_fun_def)
apply (subst lfp_unfold, simp_all add: ac_simps)
apply (auto intro!: mono_comp mono_comp_fun)
done
next
fix x y z :: "'a MonoTran" assume A: "x * z \<sqinter> y \<le> z" from A show "x ^ \<omega> * y \<le> z"
apply transfer
apply (auto simp add: lfp_omega lfp_def)
apply (rule Inf_lower)
apply (auto simp add: Omega_fun_def ac_simps)
done
next
fix x :: "'a MonoTran" show "x ^ * = x * x ^ * \<sqinter> 1"
apply transfer
apply (auto simp add: star_fun_def Omega_fun_def)
apply (subst gfp_unfold, simp_all add: ac_simps)
apply (auto intro!: mono_comp mono_comp_fun)
done
next
fix x y z :: "'a MonoTran" assume A: "z \<le> x * z \<sqinter> y" from A show "z \<le> x ^ * * y"
apply transfer
apply (auto simp add: gfp_star gfp_def)
apply (rule Sup_upper)
apply (auto simp add: Omega_fun_def)
done
next
fix x :: "'a MonoTran" show "x ^ \<otimes> = ((x ^ o) ^ *) ^ o"
by (simp add: dual_star_MonoTran_def)
qed
end
context mbt_algebra begin
lemma dual_top [simp]: "\<top> ^ o = \<bottom>"
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
by (subst dual_le, simp)
lemma dual_bot [simp]: "\<bottom> ^ o = \<top>"
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
by (subst dual_le, simp)
lemma dual_inf: "(x \<sqinter> y) ^ o = (x ^ o) \<squnion> (y ^ o)"
- apply (rule antisym, simp_all, safe)
+ apply (rule order.antisym, simp_all, safe)
apply (subst dual_le, simp, safe)
apply (subst dual_le, simp)
apply (subst dual_le, simp)
apply (subst dual_le, simp)
by (subst dual_le, simp)
lemma dual_sup: "(x \<squnion> y) ^ o = (x ^ o) \<sqinter> (y ^ o)"
- apply (rule antisym, simp_all, safe)
+ apply (rule order.antisym, simp_all, safe)
apply (subst dual_le, simp)
apply (subst dual_le, simp)
apply (subst dual_le, simp, safe)
apply (subst dual_le, simp)
by (subst dual_le, simp)
lemma sup_comp: "(x \<squnion> y) * z = (x * z) \<squnion> (y * z)"
apply (subgoal_tac "((x ^ o \<sqinter> y ^ o) * z ^ o) ^ o = ((x ^ o * z ^ o) \<sqinter> (y ^ o * z ^ o)) ^ o")
apply (simp add: dual_inf dual_comp)
by (simp add: inf_comp)
lemma dual_eq: "x ^ o = y ^ o \<Longrightarrow> x = y"
apply (subgoal_tac "(x ^ o) ^ o = (y ^ o) ^ o")
apply (subst (asm) dual_dual)
apply (subst (asm) dual_dual)
by simp_all
lemma dual_neg_top [simp]: "(x ^ o * \<bottom>) \<squnion> (x * \<top>) = \<top>"
apply (rule dual_eq)
by(simp add: dual_sup dual_comp dual_neg)
lemma bot_comp [simp]: "\<bottom> * x = \<bottom>"
by (rule dual_eq, simp add: dual_comp)
lemma [simp]: "(x * \<top>) * y = x * \<top>"
by (simp add: mult.assoc)
lemma [simp]: "(x * \<bottom>) * y = x * \<bottom>"
by (simp add: mult.assoc)
lemma gt_one_comp: "1 \<le> x \<Longrightarrow> y \<le> x * y"
by (cut_tac x = 1 and y = x and z = y in le_comp_right, simp_all)
theorem omega_comp_fix: "x ^ \<omega> * y = (x * (x ^ \<omega>) * y) \<sqinter> y"
apply (subst omega_fix)
by (simp add: inf_comp)
theorem dual_star_fix: "x^\<otimes> = (x * (x^\<otimes>)) \<squnion> 1"
by (metis dual_comp dual_dual dual_inf dual_one dual_star_def star_fix)
theorem star_comp_fix: "x ^ * * y = (x * (x ^ *) * y) \<sqinter> y"
apply (subst star_fix)
by (simp add: inf_comp)
theorem dual_star_comp_fix: "x^\<otimes> * y = (x * (x^\<otimes>) * y) \<squnion> y"
apply (subst dual_star_fix)
by (simp add: sup_comp)
theorem dual_star_least: "(x * z) \<squnion> y \<le> z \<Longrightarrow> (x^\<otimes>) * y \<le> z"
apply (subst dual_le)
apply (simp add: dual_star_def dual_comp)
apply (rule star_greatest)
apply (subst dual_le)
by (simp add: dual_inf dual_comp)
lemma omega_one [simp]: "1 ^ \<omega> = \<bottom>"
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
by (cut_tac x = "1::'a" and y = 1 and z = \<bottom> in omega_least, simp_all)
lemma omega_mono: "x \<le> y \<Longrightarrow> x ^ \<omega> \<le> y ^ \<omega>"
apply (cut_tac x = x and y = 1 and z = "y ^ \<omega>" in omega_least, simp_all)
apply (subst (2) omega_fix, simp_all)
apply (rule_tac y = "x * y ^ \<omega>" in order_trans, simp)
by (rule le_comp_right, simp)
end
sublocale mbt_algebra < conjunctive "inf" "inf" "times"
done
sublocale mbt_algebra < disjunctive "sup" "sup" "times"
done
context mbt_algebra begin
lemma dual_conjunctive: "x \<in> conjunctive \<Longrightarrow> x ^ o \<in> disjunctive"
apply (simp add: conjunctive_def disjunctive_def)
apply safe
apply (rule dual_eq)
by (simp add: dual_comp dual_sup)
lemma dual_disjunctive: "x \<in> disjunctive \<Longrightarrow> x ^ o \<in> conjunctive"
apply (simp add: conjunctive_def disjunctive_def)
apply safe
apply (rule dual_eq)
by (simp add: dual_comp dual_inf)
lemma comp_pres_conj: "x \<in> conjunctive \<Longrightarrow> y \<in> conjunctive \<Longrightarrow> x * y \<in> conjunctive"
apply (subst conjunctive_def, safe)
by (simp add: mult.assoc conjunctiveD)
lemma comp_pres_disj: "x \<in> disjunctive \<Longrightarrow> y \<in> disjunctive \<Longrightarrow> x * y \<in> disjunctive"
apply (subst disjunctive_def, safe)
by (simp add: mult.assoc disjunctiveD)
lemma start_pres_conj: "x \<in> conjunctive \<Longrightarrow> (x ^ *) \<in> conjunctive"
apply (subst conjunctive_def, safe)
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
apply (metis inf_le1 inf_le2 le_comp)
apply (rule star_greatest)
apply (subst conjunctiveD, simp)
apply (subst star_comp_fix)
apply (subst star_comp_fix)
by (metis inf.assoc inf_left_commute mult.assoc order_refl)
lemma dual_star_pres_disj: "x \<in> disjunctive \<Longrightarrow> x^\<otimes> \<in> disjunctive"
apply (simp add: dual_star_def)
apply (rule dual_conjunctive)
apply (rule start_pres_conj)
by (rule dual_disjunctive, simp)
subsection\<open>Assertions\<close>
text\<open>
Usually, in Kleene algebra with tests or in other progrm algebras, tests or assertions
or assumptions are defined using an existential quantifier. An element of the algebra
is a test if it has a complement with respect to $\bot$ and $1$. In this formalization
assertions can be defined much simpler using the dual operator.
\<close>
definition
"assertion = {x . x \<le> 1 \<and> (x * \<top>) \<sqinter> (x ^ o) = x}"
lemma assertion_prop: "x \<in> assertion \<Longrightarrow> (x * \<top>) \<sqinter> 1 = x"
apply (simp add: assertion_def)
apply safe
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
proof -
assume [simp]: "x \<le> 1"
assume A: "x * \<top> \<sqinter> x ^ o = x"
have "x * \<top> \<sqinter> 1 \<le> x * \<top> \<sqinter> x ^ o"
apply simp
apply (rule_tac y = 1 in order_trans)
apply simp
apply (subst dual_le)
by simp
also have "\<dots> = x" by (cut_tac A, simp)
finally show "x * \<top> \<sqinter> 1 \<le> x" .
next
assume A: "x * \<top> \<sqinter> x ^ o = x"
have "x = x * \<top> \<sqinter> x ^ o" by (simp add: A)
also have "\<dots> \<le> x * \<top>" by simp
finally show "x \<le> x * \<top>" .
qed
lemma dual_assertion_prop: "x \<in> assertion \<Longrightarrow> ((x ^ o) * \<bottom>) \<squnion> 1 = x ^ o"
apply (rule dual_eq)
by (simp add: dual_sup dual_comp assertion_prop)
lemma assertion_disjunctive: "x \<in> assertion \<Longrightarrow> x \<in> disjunctive"
apply (simp add: disjunctive_def, safe)
apply (drule assertion_prop)
proof -
assume A: "x * \<top> \<sqinter> 1 = x"
fix y z::"'a"
have "x * (y \<squnion> z) = (x * \<top> \<sqinter> 1) * (y \<squnion> z)" by (cut_tac A, simp)
also have "\<dots> = (x * \<top>) \<sqinter> (y \<squnion> z)" by (simp add: inf_comp)
also have "\<dots> = ((x * \<top>) \<sqinter> y) \<squnion> ((x * \<top>) \<sqinter> z)" by (simp add: inf_sup_distrib)
also have "\<dots> = (((x * \<top>) \<sqinter> 1) * y) \<squnion> (((x * \<top>) \<sqinter> 1) * z)" by (simp add: inf_comp)
also have "\<dots> = x * y \<squnion> x * z" by (cut_tac A, simp)
finally show "x * (y \<squnion> z) = x * y \<squnion> x * z" .
qed
lemma Abs_MonoTran_injective: "mono x \<Longrightarrow> mono y \<Longrightarrow> Abs_MonoTran x = Abs_MonoTran y \<Longrightarrow> x = y"
apply (subgoal_tac "Rep_MonoTran (Abs_MonoTran x) = Rep_MonoTran (Abs_MonoTran y)")
apply (subst (asm) Abs_MonoTran_inverse, simp)
by (subst (asm) Abs_MonoTran_inverse, simp_all)
end
lemma mbta_MonoTran_disjunctive: "Rep_MonoTran ` disjunctive = Apply.disjunctive"
apply (simp add: disjunctive_def Apply.disjunctive_def)
apply transfer
apply auto
proof -
fix f :: "'a \<Rightarrow> 'a" and a b
assume prem: "\<forall>y. mono y \<longrightarrow> (\<forall>z. mono z \<longrightarrow> f \<circ> y \<squnion> z = (f \<circ> y) \<squnion> (f \<circ> z))"
{ fix g h :: "'b \<Rightarrow> 'a"
assume "mono g" and "mono h"
then have "f \<circ> g \<squnion> h = (f \<circ> g) \<squnion> (f \<circ> h)"
using prem by blast
} note * = this
assume "mono f"
show "f (a \<squnion> b) = f a \<squnion> f b" (is "?P = ?Q")
proof (rule order_antisym)
show "?P \<le> ?Q"
using * [of "\<lambda>_. a" "\<lambda>_. b"] by (simp add: comp_def fun_eq_iff)
next
from \<open>mono f\<close> show "?Q \<le> ?P" by (rule Lattices.semilattice_sup_class.mono_sup)
qed
next
fix f :: "'a \<Rightarrow> 'a"
assume "\<forall>y z. f (y \<squnion> z) = f y \<squnion> f z"
then have *: "\<And>y z. f (y \<squnion> z) = f y \<squnion> f z" by blast
show "mono f"
proof
fix a b :: 'a
assume "a \<le> b"
then show "f a \<le> f b"
unfolding sup.order_iff * [symmetric] by simp
qed
qed
lemma assertion_MonoTran: "assertion = Abs_MonoTran ` assertion_fun"
apply (safe)
apply (subst assertion_fun_disj_less_one)
apply (simp add: image_def)
apply (rule_tac x = "Rep_MonoTran x" in bexI)
apply (simp add: Rep_MonoTran_inverse)
apply safe
apply (drule assertion_disjunctive)
apply (unfold mbta_MonoTran_disjunctive [THEN sym], simp)
apply (simp add: assertion_def less_eq_MonoTran_def one_MonoTran_def Abs_MonoTran_inverse)
apply (simp add: assertion_def)
by (simp_all add: inf_MonoTran_def less_eq_MonoTran_def
times_MonoTran_def dual_MonoTran_def top_MonoTran_def Abs_MonoTran_inverse one_MonoTran_def assertion_fun_dual)
context mbt_algebra begin
lemma assertion_conjunctive: "x \<in> assertion \<Longrightarrow> x \<in> conjunctive"
apply (simp add: conjunctive_def, safe)
apply (drule assertion_prop)
proof -
assume A: "x * \<top> \<sqinter> 1 = x"
fix y z::"'a"
have "x * (y \<sqinter> z) = (x * \<top> \<sqinter> 1) * (y \<sqinter> z)" by (cut_tac A, simp)
also have "\<dots> = (x * \<top>) \<sqinter> (y \<sqinter> z)" by (simp add: inf_comp)
also have "\<dots> = ((x * \<top>) \<sqinter> y) \<sqinter> ((x * \<top>) \<sqinter> z)"
- apply (rule antisym, simp_all, safe)
+ apply (rule order.antisym, simp_all, safe)
apply (rule_tac y = "y \<sqinter> z" in order_trans)
apply (rule inf_le2)
apply simp
apply (rule_tac y = "y \<sqinter> z" in order_trans)
apply (rule inf_le2)
apply simp_all
apply (simp add: inf_assoc)
apply (rule_tac y = " x * \<top> \<sqinter> y" in order_trans)
apply (rule inf_le1)
apply simp
apply (rule_tac y = " x * \<top> \<sqinter> z" in order_trans)
apply (rule inf_le2)
by simp
also have "\<dots> = (((x * \<top>) \<sqinter> 1) * y) \<sqinter> (((x * \<top>) \<sqinter> 1) * z)" by (simp add: inf_comp)
also have "\<dots> = (x * y) \<sqinter> (x * z)" by (cut_tac A, simp)
finally show "x * (y \<sqinter> z) = (x * y) \<sqinter> (x * z)" .
qed
lemma dual_assertion_conjunctive: "x \<in> assertion \<Longrightarrow> x ^ o \<in> conjunctive"
apply (drule assertion_disjunctive)
by (rule dual_disjunctive, simp)
lemma dual_assertion_disjunct: "x \<in> assertion \<Longrightarrow> x ^ o \<in> disjunctive"
apply (drule assertion_conjunctive)
by (rule dual_conjunctive, simp)
lemma [simp]: "x \<in> assertion \<Longrightarrow> y \<in> assertion \<Longrightarrow> x \<sqinter> y \<le> x * y"
apply (simp add: assertion_def, safe)
proof -
assume A: "x \<le> 1"
assume B: "x * \<top> \<sqinter> x ^ o = x"
assume C: "y \<le> 1"
assume D: "y * \<top> \<sqinter> y ^ o = y"
have "x \<sqinter> y = (x * \<top> \<sqinter> x ^ o) \<sqinter> (y * \<top> \<sqinter> y ^ o)" by (cut_tac B D, simp)
also have "\<dots> \<le> (x * \<top>) \<sqinter> (((x^o) * (y * \<top>)) \<sqinter> ((x^o) * (y^o)))"
apply (simp, safe)
apply (rule_tac y = "x * \<top> \<sqinter> x ^ o" in order_trans)
apply (rule inf_le1)
apply simp
apply (rule_tac y = "y * \<top>" in order_trans)
apply (rule_tac y = "y * \<top> \<sqinter> y ^ o" in order_trans)
apply (rule inf_le2)
apply simp
apply (rule gt_one_comp)
apply (subst dual_le, simp add: A)
apply (rule_tac y = "y ^ o" in order_trans)
apply (rule_tac y = "y * \<top> \<sqinter> y ^ o" in order_trans)
apply (rule inf_le2)
apply simp
apply (rule gt_one_comp)
by (subst dual_le, simp add: A)
also have "... = ((x * \<top>) \<sqinter> (x ^ o)) * ((y * \<top>) \<sqinter> (y ^ o))"
apply (cut_tac x = x in dual_assertion_conjunctive)
apply (cut_tac A, cut_tac B, simp add: assertion_def)
by (simp add: inf_comp conjunctiveD)
also have "... = x * y"
by (cut_tac B, cut_tac D, simp)
finally show "x \<sqinter> y \<le> x * y" .
qed
lemma [simp]: "x \<in> assertion \<Longrightarrow> x * y \<le> y"
by (unfold assertion_def, cut_tac x = x and y = 1 and z = y in le_comp_right, simp_all)
lemma [simp]: "x \<in> assertion \<Longrightarrow> y \<in> assertion \<Longrightarrow> x * y \<le> x"
apply (subgoal_tac "x * y \<le> (x * \<top>) \<sqinter> (x ^ o)")
apply (simp add: assertion_def)
apply (simp, safe)
apply (rule le_comp, simp)
apply (rule_tac y = 1 in order_trans)
apply (rule_tac y = y in order_trans)
apply simp
apply (simp add: assertion_def)
by (subst dual_le, simp add: assertion_def)
lemma assertion_inf_comp_eq: "x \<in> assertion \<Longrightarrow> y \<in> assertion \<Longrightarrow> x \<sqinter> y = x * y"
- by (rule antisym, simp_all)
+ by (rule order.antisym, simp_all)
lemma one_right_assertion [simp]: "x \<in> assertion \<Longrightarrow> x * 1 = x"
apply (drule assertion_prop)
proof -
assume A: "x * \<top> \<sqinter> 1 = x"
have "x * 1 = (x * \<top> \<sqinter> 1) * 1" by (simp add: A)
also have "\<dots> = x * \<top> \<sqinter> 1" by (simp add: inf_comp)
also have "\<dots> = x" by (simp add: A)
finally show ?thesis .
qed
lemma [simp]: "x \<in> assertion \<Longrightarrow> x \<squnion> 1 = 1"
- by (rule antisym, simp_all add: assertion_def)
+ by (rule order.antisym, simp_all add: assertion_def)
lemma [simp]: "x \<in> assertion \<Longrightarrow> 1 \<squnion> x = 1"
- by (rule antisym, simp_all add: assertion_def)
+ by (rule order.antisym, simp_all add: assertion_def)
lemma [simp]: "x \<in> assertion \<Longrightarrow> x \<sqinter> 1 = x"
- by (rule antisym, simp_all add: assertion_def)
+ by (rule order.antisym, simp_all add: assertion_def)
lemma [simp]: "x \<in> assertion \<Longrightarrow> 1 \<sqinter> x = x"
- by (rule antisym, simp_all add: assertion_def)
+ by (rule order.antisym, simp_all add: assertion_def)
lemma [simp]: "x \<in> assertion \<Longrightarrow> x \<le> x * \<top>"
by (cut_tac x = 1 and y = \<top> and z = x in le_comp, simp_all)
lemma [simp]: "x \<in> assertion \<Longrightarrow> x \<le> 1"
by (simp add: assertion_def)
definition
"neg_assert (x::'a) = (x ^ o * \<bottom>) \<sqinter> 1"
lemma sup_uminus[simp]: "x \<in> assertion \<Longrightarrow> x \<squnion> neg_assert x = 1"
apply (simp add: neg_assert_def)
apply (simp add: sup_inf_distrib)
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
apply (unfold assertion_def)
apply safe
apply (subst dual_le)
apply (simp add: dual_sup dual_comp)
apply (subst inf_commute)
by simp
lemma inf_uminus[simp]: "x \<in> assertion \<Longrightarrow> x \<sqinter> neg_assert x = \<bottom>"
apply (simp add: neg_assert_def)
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
apply (rule_tac y = "x \<sqinter> (x ^ o * \<bottom>)" in order_trans)
apply simp
apply (rule_tac y = "x ^ o * \<bottom> \<sqinter> 1" in order_trans)
apply (rule inf_le2)
apply simp
apply (rule_tac y = "(x * \<top>) \<sqinter> (x ^ o * \<bottom>)" in order_trans)
apply simp
apply (rule_tac y = x in order_trans)
apply simp_all
by (simp add: dual_neg)
lemma uminus_assertion[simp]: "x \<in> assertion \<Longrightarrow> neg_assert x \<in> assertion"
apply (subst assertion_def)
apply (simp add: neg_assert_def)
apply (simp add: inf_comp dual_inf dual_comp inf_sup_distrib)
apply (subst inf_commute)
by (simp add: dual_neg)
lemma uminus_uminus [simp]: "x \<in> assertion \<Longrightarrow> neg_assert (neg_assert x) = x"
apply (simp add: neg_assert_def)
by (simp add: dual_inf dual_comp sup_comp assertion_prop)
lemma dual_comp_neg [simp]: "x ^ o * y \<squnion> (neg_assert x) * \<top> = x ^ o * y"
apply (simp add: neg_assert_def inf_comp)
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
by (rule le_comp, simp)
lemma [simp]: "(neg_assert x) ^ o * y \<squnion> x * \<top> = (neg_assert x) ^ o * y"
apply (simp add: neg_assert_def inf_comp dual_inf dual_comp sup_comp)
- by (rule antisym, simp_all)
+ by (rule order.antisym, simp_all)
lemma [simp]: " x * \<top> \<squnion> (neg_assert x) ^ o * y= (neg_assert x) ^ o * y"
by (simp add: neg_assert_def inf_comp dual_inf dual_comp sup_comp)
lemma inf_assertion [simp]: "x \<in> assertion \<Longrightarrow> y \<in> assertion \<Longrightarrow> x \<sqinter> y \<in> assertion"
apply (subst assertion_def)
apply safe
apply (rule_tac y = x in order_trans)
apply simp_all
apply (simp add: assertion_inf_comp_eq)
proof -
assume A: "x \<in> assertion"
assume B: "y \<in> assertion"
have C: "(x * \<top>) \<sqinter> (x ^ o) = x"
by (cut_tac A, unfold assertion_def, simp)
have D: "(y * \<top>) \<sqinter> (y ^ o) = y"
by (cut_tac B, unfold assertion_def, simp)
have "x * y = ((x * \<top>) \<sqinter> (x ^ o)) * ((y * \<top>) \<sqinter> (y ^ o))" by (simp add: C D)
also have "\<dots> = x * \<top> \<sqinter> ((x ^ o) * ((y * \<top>) \<sqinter> (y ^ o)))" by (simp add: inf_comp)
also have "\<dots> = x * \<top> \<sqinter> ((x ^ o) * (y * \<top>)) \<sqinter> ((x ^ o) *(y ^ o))"
by (cut_tac A, cut_tac x = x in dual_assertion_conjunctive, simp_all add: conjunctiveD inf_assoc)
also have "\<dots> = (((x * \<top>) \<sqinter> (x ^ o)) * (y * \<top>)) \<sqinter> ((x ^ o) *(y ^ o))"
by (simp add: inf_comp)
also have "\<dots> = (x * y * \<top>) \<sqinter> ((x * y) ^ o)" by (simp add: C mult.assoc dual_comp)
finally show "(x * y * \<top>) \<sqinter> ((x * y) ^ o) = x * y" by simp
qed
lemma comp_assertion [simp]: "x \<in> assertion \<Longrightarrow> y \<in> assertion \<Longrightarrow> x * y \<in> assertion"
by (subst assertion_inf_comp_eq [THEN sym], simp_all)
lemma sup_assertion [simp]: "x \<in> assertion \<Longrightarrow> y \<in> assertion \<Longrightarrow> x \<squnion> y \<in> assertion"
apply (subst assertion_def)
apply safe
apply (unfold assertion_def)
apply simp
apply safe
proof -
assume [simp]: "x \<le> 1"
assume [simp]: "y \<le> 1"
assume A: "x * \<top> \<sqinter> x ^ o = x"
assume B: "y * \<top> \<sqinter> y ^ o = y"
have "(y * \<top>) \<sqinter> (x ^ o) \<sqinter> (y ^ o) = (x ^ o) \<sqinter> (y * \<top>) \<sqinter> (y ^ o)" by (simp add: inf_commute)
also have "\<dots> = (x ^ o) \<sqinter> ((y * \<top>) \<sqinter> (y ^ o))" by (simp add: inf_assoc)
also have "\<dots> = (x ^ o) \<sqinter> y" by (simp add: B)
also have "\<dots> = y"
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
apply (rule_tac y = 1 in order_trans)
apply simp
by (subst dual_le, simp)
finally have [simp]: "(y * \<top>) \<sqinter> (x ^ o) \<sqinter> (y ^ o) = y" .
have "x * \<top> \<sqinter> (x ^ o) \<sqinter> (y ^ o) = x \<sqinter> (y ^ o)" by (simp add: A)
also have "\<dots> = x"
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
apply (rule_tac y = 1 in order_trans)
apply simp
by (subst dual_le, simp)
finally have [simp]: "x * \<top> \<sqinter> (x ^ o) \<sqinter> (y ^ o) = x" .
have "(x \<squnion> y) * \<top> \<sqinter> (x \<squnion> y) ^ o = (x * \<top> \<squnion> y * \<top>) \<sqinter> ((x ^ o) \<sqinter> (y ^ o))" by (simp add: sup_comp dual_sup)
also have "\<dots> = x \<squnion> y" by (simp add: inf_sup_distrib inf_assoc [THEN sym])
finally show "(x \<squnion> y) * \<top> \<sqinter> (x \<squnion> y) ^ o = x \<squnion> y" .
qed
lemma [simp]: "x \<in> assertion \<Longrightarrow> x * x = x"
by (simp add: assertion_inf_comp_eq [THEN sym])
lemma [simp]: "x \<in> assertion \<Longrightarrow> (x ^ o) * (x ^ o) = x ^ o"
apply (rule dual_eq)
by (simp add: dual_comp assertion_inf_comp_eq [THEN sym])
lemma [simp]: "x \<in> assertion \<Longrightarrow> x * (x ^ o) = x"
proof -
assume A: "x \<in> assertion"
have B: "x * \<top> \<sqinter> (x ^ o) = x" by (cut_tac A, unfold assertion_def, simp)
have "x * x ^ o = (x * \<top> \<sqinter> (x ^ o)) * x ^ o" by (simp add: B)
also have "\<dots> = x * \<top> \<sqinter> (x ^ o)" by (cut_tac A, simp add: inf_comp)
also have "\<dots> = x" by (simp add: B)
finally show ?thesis .
qed
lemma [simp]: "x \<in> assertion \<Longrightarrow> (x ^ o) * x = x ^ o"
apply (rule dual_eq)
by (simp add: dual_comp)
lemma [simp]: "\<bottom> \<in> assertion"
by (unfold assertion_def, simp)
lemma [simp]: "1 \<in> assertion"
by (unfold assertion_def, simp)
subsection \<open>Weakest precondition of true\<close>
definition
"wpt x = (x * \<top>) \<sqinter> 1"
lemma wpt_is_assertion [simp]: "wpt x \<in> assertion"
apply (unfold wpt_def assertion_def, safe)
apply simp
apply (simp add: inf_comp dual_inf dual_comp inf_sup_distrib)
- apply (rule antisym)
+ apply (rule order.antisym)
by (simp_all add: dual_neg)
lemma wpt_comp: "(wpt x) * x = x"
apply (simp add: wpt_def inf_comp)
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
by (cut_tac x = 1 and y = \<top> and z = x in le_comp, simp_all)
lemma wpt_comp_2: "wpt (x * y) = wpt (x * (wpt y))"
by (simp add: wpt_def inf_comp mult.assoc)
lemma wpt_assertion [simp]: "x \<in> assertion \<Longrightarrow> wpt x = x"
by (simp add: wpt_def assertion_prop)
lemma wpt_le_assertion: "x \<in> assertion \<Longrightarrow> x * y = y \<Longrightarrow> wpt y \<le> x"
apply (simp add: wpt_def)
proof -
assume A: "x \<in> assertion"
assume B: "x * y = y"
have "y * \<top> \<sqinter> 1 = x * (y * \<top>) \<sqinter> 1" by (simp add: B mult.assoc [THEN sym])
also have "\<dots> \<le> x * \<top> \<sqinter> 1"
apply simp
apply (rule_tac y = "x * (y * \<top>)" in order_trans)
apply simp_all
by (rule le_comp, simp)
also have "\<dots> = x" by (cut_tac A, simp add: assertion_prop)
finally show "y * \<top> \<sqinter> 1 \<le> x" .
qed
lemma wpt_choice: "wpt (x \<sqinter> y) = wpt x \<sqinter> wpt y"
apply (simp add: wpt_def inf_comp)
proof -
have "x * \<top> \<sqinter> 1 \<sqinter> (y * \<top> \<sqinter> 1) = x * \<top> \<sqinter> ((y * \<top> \<sqinter> 1) \<sqinter> 1)" apply (subst inf_assoc) by (simp add: inf_commute)
also have "... = x * \<top> \<sqinter> (y * \<top> \<sqinter> 1)" by (subst inf_assoc, simp)
also have "... = (x * \<top>) \<sqinter> (y * \<top>) \<sqinter> 1" by (subst inf_assoc, simp)
finally show "x * \<top> \<sqinter> (y * \<top>) \<sqinter> 1 = x * \<top> \<sqinter> 1 \<sqinter> (y * \<top> \<sqinter> 1)" by simp
qed
end
context lattice begin
lemma [simp]: "x \<le> y \<Longrightarrow> x \<sqinter> y = x"
by (simp add: inf_absorb1)
end
context mbt_algebra begin
lemma wpt_dual_assertion_comp: "x \<in> assertion \<Longrightarrow> y \<in> assertion \<Longrightarrow> wpt ((x ^ o) * y) = (neg_assert x) \<squnion> y"
apply (simp add: wpt_def neg_assert_def)
proof -
assume A: "x \<in> assertion"
assume B: "y \<in> assertion"
have C: "((x ^ o) * \<bottom>) \<squnion> 1 = x ^ o"
by (rule dual_assertion_prop, rule A)
have "x ^ o * y * \<top> \<sqinter> 1 = (((x ^ o) * \<bottom>) \<squnion> 1) * y * \<top> \<sqinter> 1" by (simp add: C)
also have "\<dots> = ((x ^ o) * \<bottom> \<squnion> (y * \<top>)) \<sqinter> 1" by (simp add: sup_comp)
also have "\<dots> = (((x ^ o) * \<bottom>) \<sqinter> 1) \<squnion> ((y * \<top>) \<sqinter> 1)" by (simp add: inf_sup_distrib2)
also have "\<dots> = (((x ^ o) * \<bottom>) \<sqinter> 1) \<squnion> y" by (cut_tac B, drule assertion_prop, simp)
finally show "x ^ o * y * \<top> \<sqinter> 1 = (((x ^ o) * \<bottom>) \<sqinter> 1) \<squnion> y" .
qed
lemma le_comp_left_right: "x \<le> y \<Longrightarrow> u \<le> v \<Longrightarrow> x * u \<le> y * v"
apply (rule_tac y = "x * v" in order_trans)
apply (rule le_comp, simp)
by (rule le_comp_right, simp)
lemma wpt_dual_assertion: "x \<in> assertion \<Longrightarrow> wpt (x ^ o) = 1"
apply (simp add: wpt_def)
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (cut_tac x = 1 and y = "x ^ o" and u = 1 and v = \<top> in le_comp_left_right)
apply simp_all
apply (subst dual_le)
by simp
lemma assertion_commute: "x \<in> assertion \<Longrightarrow> y \<in> conjunctive \<Longrightarrow> y * x = wpt(y * x) * y"
apply (simp add: wpt_def)
apply (simp add: inf_comp)
apply (drule_tac x = y and y = "x * \<top>" and z = 1 in conjunctiveD)
by (simp add: mult.assoc [THEN sym] assertion_prop)
lemma wpt_mono: "x \<le> y \<Longrightarrow> wpt x \<le> wpt y"
apply (simp add: wpt_def)
apply (rule_tac y = "x * \<top>" in order_trans, simp_all)
by (rule le_comp_right, simp)
lemma "a \<in> conjunctive \<Longrightarrow> x * a \<le> a * y \<Longrightarrow> (x ^ \<omega>) * a \<le> a * (y ^ \<omega>)"
apply (rule omega_least)
apply (simp add: mult.assoc [THEN sym])
apply (rule_tac y = "a * y * y ^ \<omega> \<sqinter> a" in order_trans)
apply (simp)
apply (rule_tac y = "x * a * y ^ \<omega>" in order_trans, simp_all)
apply (rule le_comp_right, simp)
apply (simp add: mult.assoc)
apply (subst (2) omega_fix)
by (simp add: conjunctiveD)
lemma [simp]: "x \<le> 1 \<Longrightarrow> y * x \<le> y"
by (cut_tac x = x and y = 1 and z = y in le_comp, simp_all)
lemma [simp]: "x \<le> x * \<top>"
by (cut_tac x = 1 and y = \<top> and z = x in le_comp, simp_all)
lemma [simp]: "x * \<bottom> \<le> x"
by (cut_tac x = \<bottom> and y = 1 and z = x in le_comp, simp_all)
end
subsection\<open>Monotonic Boolean trasformers algebra with post condition statement\<close>
definition
"post_fun (p::'a::order) q = (if p \<le> q then (\<top>::'b::{order_bot,order_top}) else \<bottom>)"
lemma mono_post_fun [simp]: "mono (post_fun (p::_::{order_bot,order_top}))"
apply (simp add: post_fun_def mono_def, safe)
apply (subgoal_tac "p \<le> y", simp)
apply (rule_tac y = x in order_trans)
apply simp_all
done
lemma post_top [simp]: "post_fun p p = \<top>"
by (simp add: post_fun_def)
lemma post_refin [simp]: "mono S \<Longrightarrow> ((S p)::'a::bounded_lattice) \<sqinter> (post_fun p) x \<le> S x"
apply (simp add: le_fun_def assert_fun_def post_fun_def, safe)
by (rule_tac f = S in monoD, simp_all)
class post_mbt_algebra = mbt_algebra +
fixes post :: "'a \<Rightarrow> 'a"
assumes post_1: "(post x) * x * \<top> = \<top>"
and post_2: "y * x * \<top> \<sqinter> (post x) \<le> y"
instantiation MonoTran :: (complete_boolean_algebra) post_mbt_algebra
begin
lift_definition post_MonoTran :: "'a::complete_boolean_algebra MonoTran \<Rightarrow> 'a::complete_boolean_algebra MonoTran"
is "\<lambda>x. post_fun (x \<top>)"
by (rule mono_post_fun)
instance proof
fix x :: "'a MonoTran" show "post x * x * \<top> = \<top>"
apply transfer
apply (simp add: fun_eq_iff)
done
fix x y :: "'a MonoTran" show "y * x * \<top> \<sqinter> post x \<le> y"
apply transfer
apply (simp add: le_fun_def)
done
qed
end
subsection\<open>Complete monotonic Boolean transformers algebra\<close>
class complete_mbt_algebra = post_mbt_algebra + complete_distrib_lattice +
assumes Inf_comp: "(Inf X) * z = (INF x \<in> X . (x * z))"
instance MonoTran :: (complete_boolean_algebra) complete_mbt_algebra
apply intro_classes
apply transfer
apply (simp add: Inf_comp_fun)
done
context complete_mbt_algebra begin
lemma dual_Inf: "(Inf X) ^ o = (SUP x\<in> X . x ^ o)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (subst dual_le, simp)
apply (rule Inf_greatest)
apply (subst dual_le, simp)
apply (rule SUP_upper, simp)
apply (rule SUP_least)
apply (subst dual_le, simp)
by (rule Inf_lower, simp)
lemma dual_Sup: "(Sup X) ^ o = (INF x\<in> X . x ^ o)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule INF_greatest)
apply (subst dual_le, simp)
apply (rule Sup_upper, simp)
apply (subst dual_le, simp)
apply (rule Sup_least)
apply (subst dual_le, simp)
by (rule INF_lower, simp)
lemma INF_comp: "(\<Sqinter>(f ` A)) * z = (INF a \<in> A . (f a) * z)"
unfolding Inf_comp
apply (subgoal_tac "((\<lambda>x::'a. x * z) ` f ` A) = ((\<lambda>a::'b. f a * z) ` A)")
by auto
lemma dual_INF: "(\<Sqinter>(f ` A)) ^ o = (SUP a \<in> A . (f a) ^ o)"
unfolding Inf_comp dual_Inf
apply (subgoal_tac "(dual ` f ` A) = ((\<lambda>a::'b. f a ^ o) ` A)")
by auto
lemma dual_SUP: "(\<Squnion>(f ` A)) ^ o = (INF a \<in> A . (f a) ^ o)"
unfolding dual_Sup
apply (subgoal_tac "(dual ` f ` A) = ((\<lambda>a::'b. f a ^ o) ` A)")
by auto
lemma Sup_comp: "(Sup X) * z = (SUP x \<in> X . (x * z))"
apply (rule dual_eq)
by (simp add: dual_comp dual_Sup dual_SUP INF_comp image_comp)
lemma SUP_comp: "(\<Squnion>(f ` A)) * z = (SUP a \<in> A . (f a) * z)"
unfolding Sup_comp
apply (subgoal_tac "((\<lambda>x::'a. x * z) ` f ` A) = ((\<lambda>a::'b. f a * z) ` A)")
by auto
lemma Sup_assertion [simp]: "X \<subseteq> assertion \<Longrightarrow> Sup X \<in> assertion"
apply (unfold assertion_def)
apply safe
apply (rule Sup_least)
apply blast
apply (simp add: Sup_comp dual_Sup Sup_inf)
apply (subgoal_tac "((\<lambda>y . y \<sqinter> \<Sqinter>(dual ` X)) ` (\<lambda>x . x * \<top>) ` X) = X")
apply simp
proof -
assume A: "X \<subseteq> {x. x \<le> 1 \<and> x * \<top> \<sqinter> x ^ o = x}"
have B [simp]: "!! x . x \<in> X \<Longrightarrow> x * \<top> \<sqinter> (\<Sqinter>(dual ` X)) = x"
proof -
fix x
assume C: "x \<in> X"
have "x * \<top> \<sqinter> \<Sqinter>(dual ` X) = x * \<top> \<sqinter> (x ^ o \<sqinter> \<Sqinter>(dual ` X))"
apply (subgoal_tac "\<Sqinter>(dual ` X) = (x ^ o \<sqinter> \<Sqinter>(dual ` X))", simp)
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
apply (rule Inf_lower, cut_tac C, simp)
done
also have "\<dots> = x \<sqinter> \<Sqinter>(dual ` X)" by (unfold inf_assoc [THEN sym], cut_tac A, cut_tac C, auto)
also have "\<dots> = x"
- apply (rule antisym, simp_all)
+ apply (rule order.antisym, simp_all)
apply (rule INF_greatest)
apply (cut_tac A C)
apply (rule_tac y = 1 in order_trans)
apply auto[1]
apply (subst dual_le, auto)
done
finally show "x * \<top> \<sqinter> \<Sqinter>(dual ` X) = x" .
qed
show "(\<lambda>y. y \<sqinter> \<Sqinter>(dual ` X)) ` (\<lambda>x . x * \<top>) ` X = X"
by (simp add: image_comp)
qed
lemma Sup_range_assertion [simp]: "(!!w . p w \<in> assertion) \<Longrightarrow> Sup (range p) \<in> assertion"
by (rule Sup_assertion, auto)
lemma Sup_less_assertion [simp]: "(!!w . p w \<in> assertion) \<Longrightarrow> Sup_less p w \<in> assertion"
by (unfold Sup_less_def, rule Sup_assertion, auto)
theorem omega_lfp:
"x ^ \<omega> * y = lfp (\<lambda> z . (x * z) \<sqinter> y)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule lfp_greatest)
apply (drule omega_least, simp)
apply (rule lfp_lowerbound)
apply (subst (2) omega_fix)
by (simp add: inf_comp mult.assoc)
end
lemma [simp]: "mono (\<lambda> (t::'a::mbt_algebra) . x * t \<sqinter> y)"
apply (simp add: mono_def, safe)
apply (rule_tac y = "x * xa" in order_trans, simp)
by (rule le_comp, simp)
class mbt_algebra_fusion = mbt_algebra +
assumes fusion: "(\<forall> t . x * t \<sqinter> y \<sqinter> z \<le> u * (t \<sqinter> z) \<sqinter> v)
\<Longrightarrow> (x ^ \<omega>) * y \<sqinter> z \<le> (u ^ \<omega>) * v "
lemma
"class.mbt_algebra_fusion (1::'a::complete_mbt_algebra) ((*)) (\<sqinter>) (\<le>) (<) (\<squnion>) dual dual_star omega star \<bottom> \<top>"
apply unfold_locales
apply (cut_tac h = "\<lambda> t . t \<sqinter> z" and f = "\<lambda> t . x * t \<sqinter> y" and g = "\<lambda> t . u * t \<sqinter> v" in weak_fusion)
apply (rule inf_Disj)
apply simp_all
apply (simp add: le_fun_def)
by (simp add: omega_lfp)
context mbt_algebra_fusion
begin
lemma omega_star: "x \<in> conjunctive \<Longrightarrow> x ^ \<omega> = wpt (x ^ \<omega>) * (x ^ *)"
apply (simp add: wpt_def inf_comp)
- apply (rule antisym)
+ apply (rule order.antisym)
apply (cut_tac x = x and y = 1 and z = "x ^ \<omega> * \<top> \<sqinter> x ^ *" in omega_least)
apply (simp_all add: conjunctiveD,safe)
apply (subst (2) omega_fix)
apply (simp add: inf_comp inf_assoc mult.assoc)
apply (metis inf.commute inf_assoc inf_le1 star_fix)
apply (cut_tac x = x and y = \<top> and z = "x ^ *" and u = x and v = 1 in fusion)
apply (simp add: conjunctiveD)
apply (metis inf_commute inf_le1 le_infE star_fix)
by (metis mult.right_neutral)
lemma omega_pres_conj: "x \<in> conjunctive \<Longrightarrow> x ^ \<omega> \<in> conjunctive"
apply (subst omega_star, simp)
apply (rule comp_pres_conj)
apply (rule assertion_conjunctive, simp)
by (rule start_pres_conj, simp)
end
end
diff --git a/thys/Multirelations/C_Algebras.thy b/thys/Multirelations/C_Algebras.thy
--- a/thys/Multirelations/C_Algebras.thy
+++ b/thys/Multirelations/C_Algebras.thy
@@ -1,2187 +1,2187 @@
(*
Title: Binary Multirelations
Author: Hitoshi Furusawa, Georg Struth
Maintainer: <g.struth at sheffield.ac.uk>
*)
section \<open>C-Algebras\<close>
theory C_Algebras
imports Kleene_Algebra.Dioid
begin
no_notation
times (infixl "\<cdot>" 70)
subsection \<open>C-Monoids\<close>
text \<open>We start with the c-monoid axioms. These can be found in Section~4 of~\cite{FurusawaS15a}.\<close>
class proto_monoid =
fixes s_id :: "'a" ("1\<^sub>\<sigma>")
and s_prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<cdot>" 80)
assumes s_prod_idl [simp]: "1\<^sub>\<sigma> \<cdot> x = x"
and s_prod_idr [simp]: "x \<cdot> 1\<^sub>\<sigma> = x"
class proto_bi_monoid = proto_monoid +
fixes c_id :: "'a" ("1\<^sub>\<pi>")
and c_prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<parallel>" 80)
assumes c_prod_idl [simp]: "1\<^sub>\<pi> \<parallel> x = x"
and c_prod_assoc: "(x \<parallel> y) \<parallel> z = x \<parallel> (y \<parallel> z)"
and c_prod_comm: "x \<parallel> y = y \<parallel> x"
class c_monoid = proto_bi_monoid +
assumes c1 [simp]: "(x \<cdot> 1\<^sub>\<pi>) \<parallel> x = x"
and c2 [simp]: "((x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>) \<cdot> y = (x \<cdot> 1\<^sub>\<pi>) \<parallel> y"
and c3: "(x \<parallel> y) \<cdot> 1\<^sub>\<pi> = (x \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> 1\<^sub>\<pi>)"
and c4: "(x \<cdot> y) \<cdot> 1\<^sub>\<pi> = x \<cdot> (y \<cdot> 1\<^sub>\<pi>)"
and c5 [simp]: "1\<^sub>\<sigma> \<parallel> 1\<^sub>\<sigma> = 1\<^sub>\<sigma>"
begin
text \<open>Next we define domain explicitly as at the beginning of Section 4 in~\cite{FurusawaS15a}
and start proving the algebraic facts from Section 4. Those involving concrete multirelations, such as Proposition 4.1,
are considered in the theory file for multirelations.\<close>
definition (in c_monoid) d :: "'a \<Rightarrow> 'a" where
"d x = (x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
lemma c_prod_idr [simp]: "x \<parallel> 1\<^sub>\<pi> = x"
by (simp add: local.c_prod_comm)
text \<open>We prove the retraction properties of Lemma 4.2.\<close>
lemma c_idem [simp]: "1\<^sub>\<pi> \<cdot> 1\<^sub>\<pi> = 1\<^sub>\<pi>"
by (metis c_prod_idr local.c1)
lemma d_idem [simp]: "d (d x) = d x"
by (simp add: local.d_def)
lemma p_id_idem: "(x \<cdot> 1\<^sub>\<pi>) \<cdot> 1\<^sub>\<pi> = x \<cdot> 1\<^sub>\<pi>"
by (simp add: local.c4)
text \<open>Lemma 4.3.\<close>
lemma c2_d: "d x \<cdot> y = (x \<cdot> 1\<^sub>\<pi>) \<parallel> y"
by (simp add: local.d_def)
lemma cd_2_var: "d (x \<cdot> 1\<^sub>\<pi>) \<cdot> y = (x \<cdot> 1\<^sub>\<pi>) \<parallel> y"
by (simp add: c2_d local.c4)
lemma dc_prop1 [simp]: "d x \<cdot> 1\<^sub>\<pi> = x \<cdot> 1\<^sub>\<pi>"
by (simp add: c2_d)
lemma dc_prop2 [simp]: "d (x \<cdot> 1\<^sub>\<pi>) = d x"
by (simp add: local.c4 local.d_def)
lemma ds_prop [simp]: "d x \<parallel> 1\<^sub>\<sigma> = d x"
by (simp add: local.c_prod_assoc local.d_def)
lemma dc [simp]: "d 1\<^sub>\<pi> = 1\<^sub>\<sigma>"
by (simp add: local.d_def)
text \<open>Part (5) of this Lemma has already been verified above. The next two statements
verify the two algebraic properties mentioned in the proof of Proposition 4.4.\<close>
lemma dc_iso [simp]: "d (d x \<cdot> 1\<^sub>\<pi>) = d x"
by simp
lemma cd_iso [simp]: "d (x \<cdot> 1\<^sub>\<pi>) \<cdot> 1\<^sub>\<pi> = x \<cdot> 1\<^sub>\<pi>"
by simp
text \<open>Proposition 4.5.\<close>
lemma d_conc6: "d (x \<parallel> y) = d x \<parallel> d y"
proof -
have "d (x \<parallel> y) = ((x \<parallel> y) \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
by (simp add: local.d_def)
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
by (simp add: local.c3)
finally show ?thesis
by (metis ds_prop local.c_prod_assoc local.c_prod_comm local.d_def)
qed
lemma d_conc_s_prod_ax: "d x \<parallel> d y = d x \<cdot> d y"
proof -
have "d x \<parallel> d y = (x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma> \<parallel> d y"
using local.d_def by presburger
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> d y"
using d_conc6 local.c3 local.c_prod_assoc local.d_def by auto
also have "... = ((x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>) \<cdot> d y"
by simp
finally show ?thesis
using local.d_def by auto
qed
lemma d_rest_ax [simp]: "d x \<cdot> x = x"
by (simp add: c2_d)
lemma d_loc_ax [simp]: "d (x \<cdot> d y) = d (x \<cdot> y)"
proof -
have "d (x \<cdot> d y) = (x \<cdot> d y \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
by (simp add: local.d_def)
also have "... = (x \<cdot> y \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
by (simp add: local.c4)
finally show ?thesis
by (simp add: local.d_def)
qed
lemma d_exp_ax [simp]: "d (d x \<cdot> y) = d x \<cdot> d y"
proof -
have "d (d x \<cdot> y) = d (d x \<cdot> d y)"
by (simp add: d_conc6)
also have "... = d (d (x \<parallel> y))"
by (simp add: d_conc6 d_conc_s_prod_ax)
also have "... = d (x \<parallel> y)"
by simp
finally show ?thesis
by (simp add: d_conc6 d_conc_s_prod_ax)
qed
lemma d_comm_ax: "d x \<cdot> d y = d y \<cdot> d x"
proof -
have "(d x) \<cdot> (d y) = d (x \<parallel> y)"
by (simp add: d_conc6 d_conc_s_prod_ax)
also have "... = d (y \<parallel> x)"
using local.c_prod_comm by auto
finally show ?thesis
by (simp add: d_conc6 d_conc_s_prod_ax)
qed
lemma d_s_id_prop [simp]: "d 1\<^sub>\<sigma> = 1\<^sub>\<sigma>"
using local.d_def by auto
text \<open>Next we verify the conditions of Proposition 4.6.\<close>
lemma d_s_prod_closed [simp]: "d (d x \<cdot> d y) = d x \<cdot> d y"
by simp
lemma d_p_prod_closed [simp]: "d (d x \<parallel> d y) = d x \<parallel> d y"
using c2_d d_conc6 by auto
lemma d_idem2 [simp]: "d x \<cdot> d x = d x"
by (metis d_exp_ax d_rest_ax)
lemma d_assoc: "(d x \<cdot> d y) \<cdot> d z = d x \<cdot> (d y \<cdot> d z)"
proof -
have "\<And>x y. d x \<cdot> d y = d (x \<parallel> y)"
by (simp add: d_conc6 d_conc_s_prod_ax)
thus ?thesis
by (simp add: local.c_prod_assoc)
qed
lemma iso_1 [simp]: "(d x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma> = d x"
by (simp add: local.d_def)
text \<open>Lemma 4.7.\<close>
lemma x_c_par_idem [simp]: "(x \<cdot> 1\<^sub>\<pi>) \<parallel> (x \<cdot> 1\<^sub>\<pi>) = x \<cdot> 1\<^sub>\<pi>"
proof -
have "(x \<cdot> 1\<^sub>\<pi>) \<parallel> (x \<cdot> 1\<^sub>\<pi>) = d x \<cdot> (x \<cdot> 1\<^sub>\<pi>)"
using c2_d by auto
also have "... = d (x \<cdot> 1\<^sub>\<pi>) \<cdot> (x \<cdot> 1\<^sub>\<pi>)"
by simp
finally show ?thesis
using d_rest_ax by presburger
qed
lemma d_idem_par [simp]: "d x \<parallel> d x = d x "
by (simp add: d_conc_s_prod_ax)
lemma d_inter_r: "d x \<cdot> (y \<parallel> z) = (d x \<cdot> y) \<parallel> (d x \<cdot> z)"
proof -
have "(d x) \<cdot> (y \<parallel> z) = (x \<cdot> 1\<^sub>\<pi>) \<parallel> y \<parallel> z"
using c2_d local.c_prod_assoc by auto
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> y \<parallel> (x \<cdot> 1\<^sub>\<pi>) \<parallel> z"
using local.c_prod_assoc local.c_prod_comm by force
finally show ?thesis
by (simp add: c2_d local.c_prod_assoc)
qed
text \<open>Now we provide the counterexamples of Lemma 4.8.\<close>
lemma "(x \<parallel> y) \<cdot> d z = (x \<cdot> d z) \<parallel> (y \<cdot> d z)"
nitpick
oops
lemma "(x \<cdot> y) \<cdot> d z = x \<cdot> (y \<cdot> d z)"
nitpick
oops
lemma "1\<^sub>\<pi> \<cdot> x = 1\<^sub>\<pi>"
nitpick
oops
end
subsection \<open>C-Trioids\<close>
text \<open>We can now define the class of c-trioids and prove properties in this class. This covers
the algebraic material of Section 5 in~\cite{FurusawaS15a}.\<close>
class proto_dioid = join_semilattice_zero + proto_monoid +
assumes s_prod_distr: "(x + y) \<cdot> z = x \<cdot> z + y \<cdot> z"
and s_prod_subdistl: "x \<cdot> y + x \<cdot> z \<le> x \<cdot> (y + z)"
and s_prod_annil [simp]: "0 \<cdot> x = 0"
begin
lemma s_prod_isol: "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
by (metis join.sup.boundedE order_prop s_prod_subdistl)
lemma s_prod_isor: "x \<le> y \<Longrightarrow> x \<cdot> z \<le> y \<cdot> z"
using local.order_prop local.s_prod_distr by auto
end
class proto_trioid = proto_dioid + proto_bi_monoid +
assumes p_prod_distl: "x \<parallel> (y + z) = x \<parallel> y + x \<parallel> z"
and p_rpd_annir [simp]: "x \<parallel> 0 = 0"
sublocale proto_trioid \<subseteq> ab_semigroup_mult c_prod
proof
fix x y z
show "x \<parallel> y \<parallel> z = x \<parallel> (y \<parallel> z)"
by (rule c_prod_assoc)
show "x \<parallel> y = y \<parallel> x"
by (rule c_prod_comm)
qed
sublocale proto_trioid \<subseteq> dioid_one_zero "(+)" "(\<parallel>)" "1\<^sub>\<pi>" 0 "(\<le>)" "(<)"
proof
fix x y z
show "(x + y) \<parallel> z = x \<parallel> z + y \<parallel> z"
by (simp add: local.c_prod_comm local.p_prod_distl)
show "1\<^sub>\<pi> \<parallel> x = x"
using local.c_prod_idl by blast
show "x \<parallel> 1\<^sub>\<pi> = x"
by (simp add: local.mult_commute)
show "0 + x = x"
by (rule add.left_neutral)
show "0 \<parallel> x = 0"
by (simp add: local.mult_commute)
show "x \<parallel> 0 = 0"
by (rule p_rpd_annir)
show "x + x = x"
by (rule add_idem)
show "x \<parallel> (y + z) = x \<parallel> y + x \<parallel> z"
by (rule p_prod_distl)
qed
class c_trioid = proto_trioid + c_monoid +
assumes c6: "x \<cdot> 1\<^sub>\<pi> \<le> 1\<^sub>\<pi>"
begin
text \<open>We show that every c-trioid is a c-monoid.\<close>
subclass c_monoid ..
subclass proto_trioid ..
lemma "1\<^sub>\<pi> \<cdot> 0 = 1\<^sub>\<pi>"
nitpick
oops
lemma zero_p_id_prop [simp]: "(x \<cdot> 0) \<cdot> 1\<^sub>\<pi> = x \<cdot> 0"
by (simp add: local.c4)
text \<open>The following facts prove and refute properties related to sequential and parallel subidentities.\<close>
lemma d_subid: "d x = x \<Longrightarrow> x \<le> 1\<^sub>\<sigma>"
by (metis local.c6 local.c_idem local.d_def local.dc local.mult_isor)
lemma "x \<le> 1\<^sub>\<sigma> \<Longrightarrow> d x = x"
nitpick
oops
lemma p_id_term: "x \<cdot> 1\<^sub>\<pi> = x \<Longrightarrow> x \<le> 1\<^sub>\<pi>"
by (metis local.c6)
lemma "x \<le> 1\<^sub>\<pi> \<Longrightarrow> x \<cdot> 1\<^sub>\<pi> = x"
nitpick
oops
text \<open>Proposition 5.1. is covered by the theory file on multirelations.
We verify the remaining conditions in Proposition 5.2.\<close>
lemma dlp_ax: "x \<le> d x \<cdot> x"
by simp
lemma d_add_ax: "d (x + y) = d x + d y"
proof -
have "d (x + y) = ((x + y) \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
using local.d_def by blast
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma> + (y \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
by (simp add: local.distrib_right local.s_prod_distr)
finally show ?thesis
by (simp add: local.d_def)
qed
lemma d_sub_id_ax: "d x \<le> 1\<^sub>\<sigma>"
proof -
have "d x = (x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
by (simp add: local.d_def)
also have "... \<le> 1\<^sub>\<pi> \<parallel> 1\<^sub>\<sigma>"
using local.c6 local.mult_isor by blast
finally show ?thesis
by simp
qed
lemma d_zero_ax [simp]: "d 0 = 0"
by (simp add: local.d_def)
text\<open>We verify the algebraic conditions in Proposition 5.3.\<close>
lemma d_absorb1 [simp]: "d x + (d x \<cdot> d y) = d x"
-proof (rule antisym)
+proof (rule order.antisym)
have "d x + (d x \<cdot> d y) \<le> d x + (d x \<cdot> 1\<^sub>\<sigma>)"
by (metis d_sub_id_ax c2_d d_def join.sup.bounded_iff join.sup.semilattice_axioms join.sup_ge1 s_prod_isol semilattice.idem)
thus "d x + (d x \<cdot> d y) \<le> d x"
by simp
show "d x \<le> d x + ((d x) \<cdot> (d y))"
using join.sup_ge1 by blast
qed
lemma d_absorb2 [simp]: "d x \<cdot> (d x + d y) = d x"
proof -
have "x \<cdot> 1\<^sub>\<pi> \<parallel> d x = d x"
by (metis local.c1 local.dc_prop1)
thus ?thesis
by (metis d_absorb1 local.c2_d local.p_prod_distl)
qed
lemma d_dist1: "d x \<cdot> (d y + d z) = d x \<cdot> d y + d x \<cdot> d z"
by (simp add: local.c2_d local.p_prod_distl)
lemma d_dist2: "d x + (d y \<cdot> d z) = (d x + d y) \<cdot> (d x + d z)"
proof -
have "(d x + d y) \<cdot> (d x + d z) = d x \<cdot> d x + d x \<cdot> d z + d y \<cdot> d x + d y \<cdot> d z"
using add_assoc d_dist1 local.s_prod_distr by force
also have "... = d x + d x \<cdot> d z + d x \<cdot> d y + d y \<cdot> d z"
using local.d_comm_ax by auto
finally show ?thesis
by simp
qed
lemma d_add_prod_closed [simp]: "d (d x + d y) = d x + d y"
by (simp add: d_add_ax)
text \<open>The following properties are not covered in the article.\<close>
lemma x_zero_prop: "(x \<cdot> 0) \<parallel> y = d (x \<cdot> 0) \<cdot> y"
by (simp add: local.c2_d)
lemma cda_add_ax: "d ((x + y) \<cdot> z) = d (x \<cdot> z) + d (y \<cdot> z)"
by (simp add: d_add_ax local.s_prod_distr)
lemma d_x_zero: "d (x \<cdot> 0) = (x \<cdot> 0) \<parallel> 1\<^sub>\<sigma>"
by (simp add: x_zero_prop)
text \<open>Lemma 5.4 is verified below because its proofs are simplified by using facts from the next subsection.\<close>
subsection \<open>Results for Concurrent Dynamic Algebra\<close>
text \<open>The following proofs and refutation are related to Section 6 in~\cite{FurusawaS15a}.
We do not consider those involving Kleene algebras in this section. We also do not introduce specific
notation for diamond operators.\<close>
text \<open>First we prove Lemma 6.1. Part (1) and (3) have already been verified above. Part (2) and (4) require
additional assumptions which are present in the context of concurrent dynamic algebra~\cite{FurusawaS15b}. We
also present the counterexamples from Lemma 6.3.\<close>
lemma "(x \<cdot> y) \<cdot> d z = x \<cdot> (y \<cdot> d z)"
nitpick
oops
lemma "d((x \<cdot> y) \<cdot> z) = d (x \<cdot> d (y \<cdot> z))"
nitpick
oops
lemma cda_ax1: "(x \<cdot> y) \<cdot> d z = x \<cdot> (y \<cdot> d z) \<Longrightarrow> d((x \<cdot> y) \<cdot> z) = d (x \<cdot> d (y \<cdot> z))"
by (metis local.d_loc_ax)
lemma d_inter: "(x \<parallel> y) \<cdot> d z = (x \<cdot> d z) \<parallel> (y \<cdot> d z)"
nitpick
oops
lemma "d ((x \<parallel> y) \<cdot> z) = d (x \<cdot> z) \<cdot> d (y \<cdot> z)"
nitpick
oops
lemma cda_ax2:
assumes "(x \<parallel> y) \<cdot> d z = (x \<cdot> d z) \<parallel> (y \<cdot> d z)"
shows "d ((x \<parallel> y) \<cdot> z) = d (x \<cdot> z) \<cdot> d (y \<cdot> z)"
by (metis assms local.d_conc6 local.d_conc_s_prod_ax local.d_loc_ax)
text \<open>Next we present some results that do not feature in the article.\<close>
lemma "(x \<cdot> y) \<cdot> 0 = x \<cdot> (y \<cdot> 0)"
nitpick
oops
lemma d_x_zero_prop [simp]: "d (x \<cdot> 0) \<cdot> 1\<^sub>\<pi> = x \<cdot> 0"
by simp
lemma "x \<le> 1\<^sub>\<sigma> \<and> y \<le> 1\<^sub>\<sigma> \<longrightarrow> x \<cdot> y = x \<parallel> y"
nitpick
oops
lemma "x \<cdot> (y \<parallel> z) \<le> (x \<cdot> y) \<parallel> (x \<cdot> z)"
nitpick
oops
lemma "x \<le> x \<parallel> x"
nitpick
oops
text \<open>Lemma 5.4\<close>
lemma d_lb1: "d x \<cdot> d y \<le> d x"
by (simp add: less_eq_def add_commute)
lemma d_lb2: "d x \<cdot> d y \<le> d y"
using d_lb1 local.d_comm_ax by fastforce
lemma d_glb: "d z \<le> d x \<and> d z \<le> d y \<Longrightarrow> d z \<le> d x \<cdot> d y"
by (simp add: d_dist2 local.less_eq_def)
lemma d_glb_iff: "d z \<le> d x \<and> d z \<le> d y \<longleftrightarrow> d z \<le> d x \<cdot> d y"
using d_glb d_lb1 d_lb2 local.order_trans by blast
lemma x_zero_le_c: "x \<cdot> 0 \<le> 1\<^sub>\<pi>"
by (simp add: p_id_term)
lemma p_subid_lb1: "(x \<cdot> 0) \<parallel> (y \<cdot> 0) \<le> x \<cdot> 0"
using local.mult_isol x_zero_le_c by fastforce
lemma p_subid_lb2: "(x \<cdot> 0) \<parallel> (y \<cdot> 0) \<le> y \<cdot> 0"
using local.mult_commute p_subid_lb1 by fastforce
lemma p_subid_idem [simp]: "(x \<cdot> 0) \<parallel> (x \<cdot> 0) = x \<cdot> 0"
by (metis local.c1 zero_p_id_prop)
lemma p_subid_glb: "z \<cdot> 0 \<le> x \<cdot> 0 \<and> z \<cdot> 0 \<le> y \<cdot> 0 \<Longrightarrow> z \<cdot> 0 \<le> (x \<cdot> 0) \<parallel> (y \<cdot> 0)"
using local.mult_isol_var by force
lemma p_subid_glb_iff: "z \<cdot> 0 \<le> x \<cdot> 0 \<and> z \<cdot> 0 \<le> y \<cdot> 0 \<longleftrightarrow> z \<cdot> 0 \<le> (x \<cdot> 0) \<parallel> (y \<cdot> 0)"
using local.order_trans p_subid_glb p_subid_lb1 p_subid_lb2 by blast
lemma x_c_glb: "z \<cdot> 1\<^sub>\<pi> \<le> x \<cdot> 1\<^sub>\<pi> \<and> z \<cdot> 1\<^sub>\<pi> \<le> y \<cdot> 1\<^sub>\<pi> \<Longrightarrow> z \<cdot> 1\<^sub>\<pi> \<le> (x \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> 1\<^sub>\<pi>)"
using local.mult_isol_var by force
lemma x_c_lb1: "(x \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> 1\<^sub>\<pi>) \<le> x \<cdot> 1\<^sub>\<pi>"
using local.c6 local.mult_isol_var by force
lemma x_c_lb2: "(x \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> 1\<^sub>\<pi>) \<le> y \<cdot> 1\<^sub>\<pi>"
using local.mult_commute x_c_lb1 by fastforce
lemma x_c_glb_iff: "z \<cdot> 1\<^sub>\<pi> \<le> x \<cdot> 1\<^sub>\<pi> \<and> z \<cdot> 1\<^sub>\<pi> \<le> y \<cdot> 1\<^sub>\<pi> \<longleftrightarrow> z \<cdot> 1\<^sub>\<pi> \<le> (x \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> 1\<^sub>\<pi>)"
by (meson local.order.trans x_c_glb x_c_lb1 x_c_lb2)
end
subsection \<open>C-Lattices\<close>
text \<open>We can now define c-lattices and prove the results from Section 7 in~\cite{FurusawaS15a}.\<close>
class pbl_monoid = proto_trioid +
fixes U :: 'a
fixes meet :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70)
assumes U_def: "x \<le> U"
and meet_assoc: "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
and meet_comm: "x \<sqinter> y = y \<sqinter> x"
and meet_idem [simp]: "x \<sqinter> x = x"
and absorp1: "x \<sqinter> (x + y) = x"
and absorp2: "x + (x \<sqinter> y) = x"
begin
sublocale lattice "(\<sqinter>)" "(\<le>)" "(<)" "(+)"
proof
show a: "\<And>x y. x \<sqinter> y \<le> x"
by (simp add: local.absorp2 local.less_eq_def add_commute)
show b: " \<And>x y. x \<sqinter> y \<le> y"
using a local.meet_comm by fastforce
show " \<And>x y z. x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z"
by (metis b local.absorp1 local.less_eq_def local.meet_assoc)
qed
lemma meet_glb: "z \<le> x \<and> z \<le> y \<Longrightarrow> z \<le> x \<sqinter> y"
by simp
lemma meet_prop: "z \<le> x \<and> z \<le> y \<longleftrightarrow> z \<le> x \<sqinter> y"
by simp
end
class pbdl_monoid = pbl_monoid +
assumes lat_dist1: "x + (y \<sqinter> z) = (x + y) \<sqinter> (x + z)"
begin
lemma lat_dist2: "(x \<sqinter> y) + z = (x + z) \<sqinter> (y + z)"
by (simp add: local.lat_dist1 add_commute)
lemma lat_dist3: "x \<sqinter> (y + z) = (x \<sqinter> y) + (x \<sqinter> z)"
proof -
have "\<And>x y z. x \<sqinter> ((x + y) \<sqinter> z) = x \<sqinter> z"
by (metis local.absorp1 local.meet_assoc)
thus ?thesis
using lat_dist2 local.absorp2 add_commute by force
qed
lemma lat_dist4: "(x + y) \<sqinter> z = (x \<sqinter> z) + (y \<sqinter> z)"
using lat_dist3 local.meet_comm by auto
lemma d_equiv_prop: "(\<forall>z. z + x = z + y \<and> z \<sqinter> x = z \<sqinter> y) \<Longrightarrow> x = y"
by (metis local.add_zerol)
end
text \<open>The symbol $\overline{1}_\pi$ from~\cite{FurusawaS15a} is written nc in this theory file.\<close>
class c_lattice = pbdl_monoid +
fixes nc :: "'a"
assumes cl1 [simp]: "x \<cdot> 1\<^sub>\<pi> + x \<cdot> nc = x \<cdot> U"
and cl2 [simp]: "1\<^sub>\<pi> \<sqinter> (x + nc) = x \<cdot> 0"
and cl3: "x \<cdot> (y \<parallel> z) \<le> (x \<cdot> y) \<parallel> (x \<cdot> z)"
and cl4: "z \<parallel> z \<le> z \<Longrightarrow> (x \<parallel> y) \<cdot> z = (x \<cdot> z) \<parallel> (y \<cdot> z)"
and cl5: "x \<cdot> (y \<cdot> (z \<cdot> 0)) = (x \<cdot> y) \<cdot> (z \<cdot> 0)"
and cl6 [simp]: "(x \<cdot> 0) \<cdot> z = x \<cdot> 0"
and cl7 [simp]: "1\<^sub>\<sigma> \<parallel> 1\<^sub>\<sigma> = 1\<^sub>\<sigma>"
and cl8 [simp]: "((x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>) \<cdot> y = (x \<cdot> 1\<^sub>\<pi>) \<parallel> y"
and cl9 [simp]: "((x \<sqinter> 1\<^sub>\<sigma>) \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma> = x \<sqinter> 1\<^sub>\<sigma>"
and cl10: "((x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma> = 1\<^sub>\<sigma> \<sqinter> (x \<sqinter> nc) \<cdot> nc"
and cl11 [simp]: "((x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>) \<parallel> nc = (x \<sqinter> nc) \<cdot> nc"
begin
text \<open>We show that every c-lattice is a c-trioid (Proposition 7.1) Proposition 7.2 is again
covered by the theory for multirelations.\<close>
subclass c_trioid
proof
fix x y
show "x \<cdot> 1\<^sub>\<pi> \<parallel> 1\<^sub>\<sigma> \<cdot> y = x \<cdot> 1\<^sub>\<pi> \<parallel> y"
by auto
show "x \<parallel> y \<cdot> 1\<^sub>\<pi> = x \<cdot> 1\<^sub>\<pi> \<parallel> (y \<cdot> 1\<^sub>\<pi>)"
by (simp add: local.cl4)
show "x \<cdot> y \<cdot> 1\<^sub>\<pi> = x \<cdot> (y \<cdot> 1\<^sub>\<pi>)"
by (metis local.absorp1 local.cl2 local.cl5)
show "1\<^sub>\<sigma> \<parallel> 1\<^sub>\<sigma> = 1\<^sub>\<sigma>"
by (meson local.cl7)
show x: "x \<cdot> 1\<^sub>\<pi> \<le> 1\<^sub>\<pi>"
by (metis local.absorp1 local.cl2 local.cl5 local.inf_le1 local.s_prod_idl)
show "x \<cdot> 1\<^sub>\<pi> \<parallel> x = x"
- by (metis x local.eq_iff local.cl3 local.mult_1_right local.mult_commute local.mult_isol local.s_prod_idr)
+ by (metis x order.eq_iff local.cl3 local.mult_1_right local.mult_commute local.mult_isol local.s_prod_idr)
qed
text \<open>First we verify the complementation conditions after the definition of c-lattices.\<close>
lemma c_nc_comp1 [simp]: "1\<^sub>\<pi> + nc = U"
by (metis local.cl1 local.s_prod_idl)
lemma c_nc_comp2 [simp]: "1\<^sub>\<pi> \<sqinter> nc = 0"
by (metis local.add_zero_l local.cl2 local.s_prod_annil)
lemma c_0: "x \<sqinter> 1\<^sub>\<pi> = x \<cdot> 0"
by (metis c_nc_comp2 local.add_zeror local.cl2 local.lat_dist3 local.meet_comm)
text \<open>Next we verify the conditions in Proposition 7.2.\<close>
lemma d_s_subid: "d x = x \<longleftrightarrow> x \<le> 1\<^sub>\<sigma>"
by (metis local.cl9 local.d_def local.d_subid local.inf.absorb_iff1)
lemma term_p_subid: "x \<cdot> 1\<^sub>\<pi> = x \<longleftrightarrow> x \<le> 1\<^sub>\<pi>"
by (metis c_0 local.cl6 local.inf.absorb_iff1 local.p_id_term)
lemma term_p_subid_var: "x \<cdot> 0 = x \<longleftrightarrow> x \<le> 1\<^sub>\<pi>"
using c_0 local.inf.absorb_iff1 by auto
lemma vec_iff: "d x \<cdot> U = x \<longleftrightarrow> (x \<cdot> 1\<^sub>\<pi>) \<parallel> U = x"
by (simp add: local.c2_d)
lemma nc_iff1: "x \<le> nc \<longleftrightarrow> x \<sqinter> 1\<^sub>\<pi> = 0"
proof
fix x
assume assm: "x \<le> nc"
hence "x = x \<sqinter> nc"
by (simp add: local.inf.absorb_iff1)
hence "x \<sqinter> 1\<^sub>\<pi> = x \<sqinter> nc \<sqinter> 1\<^sub>\<pi>"
by auto
then show "x \<sqinter> 1\<^sub>\<pi> = 0"
by (metis assm c_0 c_nc_comp2 local.cl2 local.less_eq_def)
next
fix x
assume assm: "x \<sqinter> 1\<^sub>\<pi> = 0"
have "x = (x \<sqinter> nc) + (x \<sqinter> 1\<^sub>\<pi>)"
by (metis c_nc_comp1 local.U_def local.add_comm local.lat_dist3 local.inf.absorb_iff1)
hence "x = x \<sqinter> nc"
using assm by auto
thus "x \<le> nc"
using local.inf.absorb_iff1 by auto
qed
lemma nc_iff2: "x \<le> nc \<longleftrightarrow> x \<cdot> 0 = 0"
using c_0 nc_iff1 by auto
text \<open>The results of Lemma 7.3 are again at the multirelational level.
Hence we continue with Lemma 7.4.\<close>
lemma assoc_p_subid: "(x \<cdot> y) \<cdot> (z \<cdot> 1\<^sub>\<pi>) = x \<cdot> (y \<cdot> (z \<cdot> 1\<^sub>\<pi>))"
by (metis c_0 local.c6 local.cl5 local.inf.absorb_iff1)
lemma zero_assoc3: "(x \<cdot> y) \<cdot> 0 = x \<cdot> (y \<cdot> 0)"
by (metis local.cl5 local.s_prod_annil)
lemma x_zero_interr: "(x \<cdot> 0) \<parallel> (y \<cdot> 0) = (x \<parallel> y) \<cdot> 0"
by (simp add: local.cl4)
lemma p_subid_interr: "(x \<cdot> z \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> z \<cdot> 1\<^sub>\<pi>) = (x \<parallel> y) \<cdot> z \<cdot> 1\<^sub>\<pi>"
by (simp add: local.c4 local.cl4)
lemma d_interr: "(x \<cdot> d z) \<parallel> (y \<cdot> d z) = (x \<parallel> y) \<cdot> d z"
by (simp add: local.cl4)
lemma subidem_par: "x \<le> x \<parallel> x"
proof -
have "x = x \<cdot> 1\<^sub>\<sigma>"
by auto
also have "... = x \<cdot> (1\<^sub>\<sigma> \<parallel> 1\<^sub>\<sigma>)"
by auto
finally show ?thesis
by (metis local.cl3 local.cl7)
qed
lemma meet_le_par: "x \<sqinter> y \<le> x \<parallel> y"
proof -
have "x \<sqinter> y = (x \<sqinter> y) \<sqinter> (x \<sqinter> y)"
using local.meet_idem by presburger
thus ?thesis
using local.inf_le1 local.inf_le2 local.mult_isol_var local.order_trans subidem_par by blast
qed
text\<open>Next we verify Lemma 7.5 and prove some related properties.\<close>
lemma x_split [simp]: "(x \<sqinter> nc) + (x \<sqinter> 1\<^sub>\<pi>) = x"
proof -
have "x = x \<sqinter> U"
using local.U_def local.inf.absorb_iff1 by auto
also have "... = x \<sqinter> (nc + 1\<^sub>\<pi>)"
by (simp add: add_commute)
finally show ?thesis
by (metis local.lat_dist3)
qed
lemma x_split_var [simp]: "(x \<sqinter> nc) + (x \<cdot> 0) = x"
by (metis local.c_0 x_split)
lemma s_subid_closed [simp]: "x \<sqinter> nc \<sqinter> 1\<^sub>\<sigma> = x \<sqinter> 1\<^sub>\<sigma>"
proof -
have "x \<sqinter> 1\<^sub>\<sigma> = ((x \<sqinter> nc) + (x \<sqinter> 1\<^sub>\<pi>)) \<sqinter> 1\<^sub>\<sigma>"
using x_split by presburger
also have "... = (x \<sqinter> nc \<sqinter> 1\<^sub>\<sigma>) + (x \<sqinter> 1\<^sub>\<pi> \<sqinter> 1\<^sub>\<sigma>)"
by (simp add: local.lat_dist3 local.meet_comm)
also have "... = (x \<sqinter> nc \<sqinter> 1\<^sub>\<sigma>) + (x \<sqinter> 0)"
by (metis c_0 local.meet_assoc local.meet_comm local.s_prod_idl)
finally show ?thesis
by (metis local.absorp1 local.add_zeror local.lat_dist1 local.meet_comm)
qed
lemma sub_id_le_nc: "x \<sqinter> 1\<^sub>\<sigma> \<le> nc"
by (metis local.inf.absorb_iff2 local.inf_left_commute local.meet_comm s_subid_closed)
lemma s_x_c [simp]: "1\<^sub>\<sigma> \<sqinter> (x \<cdot> 1\<^sub>\<pi>) = 0"
proof -
have "1\<^sub>\<sigma> \<sqinter> 1\<^sub>\<pi> = 0"
using c_0 local.s_prod_idl by presburger
hence "1\<^sub>\<sigma> \<sqinter> x \<cdot> 1\<^sub>\<pi> \<le> 0"
using local.c6 local.inf_le1 local.inf_le2 local.meet_prop local.order.trans by blast
thus ?thesis
using local.less_eq_def local.no_trivial_inverse by blast
qed
lemma s_x_zero [simp]: "1\<^sub>\<sigma> \<sqinter> (x \<cdot> 0) = 0"
by (metis local.cl6 s_x_c)
lemma c_nc [simp]: "(x \<cdot> 1\<^sub>\<pi>) \<sqinter> nc = 0"
proof -
have "x \<cdot> 1\<^sub>\<pi> \<sqinter> nc \<le> 1\<^sub>\<pi>"
by (meson local.c6 local.dual_order.trans local.inf_le1)
thus ?thesis
by (metis local.inf_le2 nc_iff2 term_p_subid_var)
qed
lemma zero_nc [simp]: "(x \<cdot> 0) \<sqinter> nc = 0"
by (metis c_nc local.cl6)
lemma nc_zero [simp]: "(x \<sqinter> nc) \<cdot> 0 = 0"
by (meson local.inf_le2 nc_iff2)
text \<open>Lemma 7.6.\<close>
lemma c_def [simp]: "U \<cdot> 0 = 1\<^sub>\<pi>"
by (metis c_nc_comp1 c_0 local.absorp1 local.meet_comm)
lemma c_x_prop [simp]: "1\<^sub>\<pi> \<cdot> x = 1\<^sub>\<pi>"
using c_def local.cl6 by blast
lemma U_idem_s_prod [simp]: "U \<cdot> U = U"
- by (metis local.U_def local.eq_iff local.s_prod_idl local.s_prod_isor)
+ by (metis local.U_def order.eq_iff local.s_prod_idl local.s_prod_isor)
lemma U_idem_p_prod [simp]: "U \<parallel> U = U"
- using local.U_def local.eq_iff subidem_par by presburger
+ using local.U_def order.eq_iff subidem_par by presburger
lemma U_c [simp]: "U \<cdot> 1\<^sub>\<pi> = 1\<^sub>\<pi>"
by (metis U_idem_s_prod local.c_def zero_assoc3)
lemma s_le_nc: "1\<^sub>\<sigma> \<le> nc"
by (metis local.meet_idem sub_id_le_nc)
lemma nc_c [simp]: "nc \<cdot> 1\<^sub>\<pi> = 1\<^sub>\<pi>"
-proof (rule antisym)
+proof (rule order.antisym)
have "nc \<cdot> 1\<^sub>\<pi> = nc \<cdot> 1\<^sub>\<pi> \<cdot> 0"
by (simp add: zero_assoc3)
also have "... = nc \<cdot> 1\<^sub>\<pi> \<sqinter> 1\<^sub>\<pi>"
by (simp add: c_0)
finally show "nc \<cdot> 1\<^sub>\<pi> \<le> 1\<^sub>\<pi>"
using local.c6 by blast
show "1\<^sub>\<pi> \<le> nc \<cdot> 1\<^sub>\<pi>"
using local.s_prod_isor s_le_nc by fastforce
qed
lemma nc_nc [simp]: "nc \<cdot> nc = nc"
proof -
have "nc \<cdot> nc = (nc \<cdot> 1\<^sub>\<pi>) \<parallel> nc"
by (metis local.cl11 local.meet_idem)
thus ?thesis
by simp
qed
lemma U_nc [simp]: "U \<cdot> nc = U"
proof -
have "U \<cdot> nc = (1\<^sub>\<pi> + nc) \<cdot> nc"
by force
also have "... = 1\<^sub>\<pi> \<cdot> nc + nc \<cdot> nc"
using local.s_prod_distr by blast
also have "... = 1\<^sub>\<pi> + nc"
by simp
finally show ?thesis
by auto
qed
lemma nc_U [simp]: "nc \<cdot> U = U"
proof -
have "nc \<cdot> U = nc \<cdot> 1\<^sub>\<pi> + nc \<cdot> nc"
using local.cl1 by presburger
thus ?thesis
by simp
qed
lemma nc_nc_par [simp]: "nc \<parallel> nc = nc"
proof -
have "nc \<parallel> nc = (nc \<parallel> nc \<sqinter> nc) + (nc \<parallel> nc) \<cdot> 0"
by simp
also have "... = nc + (nc \<cdot> 0) \<parallel> (nc \<cdot> 0)"
by (metis local.meet_comm local.inf.absorb_iff1 subidem_par x_zero_interr)
also have "... = nc + 0 \<parallel> 0"
by (metis local.absorp1 local.meet_comm nc_zero)
finally show ?thesis
by (metis add_commute local.add_zerol local.annil)
qed
lemma U_nc_par [simp]: "U \<parallel> nc = nc"
proof -
have "U \<parallel> nc = nc \<parallel> nc + 1\<^sub>\<pi> \<parallel> nc"
by (metis c_nc_comp1 local.add_comm local.distrib_right)
also have "... = nc + nc"
by force
finally show ?thesis
by simp
qed
text \<open>We prove Lemma 7.8 and related properties.\<close>
lemma x_y_split [simp]: "(x \<sqinter> nc) \<cdot> y + x \<cdot> 0 = x \<cdot> y"
by (metis c_0 local.cl6 local.s_prod_distr x_split)
lemma x_y_prop: "1\<^sub>\<sigma> \<sqinter> (x \<sqinter> nc) \<cdot> y = 1\<^sub>\<sigma> \<sqinter> x \<cdot> y"
proof -
have "1\<^sub>\<sigma> \<sqinter> x \<cdot> y = 1\<^sub>\<sigma> \<sqinter> ((x \<sqinter> nc) \<cdot> y + x \<cdot> 0)"
using x_y_split by presburger
also have "... = (1\<^sub>\<sigma> \<sqinter> (x \<sqinter> nc) \<cdot> y) + (1\<^sub>\<sigma> \<sqinter> x \<cdot> 0)"
by (simp add: local.lat_dist3 add_commute)
finally show ?thesis
by (metis local.add_zeror s_x_zero)
qed
lemma s_nc_U: "1\<^sub>\<sigma> \<sqinter> x \<cdot> nc = 1\<^sub>\<sigma> \<sqinter> x \<cdot> U"
proof -
have "1\<^sub>\<sigma> \<sqinter> x \<cdot> U = 1\<^sub>\<sigma> \<sqinter> (x \<cdot> nc + x \<cdot> 1\<^sub>\<pi>)"
by (simp add: add_commute)
also have "... = (1\<^sub>\<sigma> \<sqinter> x \<cdot> nc) + (1\<^sub>\<sigma> \<sqinter> x \<cdot> 1\<^sub>\<pi>)"
using local.lat_dist3 by blast
finally show ?thesis
by (metis local.add_zeror s_x_c)
qed
lemma sid_le_nc_var: "1\<^sub>\<sigma> \<sqinter> x \<le> 1\<^sub>\<sigma> \<sqinter> x \<parallel> nc"
proof -
have "1\<^sub>\<sigma> \<sqinter> x = x \<sqinter> (1\<^sub>\<sigma> \<sqinter> nc)"
by (metis (no_types) local.inf.absorb1 local.inf.commute s_le_nc)
hence "1\<^sub>\<sigma> \<sqinter> x \<parallel> nc + 1\<^sub>\<sigma> \<sqinter> x = (x \<parallel> nc + x \<sqinter> nc) \<sqinter> 1\<^sub>\<sigma>"
using local.inf.commute local.inf.left_commute local.lat_dist4 by auto
thus ?thesis
by (metis (no_types) local.inf.commute local.join.sup.absorb_iff1 meet_le_par)
qed
lemma s_nc_par_U: "1\<^sub>\<sigma> \<sqinter> x \<parallel> nc = 1\<^sub>\<sigma> \<sqinter> x \<parallel> U"
proof -
have "1\<^sub>\<sigma> \<sqinter> x \<parallel> U = 1\<^sub>\<sigma> \<sqinter> (x \<parallel> nc + x)"
by (metis c_nc_comp1 local.add_comm local.distrib_left local.mult_oner)
also have "... = (1\<^sub>\<sigma> \<sqinter> x \<parallel> nc) + (x \<sqinter> 1\<^sub>\<sigma>)"
by (metis local.lat_dist3 local.meet_comm)
also have "... = 1\<^sub>\<sigma> \<sqinter> x \<parallel> nc"
by (metis local.add_comm local.less_eq_def local.meet_comm sid_le_nc_var)
finally show ?thesis
by metis
qed
lemma x_c_nc_split: "(x \<cdot> 1\<^sub>\<pi>) \<parallel> nc = (x \<sqinter> nc) \<cdot> nc + (x \<cdot> 0) \<parallel> nc"
by (metis local.cl11 local.mult_commute local.p_prod_distl x_y_split)
lemma x_c_U_split: "(x \<cdot> 1\<^sub>\<pi>) \<parallel> U = x \<cdot> U + (x \<cdot> 0) \<parallel> U"
proof -
have "x \<cdot> U + (x \<cdot> 0) \<parallel> U = (x \<sqinter> nc) \<cdot> U + (x \<cdot> 0) \<parallel> U"
by (metis U_c U_idem_s_prod U_nc local.add_assoc' local.cl1 local.distrib_left local.mult_oner x_y_split)
also have "... = (x \<sqinter> nc) \<cdot> nc + (x \<sqinter> nc) \<cdot> 1\<^sub>\<pi> + (x \<cdot> 0) \<parallel> nc + x \<cdot> 0"
by (metis add_commute c_nc_comp1 local.cl1 local.combine_common_factor local.mult_1_right local.mult_commute)
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> nc + x \<cdot> 1\<^sub>\<pi>"
by (metis local.add_ac(1) local.add_commute x_c_nc_split x_y_split)
thus ?thesis
by (metis c_nc_comp1 calculation local.add_comm local.distrib_left local.mult_oner)
qed
subsection \<open>Domain in C-Lattices\<close>
text \<open>We now prove variants of the domain axioms and verify the properties of Section 8 in~\cite{FurusawaS15a}.\<close>
lemma cl9_d [simp]: "d (x \<sqinter> 1\<^sub>\<sigma>) = x \<sqinter> 1\<^sub>\<sigma>"
by (simp add: local.d_def)
lemma cl10_d: "d (x \<sqinter> nc) = 1\<^sub>\<sigma> \<sqinter> (x \<sqinter> nc) \<cdot> nc"
using local.cl10 local.d_def by auto
lemma cl11_d [simp]: "d (x \<sqinter> nc) \<cdot> nc = (x \<sqinter> nc) \<cdot> nc"
using local.c2_d by force
lemma cl10_d_var1: "d (x \<sqinter> nc) = 1\<^sub>\<sigma> \<sqinter> x \<cdot> nc"
by (simp add: cl10_d x_y_prop)
lemma cl10_d_var2: "d (x \<sqinter> nc) = 1\<^sub>\<sigma> \<sqinter> (x \<sqinter> nc) \<cdot> U"
by (simp add: cl10_d s_nc_U)
lemma cl10_d_var3: "d (x \<sqinter> nc) = 1\<^sub>\<sigma> \<sqinter> x \<cdot> U"
by (simp add: cl10_d_var1 s_nc_U)
text \<open>We verify the remaining properties of Lemma 8.1.\<close>
lemma d_U [simp]: "d U = 1\<^sub>\<sigma>"
by (simp add: local.d_def)
lemma d_nc [simp]: "d nc = 1\<^sub>\<sigma>"
using local.d_def by auto
lemma alt_d_def_nc_nc: "d (x \<sqinter> nc) = 1\<^sub>\<sigma> \<sqinter> ((x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>) \<parallel> nc"
by (simp add: cl10_d_var1 x_y_prop)
lemma alt_d_def_nc_U: "d (x \<sqinter> nc) = 1\<^sub>\<sigma> \<sqinter> ((x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>) \<parallel> U"
by (metis alt_d_def_nc_nc local.c2_d s_nc_U)
text \<open>We verify the identity before Lemma 8.2 of~\cite{FurusawaS15a} together with variants.\<close>
lemma d_def_split [simp]: "d (x \<sqinter> nc) + d (x \<cdot> 0) = d x"
by (metis local.d_add_ax x_split_var)
lemma d_def_split_var [simp]: "d (x \<sqinter> nc) + (x \<cdot> 0) \<parallel> 1\<^sub>\<sigma> = d x"
by (metis d_def_split local.d_x_zero)
lemma ax7 [simp]: "(1\<^sub>\<sigma> \<sqinter> x \<cdot> U) + (x \<cdot> 0) \<parallel> 1\<^sub>\<sigma> = d x"
by (metis cl10_d_var3 d_def_split_var)
text \<open>Lemma 8.2.\<close>
lemma dom12_d: "d x = 1\<^sub>\<sigma> \<sqinter> (x \<cdot> 1\<^sub>\<pi>) \<parallel> nc"
proof -
have "1\<^sub>\<sigma> \<sqinter> (x \<cdot> 1\<^sub>\<pi>) \<parallel> nc = 1\<^sub>\<sigma> \<sqinter> ((x \<sqinter> nc) \<cdot> 1\<^sub>\<pi> + x \<cdot> 0) \<parallel> nc"
using x_y_split by presburger
also have "... = (1\<^sub>\<sigma> \<sqinter> ((x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>) \<parallel> nc) + (1\<^sub>\<sigma> \<sqinter> (x \<cdot> 0) \<parallel> nc)"
by (simp add: local.lat_dist3 local.mult_commute local.p_prod_distl add_commute)
also have "... = d (x \<sqinter> nc) + d (x \<cdot> 0)"
by (metis add_commute c_0 cl10_d_var1 local.add_zerol local.annil local.c2_d local.d_def local.mult_commute local.mult_onel local.zero_p_id_prop x_split)
finally show ?thesis
by (metis d_def_split)
qed
lemma dom12_d_U: "d x = 1\<^sub>\<sigma> \<sqinter> (x \<cdot> 1\<^sub>\<pi>) \<parallel> U"
by (simp add: dom12_d s_nc_par_U)
lemma dom_def_var: "d x = (x \<cdot> U \<sqinter> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<sigma>"
by (simp add: c_0 local.d_def zero_assoc3)
text\<open>Lemma 8.3.\<close>
lemma ax5_d [simp]: "d (x \<sqinter> nc) \<cdot> U = (x \<sqinter> nc) \<cdot> U"
proof -
have "d (x \<sqinter> nc) \<cdot> U = d (x \<sqinter> nc) \<cdot> nc + d (x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>"
using add_commute local.cl1 by presburger
also have "... = (x \<sqinter> nc) \<cdot> nc + (x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>"
by simp
finally show ?thesis
by (simp add: add_commute)
qed
lemma ax5_0 [simp]: "d (x \<cdot> 0) \<cdot> U = (x \<cdot> 0) \<parallel> U"
using local.x_zero_prop by presburger
lemma x_c_U_split2: "d x \<cdot> nc = (x \<sqinter> nc) \<cdot> nc + (x \<cdot> 0) \<parallel> nc"
by (simp add: local.c2_d x_c_nc_split)
lemma x_c_U_split3: "d x \<cdot> U = (x \<sqinter> nc) \<cdot> U + (x \<cdot> 0) \<parallel> U"
by (metis d_def_split local.s_prod_distr ax5_0 ax5_d)
lemma x_c_U_split_d: "d x \<cdot> U = x \<cdot> U + (x \<cdot> 0) \<parallel> U"
using local.c2_d x_c_U_split by presburger
lemma x_U_prop2: "x \<cdot> nc = d (x \<sqinter> nc) \<cdot> nc + x \<cdot> 0"
by (metis local.c2_d local.cl11 x_y_split)
lemma x_U_prop3: "x \<cdot> U = d (x \<sqinter> nc) \<cdot> U + x \<cdot> 0"
by (metis ax5_d x_y_split)
lemma d_x_nc [simp]: "d (x \<cdot> nc) = d x"
using local.c4 local.d_def by auto
lemma d_x_U [simp]: "d (x \<cdot> U) = d x"
by (simp add: local.c4 local.d_def)
text \<open>The next properties of domain are important, but do not feature in~\cite{FurusawaS15a}.
Proofs can be found in~\cite{FurusawaS15b}.\<close>
lemma d_llp1: "d x \<le> d y \<Longrightarrow> x \<le> d y \<cdot> x"
by (metis local.d_rest_ax local.s_prod_isor)
lemma d_llp2: "x \<le> d y \<cdot> x \<Longrightarrow> d x \<le> d y"
proof -
assume a1: "x \<le> d y \<cdot> x"
have "\<forall>x y. d (x \<parallel> y) = x \<cdot> 1\<^sub>\<pi> \<parallel> d y"
using local.c2_d local.d_conc6 local.d_conc_s_prod_ax by presburger
hence "d x \<le> d (y \<cdot> 1\<^sub>\<pi>)"
- using a1 by (metis (no_types) local.c2_d local.c6 local.c_prod_comm local.eq_iff local.mult_isol local.mult_oner)
+ using a1 by (metis (no_types) local.c2_d local.c6 local.c_prod_comm order.eq_iff local.mult_isol local.mult_oner)
thus ?thesis
by simp
qed
lemma demod1: "d (x \<cdot> y) \<le> d z \<Longrightarrow> x \<cdot> d y \<le> d z \<cdot> x"
proof -
assume "d (x \<cdot> y) \<le> d z"
hence "\<forall>v. x \<cdot> y \<cdot> 1\<^sub>\<pi> \<parallel> v \<le> z \<cdot> 1\<^sub>\<pi> \<parallel> v"
by (metis (no_types) local.c2_d local.s_prod_isor)
hence "\<forall>v. x \<cdot> (y \<cdot> 1\<^sub>\<pi> \<parallel> v) \<le> z \<cdot> 1\<^sub>\<pi> \<parallel> (x \<cdot> v)"
by (metis local.c4 local.cl3 local.dual_order.trans)
thus ?thesis
by (metis local.c2_d local.s_prod_idr)
qed
lemma demod2: "x \<cdot> d y \<le> d z \<cdot> x \<Longrightarrow> d (x \<cdot> y) \<le> d z"
proof -
assume "x \<cdot> d y \<le> d z \<cdot> x"
hence "d (x \<cdot> y) \<le> d (d z \<cdot> x)"
by (metis local.d_def local.d_loc_ax local.mult_isor local.s_prod_isor)
thus ?thesis
using local.d_conc6 local.d_conc_s_prod_ax local.d_glb_iff by fastforce
qed
subsection \<open>Structural Properties of C-Lattices\<close>
text \<open>Now we consider the results from Section 9 and 10 in~\cite{FurusawaS15a}.
First we verify the conditions for Proposition 9.1.\<close>
lemma d_meet_closed [simp]: "d (d x \<sqinter> d y) = d x \<sqinter> d y"
using d_s_subid local.d_sub_id_ax local.inf_le1 local.order_trans by blast
lemma d_s_prod_eq_meet: "d x \<cdot> d y = d x \<sqinter> d y"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis local.d_lb1 local.d_lb2 local.meet_glb)
by (metis d_meet_closed local.inf_le1 local.inf_le2 local.d_glb)
lemma d_p_prod_eq_meet: "d x \<parallel> d y = d x \<sqinter> d y"
by (simp add: d_s_prod_eq_meet local.d_conc_s_prod_ax)
lemma s_id_par_s_prod: "(x \<sqinter> 1\<^sub>\<sigma>) \<parallel> (y \<sqinter> 1\<^sub>\<sigma>) = (x \<sqinter> 1\<^sub>\<sigma>) \<cdot> (y \<sqinter> 1\<^sub>\<sigma>)"
by (metis cl9_d local.d_conc_s_prod_ax)
lemma s_id_par [simp]: "x \<sqinter> 1\<^sub>\<sigma> \<parallel> x \<sqinter> 1\<^sub>\<sigma> = x \<sqinter> 1\<^sub>\<sigma>"
using local.meet_assoc local.meet_comm local.inf.absorb_iff1 meet_le_par by auto
text \<open>We verify the remaining conditions in Proposition 9.2.\<close>
lemma p_subid_par_eq_meet: "(x \<cdot> 0) \<parallel> (y \<cdot> 0) = (x \<cdot> 0) \<sqinter> (y \<cdot> 0)"
by (simp add: local.meet_glb local.order.antisym local.p_subid_lb1 local.p_subid_lb2 meet_le_par)
lemma p_subid_par_eq_meet_var: "(x \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> 1\<^sub>\<pi>) = (x \<cdot> 1\<^sub>\<pi>) \<sqinter> (y \<cdot> 1\<^sub>\<pi>)"
by (metis c_x_prop p_subid_par_eq_meet zero_assoc3)
lemma x_zero_add_closed: "x \<cdot> 0 + y \<cdot> 0 = (x + y) \<cdot> 0"
by (simp add: local.s_prod_distr)
lemma x_zero_meet_closed: "(x \<cdot> 0) \<sqinter> (y \<cdot> 0) = (x \<sqinter> y) \<cdot> 0"
by (metis c_0 local.cl6 local.meet_assoc local.meet_comm)
text \<open>The following set of lemmas investigates the closure properties of vectors, including Lemma 9,3.\<close>
lemma U_par_zero [simp]: "(0 \<cdot> c) \<parallel> U = 0"
by fastforce
lemma U_par_s_id [simp]: "(1\<^sub>\<sigma> \<cdot> 1\<^sub>\<pi>) \<parallel> U = U"
by auto
lemma U_par_p_id [simp]: "(1\<^sub>\<pi> \<cdot> 1\<^sub>\<pi>) \<parallel> U = U"
by auto
lemma U_par_nc [simp]: "(nc \<cdot> 1\<^sub>\<pi>) \<parallel> U = U"
by auto
lemma d_add_var: "d x \<cdot> z + d y \<cdot> z = d (x + y) \<cdot> z"
by (simp add: local.d_add_ax local.s_prod_distr)
lemma d_interr_U: "(d x \<cdot> U) \<parallel> (d y \<cdot> U) = d (x \<parallel> y) \<cdot> U"
by (simp add: local.cl4 local.d_conc6)
lemma d_meet:
assumes "\<And> x y z. (x \<sqinter> y \<sqinter> 1\<^sub>\<sigma>) \<cdot> z = (x \<sqinter> 1\<^sub>\<sigma>) \<cdot> z \<sqinter> (y \<sqinter> 1\<^sub>\<sigma>) \<cdot> z"
shows "d x \<cdot> z \<sqinter> d y \<cdot> z = (d x \<sqinter> d y) \<cdot> z"
proof -
have "(d x \<sqinter> d y) \<cdot> z = (d x \<sqinter> d y \<sqinter> 1\<^sub>\<sigma>) \<cdot> z"
using local.d_sub_id_ax local.meet_assoc local.inf.absorb_iff1 by fastforce
also have "... = (d x \<sqinter> 1\<^sub>\<sigma>) \<cdot> z \<sqinter> (d y \<sqinter> 1\<^sub>\<sigma>) \<cdot> z"
using assms by auto
finally show ?thesis
by (metis local.d_sub_id_ax local.inf.absorb_iff1)
qed
text \<open>Proposition 9.4\<close>
lemma nc_zero_closed [simp]: "0 \<sqinter> nc = 0"
by (simp add: local.inf.commute local.inf_absorb2)
lemma nc_s [simp]: "1\<^sub>\<sigma> \<sqinter> nc = 1\<^sub>\<sigma>"
using local.inf.absorb_iff1 s_le_nc by blast
lemma nc_add_closed: "(x \<sqinter> nc) + (y \<sqinter> nc) = (x + y) \<sqinter> nc"
using local.lat_dist4 by force
lemma nc_meet_closed: "(x \<sqinter> nc) \<sqinter> (y \<sqinter> nc) = x \<sqinter> y \<sqinter> nc"
using local.meet_assoc local.meet_comm local.inf_le1 local.inf.absorb_iff1 by fastforce
lemma nc_scomp_closed: "((x \<sqinter> nc) \<cdot> (y \<sqinter> nc)) \<le> nc"
by (simp add: c_0 nc_iff1 zero_assoc3)
lemma nc_scomp_closed_alt [simp]: "((x \<sqinter> nc) \<cdot> (y \<sqinter> nc)) \<sqinter> nc = (x \<sqinter> nc) \<cdot> (y \<sqinter> nc)"
using local.inf.absorb_iff1 nc_scomp_closed by blast
lemma nc_ccomp_closed: "(x \<sqinter> nc) \<parallel> (y \<sqinter> nc) \<le> nc"
proof -
have "(x \<sqinter> nc) \<parallel> (y \<sqinter> nc) \<le> nc \<parallel> nc"
by (meson local.inf_le2 local.mult_isol_var)
thus ?thesis
by auto
qed
lemma nc_ccomp_closed_alt [simp]: "(x \<parallel> (y \<sqinter> nc)) \<sqinter> nc = x \<parallel> (y \<sqinter> nc)"
by (metis U_nc_par local.U_def local.inf_le2 local.mult_isol_var local.inf.absorb_iff1)
text \<open>Lemma 9.6.\<close>
lemma tarski_prod:
assumes "\<And>x. x \<sqinter> nc \<noteq> 0 \<Longrightarrow> nc \<cdot> ((x \<sqinter> nc) \<cdot> nc) = nc"
and "\<And>x y z. d x \<cdot> (y \<cdot> z) = (d x \<cdot> y) \<cdot> z"
shows "((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc) = (if (y \<sqinter> nc) = 0 then 0 else (x \<sqinter> nc) \<cdot> nc)"
proof (cases "y \<sqinter> nc = 0")
fix x y
assume assm: "y \<sqinter> nc = 0"
show "(x \<sqinter> nc) \<cdot> nc \<cdot> ((y \<sqinter> nc) \<cdot> nc) = (if y \<sqinter> nc = 0 then 0 else (x \<sqinter> nc) \<cdot> nc)"
by (metis assm c_0 local.cl6 local.meet_comm nc_zero zero_assoc3)
next
fix x y
assume assm: "y \<sqinter> nc \<noteq> 0"
have "((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc) = (d (x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc)"
by simp
also have "... = d (x \<sqinter> nc) \<cdot> (nc \<cdot> ((y \<sqinter> nc) \<cdot> nc))"
by (simp add: assms(2))
also have "... = d (x \<sqinter> nc) \<cdot> nc"
by (simp add: assm assms(1))
finally show "(x \<sqinter> nc) \<cdot> nc \<cdot> ((y \<sqinter> nc) \<cdot> nc) = (if y \<sqinter> nc = 0 then 0 else (x \<sqinter> nc) \<cdot> nc)"
by (simp add: assm)
qed
text \<open>We show the remaining conditions of Proposition 9.8.\<close>
lemma nc_prod_aux [simp]: "((x \<sqinter> nc) \<cdot> nc) \<cdot> nc = (x \<sqinter> nc) \<cdot> nc"
proof -
have "((x \<sqinter> nc) \<cdot> nc) \<cdot> nc = (d (x \<sqinter> nc) \<cdot> nc) \<cdot> nc"
by simp
also have "... = d (x \<sqinter> nc) \<cdot> (nc \<cdot> nc)"
by (metis cl11_d d_x_nc local.cl11 local.meet_idem nc_ccomp_closed_alt nc_nc)
also have "... = d (x \<sqinter> nc) \<cdot> nc"
by auto
finally show ?thesis
by simp
qed
lemma nc_vec_add_closed: "((x \<sqinter> nc) \<cdot> nc + (y \<sqinter> nc) \<cdot> nc) \<cdot> nc = (x \<sqinter> nc) \<cdot> nc + (y \<sqinter> nc) \<cdot> nc"
by (simp add: local.s_prod_distr)
lemma nc_vec_par_closed: "(((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc)) \<cdot> nc = ((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc)"
by (simp add: local.cl4)
lemma nc_vec_par_is_meet:
assumes "\<And> x y z. (d x \<sqinter> d y) \<cdot> z = d x \<cdot> z \<sqinter> d y \<cdot> z"
shows "((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc) = ((x \<sqinter> nc) \<cdot> nc) \<sqinter> ((y \<sqinter> nc) \<cdot> nc)"
proof -
have "((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc) = (d (x \<sqinter> nc) \<cdot> nc) \<parallel> (d (y \<sqinter> nc) \<cdot> nc)"
by auto
also have "... = (d (x \<sqinter> nc) \<parallel> d (y \<sqinter> nc)) \<cdot> nc"
by (simp add: local.cl4)
also have "... = (d (x \<sqinter> nc) \<sqinter> d (y \<sqinter> nc)) \<cdot> nc"
by (simp add: d_p_prod_eq_meet)
finally show ?thesis
by (simp add: assms)
qed
lemma nc_vec_meet_closed:
assumes "\<And> x y z. (d x \<sqinter> d y) \<cdot> z = d x \<cdot> z \<sqinter> d y \<cdot> z"
shows "((x \<sqinter> nc) \<cdot> nc \<sqinter> (y \<sqinter> nc) \<cdot> nc) \<cdot> nc = (x \<sqinter> nc) \<cdot> nc \<sqinter> (y \<sqinter> nc) \<cdot> nc"
proof -
have "((x \<sqinter> nc) \<cdot> nc \<sqinter> (y \<sqinter> nc) \<cdot> nc) \<cdot> nc = (((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc)) \<cdot> nc"
by (simp add: assms nc_vec_par_is_meet)
also have "... = ((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc)"
by (simp add: nc_vec_par_closed)
finally show ?thesis
by (simp add: assms nc_vec_par_is_meet)
qed
lemma nc_vec_seq_closed:
assumes "\<And>x. x \<sqinter> nc \<noteq> 0 \<Longrightarrow> nc \<cdot> ((x \<sqinter> nc) \<cdot> nc) = nc"
and "\<And>x y z. d x \<cdot> (y \<cdot> z) = (d x \<cdot> y) \<cdot> z"
shows "(((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc)) \<cdot> nc = ((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc)"
proof -
have one : "y \<sqinter> nc = 0 \<Longrightarrow> (((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc)) \<cdot> nc = ((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc)"
by simp
have "y \<sqinter> nc \<noteq> 0 \<Longrightarrow> (((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc)) \<cdot> nc = ((x \<sqinter> nc) \<cdot> nc) \<cdot> ((y \<sqinter> nc) \<cdot> nc)"
by (simp add: assms(1) assms(2) tarski_prod)
thus ?thesis
using one by blast
qed
text \<open>Proposition 10.1 and 10.2.\<close>
lemma iso3 [simp]: "d (d x \<cdot> U) = d x "
by simp
lemma iso4 [simp]: "d ((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> U = (x \<cdot> 1\<^sub>\<pi>) \<parallel> U"
by (simp add: local.c3 local.c4 vec_iff)
lemma iso5 [simp]: "((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> 1\<^sub>\<pi> = x \<cdot> 1\<^sub>\<pi>"
by (simp add: local.c3 local.c4)
lemma iso6 [simp]: "(((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> 1\<^sub>\<pi>) \<parallel> U = (x \<cdot> 1\<^sub>\<pi>) \<parallel> U"
by simp
lemma iso3_sharp [simp]: "d (d (x \<sqinter> nc) \<cdot> nc) = d (x \<sqinter> nc)"
using d_s_subid local.c4 local.d_def local.inf_le1 by auto
lemma iso4_sharp [simp]: "d ((x \<sqinter> nc) \<cdot> nc) \<cdot> nc = (x \<sqinter> nc) \<cdot> nc"
by (simp add: local.c2_d local.c4)
lemma iso5_sharp [simp]: "(((x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>) \<parallel> nc) \<cdot> 1\<^sub>\<pi> = (x \<sqinter> nc) \<cdot> 1\<^sub>\<pi>"
by (simp add: local.c3 local.c4)
lemma iso6_sharp [simp]: "(((x \<sqinter> nc) \<cdot> nc) \<cdot> 1\<^sub>\<pi>) \<parallel> nc = (x \<sqinter> nc) \<cdot> nc"
using local.c4 local.cl11 nc_c by presburger
text\<open>We verify Lemma 15.2 at this point, because it is helpful for the following proofs.\<close>
lemma uc_par_meet: "x \<parallel> U \<sqinter> y \<parallel> U = x \<parallel> U \<parallel> y \<parallel> U"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis local.c_prod_assoc meet_le_par)
by (metis U_idem_p_prod local.U_def local.c_prod_assoc local.meet_prop local.mult.left_commute local.mult_double_iso)
lemma uc_unc [simp]: "x \<parallel> U \<parallel> x \<parallel> U = x \<parallel> U"
by (metis local.meet_idem uc_par_meet)
lemma uc_interr: "(x \<parallel> y) \<cdot> (z \<parallel> U) = (x \<cdot> (z \<parallel> U)) \<parallel> (y \<cdot> (z \<parallel> U))"
proof -
have "(z \<parallel> U) \<parallel> (z \<parallel> U) = z \<parallel> U"
by (metis local.c_prod_assoc uc_unc)
thus ?thesis
by (simp add: local.cl4)
qed
text\<open>We verify the remaining cases of Proposition 10.3.\<close>
lemma sc_hom_meet: "(d x \<sqinter> d y) \<cdot> 1\<^sub>\<pi> = (d x) \<cdot> 1\<^sub>\<pi> \<sqinter> (d y) \<cdot> 1\<^sub>\<pi>"
by (metis d_p_prod_eq_meet local.c3 p_subid_par_eq_meet_var)
lemma sc_hom_seq: "(d x \<cdot> d y) \<cdot> 1\<^sub>\<pi> = (d x \<sqinter> d y) \<cdot> 1\<^sub>\<pi>"
by (simp add: d_s_prod_eq_meet)
lemma cs_hom_meet: "d (x \<cdot> 1\<^sub>\<pi> \<sqinter> y \<cdot> 1\<^sub>\<pi>) = d (x \<cdot> 1\<^sub>\<pi>) \<sqinter> d (y \<cdot> 1\<^sub>\<pi>)"
by (metis d_p_prod_eq_meet local.d_conc6 p_subid_par_eq_meet_var)
lemma sv_hom_meet: "(d x \<sqinter> d y) \<cdot> U = (d x) \<cdot> U \<sqinter> (d y) \<cdot> U"
proof -
have "(d x \<sqinter> d y) \<cdot> U = ((d x) \<cdot> U) \<parallel> ((d y) \<cdot> U)"
by (simp add: d_interr_U d_p_prod_eq_meet local.d_conc6)
thus ?thesis
by (simp add: local.c2_d local.c_prod_assoc uc_par_meet)
qed
lemma sv_hom_par: "(x \<parallel> y) \<cdot> U = (x \<cdot> U) \<parallel> (y \<cdot> U)"
by (simp add: local.cl4)
lemma vs_hom_meet: "d (((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<sqinter> ((y \<cdot> 1\<^sub>\<pi>) \<parallel> U)) = d ((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<sqinter> d ((y \<cdot> 1\<^sub>\<pi>) \<parallel> U)"
proof -
have f1: "\<And>x y. x \<cdot> 1\<^sub>\<pi> \<parallel> 1\<^sub>\<sigma> \<sqinter> y \<cdot> 1\<^sub>\<pi> \<parallel> 1\<^sub>\<sigma> = x \<parallel> y \<cdot> 1\<^sub>\<pi> \<parallel> 1\<^sub>\<sigma>"
using d_p_prod_eq_meet local.d_conc6 local.d_def by auto
hence "\<And>x y. x \<cdot> 1\<^sub>\<pi> \<parallel> U \<sqinter> y \<cdot> 1\<^sub>\<pi> \<parallel> U = x \<parallel> y \<cdot> 1\<^sub>\<pi> \<parallel> U"
using local.d_def sv_hom_meet by force
thus ?thesis
using f1 by (simp add: local.d_def)
qed
lemma cv_hom_meet: "(x \<cdot> 1\<^sub>\<pi> \<sqinter> y \<cdot> 1\<^sub>\<pi>) \<parallel> U = (x \<cdot> 1\<^sub>\<pi>) \<parallel> U \<sqinter> (y \<cdot> 1\<^sub>\<pi>) \<parallel> U"
proof -
have "d (x \<parallel> y) \<cdot> U = x \<cdot> 1\<^sub>\<pi> \<parallel> U \<sqinter> y \<cdot> 1\<^sub>\<pi> \<parallel> U"
by (simp add: d_p_prod_eq_meet local.c2_d local.d_conc6 sv_hom_meet)
thus ?thesis
using local.c2_d local.c3 p_subid_par_eq_meet_var by auto
qed
lemma cv_hom_par [simp]: " x \<parallel> U \<parallel> y \<parallel> U = (x \<parallel> y) \<parallel> U"
by (metis U_idem_p_prod local.mult.left_commute local.mult_assoc)
lemma vc_hom_meet: "((x \<cdot> 1\<^sub>\<pi>) \<parallel> U \<sqinter> (y \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> 1\<^sub>\<pi> = ((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> 1\<^sub>\<pi> \<sqinter> ((y \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> 1\<^sub>\<pi>"
by (metis cv_hom_meet iso5 local.c3 p_subid_par_eq_meet_var)
lemma vc_hom_seq: "(((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> ((y \<cdot> 1\<^sub>\<pi>) \<parallel> U)) \<cdot> 1\<^sub>\<pi> = (((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> 1\<^sub>\<pi>) \<cdot> (((y \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> 1\<^sub>\<pi>)"
proof -
have "(((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> ((y \<cdot> 1\<^sub>\<pi>) \<parallel> U)) \<cdot> 1\<^sub>\<pi> = ((x \<cdot> 1\<^sub>\<pi>) \<parallel> U) \<cdot> (y \<cdot> 1\<^sub>\<pi>)"
by (simp add: local.c4)
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> (U \<cdot> (y \<cdot> 1\<^sub>\<pi>))"
by (metis assoc_p_subid local.cl8)
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> (nc \<cdot> (y \<cdot> 1\<^sub>\<pi>) + 1\<^sub>\<pi> \<cdot> (y \<cdot> 1\<^sub>\<pi>))"
by (metis add_commute c_nc_comp1 local.s_prod_distr)
also have "... = (x \<cdot> 1\<^sub>\<pi>) \<parallel> 1\<^sub>\<pi>"
by (metis add_commute c_x_prop local.absorp2 local.c4 local.meet_comm local.mult_oner p_subid_par_eq_meet_var)
thus ?thesis
by (simp add: assoc_p_subid calculation)
qed
text \<open>Proposition 10.4.\<close>
lemma nsv_hom_meet: "(d x \<sqinter> d y) \<cdot> nc = (d x) \<cdot> nc \<sqinter> (d y) \<cdot> nc"
-proof (rule antisym)
+proof (rule order.antisym)
have "(d x \<sqinter> d y) \<cdot> nc \<le> (d x) \<cdot> nc"
by (simp add: local.s_prod_isor)
hence "(d x \<sqinter> d y) \<cdot> nc \<le> (d x) \<cdot> nc"
by blast
thus "(d x \<sqinter> d y) \<cdot> nc \<le> (d x) \<cdot> nc \<sqinter> (d y) \<cdot> nc"
by (simp add: local.s_prod_isor)
have "(d x) \<cdot> nc \<sqinter> (d y) \<cdot> nc \<le> ((d x) \<cdot> nc) \<parallel> ((d y) \<cdot> nc)"
by (simp add: meet_le_par)
also have "... = (d x \<parallel> d y) \<cdot> nc"
by (metis local.cl4 nc_nc_par subidem_par)
finally show "(d x) \<cdot> nc \<sqinter> (d y) \<cdot> nc \<le> (d x \<sqinter> d y) \<cdot> nc"
by (simp add: d_p_prod_eq_meet)
qed
lemma nsv_hom_par: "(x \<parallel> y) \<cdot> nc = (x \<cdot> nc) \<parallel> (y \<cdot> nc)"
by (simp add: local.cl4)
lemma vec_p_prod_meet: "((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc) = ((x \<sqinter> nc) \<cdot> nc) \<sqinter> ((y \<sqinter> nc) \<cdot> nc)"
proof -
have "((x \<sqinter> nc) \<cdot> nc) \<parallel> ((y \<sqinter> nc) \<cdot> nc) = (d (x \<sqinter> nc) \<cdot> nc) \<parallel> (d (y \<sqinter> nc) \<cdot> nc)"
by (metis cl11_d)
also have "... = (d (x \<sqinter> nc) \<parallel> d (y \<sqinter> nc)) \<cdot> nc"
by (simp add: nsv_hom_par)
also have "... = (d (x \<sqinter> nc) \<sqinter> d (y \<sqinter> nc)) \<cdot> nc"
by (simp add: d_p_prod_eq_meet)
also have "... = (d (x \<sqinter> nc) \<cdot> nc) \<sqinter> (d (y \<sqinter> nc) \<cdot> nc)"
by (simp add: nsv_hom_meet)
thus ?thesis
by (simp add: calculation)
qed
lemma nvs_hom_meet: "d (((x \<sqinter> nc) \<cdot> nc) \<sqinter> ((y \<sqinter> nc) \<cdot> nc)) = d ((x \<sqinter> nc) \<cdot> nc) \<sqinter> d ((y \<sqinter> nc) \<cdot> nc)"
by (metis d_p_prod_eq_meet local.d_conc6 vec_p_prod_meet)
lemma ncv_hom_meet: "(x \<cdot> 1\<^sub>\<pi> \<sqinter> y \<cdot> 1\<^sub>\<pi>) \<parallel> nc = (x \<cdot> 1\<^sub>\<pi>) \<parallel> nc \<sqinter> (y \<cdot> 1\<^sub>\<pi>) \<parallel> nc"
by (metis d_p_prod_eq_meet local.c2_d local.c3 local.d_conc6 nsv_hom_meet p_subid_par_eq_meet_var)
lemma ncv_hom_par: "(x \<parallel> y) \<parallel> nc = x \<parallel> nc \<parallel> y \<parallel> nc"
by (metis local.mult_assoc local.mult_commute nc_nc_par)
lemma nvc_hom_meet: "((x \<sqinter> nc) \<cdot> nc \<sqinter> (y \<sqinter> nc) \<cdot> nc) \<cdot> 1\<^sub>\<pi> = ((x \<sqinter> nc) \<cdot> nc) \<cdot> 1\<^sub>\<pi> \<sqinter> ((y \<sqinter> nc) \<cdot> nc) \<cdot> 1\<^sub>\<pi>"
by (metis local.c3 p_subid_par_eq_meet_var vec_p_prod_meet)
subsection \<open>Terminal and Nonterminal Elements\<close>
text \<open>Now we define the projection functions on terminals and nonterminal parts and verify the properties
of Section 11 in~\cite{FurusawaS15a}.\<close>
definition tau :: "'a \<Rightarrow> 'a" ("\<tau>") where
"\<tau> x = x \<cdot> 0"
definition nu :: "'a \<Rightarrow> 'a" ("\<nu>") where
"\<nu> x = x \<sqinter> nc"
text \<open>Lemma 11.1.\<close>
lemma tau_int: "\<tau> x \<le> x"
by (metis c_0 local.inf_le1 tau_def)
lemma nu_int: "\<nu> x \<le> x"
by (simp add: nu_def)
lemma tau_ret [simp]: "\<tau> (\<tau> x) = \<tau> x"
by (simp add: tau_def)
lemma nu_ret [simp]: "\<nu> (\<nu> x) = \<nu> x"
by (simp add: local.meet_assoc nu_def)
lemma tau_iso: "x \<le> y \<Longrightarrow> \<tau> x \<le> \<tau> y"
using local.order_prop local.s_prod_distr tau_def by auto
lemma nu_iso: "x \<le> y \<Longrightarrow> \<nu> x \<le> \<nu> y"
using local.inf_mono nu_def by auto
text \<open>Lemma 11.2.\<close>
lemma tau_zero [simp]: "\<tau> 0 = 0"
by (simp add: tau_def)
lemma nu_zero [simp]: "\<nu> 0 = 0"
using nu_def by auto
lemma tau_s [simp]: "\<tau> 1\<^sub>\<sigma> = 0"
using tau_def by auto
lemma nu_s [simp]: "\<nu> 1\<^sub>\<sigma> = 1\<^sub>\<sigma>"
using nu_def by auto
lemma tau_c [simp]: "\<tau> 1\<^sub>\<pi> = 1\<^sub>\<pi>"
using c_x_prop tau_def by presburger
lemma nu_c [simp]: "\<nu> 1\<^sub>\<pi> = 0"
using c_nc_comp2 nu_def by presburger
lemma tau_nc [simp]: "\<tau> nc = 0"
using nc_iff2 tau_def by auto
lemma nu_nc [simp]: "\<nu> nc = nc"
using nu_def by auto
lemma tau_U [simp]: "\<tau> U = 1\<^sub>\<pi>"
using c_def tau_def by presburger
lemma nu_U [simp]: "\<nu> U = nc"
using local.U_def local.meet_comm local.inf.absorb_iff1 nu_def by fastforce
text \<open>Lemma 11.3.\<close>
lemma tau_add [simp]: "\<tau> (x + y) = \<tau> x + \<tau> y"
by (simp add: tau_def x_zero_add_closed)
lemma nu_add [simp]: "\<nu> (x + y) = \<nu> x + \<nu> y"
by (simp add: local.lat_dist3 local.meet_comm nu_def)
lemma tau_meet [simp]: "\<tau> (x \<sqinter> y) = \<tau> x \<sqinter> \<tau> y"
using tau_def x_zero_meet_closed by auto
lemma nu_meet [simp]: "\<nu> (x \<sqinter> y) = \<nu> x \<sqinter> \<nu> y"
using nc_meet_closed nu_def by auto
lemma tau_seq: "\<tau> (x \<cdot> y) = \<tau> x + \<nu> x \<cdot> \<tau> y"
using local.add_comm nu_def tau_def x_y_split zero_assoc3 by presburger
lemma tau_par [simp]: "\<tau> (x \<parallel> y) = \<tau> x \<parallel> \<tau> y"
using tau_def x_zero_interr by presburger
lemma nu_par_aux1: "x \<parallel> \<tau> y = d (\<tau> y) \<cdot> x"
by (simp add: local.c2_d local.mult_commute tau_def)
lemma nu_par_aux2 [simp]: "\<nu> (\<nu> x \<parallel> \<nu> y) = \<nu> x \<parallel> \<nu> y"
by (simp add: nu_def)
lemma nu_par_aux3 [simp]: "\<nu> (\<nu> x \<parallel> \<tau> y) = \<nu> x \<parallel> \<tau> y"
by (metis local.mult_commute nc_ccomp_closed_alt nu_def)
lemma nu_par_aux4 [simp]: "\<nu> (\<tau> x \<parallel> \<tau> y) = 0"
by (metis nu_def tau_def tau_par zero_nc)
lemma nu_par: "\<nu> (x \<parallel> y) = d (\<tau> x) \<cdot> \<nu> y + d (\<tau> y) \<cdot> \<nu> x + \<nu> x \<parallel> \<nu> y"
proof -
have "\<nu> (x \<parallel> y) = \<nu> (\<nu> x \<parallel> \<nu> y) + \<nu> (\<nu> x \<parallel> \<tau> y) + \<nu> (\<tau> x \<parallel> \<nu> y) + \<nu> (\<tau> x \<parallel> \<tau> y)"
by (metis local.distrib_left local.distrib_right nu_add nu_def tau_def x_split_var join.sup.commute join.sup.left_commute)
also have "\<nu> (x \<parallel> y) = \<nu> x \<parallel> \<nu> y + \<nu> x \<parallel> \<tau> y + \<tau> x \<parallel> \<nu> y"
by (simp add: calculation local.c_prod_comm)
thus ?thesis
using local.join.sup_assoc local.join.sup_commute local.mult_commute nu_par_aux1 by auto
qed
text \<open>Lemma 11.5.\<close>
lemma sprod_tau_nu: "x \<cdot> y = \<tau> x + \<nu> x \<cdot> y"
by (metis local.add_comm nu_def tau_def x_y_split)
lemma pprod_tau_nu: "x \<parallel> y = \<nu> x \<parallel> \<nu> y + d (\<tau> x) \<cdot> \<nu> y + d (\<tau> y) \<cdot> \<nu> x + \<tau> x \<parallel> \<tau> y"
proof -
have "x \<parallel> y = \<nu> (x \<parallel> y) + \<tau> (x \<parallel> y)"
by (simp add: nu_def tau_def)
also have "... = (d (\<tau> x) \<cdot> \<nu> y + d (\<tau> y) \<cdot> \<nu> x + \<nu> x \<parallel> \<nu> y) + \<tau> x \<parallel> \<tau> y"
by (simp add: nu_par)
thus ?thesis
using add_assoc add_commute calculation by force
qed
text \<open>We now verify some additional properties which are not mentioned in the paper.\<close>
lemma tau_idem [simp]: "\<tau> x \<cdot> \<tau> x = \<tau> x"
by (simp add: tau_def)
lemma tau_interr: "(x \<parallel> y) \<cdot> \<tau> z = (x \<cdot> \<tau> z) \<parallel> (y \<cdot> \<tau> z)"
by (simp add: local.cl4 tau_def)
lemma tau_le_c: "\<tau> x \<le> 1\<^sub>\<pi>"
by (simp add: local.x_zero_le_c tau_def)
lemma c_le_tauc: "1\<^sub>\<pi> \<le> \<tau> 1\<^sub>\<pi>"
using local.eq_refl tau_c by presburger
lemma x_alpha_tau [simp]: "\<nu> x + \<tau> x = x"
using nu_def tau_def x_split_var by presburger
lemma alpha_tau_zero [simp]: "\<nu> (\<tau> x) = 0"
by (simp add: nu_def tau_def)
lemma tau_alpha_zero [simp]: "\<tau> (\<nu> x) = 0"
by (simp add: nu_def tau_def)
lemma sprod_tau_nu_var [simp]: "\<nu> (\<nu> x \<cdot> y) = \<nu> (x \<cdot> y)"
proof -
have "\<nu> (x \<cdot> y) = \<nu> (\<tau> x) + \<nu> (\<nu> x \<cdot> y)"
by (metis nu_add sprod_tau_nu)
thus ?thesis
by simp
qed
lemma tau_s_prod [simp]: "\<tau> (x \<cdot> y) = x \<cdot> \<tau> y"
by (simp add: tau_def zero_assoc3)
lemma alpha_fp: "\<nu> x = x \<longleftrightarrow> x \<cdot> 0 = 0"
by (metis local.add_zeror tau_alpha_zero tau_def x_alpha_tau)
lemma alpha_prod_closed [simp]: "\<nu> (\<nu> x \<cdot> \<nu> y) = \<nu> x \<cdot> \<nu> y"
by (simp add: nu_def)
lemma alpha_par_prod [simp]: "\<nu> (x \<parallel> \<nu> y) = x \<parallel> \<nu> y"
by (simp add: nu_def)
lemma p_prod_tau_alpha: "x \<parallel> y = x \<parallel> \<nu> y + \<nu> x \<parallel> y + \<tau> x \<parallel> \<tau> y"
proof -
have "x \<parallel> y = (\<nu> x + \<tau> x) \<parallel> (\<nu> y + \<tau> y)"
using x_alpha_tau by simp
also have "... = \<nu> x \<parallel> \<nu> y + \<nu> x \<parallel> \<tau> y + \<tau> x \<parallel> \<nu> y + \<tau> x \<parallel> \<tau> y"
by (metis add_commute local.combine_common_factor local.p_prod_distl)
also have "... = (\<nu> x \<parallel> \<nu> y + \<nu> x \<parallel> \<tau> y) + (\<nu> x \<parallel> \<nu> y + \<tau> x \<parallel> \<nu> y) + \<tau> x \<parallel> \<tau> y"
by (simp add: add_ac)
thus ?thesis
by (metis calculation local.add_comm local.distrib_left local.distrib_right x_alpha_tau)
qed
lemma p_prod_tau_alpha_var: "x \<parallel> y = x \<parallel> \<nu> y + \<nu> x \<parallel> y + \<tau> (x \<parallel> y)"
by (metis p_prod_tau_alpha tau_par)
lemma alpha_par: "\<nu> (x \<parallel> y) = \<nu> x \<parallel> y + x \<parallel> \<nu> y"
proof -
have "\<nu> (x \<parallel> y) = \<nu> (x \<parallel> \<nu> y) + \<nu> (\<nu> x \<parallel> y) + \<nu> (\<tau> (x \<parallel> y))"
by (metis nu_add p_prod_tau_alpha_var)
thus ?thesis
by (simp add: local.mult_commute add_ac)
qed
lemma alpha_tau [simp]: "\<nu> (x \<cdot> \<tau> y) = 0"
by (metis alpha_tau_zero tau_s_prod)
lemma nu_par_prop: "\<nu> x = x \<Longrightarrow> \<nu> (x \<parallel> y) = x \<parallel> y"
by (metis alpha_par_prod local.mult_commute)
lemma tau_seq_prop: "\<tau> x = x \<Longrightarrow> x \<cdot> y = x"
by (metis local.cl6 tau_def)
lemma tau_seq_prop2: "\<tau> y = y \<Longrightarrow> \<tau> (x \<cdot> y) = x \<cdot> y"
by auto
lemma d_nu: "\<nu> (d x \<cdot> y) = d x \<cdot> \<nu> y"
proof -
have "\<nu> (d x \<cdot> y) = \<nu> ((x \<cdot> 1\<^sub>\<pi>) \<parallel> y)"
by (simp add: local.c2_d)
also have "... = d (\<tau> (x \<cdot> 1\<^sub>\<pi>)) \<cdot> \<nu> y + d (\<tau> y) \<cdot> \<nu> (x \<cdot> 1\<^sub>\<pi>) + \<nu> (x \<cdot> 1\<^sub>\<pi>) \<parallel> \<nu> y"
by (simp add: nu_par)
thus ?thesis
using alpha_par local.c2_d nu_def by force
qed
text \<open>Lemma 11.6 and 11.7.\<close>
lemma nu_ideal1: "\<lbrakk>\<nu> x = x; y \<le> x\<rbrakk> \<Longrightarrow> \<nu> y = y"
by (metis local.meet_prop local.inf.absorb_iff1 nu_def)
lemma tau_ideal1: "\<lbrakk>\<tau> x = x; y \<le> x\<rbrakk> \<Longrightarrow> \<tau> y = y"
by (metis local.dual_order.trans tau_def term_p_subid_var)
lemma nu_ideal2: "\<lbrakk>\<nu> x = x; \<nu> y = y\<rbrakk> \<Longrightarrow> \<nu> (x + y) = x + y"
by (simp add: local.lat_dist3 local.meet_comm)
lemma tau_ideal2: "\<lbrakk>\<tau> x = x; \<tau> y = y\<rbrakk> \<Longrightarrow> \<tau> (x + y) = x + y"
by simp
lemma tau_ideal3: "\<tau> x = x \<Longrightarrow> \<tau> (x \<cdot> y) = x \<cdot> y"
by (simp add: tau_seq_prop)
text \<open>We prove the precongruence properties of Lemma 11.9.\<close>
lemma tau_add_precong: "\<tau> x \<le> \<tau> y \<Longrightarrow> \<tau> (x + z) \<le> \<tau> (y + z)"
proof -
assume "\<tau> x \<le> \<tau> y"
hence "(x + y) \<cdot> 0 = y \<cdot> 0" using local.less_eq_def local.s_prod_distr tau_def
by auto
hence "(x + z + y) \<cdot> 0 = (y + z) \<cdot> 0"
by (metis (no_types) add_assoc add_commute local.s_prod_distr)
thus "\<tau> (x + z) \<le> \<tau> (y + z)" using local.order_prop local.s_prod_distr tau_def
by metis
qed
lemma tau_meet_precong: "\<tau> x \<le> \<tau> y \<Longrightarrow> \<tau> (x \<sqinter> z) \<le> \<tau> (y \<sqinter> z)"
proof -
assume "\<tau> x \<le> \<tau> y"
hence "\<And>z. (x \<sqinter> y \<sqinter> z) \<cdot> 0 = (x \<sqinter> z) \<cdot> 0"
by (metis local.le_iff_inf tau_def x_zero_meet_closed)
thus ?thesis
using local.inf_left_commute local.le_iff_inf local.meet_comm tau_def x_zero_meet_closed by fastforce
qed
lemma tau_par_precong: "\<tau> x \<le> \<tau> y \<Longrightarrow> \<tau> (x \<parallel> z) \<le> \<tau> (y \<parallel> z)"
proof -
assume "\<tau> x \<le> \<tau> y"
hence "x \<parallel> z \<cdot> 0 \<le> y \<cdot> 0"
by (metis (no_types) local.dual_order.trans local.p_subid_lb1 tau_def tau_par)
thus "\<tau> (x \<parallel> z) \<le> \<tau> (y \<parallel> z)"
by (simp add: \<open>\<tau> x \<le> \<tau> y\<close> local.mult_isor)
qed
lemma tau_seq_precongl: "\<tau> x \<le> \<tau> y \<Longrightarrow> \<tau> (z \<cdot> x) \<le> \<tau> (z \<cdot> y)"
by (simp add: local.s_prod_isol)
lemma nu_add_precong: "\<nu> x \<le> \<nu> y \<Longrightarrow> \<nu> (x + z) \<le> \<nu> (y + z)"
proof -
assume "\<nu> x \<le> \<nu> y"
hence "\<nu> x = \<nu> x \<sqinter> \<nu> y"
using local.inf.absorb_iff1 by auto
hence "\<forall>a. \<nu> (x + a) = \<nu> (x + a) \<sqinter> \<nu> (y + a)"
by (metis (no_types) local.lat_dist2 nu_add)
thus ?thesis
using local.inf.absorb_iff1 by presburger
qed
lemma nu_meet_precong: "\<nu> x \<le> \<nu> y \<Longrightarrow> \<nu> (x \<sqinter> z) \<le> \<nu> (y \<sqinter> z)"
proof -
assume "\<nu> x \<le> \<nu> y"
hence "\<nu> y = \<nu> x + \<nu> y"
using local.less_eq_def by auto
hence "\<nu> (y \<sqinter> z) = \<nu> (x \<sqinter> z) + \<nu> (y \<sqinter> z)"
by (metis (no_types) local.lat_dist4 nu_meet)
thus ?thesis
using local.less_eq_def by presburger
qed
lemma nu_seq_precongr: "\<nu> x \<le> \<nu> y \<Longrightarrow> \<nu> (x \<cdot> z) \<le> \<nu> (y \<cdot> z)"
proof -
assume a: "\<nu> x \<le> \<nu> y"
have "\<nu> (x \<cdot> z) = \<nu> (\<nu> x \<cdot> z)"
by simp
also have "... \<le> \<nu> (\<nu> y \<cdot> z)"
by (metis a local.less_eq_def local.s_prod_distr nu_iso)
thus ?thesis
by simp
qed
text\<open>We prove the congruence properties of Corollary~11.11.\<close>
definition tcg :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
"tcg x y = (\<tau> x \<le> \<tau> y \<and> \<tau> y \<le> \<tau> x)"
definition ncg :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
"ncg x y = (\<nu> x \<le> \<nu> y \<and> \<nu> y \<le> \<nu> x)"
lemma tcg_refl: "tcg x x"
by (simp add: tcg_def)
lemma tcg_trans: "\<lbrakk>tcg x y; tcg y z\<rbrakk> \<Longrightarrow> tcg x z"
using tcg_def by force
lemma tcg_sym: "tcg x y \<Longrightarrow> tcg y x"
by (simp add: tcg_def)
lemma ncg_refl: "ncg x x"
using ncg_def by blast
lemma ncg_trans: "\<lbrakk>ncg x y; ncg y z\<rbrakk> \<Longrightarrow> ncg x z"
using ncg_def by force
lemma ncg_sym: "ncg x y \<Longrightarrow> ncg y x"
using ncg_def by auto
lemma tcg_alt: "tcg x y = (\<tau> x = \<tau> y)"
using tcg_def by auto
lemma ncg_alt: "ncg x y = (\<nu> x = \<nu> y)"
- by (simp add: local.eq_iff ncg_def)
+ by (simp add: order.eq_iff ncg_def)
lemma tcg_add: "\<tau> x = \<tau> y \<Longrightarrow> \<tau> (x + z) = \<tau> (y + z)"
by simp
lemma tcg_meet: "\<tau> x = \<tau> y \<Longrightarrow> \<tau> (x \<sqinter> z) = \<tau> (y \<sqinter> z)"
by simp
lemma tcg_par: "\<tau> x = \<tau> y \<Longrightarrow> \<tau> (x \<parallel> z) = \<tau> (y \<parallel> z)"
by simp
lemma tcg_seql: "\<tau> x = \<tau> y \<Longrightarrow> \<tau> (z \<cdot> x) = \<tau> (z \<cdot> y)"
by simp
lemma ncg_add: "\<nu> x = \<nu> y \<Longrightarrow> \<nu> (x + z) = \<nu> (y + z)"
by simp
lemma ncg_meet: "\<nu> x = \<nu> y \<Longrightarrow> \<nu> (x \<sqinter> z) = \<nu> (y \<sqinter> z)"
by simp
lemma ncg_seqr: "\<nu> x = \<nu> y \<Longrightarrow> \<nu> (x \<cdot> z) = \<nu> (y \<cdot> z)"
- by (simp add: local.eq_iff nu_seq_precongr)
+ by (simp add: order.eq_iff nu_seq_precongr)
end
subsection \<open>Powers in C-Algebras\<close>
text \<open>We define the power functions from Section~6 in~\cite{FurusawaS15a} after Lemma~12.4.\<close>
context proto_dioid
begin
primrec p_power :: "'a \<Rightarrow> nat \<Rightarrow> 'a" where
"p_power x 0 = 1\<^sub>\<sigma>" |
"p_power x (Suc n) = x \<cdot> p_power x n"
primrec power_rd :: "'a \<Rightarrow> nat \<Rightarrow> 'a" where
"power_rd x 0 = 0" |
"power_rd x (Suc n) = 1\<^sub>\<sigma> + x \<cdot> power_rd x n"
primrec power_sq :: "'a \<Rightarrow> nat \<Rightarrow> 'a" where
"power_sq x 0 = 1\<^sub>\<sigma>" |
"power_sq x (Suc n) = 1\<^sub>\<sigma> + x \<cdot> power_sq x n"
text \<open>Lemma~12.5\<close>
lemma power_rd_chain: "power_rd x n \<le> power_rd x (n + 1)"
by (induct n, simp, metis Suc_eq_plus1 local.add_comm local.add_iso local.s_prod_isol power_rd.simps(2))
lemma power_sq_chain: "power_sq x n \<le> power_sq x (n + 1)"
by (induct n, clarsimp, metis Suc_eq_plus1 local.add_comm local.add_iso local.s_prod_isol power_sq.simps(2))
lemma pow_chain: "p_power (1\<^sub>\<sigma> + x) n \<le> p_power (1\<^sub>\<sigma> + x) (n + 1)"
by (induct n, simp, metis Suc_eq_plus1 local.p_power.simps(2) local.s_prod_isol)
lemma pow_prop: "p_power (1\<^sub>\<sigma> + x) (n + 1) = 1\<^sub>\<sigma> + x \<cdot> p_power (1\<^sub>\<sigma> + x) n"
proof (induct n)
case 0
show ?case by simp
next
case (Suc n)
have f1: "p_power (1\<^sub>\<sigma> + x) (Suc n + 1) = 1\<^sub>\<sigma> + x \<cdot> p_power (1\<^sub>\<sigma> + x) n + x \<cdot> p_power (1\<^sub>\<sigma> + x) (n + 1)"
proof -
have "p_power (1\<^sub>\<sigma> + x) (Suc (n + 1)) = (1\<^sub>\<sigma> + x) \<cdot> p_power (1\<^sub>\<sigma> + x) (n + 1)"
using local.p_power.simps(2) by blast
also have "... = p_power (1\<^sub>\<sigma> + x) (n + 1) + x \<cdot> p_power (1\<^sub>\<sigma> + x) (n + 1)"
by (metis local.s_prod_distr local.s_prod_idl)
also have "... = 1\<^sub>\<sigma> + x \<cdot> p_power (1\<^sub>\<sigma> + x) n + x \<cdot> p_power (1\<^sub>\<sigma> + x) (n + 1)"
using Suc.hyps by auto
finally show ?thesis
by simp
qed
have "x \<cdot> p_power (1\<^sub>\<sigma> + x) (Suc n) = x \<cdot> p_power (1\<^sub>\<sigma> + x) n + x \<cdot> p_power (1\<^sub>\<sigma> + x) (n + 1)"
using Suc_eq_plus1 local.less_eq_def local.s_prod_isol pow_chain by simp
with f1 show ?case by (simp add: add_ac)
qed
text \<open>Next we verify facts from the proofs of Lemma~12.6.\<close>
lemma power_rd_le_sq: "power_rd x n \<le> power_sq x n"
by (induct n, simp, simp add: local.join.le_supI2 local.s_prod_isol)
lemma power_sq_le_rd: "power_sq x n \<le> power_rd x (Suc n)"
by (induct n, simp, simp add: local.join.le_supI2 local.s_prod_isol)
lemma power_sq_power: "power_sq x n = p_power (1\<^sub>\<sigma> + x) n"
apply (induct n)
apply (simp)
using Suc_eq_plus1 pow_prop power_sq.simps(2) by simp
end
subsection \<open>C-Kleene Algebras\<close>
text \<open>The definition of c-Kleene algebra is slightly different from that in Section~6
of~\cite{FurusawaS15a}. It is used to prove properties from Section~6 and Section~12.\<close>
class c_kleene_algebra = c_lattice + star_op +
assumes star_unfold: "1\<^sub>\<sigma> + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
and star_induct: "1\<^sub>\<sigma> + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y"
begin
lemma star_irr: "1\<^sub>\<sigma> \<le> x\<^sup>\<star>"
using local.star_unfold by auto
lemma star_unfold_part: "x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
using local.star_unfold by auto
lemma star_ext_aux: "x \<le> x \<cdot> x\<^sup>\<star>"
using local.s_prod_isol star_irr by fastforce
lemma star_ext: "x \<le> x\<^sup>\<star>"
using local.order_trans star_ext_aux star_unfold_part by blast
lemma star_co_trans: "x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> x\<^sup>\<star>"
using local.s_prod_isol star_irr by fastforce
lemma star_iso: "x \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y\<^sup>\<star>"
proof -
assume a1: "x \<le> y"
have f2: "y \<cdot> y\<^sup>\<star> + y\<^sup>\<star> = y\<^sup>\<star>"
by (meson local.less_eq_def star_unfold_part)
have "x + y = y"
using a1 by (meson local.less_eq_def)
hence "x \<cdot> y\<^sup>\<star> + y\<^sup>\<star> = y\<^sup>\<star>"
using f2 by (metis (no_types) local.add_assoc' local.s_prod_distr)
thus ?thesis
using local.add_assoc' local.less_eq_def local.star_induct star_irr by presburger
qed
lemma star_unfold_eq [simp]: "1\<^sub>\<sigma> + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show a: "1\<^sub>\<sigma> + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
using local.star_unfold by blast
have "1\<^sub>\<sigma> + x \<cdot> (1\<^sub>\<sigma> + x \<cdot> x\<^sup>\<star>) \<le> 1\<^sub>\<sigma> + x \<cdot> x\<^sup>\<star>"
by (meson local.eq_refl local.join.sup_mono local.s_prod_isol local.star_unfold)
thus "x\<^sup>\<star> \<le> 1\<^sub>\<sigma> + x \<cdot> x\<^sup>\<star>"
by (simp add: local.star_induct)
qed
text \<open>Lemma 12.2.\<close>
lemma nu_star1:
assumes "\<And>x y z. x \<cdot> (y \<cdot> z) = (x \<cdot> y) \<cdot> z"
shows "x\<^sup>\<star> \<le> (\<nu> x)\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x)"
proof -
have "1\<^sub>\<sigma> + x \<cdot> ((\<nu> x)\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x)) = 1\<^sub>\<sigma> + \<tau> x + \<nu> x \<cdot> ((\<nu> x)\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x))"
by (metis add_assoc local.sprod_tau_nu)
also have "... = (1\<^sub>\<sigma> + \<nu> x \<cdot> (\<nu> x)\<^sup>\<star>) \<cdot> (1\<^sub>\<sigma> + \<tau> x)"
using assms local.s_prod_distr local.s_prod_idl by presburger
also have "... \<le> (\<nu> x)\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x)"
using local.s_prod_isor local.star_unfold by auto
thus ?thesis
by (simp add: calculation local.star_induct)
qed
lemma nu_star2:
assumes "\<And>x. x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
shows "(\<nu> x)\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x) \<le> x\<^sup>\<star>"
proof -
have "(\<nu> x)\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x) \<le> x\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x)"
using local.nu_int local.s_prod_isor star_iso by blast
also have "... \<le> x\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + x)"
using local.s_prod_isol local.join.sup_mono local.tau_int by blast
also have "... \<le> x\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (simp add: local.s_prod_isol star_ext star_irr)
finally show ?thesis
using assms local.order_trans by blast
qed
lemma nu_star:
assumes "\<And>x. x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
and "\<And>x y z. x \<cdot> (y \<cdot> z) = (x \<cdot> y) \<cdot> z"
shows "(\<nu> x)\<^sup>\<star> \<cdot> (1\<^sub>\<sigma> + \<tau> x) = x\<^sup>\<star>"
by (simp add: assms(1) assms(2) local.dual_order.antisym nu_star1 nu_star2)
text \<open>Lemma 12.3.\<close>
lemma tau_star: "(\<tau> x)\<^sup>\<star> = 1\<^sub>\<sigma> + \<tau> x"
by (metis local.cl6 local.tau_def star_unfold_eq)
lemma tau_star_var:
assumes "\<And>x y z. x \<cdot> (y \<cdot> z) = (x \<cdot> y) \<cdot> z"
and "\<And>x. x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
shows "\<tau> (x\<^sup>\<star>) = (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x"
by (metis (no_types, lifting) assms(1) assms(2) local.add_0_right local.add_comm local.s_prod_distr local.s_prod_idl local.tau_def local.tau_zero nu_star)
lemma nu_star_sub: "(\<nu> x)\<^sup>\<star> \<le> \<nu> (x\<^sup>\<star>)"
by (metis add_commute local.less_eq_def local.meet_prop local.nc_nc local.nu_def local.order.refl local.s_le_nc local.star_induct star_iso)
lemma nu_star_nu [simp]: "\<nu> ((\<nu> x)\<^sup>\<star>) = (\<nu> x)\<^sup>\<star>"
using local.nu_ideal1 local.nu_ret nu_star_sub by blast
lemma nu_star_tau [simp]: "\<nu> ((\<tau> x)\<^sup>\<star>) = 1\<^sub>\<sigma>"
using tau_star by fastforce
lemma tau_star_tau [simp]: "\<tau> ((\<tau> x)\<^sup>\<star>) = \<tau> x"
using local.s_prod_distr tau_star by auto
lemma tau_star_nu [simp]: "\<tau> ((\<nu> x)\<^sup>\<star>) = 0"
using local.alpha_fp local.tau_def nu_star_nu by presburger
text \<open>Finally we verify Lemma 6.2. Proofs can be found in~\cite{FurusawaS15b}.\<close>
lemma d_star_unfold [simp]:
assumes "\<And>x y z. (x \<cdot> y) \<cdot> d z = x \<cdot> (y \<cdot> d z)"
shows "d y + d (x \<cdot> d (x\<^sup>\<star> \<cdot> y)) = d (x\<^sup>\<star> \<cdot> y)"
proof -
have "d y + d (x \<cdot> d (x\<^sup>\<star> \<cdot> y)) = d y + d (x \<cdot> (x\<^sup>\<star> \<cdot> d y))"
by (metis local.c4 local.d_def local.dc_prop1)
moreover have "... = d (1\<^sub>\<sigma> \<cdot> d y + (x \<cdot> (x\<^sup>\<star> \<cdot> d y)))"
by (simp add: local.d_add_ax)
moreover have "... = d (1\<^sub>\<sigma> \<cdot> d y + (x \<cdot> x\<^sup>\<star>) \<cdot> d y)"
by (simp add: assms)
moreover have "... = d ((1\<^sub>\<sigma> + x \<cdot> x\<^sup>\<star>) \<cdot> d y)"
using local.s_prod_distr by presburger
ultimately show ?thesis
by simp
qed
lemma d_star_sim1:
assumes "\<And> x y z. d z + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> d z \<le> y"
and "\<And> x y z. (x \<cdot> d y) \<cdot> z = x \<cdot> (d y \<cdot> z)"
and "\<And> x y z. (d x \<cdot> y) \<cdot> z = d x \<cdot> (y \<cdot> z)"
shows "x \<cdot> d z \<le> d z \<cdot> y \<Longrightarrow> x\<^sup>\<star> \<cdot> d z \<le> d z \<cdot> y\<^sup>\<star>"
proof -
fix x y z
assume a: "x \<cdot> d z \<le> d z \<cdot> y"
have b: "x \<cdot> (d z \<cdot> y\<^sup>\<star>) \<le> d z \<cdot> (y \<cdot> y\<^sup>\<star>)"
by (metis a assms(2) assms(3) local.s_prod_isor)
hence "x \<cdot> (d z \<cdot> y\<^sup>\<star>) \<le> d z \<cdot> y\<^sup>\<star>"
proof -
have f1: "x \<cdot> (y\<^sup>\<star> \<parallel> (z \<cdot> 1\<^sub>\<pi>)) \<le> z \<cdot> 1\<^sub>\<pi> \<parallel> (y \<cdot> y\<^sup>\<star>)"
using b local.c2_d local.mult_commute by auto
have "\<exists>a. (a + z \<cdot> 1\<^sub>\<pi>) \<parallel> (y \<cdot> y\<^sup>\<star>) \<le> y\<^sup>\<star> \<parallel> (z \<cdot> 1\<^sub>\<pi>)"
by (metis (no_types) local.eq_refl local.mult_commute local.mult_isol_var local.join.sup_idem star_unfold_part)
hence "x \<cdot> (y\<^sup>\<star> \<parallel> (z \<cdot> 1\<^sub>\<pi>)) \<le> y\<^sup>\<star> \<parallel> (z \<cdot> 1\<^sub>\<pi>)"
using f1 by (metis (no_types) local.distrib_right' local.dual_order.trans local.join.sup.cobounded2)
thus ?thesis
using local.c2_d local.mult_commute by auto
qed
hence "d z + x \<cdot> (d z \<cdot> y\<^sup>\<star>)\<le> d z \<cdot> y\<^sup>\<star>"
using local.s_prod_isol star_irr by fastforce
thus "x\<^sup>\<star> \<cdot> d z \<le> d z \<cdot> y\<^sup>\<star>"
using assms(1) by force
qed
lemma d_star_induct:
assumes "\<And> x y z. d z + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> d z \<le> y"
and "\<And> x y z. (x \<cdot> d y) \<cdot> z = x \<cdot> (d y \<cdot> z)"
and "\<And> x y z. (d x \<cdot> y) \<cdot> z = d x \<cdot> (y \<cdot> z)"
shows "d (x \<cdot> y) \<le> d y \<Longrightarrow> d (x\<^sup>\<star> \<cdot> y) \<le> d y"
proof -
fix x y
assume "d (x \<cdot> y) \<le> d y"
hence "x \<cdot> d y \<le> d y \<cdot> x"
by (simp add: demod1)
hence "x\<^sup>\<star> \<cdot> d y \<le> d y \<cdot> x\<^sup>\<star>"
using assms(1) assms(2) assms(3) d_star_sim1 by blast
thus "d (x\<^sup>\<star> \<cdot> y) \<le> d y"
by (simp add: demod2)
qed
end
subsection \<open>C-Omega Algebras\<close>
text \<open>These structures do not feature in~\cite{FurusawaS15a}, but in fact,
many lemmas from Section 13 can be proved in this setting. The proto-quantales and c-quantales
using in~\cite{FurusawaS15a} provide a more expressive setting in which least and greatest fixpoints
need not be postulated; they exists due to properties of sequential composition and addition over
complete lattices.\<close>
class c_omega_algebra = c_kleene_algebra + omega_op +
assumes om_unfold: "x\<^sup>\<omega> \<le> x \<cdot> x\<^sup>\<omega>"
and om_coinduct: "y \<le> x \<cdot> y \<Longrightarrow> y \<le> x\<^sup>\<omega>"
begin
text \<open>Lemma 13.4.\<close>
lemma om_unfold_eq [simp]: "x \<cdot> x\<^sup>\<omega> = x\<^sup>\<omega>"
- apply (rule antisym)
+ apply (rule order.antisym)
using local.om_coinduct local.om_unfold local.s_prod_isol by auto
lemma om_iso: "x \<le> y \<Longrightarrow> x\<^sup>\<omega> \<le> y\<^sup>\<omega>"
by (metis local.om_coinduct local.s_prod_isor om_unfold_eq)
text \<open>Lemma 13.5.\<close>
lemma zero_om [simp]: "0\<^sup>\<omega> = 0"
by (metis local.s_prod_annil om_unfold_eq)
lemma s_id_om [simp]: "1\<^sub>\<sigma>\<^sup>\<omega> = U"
- by (simp add: local.U_def local.eq_iff local.om_coinduct)
+ by (simp add: local.U_def order.eq_iff local.om_coinduct)
lemma p_id_om [simp]: "1\<^sub>\<pi>\<^sup>\<omega> = 1\<^sub>\<pi>"
by (metis local.c_x_prop om_unfold_eq)
lemma nc_om [simp]: "nc\<^sup>\<omega> = U"
- using local.U_def local.eq_iff local.s_le_nc om_iso s_id_om by blast
+ using local.U_def order.eq_iff local.s_le_nc om_iso s_id_om by blast
lemma U_om [simp]: "U\<^sup>\<omega> = U"
- by (simp add: local.U_def local.eq_iff local.om_coinduct)
+ by (simp add: local.U_def order.eq_iff local.om_coinduct)
text \<open>Lemma 13.6.\<close>
lemma tau_om1: "\<tau> x \<le> \<tau> (x\<^sup>\<omega>)"
using local.om_coinduct local.s_prod_isor local.tau_def local.tau_int by fastforce
lemma tau_om2 [simp]: "\<tau> x\<^sup>\<omega> = \<tau> x"
by (metis local.cl6 local.tau_def om_unfold_eq)
lemma tau_om3: "(\<tau> x)\<^sup>\<omega> \<le> \<tau> (x\<^sup>\<omega>)"
by (simp add: tau_om1)
text \<open>Lemma 13.7.\<close>
lemma om_nu_tau: "(\<nu> x)\<^sup>\<omega> + (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x \<le> x\<^sup>\<omega>"
proof -
have "(\<nu> x)\<^sup>\<omega> + (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x = (\<nu> x)\<^sup>\<omega> + (1\<^sub>\<sigma> + \<nu> x \<cdot> (\<nu> x)\<^sup>\<star>) \<cdot> \<tau> x"
by auto
also have "... = (\<nu> x)\<^sup>\<omega> + \<tau> x + \<nu> x \<cdot> (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x"
using add_assoc local.s_prod_distr local.s_prod_idl by presburger
also have "... = \<tau> x + \<nu> x \<cdot> (\<nu> x)\<^sup>\<omega> + \<nu> x \<cdot> (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x"
by (simp add: add_ac)
also have "... \<le> \<tau> x + \<nu> x \<cdot> ((\<nu> x)\<^sup>\<omega> + (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x)"
by (metis add_assoc local.cl5 local.lat_dist1 local.inf.absorb_iff1 local.s_prod_subdistl local.tau_def)
also have "... = x \<cdot> ((\<nu> x)\<^sup>\<omega> + (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x)"
by (metis local.sprod_tau_nu)
finally show ?thesis
using local.om_coinduct by blast
qed
end
subsection \<open>C-Nabla Algebras\<close>
text \<open>Nabla-algebras provide yet another way of formalising non-terminating behaviour in Section 13.\<close>
class c_nabla_algebra = c_omega_algebra +
fixes nabla :: "'a \<Rightarrow> 'a" ("\<nabla>")
assumes nabla_unfold: "\<nabla> x \<le> d (x \<cdot> \<nabla> x)"
and nabla_coinduct: "d y \<le> d (x \<cdot> y) \<Longrightarrow> d y \<le> \<nabla> x"
begin
lemma nabla_unfold_eq [simp]: "\<nabla> x = d (x \<cdot> \<nabla> x)"
-proof (rule antisym)
+proof (rule order.antisym)
show "\<nabla> x \<le> d (x \<cdot> \<nabla> x)"
using local.nabla_unfold by blast
have "d (x \<cdot> \<nabla> x) \<le> d (x \<cdot> d (x \<cdot> \<nabla> x))"
by (metis local.d_def local.mult_commute local.mult_isol local.nabla_unfold local.s_prod_isol local.s_prod_isor)
also have "... = d (x \<cdot> (x \<cdot> \<nabla> x))"
using local.d_loc_ax by blast
finally show "d (x \<cdot> \<nabla> x) \<le> \<nabla> x"
by (simp add: local.nabla_coinduct)
qed
lemma nabla_le_s: "\<nabla> x \<le> 1\<^sub>\<sigma>"
by (metis local.d_sub_id_ax nabla_unfold_eq)
lemma nabla_nu [simp]: "\<nu> (\<nabla> x) = \<nabla> x"
using local.nu_ideal1 local.nu_s nabla_le_s by blast
text \<open>Proposition 13.9.\<close>
lemma nabla_omega_U:
assumes "\<And>x y z. x \<cdot> (d y \<cdot> z) = (x \<cdot> d y ) \<cdot> z"
shows "(\<nu> x)\<^sup>\<omega> = \<nabla> (\<nu> x) \<cdot> U"
-proof (rule antisym)
+proof (rule order.antisym)
have "d ((\<nu> x)\<^sup>\<omega>) \<le> \<nabla> (\<nu> x)"
using local.nabla_coinduct local.om_unfold_eq local.order_refl by presburger
hence "(\<nu> x)\<^sup>\<omega> \<le> \<nabla> (\<nu> x) \<cdot> (\<nu> x)\<^sup>\<omega>"
using local.dlp_ax local.dual_order.trans local.s_prod_isor by blast
thus "(\<nu> x)\<^sup>\<omega> \<le> \<nabla> (\<nu> x) \<cdot> U"
using local.U_def local.dual_order.trans local.s_prod_isol by blast
have "\<nu> x \<cdot> (\<nabla> (\<nu> x) \<cdot> U) = (\<nu> x \<cdot> d (\<nabla> (\<nu> x))) \<cdot> U"
by (metis assms local.d_s_subid nabla_le_s)
also have "... = (\<nu> (\<nu> x \<cdot> \<nu> (\<nabla> (\<nu> x)))) \<cdot> U"
by (metis local.d_s_subid nabla_le_s nabla_nu local.alpha_prod_closed)
also have "... = d (\<nu> (\<nu> x \<cdot> \<nu> (\<nabla> (\<nu> x)))) \<cdot> U"
using local.ax5_d local.nu_def by presburger
also have "... = d (\<nu> x \<cdot> \<nabla> (\<nu> x)) \<cdot> U"
by (metis local.alpha_prod_closed nabla_nu)
finally show "\<nabla> (\<nu> x) \<cdot> U \<le> (\<nu> x)\<^sup>\<omega>"
using local.nabla_unfold local.om_coinduct local.s_prod_isor by presburger
qed
text \<open>Corollary 13.10.\<close>
lemma nabla_omega_U_cor:
assumes "\<And>x y z. x \<cdot> (d y \<cdot> z) = (x \<cdot> d y ) \<cdot> z"
shows "\<nabla> (\<nu> x) \<cdot> U + (\<nu> x)\<^sup>\<star> \<cdot> \<tau> x \<le> x\<^sup>\<omega>"
by (metis assms nabla_omega_U local.om_nu_tau)
text \<open>Lemma 13.11.\<close>
lemma nu_om_nu:
assumes "\<And>x y z. x \<cdot> (d y \<cdot> z) = (x \<cdot> d y ) \<cdot> z"
shows "\<nu> ((\<nu> x)\<^sup>\<omega>) = \<nabla> (\<nu> x) \<cdot> nc"
proof -
have "\<nu> ((\<nu> x)\<^sup>\<omega>) = \<nu> (\<nabla> (\<nu> x) \<cdot> U)"
using assms nabla_omega_U by presburger
also have "... = \<nu> (d (\<nabla> (\<nu> x)) \<cdot> U)"
by (metis local.d_s_subid nabla_le_s)
also have "... = (\<nabla> (\<nu> x)) \<cdot> \<nu> U"
by (metis local.d_nu local.d_s_subid nabla_le_s)
finally show ?thesis
using local.nu_U by presburger
qed
lemma tau_om_nu:
assumes "\<And>x y z. x \<cdot> (d y \<cdot> z) = (x \<cdot> d y ) \<cdot> z"
shows "\<tau> ((\<nu> x)\<^sup>\<omega>) = \<nabla> (\<nu> x) \<cdot> 1\<^sub>\<pi>"
proof -
have "\<tau> ((\<nu> x)\<^sup>\<omega>) = \<tau> (\<nabla> (\<nu> x) \<cdot> U)"
by (metis assms nabla_omega_U)
also have "... = \<nabla> (\<nu> x) \<cdot> \<tau> U"
using local.tau_s_prod by blast
finally show ?thesis
using local.tau_U by blast
qed
text \<open>Proposition 13.12.\<close>
lemma wf_eq_defl: "(\<forall>y. d y \<le> d (x \<cdot> y) \<longrightarrow> d y = 0) \<longleftrightarrow> (\<forall>y. y \<le> x \<cdot> y \<longrightarrow> y = 0)"
apply standard
apply (metis local.d_add_ax local.d_rest_ax local.less_eq_def local.s_prod_annil)
by (metis local.c2_d local.c4 local.d_def local.mult_commute local.mult_onel local.p_rpd_annir local.s_prod_isor)
lemma defl_eq_om_trivial: "x\<^sup>\<omega> = 0 \<longleftrightarrow> (\<forall>y. y \<le> x \<cdot> y \<longrightarrow> y = 0)"
using local.join.bot_unique local.om_coinduct by auto
lemma wf_eq_om_trivial: "x\<^sup>\<omega> = 0 \<longleftrightarrow> (\<forall>y. d y \<le> d (x \<cdot> y) \<longrightarrow> d y = 0)"
by (simp add: defl_eq_om_trivial wf_eq_defl)
end
subsection \<open>Proto-Quantales\<close>
text \<open>Finally we define the class of proto-quantales and prove some of the
remaining facts from the article. Full c-quantales, as defined there, are not needed
for these proofs.\<close>
class proto_quantale = complete_lattice + proto_monoid +
assumes Sup_mult_distr: "Sup X \<cdot> y = Sup {x \<cdot> y | x. x \<in> X}"
and isol: "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
begin
sublocale pd?: proto_dioid "1\<^sub>\<sigma>" "(\<cdot>)" sup "(\<le>)" "(<)" "Sup {}"
proof
show "\<And>x y. (x \<le> y) = (sup x y = y)"
by (simp add: local.le_iff_sup)
show "\<And>x y. (x < y) = (x \<le> y \<and> x \<noteq> y)"
by (simp add: local.order.strict_iff_order)
show "\<And>x y z. sup (sup x y) z = sup x (sup y z)"
by (simp add: local.sup_assoc)
show "\<And>x y. sup x y = sup y x"
by (simp add: local.sup_commute)
show "\<And>x. sup x x = x"
by (simp add: insert_commute)
show "\<And>x. sup (Sup {}) x = x"
by simp
show "\<And>x y z. sup (x \<cdot> y) (x \<cdot> z) \<le> x \<cdot> (sup y z)"
by (simp add: local.isol)
show " \<And>x. Sup {} \<cdot> x = Sup {}"
proof -
fix x :: 'a
have "\<forall>A a. {} \<noteq> A \<or> {} = {aa. \<exists>ab. (aa::'a) = ab \<cdot> a \<and> ab \<in> A}"
by fastforce
thus "Sup {} \<cdot> x = Sup {}"
using local.Sup_mult_distr by presburger
qed
show "\<And>x y z. (sup x y) \<cdot> z = sup (x \<cdot> z) (y \<cdot> z)"
proof -
fix x y z
have "(sup x y) \<cdot> z = Sup {x,y} \<cdot> z"
by simp
moreover have "... = sup (x \<cdot> z) (y \<cdot> z)"
by (subst Sup_mult_distr, rule Sup_eqI, auto)
thus "(sup x y) \<cdot> z = sup (x \<cdot> z) (y \<cdot> z)"
using calculation by presburger
qed
qed
definition star_rd :: "'a \<Rightarrow> 'a" where
"star_rd x = Sup {power_rd x i |i. i \<in> \<nat>}"
definition star_sq :: "'a \<Rightarrow> 'a" where
"star_sq x = Sup {power_sq x i |i. i \<in> \<nat>}"
text \<open>Now we prove Lemma 12.6.\<close>
lemma star_rd_le_sq: "star_rd x \<le> star_sq x"
apply (auto simp: star_rd_def star_sq_def)
apply (rule Sup_mono)
using pd.power_rd_le_sq by auto
lemma star_sq_le_rd: "star_sq x \<le> star_rd x"
apply (auto simp: star_rd_def star_sq_def)
apply (rule Sup_mono)
apply auto
by (metis Nats_1 Nats_add Suc_eq_plus1 local.Sup_empty pd.power_sq_le_rd)
lemma star_rd_sq: "star_rd x = star_sq x"
by (simp add: local.dual_order.antisym star_rd_le_sq star_sq_le_rd)
lemma star_sq_power: "star_sq x = Sup {pd.p_power (sup 1\<^sub>\<sigma> x) i | i. i \<in> \<nat>}"
by (auto simp: star_sq_def pd.power_sq_power [symmetric] intro: Sup_eqI)
text \<open>The following lemma should be somewhere close to complete lattices.\<close>
end
lemma mono_aux: "mono (\<lambda>y. sup (z:: 'a :: proto_quantale) (x \<cdot> y))"
by (meson mono_def order_refl pd.s_prod_isol sup.mono)
lemma gfp_lfp_prop: "sup (gfp (\<lambda>(y :: 'a :: proto_quantale). x \<cdot> y)) (lfp (\<lambda>y. sup z (x \<cdot> y))) \<le> gfp (\<lambda>y. sup z (x \<cdot> y))"
apply (simp, rule conjI)
apply (simp add: gfp_mono)
by (simp add: lfp_le_gfp mono_aux)
end
diff --git a/thys/Nested_Multisets_Ordinals/Multiset_More.thy b/thys/Nested_Multisets_Ordinals/Multiset_More.thy
--- a/thys/Nested_Multisets_Ordinals/Multiset_More.thy
+++ b/thys/Nested_Multisets_Ordinals/Multiset_More.thy
@@ -1,1025 +1,1025 @@
(* Title: More about Multisets
Author: Mathias Fleury <mathias.fleury at mpi-inf.mpg.de>, 2015
Author: Jasmin Blanchette <blanchette at in.tum.de>, 2014, 2015
Author: Anders Schlichtkrull <andschl at dtu.dk>, 2017
Author: Dmitriy Traytel <traytel at in.tum.de>, 2014
Maintainer: Mathias Fleury <mathias.fleury at mpi-inf.mpg.de>
*)
section \<open>More about Multisets\<close>
theory Multiset_More
imports
"HOL-Library.Multiset_Order"
"HOL-Library.Sublist"
begin
text \<open>
Isabelle's theory of finite multisets is not as developed as other areas, such as lists and sets.
The present theory introduces some missing concepts and lemmas. Some of it is expected to move to
Isabelle's library.
\<close>
subsection \<open>Basic Setup\<close>
declare
diff_single_trivial [simp]
in_image_mset [iff]
image_mset.compositionality [simp]
(*To have the same rules as the set counter-part*)
mset_subset_eqD[dest, intro?] (*@{thm subsetD}*)
Multiset.in_multiset_in_set[simp]
inter_add_left1[simp]
inter_add_left2[simp]
inter_add_right1[simp]
inter_add_right2[simp]
sum_mset_sum_list[simp]
subsection \<open>Lemmas about Intersection, Union and Pointwise Inclusion\<close>
lemma subset_mset_imp_subset_add_mset: "A \<subseteq># B \<Longrightarrow> A \<subseteq># add_mset x B"
by (auto simp add: subseteq_mset_def le_SucI)
lemma subset_add_mset_notin_subset_mset: \<open>A \<subseteq># add_mset b B \<Longrightarrow> b \<notin># A \<Longrightarrow> A \<subseteq># B\<close>
by (simp add: subset_mset.le_iff_sup)
lemma subset_msetE: "\<lbrakk>A \<subset># B; \<lbrakk>A \<subseteq># B; \<not> B \<subseteq># A\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by (simp add: subset_mset.less_le_not_le)
lemma Diff_triv_mset: "M \<inter># N = {#} \<Longrightarrow> M - N = M"
by (metis diff_intersect_left_idem diff_zero)
lemma diff_intersect_sym_diff: "(A - B) \<inter># (B - A) = {#}"
by (rule multiset_eqI) simp
declare subset_msetE [elim!]
lemma subseq_mset_subseteq_mset: "subseq xs ys \<Longrightarrow> mset xs \<subseteq># mset ys"
proof (induct xs arbitrary: ys)
case (Cons x xs)
note Outer_Cons = this
then show ?case
proof (induct ys)
case (Cons y ys)
have "subseq xs ys"
by (metis Cons.prems(2) subseq_Cons' subseq_Cons2_iff)
then show ?case
using Cons by (metis mset.simps(2) mset_subset_eq_add_mset_cancel subseq_Cons2_iff
subset_mset_imp_subset_add_mset)
qed simp
qed simp
subsection \<open>Lemmas about Filter and Image\<close>
lemma count_image_mset_ge_count: "count (image_mset f A) (f b) \<ge> count A b"
by (induction A) auto
lemma count_image_mset_inj:
assumes \<open>inj f\<close>
shows \<open>count (image_mset f M) (f x) = count M x\<close>
by (induct M) (use assms in \<open>auto simp: inj_on_def\<close>)
lemma count_image_mset_le_count_inj_on:
"inj_on f (set_mset M) \<Longrightarrow> count (image_mset f M) y \<le> count M (inv_into (set_mset M) f y)"
proof (induct M)
case (add x M)
note ih = this(1) and inj_xM = this(2)
have inj_M: "inj_on f (set_mset M)"
using inj_xM by simp
show ?case
proof (cases "x \<in># M")
case x_in_M: True
show ?thesis
proof (cases "y = f x")
case y_eq_fx: True
show ?thesis
using x_in_M ih[OF inj_M] unfolding y_eq_fx by (simp add: inj_M insert_absorb)
next
case y_ne_fx: False
show ?thesis
using x_in_M ih[OF inj_M] y_ne_fx insert_absorb by fastforce
qed
next
case x_ni_M: False
show ?thesis
proof (cases "y = f x")
case y_eq_fx: True
have "f x \<notin># image_mset f M"
using x_ni_M inj_xM by force
thus ?thesis
unfolding y_eq_fx
by (metis (no_types) inj_xM count_add_mset count_greater_eq_Suc_zero_iff count_inI
image_mset_add_mset inv_into_f_f union_single_eq_member)
next
case y_ne_fx: False
show ?thesis
proof (rule ccontr)
assume neg_conj: "\<not> count (image_mset f (add_mset x M)) y
\<le> count (add_mset x M) (inv_into (set_mset (add_mset x M)) f y)"
have cnt_y: "count (add_mset (f x) (image_mset f M)) y = count (image_mset f M) y"
using y_ne_fx by simp
have "inv_into (set_mset M) f y \<in># add_mset x M \<Longrightarrow>
inv_into (set_mset (add_mset x M)) f (f (inv_into (set_mset M) f y)) =
inv_into (set_mset M) f y"
by (meson inj_xM inv_into_f_f)
hence "0 < count (image_mset f (add_mset x M)) y \<Longrightarrow>
count M (inv_into (set_mset M) f y) = 0 \<or> x = inv_into (set_mset M) f y"
using neg_conj cnt_y ih[OF inj_M]
by (metis (no_types) count_add_mset count_greater_zero_iff count_inI f_inv_into_f
image_mset_add_mset set_image_mset)
thus False
using neg_conj cnt_y x_ni_M ih[OF inj_M]
by (metis (no_types) count_greater_zero_iff count_inI eq_iff image_mset_add_mset
less_imp_le)
qed
qed
qed
qed simp
lemma mset_filter_compl: "mset (filter p xs) + mset (filter (Not \<circ> p) xs) = mset xs"
by (induction xs) (auto simp: ac_simps)
text \<open>Near duplicate of @{thm [source] filter_eq_replicate_mset}: @{thm filter_eq_replicate_mset}.\<close>
lemma filter_mset_eq: "filter_mset ((=) L) A = replicate_mset (count A L) L"
by (auto simp: multiset_eq_iff)
lemma filter_mset_cong[fundef_cong]:
assumes "M = M'" "\<And>a. a \<in># M \<Longrightarrow> P a = Q a"
shows "filter_mset P M = filter_mset Q M"
proof -
have "M - filter_mset Q M = filter_mset (\<lambda>a. \<not>Q a) M"
by (metis multiset_partition add_diff_cancel_left')
then show ?thesis
by (auto simp: filter_mset_eq_conv assms)
qed
lemma image_mset_filter_swap: "image_mset f {# x \<in># M. P (f x)#} = {# x \<in># image_mset f M. P x#}"
by (induction M) auto
lemma image_mset_cong2:
"(\<And>x. x \<in># M \<Longrightarrow> f x = g x) \<Longrightarrow> M = N \<Longrightarrow> image_mset f M = image_mset g N"
by (hypsubst, rule image_mset_cong)
lemma filter_mset_empty_conv: \<open>(filter_mset P M = {#}) = (\<forall>L\<in>#M. \<not> P L)\<close>
by (induction M) auto
lemma multiset_filter_mono2: \<open>filter_mset P A \<subseteq># filter_mset Q A \<longleftrightarrow> (\<forall>a\<in>#A. P a \<longrightarrow> Q a)\<close>
- by (induction A) (auto intro: subset_mset.order.trans)
+ by (induction A) (auto intro: subset_mset.trans)
lemma image_filter_cong:
assumes \<open>\<And>C. C \<in># M \<Longrightarrow> P C \<Longrightarrow> f C = g C\<close>
shows \<open>{#f C. C \<in># {#C \<in># M. P C#}#} = {#g C | C\<in># M. P C#}\<close>
using assms by (induction M) auto
lemma image_mset_filter_swap2: \<open>{#C \<in># {#P x. x \<in># D#}. Q C #} = {#P x. x \<in># {#C| C \<in># D. Q (P C)#}#}\<close>
by (simp add: image_mset_filter_swap)
declare image_mset_cong2 [cong]
lemma filter_mset_empty_if_finite_and_filter_set_empty:
assumes
"{x \<in> X. P x} = {}" and
"finite X"
shows "{#x \<in># mset_set X. P x#} = {#}"
proof -
have empty_empty: "\<And>Y. set_mset Y = {} \<Longrightarrow> Y = {#}"
by auto
from assms have "set_mset {#x \<in># mset_set X. P x#} = {}"
by auto
then show ?thesis
by (rule empty_empty)
qed
subsection \<open>Lemmas about Sum\<close>
lemma sum_image_mset_sum_map[simp]: "sum_mset (image_mset f (mset xs)) = sum_list (map f xs)"
by (metis mset_map sum_mset_sum_list)
lemma sum_image_mset_mono:
fixes f :: "'a \<Rightarrow> 'b::canonically_ordered_monoid_add"
assumes sub: "A \<subseteq># B"
shows "(\<Sum>m \<in># A. f m) \<le> (\<Sum>m \<in># B. f m)"
by (metis image_mset_union le_iff_add sub subset_mset.add_diff_inverse sum_mset.union)
lemma sum_image_mset_mono_mem:
"n \<in># M \<Longrightarrow> f n \<le> (\<Sum>m \<in># M. f m)" for f :: "'a \<Rightarrow> 'b::canonically_ordered_monoid_add"
using le_iff_add multi_member_split by fastforce
lemma count_sum_mset_if_1_0: \<open>count M a = (\<Sum>x\<in>#M. if x = a then 1 else 0)\<close>
by (induction M) auto
lemma sum_mset_dvd:
fixes k :: "'a::comm_semiring_1_cancel"
assumes "\<forall>m \<in># M. k dvd f m"
shows "k dvd (\<Sum>m \<in># M. f m)"
using assms by (induct M) auto
lemma sum_mset_distrib_div_if_dvd:
fixes k :: "'a::unique_euclidean_semiring"
assumes "\<forall>m \<in># M. k dvd f m"
shows "(\<Sum>m \<in># M. f m) div k = (\<Sum>m \<in># M. f m div k)"
using assms by (induct M) (auto simp: div_plus_div_distrib_dvd_left)
subsection \<open>Lemmas about Remove\<close>
lemma set_mset_minus_replicate_mset[simp]:
"n \<ge> count A a \<Longrightarrow> set_mset (A - replicate_mset n a) = set_mset A - {a}"
"n < count A a \<Longrightarrow> set_mset (A - replicate_mset n a) = set_mset A"
unfolding set_mset_def by (auto split: if_split simp: not_in_iff)
abbreviation removeAll_mset :: "'a \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where
"removeAll_mset C M \<equiv> M - replicate_mset (count M C) C"
lemma mset_removeAll[simp, code]: "removeAll_mset C (mset L) = mset (removeAll C L)"
by (induction L) (auto simp: ac_simps multiset_eq_iff split: if_split_asm)
lemma removeAll_mset_filter_mset: "removeAll_mset C M = filter_mset ((\<noteq>) C) M"
by (induction M) (auto simp: ac_simps multiset_eq_iff)
abbreviation remove1_mset :: "'a \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where
"remove1_mset C M \<equiv> M - {#C#}"
lemma removeAll_subseteq_remove1_mset: "removeAll_mset x M \<subseteq># remove1_mset x M"
by (auto simp: subseteq_mset_def)
lemma in_remove1_mset_neq:
assumes ab: "a \<noteq> b"
shows "a \<in># remove1_mset b C \<longleftrightarrow> a \<in># C"
by (metis assms diff_single_trivial in_diffD insert_DiffM insert_noteq_member)
lemma size_mset_removeAll_mset_le_iff: "size (removeAll_mset x M) < size M \<longleftrightarrow> x \<in># M"
by (auto intro: count_inI mset_subset_size simp: subset_mset_def multiset_eq_iff)
lemma size_remove1_mset_If: \<open>size (remove1_mset x M) = size M - (if x \<in># M then 1 else 0)\<close>
by (auto simp: size_Diff_subset_Int)
lemma size_mset_remove1_mset_le_iff: "size (remove1_mset x M) < size M \<longleftrightarrow> x \<in># M"
using less_irrefl
by (fastforce intro!: mset_subset_size elim: in_countE simp: subset_mset_def multiset_eq_iff)
lemma remove_1_mset_id_iff_notin: "remove1_mset a M = M \<longleftrightarrow> a \<notin># M"
by (meson diff_single_trivial multi_drop_mem_not_eq)
lemma id_remove_1_mset_iff_notin: "M = remove1_mset a M \<longleftrightarrow> a \<notin># M"
using remove_1_mset_id_iff_notin by metis
lemma remove1_mset_eqE:
"remove1_mset L x1 = M \<Longrightarrow>
(L \<in># x1 \<Longrightarrow> x1 = M + {#L#} \<Longrightarrow> P) \<Longrightarrow>
(L \<notin># x1 \<Longrightarrow> x1 = M \<Longrightarrow> P) \<Longrightarrow>
P"
by (cases "L \<in># x1") auto
lemma image_filter_ne_mset[simp]:
"image_mset f {#x \<in># M. f x \<noteq> y#} = removeAll_mset y (image_mset f M)"
by (induction M) simp_all
lemma image_mset_remove1_mset_if:
"image_mset f (remove1_mset a M) =
(if a \<in># M then remove1_mset (f a) (image_mset f M) else image_mset f M)"
by (auto simp: image_mset_Diff)
lemma filter_mset_neq: "{#x \<in># M. x \<noteq> y#} = removeAll_mset y M"
by (metis add_diff_cancel_left' filter_eq_replicate_mset multiset_partition)
lemma filter_mset_neq_cond: "{#x \<in># M. P x \<and> x \<noteq> y#} = removeAll_mset y {# x\<in>#M. P x#}"
by (metis filter_filter_mset filter_mset_neq)
lemma remove1_mset_add_mset_If:
"remove1_mset L (add_mset L' C) = (if L = L' then C else remove1_mset L C + {#L'#})"
by (auto simp: multiset_eq_iff)
lemma minus_remove1_mset_if:
"A - remove1_mset b B = (if b \<in># B \<and> b \<in># A \<and> count A b \<ge> count B b then {#b#} + (A - B) else A - B)"
by (auto simp: multiset_eq_iff count_greater_zero_iff[symmetric]
simp del: count_greater_zero_iff)
lemma add_mset_eq_add_mset_ne:
"a \<noteq> b \<Longrightarrow> add_mset a A = add_mset b B \<longleftrightarrow> a \<in># B \<and> b \<in># A \<and> A = add_mset b (B - {#a#})"
by (metis (no_types, lifting) diff_single_eq_union diff_union_swap multi_self_add_other_not_self
remove_1_mset_id_iff_notin union_single_eq_diff)
lemma add_mset_eq_add_mset: \<open>add_mset a M = add_mset b M' \<longleftrightarrow>
(a = b \<and> M = M') \<or> (a \<noteq> b \<and> b \<in># M \<and> add_mset a (M - {#b#}) = M')\<close>
by (metis add_mset_eq_add_mset_ne add_mset_remove_trivial union_single_eq_member)
(* TODO move to Multiset: could replace add_mset_remove_trivial_eq? *)
lemma add_mset_remove_trivial_iff: \<open>N = add_mset a (N - {#b#}) \<longleftrightarrow> a \<in># N \<and> a = b\<close>
by (metis add_left_cancel add_mset_remove_trivial insert_DiffM2 single_eq_single
size_mset_remove1_mset_le_iff union_single_eq_member)
lemma trivial_add_mset_remove_iff: \<open>add_mset a (N - {#b#}) = N \<longleftrightarrow> a \<in># N \<and> a = b\<close>
by (subst eq_commute) (fact add_mset_remove_trivial_iff)
lemma remove1_single_empty_iff[simp]: \<open>remove1_mset L {#L'#} = {#} \<longleftrightarrow> L = L'\<close>
using add_mset_remove_trivial_iff by fastforce
lemma add_mset_less_imp_less_remove1_mset:
assumes xM_lt_N: "add_mset x M < N"
shows "M < remove1_mset x N"
proof -
have "M < N"
using assms le_multiset_right_total mset_le_trans by blast
then show ?thesis
by (metis add_less_cancel_right add_mset_add_single diff_single_trivial insert_DiffM2 xM_lt_N)
qed
subsection \<open>Lemmas about Replicate\<close>
lemma replicate_mset_minus_replicate_mset_same[simp]:
"replicate_mset m x - replicate_mset n x = replicate_mset (m - n) x"
by (induct m arbitrary: n, simp, metis left_diff_repeat_mset_distrib' repeat_mset_replicate_mset)
lemma replicate_mset_subset_iff_lt[simp]: "replicate_mset m x \<subset># replicate_mset n x \<longleftrightarrow> m < n"
by (induct n m rule: diff_induct) (auto intro: subset_mset.gr_zeroI)
lemma replicate_mset_subseteq_iff_le[simp]: "replicate_mset m x \<subseteq># replicate_mset n x \<longleftrightarrow> m \<le> n"
by (induct n m rule: diff_induct) auto
lemma replicate_mset_lt_iff_lt[simp]: "replicate_mset m x < replicate_mset n x \<longleftrightarrow> m < n"
by (induct n m rule: diff_induct) (auto intro: subset_mset.gr_zeroI gr_zeroI)
lemma replicate_mset_le_iff_le[simp]: "replicate_mset m x \<le> replicate_mset n x \<longleftrightarrow> m \<le> n"
by (induct n m rule: diff_induct) auto
lemma replicate_mset_eq_iff[simp]:
"replicate_mset m x = replicate_mset n y \<longleftrightarrow> m = n \<and> (m \<noteq> 0 \<longrightarrow> x = y)"
by (cases m; cases n; simp)
(metis in_replicate_mset insert_noteq_member size_replicate_mset union_single_eq_diff)
lemma replicate_mset_plus: "replicate_mset (a + b) C = replicate_mset a C + replicate_mset b C"
by (induct a) (auto simp: ac_simps)
lemma mset_replicate_replicate_mset: "mset (replicate n L) = replicate_mset n L"
by (induction n) auto
lemma set_mset_single_iff_replicate_mset: "set_mset U = {a} \<longleftrightarrow> (\<exists>n > 0. U = replicate_mset n a)"
by (rule, metis count_greater_zero_iff count_replicate_mset insertI1 multi_count_eq singletonD
zero_less_iff_neq_zero, force)
lemma ex_replicate_mset_if_all_elems_eq:
assumes "\<forall>x \<in># M. x = y"
shows "\<exists>n. M = replicate_mset n y"
using assms by (metis count_replicate_mset mem_Collect_eq multiset_eqI neq0_conv set_mset_def)
subsection \<open>Multiset and Set Conversions\<close>
lemma count_mset_set_if: "count (mset_set A) a = (if a \<in> A \<and> finite A then 1 else 0)"
by auto
lemma mset_set_set_mset_empty_mempty[iff]: "mset_set (set_mset D) = {#} \<longleftrightarrow> D = {#}"
by (simp add: mset_set_empty_iff)
lemma count_mset_set_le_one: "count (mset_set A) x \<le> 1"
by (simp add: count_mset_set_if)
lemma mset_set_set_mset_subseteq[simp]: "mset_set (set_mset A) \<subseteq># A"
by (simp add: mset_set_set_mset_msubset)
lemma mset_sorted_list_of_set[simp]: "mset (sorted_list_of_set A) = mset_set A"
by (metis mset_sorted_list_of_multiset sorted_list_of_mset_set)
lemma sorted_sorted_list_of_multiset[simp]:
"sorted (sorted_list_of_multiset (M :: 'a::linorder multiset))"
by (metis mset_sorted_list_of_multiset sorted_list_of_multiset_mset sorted_sort)
lemma mset_take_subseteq: "mset (take n xs) \<subseteq># mset xs"
apply (induct xs arbitrary: n)
apply simp
by (case_tac n) simp_all
lemma sorted_list_of_multiset_eq_Nil[simp]: "sorted_list_of_multiset M = [] \<longleftrightarrow> M = {#}"
by (metis mset_sorted_list_of_multiset sorted_list_of_multiset_empty)
subsection \<open>Duplicate Removal\<close>
(* TODO: use abbreviation? *)
definition remdups_mset :: "'v multiset \<Rightarrow> 'v multiset" where
"remdups_mset S = mset_set (set_mset S)"
lemma set_mset_remdups_mset[simp]: \<open>set_mset (remdups_mset A) = set_mset A\<close>
unfolding remdups_mset_def by auto
lemma count_remdups_mset_eq_1: "a \<in># remdups_mset A \<longleftrightarrow> count (remdups_mset A) a = 1"
unfolding remdups_mset_def by (auto simp: count_eq_zero_iff intro: count_inI)
lemma remdups_mset_empty[simp]: "remdups_mset {#} = {#}"
unfolding remdups_mset_def by auto
lemma remdups_mset_singleton[simp]: "remdups_mset {#a#} = {#a#}"
unfolding remdups_mset_def by auto
lemma remdups_mset_eq_empty[iff]: "remdups_mset D = {#} \<longleftrightarrow> D = {#}"
unfolding remdups_mset_def by blast
lemma remdups_mset_singleton_sum[simp]:
"remdups_mset (add_mset a A) = (if a \<in># A then remdups_mset A else add_mset a (remdups_mset A))"
unfolding remdups_mset_def by (simp_all add: insert_absorb)
lemma mset_remdups_remdups_mset[simp]: "mset (remdups D) = remdups_mset (mset D)"
by (induction D) (auto simp add: ac_simps)
declare mset_remdups_remdups_mset[symmetric, code]
definition distinct_mset :: "'a multiset \<Rightarrow> bool" where
"distinct_mset S \<longleftrightarrow> (\<forall>a. a \<in># S \<longrightarrow> count S a = 1)"
lemma distinct_mset_count_less_1: "distinct_mset S \<longleftrightarrow> (\<forall>a. count S a \<le> 1)"
using eq_iff nat_le_linear unfolding distinct_mset_def by fastforce
lemma distinct_mset_empty[simp]: "distinct_mset {#}"
unfolding distinct_mset_def by auto
lemma distinct_mset_singleton: "distinct_mset {#a#}"
unfolding distinct_mset_def by auto
lemma distinct_mset_union:
assumes dist: "distinct_mset (A + B)"
shows "distinct_mset A"
unfolding distinct_mset_count_less_1
proof (rule allI)
fix a
have \<open>count A a \<le> count (A + B) a\<close> by auto
moreover have \<open>count (A + B) a \<le> 1\<close>
using dist unfolding distinct_mset_count_less_1 by auto
ultimately show \<open>count A a \<le> 1\<close>
by simp
qed
lemma distinct_mset_minus[simp]: "distinct_mset A \<Longrightarrow> distinct_mset (A - B)"
by (metis diff_subset_eq_self mset_subset_eq_exists_conv distinct_mset_union)
lemma count_remdups_mset_If: \<open>count (remdups_mset A) a = (if a \<in># A then 1 else 0)\<close>
unfolding remdups_mset_def by auto
lemma distinct_mset_rempdups_union_mset:
assumes "distinct_mset A" and "distinct_mset B"
shows "A \<union># B = remdups_mset (A + B)"
using assms nat_le_linear unfolding remdups_mset_def
by (force simp add: multiset_eq_iff max_def count_mset_set_if distinct_mset_def not_in_iff)
lemma distinct_mset_add_mset[simp]: "distinct_mset (add_mset a L) \<longleftrightarrow> a \<notin># L \<and> distinct_mset L"
unfolding distinct_mset_def
apply (rule iffI)
apply (auto split: if_split_asm; fail)[]
by (auto simp: not_in_iff; fail)
lemma distinct_mset_size_eq_card: "distinct_mset C \<Longrightarrow> size C = card (set_mset C)"
by (induction C) auto
lemma distinct_mset_add:
"distinct_mset (L + L') \<longleftrightarrow> distinct_mset L \<and> distinct_mset L' \<and> L \<inter># L' = {#}"
by (induction L arbitrary: L') auto
lemma distinct_mset_set_mset_ident[simp]: "distinct_mset M \<Longrightarrow> mset_set (set_mset M) = M"
by (induction M) auto
lemma distinct_finite_set_mset_subseteq_iff[iff]:
assumes "distinct_mset M" "finite N"
shows "set_mset M \<subseteq> N \<longleftrightarrow> M \<subseteq># mset_set N"
by (metis assms distinct_mset_set_mset_ident finite_set_mset msubset_mset_set_iff)
lemma distinct_mem_diff_mset:
assumes dist: "distinct_mset M" and mem: "x \<in> set_mset (M - N)"
shows "x \<notin> set_mset N"
proof -
have "count M x = 1"
using dist mem by (meson distinct_mset_def in_diffD)
then show ?thesis
using mem by (metis count_greater_eq_one_iff in_diff_count not_less)
qed
lemma distinct_set_mset_eq:
assumes "distinct_mset M" "distinct_mset N" "set_mset M = set_mset N"
shows "M = N"
using assms distinct_mset_set_mset_ident by fastforce
lemma distinct_mset_union_mset[simp]:
\<open>distinct_mset (D \<union># C) \<longleftrightarrow> distinct_mset D \<and> distinct_mset C\<close>
unfolding distinct_mset_count_less_1 by force
lemma distinct_mset_inter_mset:
"distinct_mset C \<Longrightarrow> distinct_mset (C \<inter># D)"
"distinct_mset D \<Longrightarrow> distinct_mset (C \<inter># D)"
by (auto simp add: distinct_mset_def min_def count_eq_zero_iff elim!: le_SucE)
lemma distinct_mset_remove1_All: "distinct_mset C \<Longrightarrow> remove1_mset L C = removeAll_mset L C"
by (auto simp: multiset_eq_iff distinct_mset_count_less_1)
lemma distinct_mset_size_2: "distinct_mset {#a, b#} \<longleftrightarrow> a \<noteq> b"
by auto
lemma distinct_mset_filter: "distinct_mset M \<Longrightarrow> distinct_mset {# L \<in># M. P L#}"
by (simp add: distinct_mset_def)
lemma distinct_mset_mset_distinct[simp]: \<open>distinct_mset (mset xs) = distinct xs\<close>
by (induction xs) auto
lemma distinct_image_mset_inj:
\<open>inj_on f (set_mset M) \<Longrightarrow> distinct_mset (image_mset f M) \<longleftrightarrow> distinct_mset M\<close>
by (induction M) (auto simp: inj_on_def)
subsection \<open>Repeat Operation\<close>
lemma repeat_mset_compower: "repeat_mset n A = (((+) A) ^^ n) {#}"
by (induction n) auto
lemma repeat_mset_prod: "repeat_mset (m * n) A = (((+) (repeat_mset n A)) ^^ m) {#}"
by (induction m) (auto simp: repeat_mset_distrib)
subsection \<open>Cartesian Product\<close>
text \<open>Definition of the cartesian products over multisets. The construction mimics of the cartesian
product on sets and use the same theorem names (adding only the suffix \<open>_mset\<close> to Sigma
and Times). See file @{file \<open>~~/src/HOL/Product_Type.thy\<close>}\<close>
definition Sigma_mset :: "'a multiset \<Rightarrow> ('a \<Rightarrow> 'b multiset) \<Rightarrow> ('a \<times> 'b) multiset" where
"Sigma_mset A B \<equiv> \<Sum>\<^sub># {#{#(a, b). b \<in># B a#}. a \<in># A #}"
abbreviation Times_mset :: "'a multiset \<Rightarrow> 'b multiset \<Rightarrow> ('a \<times> 'b) multiset" (infixr "\<times>#" 80) where
"Times_mset A B \<equiv> Sigma_mset A (\<lambda>_. B)"
hide_const (open) Times_mset
text \<open>Contrary to the set version @{term \<open>SIGMA x:A. B\<close>}, we use the non-ASCII symbol \<open>\<in>#\<close>.\<close>
syntax
"_Sigma_mset" :: "[pttrn, 'a multiset, 'b multiset] => ('a * 'b) multiset"
("(3SIGMAMSET _\<in>#_./ _)" [0, 0, 10] 10)
translations
"SIGMAMSET x\<in>#A. B" == "CONST Sigma_mset A (\<lambda>x. B)"
text \<open>Link between the multiset and the set cartesian product:\<close>
lemma Times_mset_Times: "set_mset (A \<times># B) = set_mset A \<times> set_mset B"
unfolding Sigma_mset_def by auto
lemma Sigma_msetI [intro!]: "\<lbrakk>a \<in># A; b \<in># B a\<rbrakk> \<Longrightarrow> (a, b) \<in># Sigma_mset A B"
by (unfold Sigma_mset_def) auto
lemma Sigma_msetE[elim!]: "\<lbrakk>c \<in># Sigma_mset A B; \<And>x y. \<lbrakk>x \<in># A; y \<in># B x; c = (x, y)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (unfold Sigma_mset_def) auto
text \<open>Elimination of @{term "(a, b) \<in># A \<times># B"} -- introduces no eigenvariables.\<close>
lemma Sigma_msetD1: "(a, b) \<in># Sigma_mset A B \<Longrightarrow> a \<in># A"
by blast
lemma Sigma_msetD2: "(a, b) \<in># Sigma_mset A B \<Longrightarrow> b \<in># B a"
by blast
lemma Sigma_msetE2: "\<lbrakk>(a, b) \<in># Sigma_mset A B; \<lbrakk>a \<in># A; b \<in># B a\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by blast
lemma Sigma_mset_cong:
"\<lbrakk>A = B; \<And>x. x \<in># B \<Longrightarrow> C x = D x\<rbrakk> \<Longrightarrow> (SIGMAMSET x \<in># A. C x) = (SIGMAMSET x \<in># B. D x)"
by (metis (mono_tags, lifting) Sigma_mset_def image_mset_cong)
lemma count_sum_mset: "count (\<Sum>\<^sub># M) b = (\<Sum>P \<in># M. count P b)"
by (induction M) auto
lemma Sigma_mset_plus_distrib1[simp]: "Sigma_mset (A + B) C = Sigma_mset A C + Sigma_mset B C"
unfolding Sigma_mset_def by auto
lemma Sigma_mset_plus_distrib2[simp]:
"Sigma_mset A (\<lambda>i. B i + C i) = Sigma_mset A B + Sigma_mset A C"
unfolding Sigma_mset_def by (induction A) (auto simp: multiset_eq_iff)
lemma Times_mset_single_left: "{#a#} \<times># B = image_mset (Pair a) B"
unfolding Sigma_mset_def by auto
lemma Times_mset_single_right: "A \<times># {#b#} = image_mset (\<lambda>a. Pair a b) A"
unfolding Sigma_mset_def by (induction A) auto
lemma Times_mset_single_single[simp]: "{#a#} \<times># {#b#} = {#(a, b)#}"
unfolding Sigma_mset_def by simp
lemma count_image_mset_Pair:
"count (image_mset (Pair a) B) (x, b) = (if x = a then count B b else 0)"
by (induction B) auto
lemma count_Sigma_mset: "count (Sigma_mset A B) (a, b) = count A a * count (B a) b"
by (induction A) (auto simp: Sigma_mset_def count_image_mset_Pair)
lemma Sigma_mset_empty1[simp]: "Sigma_mset {#} B = {#}"
unfolding Sigma_mset_def by auto
lemma Sigma_mset_empty2[simp]: "A \<times># {#} = {#}"
by (auto simp: multiset_eq_iff count_Sigma_mset)
lemma Sigma_mset_mono:
assumes "A \<subseteq># C" and "\<And>x. x \<in># A \<Longrightarrow> B x \<subseteq># D x"
shows "Sigma_mset A B \<subseteq># Sigma_mset C D"
proof -
have "count A a * count (B a) b \<le> count C a * count (D a) b" for a b
using assms unfolding subseteq_mset_def by (metis count_inI eq_iff mult_eq_0_iff mult_le_mono)
then show ?thesis
by (auto simp: subseteq_mset_def count_Sigma_mset)
qed
lemma mem_Sigma_mset_iff[iff]: "((a,b) \<in># Sigma_mset A B) = (a \<in># A \<and> b \<in># B a)"
by blast
lemma mem_Times_mset_iff: "x \<in># A \<times># B \<longleftrightarrow> fst x \<in># A \<and> snd x \<in># B"
by (induct x) simp
lemma Sigma_mset_empty_iff: "(SIGMAMSET i\<in>#I. X i) = {#} \<longleftrightarrow> (\<forall>i\<in>#I. X i = {#})"
by (auto simp: Sigma_mset_def)
lemma Times_mset_subset_mset_cancel1: "x \<in># A \<Longrightarrow> (A \<times># B \<subseteq># A \<times># C) = (B \<subseteq># C)"
by (auto simp: subseteq_mset_def count_Sigma_mset)
lemma Times_mset_subset_mset_cancel2: "x \<in># C \<Longrightarrow> (A \<times># C \<subseteq># B \<times># C) = (A \<subseteq># B)"
by (auto simp: subseteq_mset_def count_Sigma_mset)
lemma Times_mset_eq_cancel2: "x \<in># C \<Longrightarrow> (A \<times># C = B \<times># C) = (A = B)"
by (auto simp: multiset_eq_iff count_Sigma_mset dest!: in_countE)
lemma split_paired_Ball_mset_Sigma_mset[simp]:
"(\<forall>z\<in>#Sigma_mset A B. P z) \<longleftrightarrow> (\<forall>x\<in>#A. \<forall>y\<in>#B x. P (x, y))"
by blast
lemma split_paired_Bex_mset_Sigma_mset[simp]:
"(\<exists>z\<in>#Sigma_mset A B. P z) \<longleftrightarrow> (\<exists>x\<in>#A. \<exists>y\<in>#B x. P (x, y))"
by blast
lemma sum_mset_if_eq_constant:
"(\<Sum>x\<in>#M. if a = x then (f x) else 0) = (((+) (f a)) ^^ (count M a)) 0"
by (induction M) (auto simp: ac_simps)
lemma iterate_op_plus: "(((+) k) ^^ m) 0 = k * m"
by (induction m) auto
lemma untion_image_mset_Pair_distribute:
"\<Sum>\<^sub>#{#image_mset (Pair x) (C x). x \<in># J - I#} =
\<Sum>\<^sub># {#image_mset (Pair x) (C x). x \<in># J#} - \<Sum>\<^sub>#{#image_mset (Pair x) (C x). x \<in># I#}"
by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant
iterate_op_plus diff_mult_distrib2)
lemma Sigma_mset_Un_distrib1: "Sigma_mset (I \<union># J) C = Sigma_mset I C \<union># Sigma_mset J C"
by (auto simp add: Sigma_mset_def union_mset_def untion_image_mset_Pair_distribute)
lemma Sigma_mset_Un_distrib2: "(SIGMAMSET i\<in>#I. A i \<union># B i) = Sigma_mset I A \<union># Sigma_mset I B"
by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant
Sigma_mset_def diff_mult_distrib2 iterate_op_plus max_def not_in_iff)
lemma Sigma_mset_Int_distrib1: "Sigma_mset (I \<inter># J) C = Sigma_mset I C \<inter># Sigma_mset J C"
by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant
Sigma_mset_def iterate_op_plus min_def not_in_iff)
lemma Sigma_mset_Int_distrib2: "(SIGMAMSET i\<in>#I. A i \<inter># B i) = Sigma_mset I A \<inter># Sigma_mset I B"
by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant
Sigma_mset_def iterate_op_plus min_def not_in_iff)
lemma Sigma_mset_Diff_distrib1: "Sigma_mset (I - J) C = Sigma_mset I C - Sigma_mset J C"
by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant
Sigma_mset_def iterate_op_plus min_def not_in_iff diff_mult_distrib2)
lemma Sigma_mset_Diff_distrib2: "(SIGMAMSET i\<in>#I. A i - B i) = Sigma_mset I A - Sigma_mset I B"
by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant
Sigma_mset_def iterate_op_plus min_def not_in_iff diff_mult_distrib)
lemma Sigma_mset_Union: "Sigma_mset (\<Sum>\<^sub>#X) B = (\<Sum>\<^sub># (image_mset (\<lambda>A. Sigma_mset A B) X))"
by (auto simp: multiset_eq_iff count_sum_mset count_image_mset_Pair sum_mset_if_eq_constant
Sigma_mset_def iterate_op_plus min_def not_in_iff sum_mset_distrib_left)
lemma Times_mset_Un_distrib1: "(A \<union># B) \<times># C = A \<times># C \<union># B \<times># C"
by (fact Sigma_mset_Un_distrib1)
lemma Times_mset_Int_distrib1: "(A \<inter># B) \<times># C = A \<times># C \<inter># B \<times># C"
by (fact Sigma_mset_Int_distrib1)
lemma Times_mset_Diff_distrib1: "(A - B) \<times># C = A \<times># C - B \<times># C"
by (fact Sigma_mset_Diff_distrib1)
lemma Times_mset_empty[simp]: "A \<times># B = {#} \<longleftrightarrow> A = {#} \<or> B = {#}"
by (auto simp: Sigma_mset_empty_iff)
lemma Times_insert_left: "A \<times># add_mset x B = A \<times># B + image_mset (\<lambda>a. Pair a x) A"
unfolding add_mset_add_single[of x B] Sigma_mset_plus_distrib2
by (simp add: Times_mset_single_right)
lemma Times_insert_right: "add_mset a A \<times># B = A \<times># B + image_mset (Pair a) B"
unfolding add_mset_add_single[of a A] Sigma_mset_plus_distrib1
by (simp add: Times_mset_single_left)
lemma fst_image_mset_times_mset [simp]:
"image_mset fst (A \<times># B) = (if B = {#} then {#} else repeat_mset (size B) A)"
by (induct B) (auto simp: Times_mset_single_right ac_simps Times_insert_left)
lemma snd_image_mset_times_mset [simp]:
"image_mset snd (A \<times># B) = (if A = {#} then {#} else repeat_mset (size A) B)"
by (induct B) (auto simp add: Times_mset_single_right Times_insert_left image_mset_const_eq)
lemma product_swap_mset: "image_mset prod.swap (A \<times># B) = B \<times># A"
by (induction A) (auto simp add: Times_mset_single_left Times_mset_single_right
Times_insert_right Times_insert_left)
context
begin
qualified definition product_mset :: "'a multiset \<Rightarrow> 'b multiset \<Rightarrow> ('a \<times> 'b) multiset" where
[code_abbrev]: "product_mset A B = A \<times># B"
lemma member_product_mset: "x \<in># product_mset A B \<longleftrightarrow> x \<in># A \<times># B"
by (simp add: Multiset_More.product_mset_def)
end
lemma count_Sigma_mset_abs_def: "count (Sigma_mset A B) = (\<lambda>(a, b) \<Rightarrow> count A a * count (B a) b)"
by (auto simp: fun_eq_iff count_Sigma_mset)
lemma Times_mset_image_mset1: "image_mset f A \<times># B = image_mset (\<lambda>(a, b). (f a, b)) (A \<times># B)"
by (induct B) (auto simp: Times_insert_left)
lemma Times_mset_image_mset2: "A \<times># image_mset f B = image_mset (\<lambda>(a, b). (a, f b)) (A \<times># B)"
by (induct A) (auto simp: Times_insert_right)
lemma sum_le_singleton: "A \<subseteq> {x} \<Longrightarrow> sum f A = (if x \<in> A then f x else 0)"
by (auto simp: subset_singleton_iff elim: finite_subset)
lemma Times_mset_assoc: "(A \<times># B) \<times># C = image_mset (\<lambda>(a, b, c). ((a, b), c)) (A \<times># B \<times># C)"
by (auto simp: multiset_eq_iff count_Sigma_mset count_image_mset vimage_def Times_mset_Times
Int_commute count_eq_zero_iff intro!: trans[OF _ sym[OF sum_le_singleton[of _ "(_, _, _)"]]]
cong: sum.cong if_cong)
subsection \<open>Transfer Rules\<close>
lemma plus_multiset_transfer[transfer_rule]:
"(rel_fun (rel_mset R) (rel_fun (rel_mset R) (rel_mset R))) (+) (+)"
by (unfold rel_fun_def rel_mset_def)
(force dest: list_all2_appendI intro: exI[of _ "_ @ _"] conjI[rotated])
lemma minus_multiset_transfer[transfer_rule]:
assumes [transfer_rule]: "bi_unique R"
shows "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (rel_mset R))) (-) (-)"
proof (unfold rel_fun_def rel_mset_def, safe)
fix xs ys xs' ys'
assume [transfer_rule]: "list_all2 R xs ys" "list_all2 R xs' ys'"
have "list_all2 R (fold remove1 xs' xs) (fold remove1 ys' ys)"
by transfer_prover
moreover have "mset (fold remove1 xs' xs) = mset xs - mset xs'"
by (induct xs' arbitrary: xs) auto
moreover have "mset (fold remove1 ys' ys) = mset ys - mset ys'"
by (induct ys' arbitrary: ys) auto
ultimately show "\<exists>xs'' ys''.
mset xs'' = mset xs - mset xs' \<and> mset ys'' = mset ys - mset ys' \<and> list_all2 R xs'' ys''"
by blast
qed
declare rel_mset_Zero[transfer_rule]
lemma count_transfer[transfer_rule]:
assumes "bi_unique R"
shows "(rel_fun (rel_mset R) (rel_fun R (=))) count count"
unfolding rel_fun_def rel_mset_def proof safe
fix x y xs ys
assume "list_all2 R xs ys" "R x y"
then show "count (mset xs) x = count (mset ys) y"
proof (induct xs ys rule: list.rel_induct)
case (Cons x' xs y' ys)
then show ?case
using assms unfolding bi_unique_alt_def2 by (auto simp: rel_fun_def)
qed simp
qed
lemma subseteq_multiset_transfer[transfer_rule]:
assumes [transfer_rule]: "bi_unique R" "right_total R"
shows "(rel_fun (rel_mset R) (rel_fun (rel_mset R) (=)))
(\<lambda>M N. filter_mset (Domainp R) M \<subseteq># filter_mset (Domainp R) N) (\<subseteq>#)"
proof -
have count_filter_mset_less:
"(\<forall>a. count (filter_mset (Domainp R) M) a \<le> count (filter_mset (Domainp R) N) a) \<longleftrightarrow>
(\<forall>a \<in> {x. Domainp R x}. count M a \<le> count N a)" for M and N by auto
show ?thesis unfolding subseteq_mset_def count_filter_mset_less
by transfer_prover
qed
lemma sum_mset_transfer[transfer_rule]:
"R 0 0 \<Longrightarrow> rel_fun R (rel_fun R R) (+) (+) \<Longrightarrow> (rel_fun (rel_mset R) R) sum_mset sum_mset"
using sum_list_transfer[of R] unfolding rel_fun_def rel_mset_def by auto
lemma Sigma_mset_transfer[transfer_rule]:
"(rel_fun (rel_mset R) (rel_fun (rel_fun R (rel_mset S)) (rel_mset (rel_prod R S))))
Sigma_mset Sigma_mset"
by (unfold Sigma_mset_def) transfer_prover
subsection \<open>Even More about Multisets\<close>
subsubsection \<open>Multisets and Functions\<close>
lemma range_image_mset:
assumes "set_mset Ds \<subseteq> range f"
shows "Ds \<in> range (image_mset f)"
proof -
have "\<forall>D. D \<in># Ds \<longrightarrow> (\<exists>C. f C = D)"
using assms by blast
then obtain f_i where
f_p: "\<forall>D. D \<in># Ds \<longrightarrow> (f (f_i D) = D)"
by metis
define Cs where
"Cs \<equiv> image_mset f_i Ds"
from f_p Cs_def have "image_mset f Cs = Ds"
by auto
then show ?thesis
by blast
qed
subsubsection \<open>Multisets and Lists\<close>
lemma length_sorted_list_of_multiset[simp]: "length (sorted_list_of_multiset A) = size A"
by (metis mset_sorted_list_of_multiset size_mset)
definition list_of_mset :: "'a multiset \<Rightarrow> 'a list" where
"list_of_mset m = (SOME l. m = mset l)"
lemma list_of_mset_exi: "\<exists>l. m = mset l"
using ex_mset by metis
lemma mset_list_of_mset[simp]: "mset (list_of_mset m) = m"
by (metis (mono_tags, lifting) ex_mset list_of_mset_def someI_ex)
lemma length_list_of_mset[simp]: "length (list_of_mset A) = size A"
unfolding list_of_mset_def by (metis (mono_tags) ex_mset size_mset someI_ex)
lemma range_mset_map:
assumes "set_mset Ds \<subseteq> range f"
shows "Ds \<in> range (\<lambda>Cl. mset (map f Cl))"
proof -
have "Ds \<in> range (image_mset f)"
by (simp add: assms range_image_mset)
then obtain Cs where Cs_p: "image_mset f Cs = Ds"
by auto
define Cl where "Cl = list_of_mset Cs"
then have "mset Cl = Cs"
by auto
then have "image_mset f (mset Cl) = Ds"
using Cs_p by auto
then have "mset (map f Cl) = Ds"
by auto
then show ?thesis
by auto
qed
lemma list_of_mset_empty[iff]: "list_of_mset m = [] \<longleftrightarrow> m = {#}"
by (metis (mono_tags, lifting) ex_mset list_of_mset_def mset_zero_iff_right someI_ex)
lemma in_mset_conv_nth: "(x \<in># mset xs) = (\<exists>i<length xs. xs ! i = x)"
by (auto simp: in_set_conv_nth)
lemma in_mset_sum_list:
assumes "L \<in># LL"
assumes "LL \<in> set Ci"
shows "L \<in># sum_list Ci"
using assms by (induction Ci) auto
lemma in_mset_sum_list2:
assumes "L \<in># sum_list Ci"
obtains LL where
"LL \<in> set Ci"
"L \<in># LL"
using assms by (induction Ci) auto
(* TODO: Make [simp]. *)
lemma in_mset_sum_list_iff: "a \<in># sum_list \<A> \<longleftrightarrow> (\<exists>A \<in> set \<A>. a \<in># A)"
by (metis in_mset_sum_list in_mset_sum_list2)
lemma subseteq_list_Union_mset:
assumes "length Ci = n"
assumes "length CAi = n"
assumes "\<forall>i<n. Ci ! i \<subseteq># CAi ! i "
shows "\<Sum>\<^sub># (mset Ci) \<subseteq># \<Sum>\<^sub># (mset CAi)"
using assms proof (induction n arbitrary: Ci CAi)
case 0
then show ?case by auto
next
case (Suc n)
from Suc have "\<forall>i<n. tl Ci ! i \<subseteq># tl CAi ! i"
by (simp add: nth_tl)
hence "\<Sum>\<^sub>#(mset (tl Ci)) \<subseteq># \<Sum>\<^sub>#(mset (tl CAi))" using Suc by auto
moreover
have "hd Ci \<subseteq># hd CAi" using Suc
by (metis hd_conv_nth length_greater_0_conv zero_less_Suc)
ultimately
show "\<Sum>\<^sub>#(mset Ci) \<subseteq># \<Sum>\<^sub>#(mset CAi)"
using Suc by (cases Ci; cases CAi) (auto intro: subset_mset.add_mono)
qed
subsubsection \<open>More on Multisets and Functions\<close>
lemma subseteq_mset_size_eql: "X \<subseteq># Y \<Longrightarrow> size Y = size X \<Longrightarrow> X = Y"
using mset_subset_size subset_mset_def by fastforce
lemma image_mset_of_subset_list:
assumes "image_mset \<eta> C' = mset lC"
shows "\<exists>qC'. map \<eta> qC' = lC \<and> mset qC' = C'"
using assms apply (induction lC arbitrary: C')
subgoal by simp
subgoal by (fastforce dest!: msed_map_invR intro: exI[of _ \<open>_ # _\<close>])
done
lemma image_mset_of_subset:
assumes "A \<subseteq># image_mset \<eta> C'"
shows "\<exists>A'. image_mset \<eta> A' = A \<and> A' \<subseteq># C'"
proof -
define C where "C = image_mset \<eta> C'"
define lA where "lA = list_of_mset A"
define lD where "lD = list_of_mset (C-A)"
define lC where "lC = lA @ lD"
have "mset lC = C"
using C_def assms unfolding lD_def lC_def lA_def by auto
then have "\<exists>qC'. map \<eta> qC' = lC \<and> mset qC' = C'"
using assms image_mset_of_subset_list unfolding C_def by metis
then obtain qC' where qC'_p: "map \<eta> qC' = lC \<and> mset qC' = C'"
by auto
let ?lA' = "take (length lA) qC'"
have m: "map \<eta> ?lA' = lA"
using qC'_p lC_def
by (metis append_eq_conv_conj take_map)
let ?A' = "mset ?lA'"
have "image_mset \<eta> ?A' = A"
using m using lA_def
by (metis (full_types) ex_mset list_of_mset_def mset_map someI_ex)
moreover have "?A' \<subseteq># C'"
using qC'_p unfolding lA_def
using mset_take_subseteq by blast
ultimately show ?thesis by blast
qed
lemma all_the_same: "\<forall>x \<in># X. x = y \<Longrightarrow> card (set_mset X) \<le> Suc 0"
by (metis card.empty card.insert card_mono finite.intros(1) finite_insert le_SucI singletonI subsetI)
lemma Melem_subseteq_Union_mset[simp]:
assumes "x \<in># T"
shows "x \<subseteq># \<Sum>\<^sub>#T"
using assms sum_mset.remove by force
lemma Melem_subset_eq_sum_list[simp]:
assumes "x \<in># mset T"
shows "x \<subseteq># sum_list T"
using assms by (metis mset_subset_eq_add_left sum_mset.remove sum_mset_sum_list)
lemma less_subset_eq_Union_mset[simp]:
assumes "i < length CAi"
shows "CAi ! i \<subseteq># \<Sum>\<^sub>#(mset CAi)"
proof -
from assms have "CAi ! i \<in># mset CAi"
by auto
then show ?thesis
by auto
qed
lemma less_subset_eq_sum_list[simp]:
assumes "i < length CAi"
shows "CAi ! i \<subseteq># sum_list CAi"
proof -
from assms have "CAi ! i \<in># mset CAi"
by auto
then show ?thesis
by auto
qed
subsubsection \<open>More on Multiset Order\<close>
lemma less_multiset_doubletons:
assumes
"y < t \<or> y < s"
"x < t \<or> x < s"
shows
"{#y, x#} < {#t, s#}"
unfolding less_multiset\<^sub>D\<^sub>M
proof (intro exI)
let ?X = "{#t, s#}"
let ?Y = "{#y, x#}"
show "?X \<noteq> {#} \<and> ?X \<subseteq># {#t, s#} \<and> {#y, x#} = {#t, s#} - ?X + ?Y
\<and> (\<forall>k. k \<in># ?Y \<longrightarrow> (\<exists>a. a \<in># ?X \<and> k < a))"
using add_eq_conv_diff assms by auto
qed
end
diff --git a/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy b/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy
--- a/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy
+++ b/thys/Nested_Multisets_Ordinals/Nested_Multiset.thy
@@ -1,408 +1,408 @@
(* Title: Nested Multisets
Author: Dmitriy Traytel <traytel at inf.ethz.ch>, 2016
Author: Jasmin Blanchette <jasmin.blanchette at inria.fr>, 2016
Maintainer: Dmitriy Traytel <traytel at inf.ethz.ch>
*)
section \<open>Nested Multisets\<close>
theory Nested_Multiset
imports "HOL-Library.Multiset_Order"
begin
declare multiset.map_comp [simp]
declare multiset.map_cong [cong]
subsection \<open>Type Definition\<close>
datatype 'a nmultiset =
Elem 'a
| MSet "'a nmultiset multiset"
inductive no_elem :: "'a nmultiset \<Rightarrow> bool" where
"(\<And>X. X \<in># M \<Longrightarrow> no_elem X) \<Longrightarrow> no_elem (MSet M)"
inductive_set sub_nmset :: "('a nmultiset \<times> 'a nmultiset) set" where
"X \<in># M \<Longrightarrow> (X, MSet M) \<in> sub_nmset"
lemma wf_sub_nmset[simp]: "wf sub_nmset"
proof (rule wfUNIVI)
fix P :: "'a nmultiset \<Rightarrow> bool" and M :: "'a nmultiset"
assume IH: "\<forall>M. (\<forall>N. (N, M) \<in> sub_nmset \<longrightarrow> P N) \<longrightarrow> P M"
show "P M"
by (induct M; rule IH[rule_format]) (auto simp: sub_nmset.simps)
qed
primrec depth_nmset :: "'a nmultiset \<Rightarrow> nat" ("|_|") where
"|Elem a| = 0"
| "|MSet M| = (let X = set_mset (image_mset depth_nmset M) in if X = {} then 0 else Suc (Max X))"
lemma depth_nmset_MSet: "x \<in># M \<Longrightarrow> |x| < |MSet M|"
by (auto simp: less_Suc_eq_le)
declare depth_nmset.simps(2)[simp del]
subsection \<open>Dershowitz and Manna's Nested Multiset Order\<close>
text \<open>The Dershowitz--Manna extension:\<close>
definition less_multiset_ext\<^sub>D\<^sub>M :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
"less_multiset_ext\<^sub>D\<^sub>M R M N \<longleftrightarrow>
(\<exists>X Y. X \<noteq> {#} \<and> X \<subseteq># N \<and> M = (N - X) + Y \<and> (\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> R k a)))"
lemma less_multiset_ext\<^sub>D\<^sub>M_imp_mult:
assumes
N_A: "set_mset N \<subseteq> A" and M_A: "set_mset M \<subseteq> A" and less: "less_multiset_ext\<^sub>D\<^sub>M R M N"
shows "(M, N) \<in> mult {(x, y). x \<in> A \<and> y \<in> A \<and> R x y}"
proof -
from less obtain X Y where
"X \<noteq> {#}" and "X \<subseteq># N" and "M = N - X + Y" and "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> R k a)"
unfolding less_multiset_ext\<^sub>D\<^sub>M_def by blast
with N_A M_A have "(N - X + Y, N - X + X) \<in> mult {(x, y). x \<in> A \<and> y \<in> A \<and> R x y}"
by (intro one_step_implies_mult, blast,
metis (mono_tags, lifting) case_prodI mem_Collect_eq mset_subset_eqD mset_subset_eq_add_right
subsetCE)
with \<open>M = N - X + Y\<close> \<open>X \<subseteq># N\<close> show "(M, N) \<in> mult {(x, y). x \<in> A \<and> y \<in> A \<and> R x y}"
by (simp add: subset_mset.diff_add)
qed
lemma mult_imp_less_multiset_ext\<^sub>D\<^sub>M:
assumes
N_A: "set_mset N \<subseteq> A" and M_A: "set_mset M \<subseteq> A" and
trans: "\<forall>x \<in> A. \<forall>y \<in> A. \<forall>z \<in> A. R x y \<longrightarrow> R y z \<longrightarrow> R x z" and
in_mult: "(M, N) \<in> mult {(x, y). x \<in> A \<and> y \<in> A \<and> R x y}"
shows "less_multiset_ext\<^sub>D\<^sub>M R M N"
using in_mult N_A M_A unfolding mult_def less_multiset_ext\<^sub>D\<^sub>M_def
proof induct
case (base N)
then obtain y M0 X where "N = add_mset y M0" and "M = M0 + X" and "\<forall>a. a \<in># X \<longrightarrow> R a y"
unfolding mult1_def by auto
thus ?case
by (auto intro: exI[of _ "{#y#}"])
next
case (step N N')
note N_N'_in_mult1 = this(2) and ih = this(3) and N'_A = this(4) and M_A = this(5)
have N_A: "set_mset N \<subseteq> A"
using N_N'_in_mult1 N'_A unfolding mult1_def by auto
obtain Y X where y_nemp: "Y \<noteq> {#}" and y_sub_N: "Y \<subseteq># N" and M_eq: "M = N - Y + X" and
ex_y: "\<forall>x. x \<in># X \<longrightarrow> (\<exists>y. y \<in># Y \<and> R x y)"
using ih[OF N_A M_A] by blast
obtain z M0 Ya where N'_eq: "N' = M0 + {#z#}" and N_eq: "N = M0 + Ya" and
z_gt: "\<forall>y. y \<in># Ya \<longrightarrow> y \<in> A \<and> z \<in> A \<and> R y z"
using N_N'_in_mult1[unfolded mult1_def] by auto
let ?Za = "Y - Ya + {#z#}"
let ?Xa = "X + Ya + (Y - Ya) - Y"
have xa_sub_x_ya: "set_mset ?Xa \<subseteq> set_mset (X + Ya)"
by (metis diff_subset_eq_self in_diffD subsetI subset_mset.diff_diff_right)
have x_A: "set_mset X \<subseteq> A"
using M_A M_eq by auto
have ya_A: "set_mset Ya \<subseteq> A"
by (simp add: subsetI z_gt)
have ex_y': "\<exists>y. y \<in># Y - Ya + {#z#} \<and> R x y" if x_in: "x \<in># X + Ya" for x
proof (cases "x \<in># X")
case True
then obtain y where y_in: "y \<in># Y" and y_gt_x: "R x y"
using ex_y by blast
show ?thesis
proof (cases "y \<in># Ya")
case False
hence "y \<in># Y - Ya + {#z#}"
using y_in count_greater_zero_iff in_diff_count by fastforce
thus ?thesis
using y_gt_x by blast
next
case True
hence "y \<in> A" and "z \<in> A" and "R y z"
using z_gt by blast+
hence "R x z"
using trans y_gt_x x_A ya_A x_in by (meson subsetCE union_iff)
thus ?thesis
by auto
qed
next
case False
hence "x \<in># Ya"
using x_in by auto
hence "x \<in> A" and "z \<in> A" and "R x z"
using z_gt by blast+
thus ?thesis
by auto
qed
show ?case
proof (rule exI[of _ ?Za], rule exI[of _ ?Xa], intro conjI)
show "Y - Ya + {#z#} \<subseteq># N'"
using mset_subset_eq_mono_add subset_eq_diff_conv y_sub_N N_eq N'_eq
by (simp add: subset_eq_diff_conv)
next
show "M = N' - (Y - Ya + {#z#}) + (X + Ya + (Y - Ya) - Y)"
unfolding M_eq N_eq N'_eq by (auto simp: multiset_eq_iff)
next
show "\<forall>x. x \<in># X + Ya + (Y - Ya) - Y \<longrightarrow> (\<exists>y. y \<in># Y - Ya + {#z#} \<and> R x y)"
using ex_y' xa_sub_x_ya by blast
qed auto
qed
lemma less_multiset_ext\<^sub>D\<^sub>M_iff_mult:
assumes
N_A: "set_mset N \<subseteq> A" and M_A: "set_mset M \<subseteq> A" and
trans: "\<forall>x \<in> A. \<forall>y \<in> A. \<forall>z \<in> A. R x y \<longrightarrow> R y z \<longrightarrow> R x z"
shows "less_multiset_ext\<^sub>D\<^sub>M R M N \<longleftrightarrow> (M, N) \<in> mult {(x, y). x \<in> A \<and> y \<in> A \<and> R x y}"
using mult_imp_less_multiset_ext\<^sub>D\<^sub>M[OF assms] less_multiset_ext\<^sub>D\<^sub>M_imp_mult[OF N_A M_A] by blast
instantiation nmultiset :: (preorder) preorder
begin
lemma less_multiset_ext\<^sub>D\<^sub>M_cong[fundef_cong]:
"(\<And>X Y k a. X \<noteq> {#} \<Longrightarrow> X \<subseteq># N \<Longrightarrow> M = (N - X) + Y \<Longrightarrow> k \<in># Y \<Longrightarrow> R k a = S k a) \<Longrightarrow>
less_multiset_ext\<^sub>D\<^sub>M R M N = less_multiset_ext\<^sub>D\<^sub>M S M N"
unfolding less_multiset_ext\<^sub>D\<^sub>M_def by metis
function less_nmultiset :: "'a nmultiset \<Rightarrow> 'a nmultiset \<Rightarrow> bool" where
"less_nmultiset (Elem a) (Elem b) \<longleftrightarrow> a < b"
| "less_nmultiset (Elem a) (MSet M) \<longleftrightarrow> True"
| "less_nmultiset (MSet M) (Elem a) \<longleftrightarrow> False"
| "less_nmultiset (MSet M) (MSet N) \<longleftrightarrow> less_multiset_ext\<^sub>D\<^sub>M less_nmultiset M N"
by pat_completeness auto
termination
by (relation "sub_nmset <*lex*> sub_nmset", fastforce,
metis sub_nmset.simps in_lex_prod mset_subset_eqD mset_subset_eq_add_right)
lemmas less_nmultiset_induct =
less_nmultiset.induct[case_names Elem_Elem Elem_MSet MSet_Elem MSet_MSet]
lemmas less_nmultiset_cases =
less_nmultiset.cases[case_names Elem_Elem Elem_MSet MSet_Elem MSet_MSet]
lemma trans_less_nmultiset: "X < Y \<Longrightarrow> Y < Z \<Longrightarrow> X < Z" for X Y Z :: "'a nmultiset"
proof (induct "Max {|X|, |Y|, |Z|}" arbitrary: X Y Z
rule: less_induct)
case less
from less(2,3) show ?case
proof (cases X; cases Y; cases Z)
fix M N N' :: "'a nmultiset multiset"
define A where "A = set_mset M \<union> set_mset N \<union> set_mset N'"
assume XYZ: "X = MSet M" "Y = MSet N" "Z = MSet N'"
then have trans: "\<forall>x \<in> A. \<forall>y \<in> A. \<forall>z \<in> A. x < y \<longrightarrow> y < z \<longrightarrow> x < z"
by (auto elim!: less(1)[rotated -1] dest!: depth_nmset_MSet simp add: A_def)
have "set_mset M \<subseteq> A" "set_mset N \<subseteq> A" "set_mset N' \<subseteq> A"
unfolding A_def by auto
with less(2,3) XYZ show "X < Z"
by (auto simp: less_multiset_ext\<^sub>D\<^sub>M_iff_mult[OF _ _ trans] mult_def)
qed (auto elim: less_trans)
qed
lemma irrefl_less_nmultiset:
fixes X :: "'a nmultiset"
shows "X < X \<Longrightarrow> False"
proof (induct X)
case (MSet M)
from MSet(2) show ?case
unfolding less_nmultiset.simps less_multiset_ext\<^sub>D\<^sub>M_def
proof safe
fix X Y :: "'a nmultiset multiset"
define XY where "XY = {(x, y). x \<in># X \<and> y \<in># Y \<and> y < x}"
then have fin: "finite XY" and trans: "trans XY"
by (auto simp: trans_def intro: trans_less_nmultiset
finite_subset[OF _ finite_cartesian_product])
assume "X \<noteq> {#}" "X \<subseteq># M" "M = M - X + Y"
then have "X = Y"
by (auto simp: mset_subset_eq_exists_conv)
with MSet(1) \<open>X \<subseteq># M\<close> have "irrefl XY"
unfolding XY_def by (force dest: mset_subset_eqD simp: irrefl_def)
with trans have "acyclic XY"
by (simp add: acyclic_irrefl)
moreover
assume "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)"
with \<open>X = Y\<close> \<open>X \<noteq> {#}\<close> have "\<not> acyclic XY"
by (intro notI, elim finite_acyclic_wf[OF fin, elim_format])
(auto dest!: spec[of _ "set_mset Y"] simp: wf_eq_minimal XY_def)
ultimately show False by blast
qed
qed simp
lemma antisym_less_nmultiset:
fixes X Y :: "'a nmultiset"
shows "X < Y \<Longrightarrow> Y < X \<Longrightarrow> False"
using trans_less_nmultiset irrefl_less_nmultiset by blast
definition less_eq_nmultiset :: "'a nmultiset \<Rightarrow> 'a nmultiset \<Rightarrow> bool" where
"less_eq_nmultiset X Y = (X < Y \<or> X = Y)"
instance
proof (intro_classes, goal_cases less_def refl trans)
case (less_def x y)
then show ?case
unfolding less_eq_nmultiset_def by (metis irrefl_less_nmultiset antisym_less_nmultiset)
next
case (refl x)
then show ?case
unfolding less_eq_nmultiset_def by blast
next
case (trans x y z)
then show ?case
unfolding less_eq_nmultiset_def by (metis trans_less_nmultiset)
qed
lemma less_multiset_ext\<^sub>D\<^sub>M_less: "less_multiset_ext\<^sub>D\<^sub>M (<) = (<)"
unfolding fun_eq_iff less_multiset_ext\<^sub>D\<^sub>M_def less_multiset\<^sub>D\<^sub>M by blast
end
instantiation nmultiset :: (order) order
begin
instance
proof (intro_classes, goal_cases antisym)
case (antisym x y)
then show ?case
unfolding less_eq_nmultiset_def by (metis trans_less_nmultiset irrefl_less_nmultiset)
qed
end
instantiation nmultiset :: (linorder) linorder
begin
lemma total_less_nmultiset:
fixes X Y :: "'a nmultiset"
shows "\<not> X < Y \<Longrightarrow> Y \<noteq> X \<Longrightarrow> Y < X"
proof (induct X Y rule: less_nmultiset_induct)
case (MSet_MSet M N)
then show ?case
unfolding nmultiset.inject less_nmultiset.simps less_multiset_ext\<^sub>D\<^sub>M_less less_multiset\<^sub>H\<^sub>O
by (metis add_diff_cancel_left' count_inI diff_add_zero in_diff_count less_imp_not_less
- mset_subset_eq_multiset_union_diff_commute subset_mset.order.refl)
+ mset_subset_eq_multiset_union_diff_commute subset_mset.refl)
qed auto
instance
proof (intro_classes, goal_cases total)
case (total x y)
then show ?case
unfolding less_eq_nmultiset_def by (metis total_less_nmultiset)
qed
end
lemma less_depth_nmset_imp_less_nmultiset: "|X| < |Y| \<Longrightarrow> X < Y"
proof (induct X Y rule: less_nmultiset_induct)
case (MSet_MSet M N)
then show ?case
proof (cases "M = {#}")
case False
with MSet_MSet show ?thesis
by (auto 0 4 simp: depth_nmset.simps(2) less_multiset_ext\<^sub>D\<^sub>M_def not_le Max_gr_iff
intro: exI[of _ N] split: if_splits)
qed (auto simp: depth_nmset.simps(2) less_multiset_ext\<^sub>D\<^sub>M_less split: if_splits)
qed simp_all
lemma less_nmultiset_imp_le_depth_nmset: "X < Y \<Longrightarrow> |X| \<le> |Y|"
proof (induct X Y rule: less_nmultiset_induct)
case (MSet_MSet M N)
then have "M < N" by (simp add: less_multiset_ext\<^sub>D\<^sub>M_less)
then show ?case
proof (cases "M = {#}" "N = {#}" rule: bool.exhaust[case_product bool.exhaust])
case [simp]: False_False
show ?thesis
unfolding depth_nmset.simps(2) Let_def False_False Suc_le_mono set_image_mset image_is_empty
set_mset_eq_empty_iff if_False
proof (intro iffD2[OF Max_le_iff] ballI iffD2[OF Max_ge_iff]; (elim imageE)?; simp)
fix X
assume [simp]: "X \<in># M"
with MSet_MSet(1)[of N M X, simplified] \<open>M < N\<close> show "\<exists>Y\<in>#N. |X| \<le> |Y|"
by (meson ex_gt_imp_less_multiset less_asym' less_depth_nmset_imp_less_nmultiset
not_le_imp_less)
qed
qed (auto simp: depth_nmset.simps(2))
qed simp_all
lemma eq_mlex_I:
fixes f :: "'a \<Rightarrow> nat" and R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
assumes "\<And>X Y. f X < f Y \<Longrightarrow> R X Y" and "antisymp R"
shows "{(X, Y). R X Y} = f <*mlex*> {(X, Y). f X = f Y \<and> R X Y}"
proof safe
fix X Y
assume "R X Y"
show "(X, Y) \<in> f <*mlex*> {(X, Y). f X = f Y \<and> R X Y}"
proof (cases "f X" "f Y" rule: linorder_cases)
case less
with \<open>R X Y\<close> show ?thesis
by (elim mlex_less)
next
case equal
with \<open>R X Y\<close> show ?thesis
by (intro mlex_leq) auto
next
case greater
from \<open>R X Y\<close> assms(1)[OF greater] \<open>antisymp R\<close> greater show ?thesis
unfolding antisymp_def by auto
qed
next
fix X Y
assume "(X, Y) \<in> f <*mlex*> {(X, Y). f X = f Y \<and> R X Y}"
then show "R X Y"
unfolding mlex_prod_def by (auto simp: assms(1))
qed
instantiation nmultiset :: (wellorder) wellorder
begin
lemma depth_nmset_eq_0[simp]: "|X| = 0 \<longleftrightarrow> (X = MSet {#} \<or> (\<exists>x. X = Elem x))"
by (cases X; simp add: depth_nmset.simps(2))
lemma depth_nmset_eq_Suc[simp]: "|X| = Suc n \<longleftrightarrow>
(\<exists>N. X = MSet N \<and> (\<exists>Y \<in># N. |Y| = n) \<and> (\<forall>Y \<in># N. |Y| \<le> n))"
by (cases X; auto simp add: depth_nmset.simps(2) intro!: Max_eqI)
(metis (no_types, lifting) Max_in finite_imageI finite_set_mset imageE image_is_empty
set_mset_eq_empty_iff)
lemma wf_less_nmultiset_depth:
"wf {(X :: 'a nmultiset, Y). |X| = i \<and> |Y| = i \<and> X < Y}"
proof (induct i rule: less_induct)
case (less i)
define A :: "'a nmultiset set" where "A = {X. |X| < i}"
from less have "wf ((depth_nmset :: 'a nmultiset \<Rightarrow> nat) <*mlex*>
(\<Union>j < i. {(X, Y). |X| = j \<and> |Y| = j \<and> X < Y}))"
by (intro wf_UN wf_mlex) auto
then have *: "wf (mult {(X :: 'a nmultiset, Y). X \<in> A \<and> Y \<in> A \<and> X < Y})"
by (intro wf_mult, elim wf_subset) (force simp: A_def mlex_prod_def not_less_iff_gr_or_eq
dest!: less_depth_nmset_imp_less_nmultiset)
show ?case
proof (cases i)
case 0
then show ?thesis
by (auto simp: inj_on_def intro!: wf_subset[OF
wf_Un[OF wf_map_prod_image[OF wf, of Elem] wf_UN[of "Elem ` UNIV" "\<lambda>x. {(x, MSet {#})}"]]])
next
case (Suc n)
then show ?thesis
by (intro wf_subset[OF wf_map_prod_image[OF *, of MSet]])
(auto 0 4 simp: map_prod_def image_iff inj_on_def A_def
dest!: less_multiset_ext\<^sub>D\<^sub>M_imp_mult[of _ A, rotated -1] split: prod.splits)
qed
qed
lemma wf_less_nmultiset: "wf {(X :: 'a nmultiset, Y :: 'a nmultiset). X < Y}" (is "wf ?R")
proof -
have "?R = depth_nmset <*mlex*> {(X, Y). |X| = |Y| \<and> X < Y}"
by (rule eq_mlex_I) (auto simp: antisymp_def less_depth_nmset_imp_less_nmultiset)
also have "{(X, Y). |X| = |Y| \<and> X < Y} = (\<Union>i. {(X, Y). |X| = i \<and> |Y| = i \<and> X < Y})"
by auto
finally show ?thesis
by (fastforce intro: wf_mlex wf_Union wf_less_nmultiset_depth)
qed
instance using wf_less_nmultiset unfolding wf_def mem_Collect_eq prod.case by intro_classes metis
end
end
diff --git a/thys/Order_Lattice_Props/Closure_Operators.thy b/thys/Order_Lattice_Props/Closure_Operators.thy
--- a/thys/Order_Lattice_Props/Closure_Operators.thy
+++ b/thys/Order_Lattice_Props/Closure_Operators.thy
@@ -1,514 +1,514 @@
(*
Title: Closure and Co-Closure Operators
Author: Georg Struth
Maintainer: Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Closure and Co-Closure Operators\<close>
theory Closure_Operators
imports Galois_Connections
begin
subsection \<open>Closure Operators\<close>
text \<open>Closure and coclosure operators in orders and complete lattices are defined in this section,
and some basic properties are proved. Isabelle infers the appropriate types. Facts are
taken mainly from the Compendium of Continuous Lattices~\cite{GierzHKLMS80} and
Rosenthal's book on quantales~\cite{Rosenthal90}.\<close>
definition clop :: "('a::order \<Rightarrow> 'a) \<Rightarrow> bool" where
"clop f = (id \<le> f \<and> mono f \<and> f \<circ> f \<le> f)"
lemma clop_extensive: "clop f \<Longrightarrow> id \<le> f"
by (simp add: clop_def)
lemma clop_extensive_var: "clop f \<Longrightarrow> x \<le> f x"
by (simp add: clop_def le_fun_def)
lemma clop_iso: "clop f \<Longrightarrow> mono f"
by (simp add: clop_def)
lemma clop_iso_var: "clop f \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
by (simp add: clop_def mono_def)
lemma clop_idem: "clop f \<Longrightarrow> f \<circ> f = f"
by (simp add: antisym clop_def le_fun_def)
lemma clop_Fix_range: "clop f \<Longrightarrow> (Fix f = range f)"
by (simp add: clop_idem retraction_prop_fix)
lemma clop_idem_var: "clop f \<Longrightarrow> f (f x) = f x"
by (simp add: clop_idem retraction_prop)
lemma clop_Inf_closed_var:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "clop f \<Longrightarrow> f \<circ> Inf \<circ> (`) f = Inf \<circ> (`) f"
unfolding clop_def mono_def comp_def le_fun_def
by (metis (mono_tags, lifting) antisym id_apply le_INF_iff order_refl)
lemma clop_top:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "clop f \<Longrightarrow> f \<top> = \<top>"
by (simp add: clop_extensive_var top.extremum_uniqueI)
lemma "clop (f::'a::complete_lattice \<Rightarrow> 'a) \<Longrightarrow> f (\<Squnion>x \<in> X. f x) = (\<Squnion>x \<in> X. f x)" (*nitpick*)
oops
lemma "clop (f::'a::complete_lattice \<Rightarrow> 'a) \<Longrightarrow> f (f x \<squnion> f y) = f x \<squnion> f y" (*nitpick*)
oops
lemma "clop (f::'a::complete_lattice \<Rightarrow> 'a) \<Longrightarrow> f \<bottom> = \<bottom>" (*nitpick *)
oops
lemma "clop (f::'a set \<Rightarrow> 'a set) \<Longrightarrow> f (\<Squnion>x \<in> X. f x) = (\<Squnion>x \<in> X. f x)" (*nitpick*)
oops
lemma "clop (f::'a set \<Rightarrow> 'a set) \<Longrightarrow> f (f x \<squnion> f y) = f x \<squnion> f y" (*nitpick*)
oops
lemma "clop (f::'a set \<Rightarrow> 'a set) \<Longrightarrow> f \<bottom> = \<bottom>" (*nitpick *)
oops
lemma clop_closure: "clop f \<Longrightarrow> (x \<in> range f) = (f x = x)"
by (simp add: clop_idem retraction_prop)
lemma clop_closure_set: "clop f \<Longrightarrow> range f = Fix f"
by (simp add: clop_Fix_range)
lemma clop_closure_prop: "(clop::('a::complete_lattice_with_dual\<Rightarrow> 'a) \<Rightarrow> bool) (Inf \<circ> \<up>)"
by (simp add: clop_def mono_def)
lemma clop_closure_prop_var: "clop (\<lambda>x::'a::complete_lattice. \<Sqinter>{y. x \<le> y})"
unfolding clop_def comp_def le_fun_def mono_def by (simp add: Inf_lower le_Inf_iff)
lemma clop_alt: "(clop f) = (\<forall>x y. x \<le> f y \<longleftrightarrow> f x \<le> f y)"
unfolding clop_def mono_def le_fun_def comp_def id_def by (meson dual_order.refl order_trans)
text \<open>Finally it is shown that adjoints in a Galois connection yield closure operators.\<close>
lemma clop_adj:
fixes f :: "'a::order \<Rightarrow> 'b::order"
shows "f \<stileturn> g \<Longrightarrow> clop (g \<circ> f)"
by (simp add: adj_cancel2 adj_idem2 adj_iso4 clop_def)
text \<open>Closure operators are monads for posets, and monads arise from adjunctions.
This fact is not formalised at this point. But here is the first step: every function
can be decomposed into a surjection followed by an injection.\<close>
definition "surj_on f Y = (\<forall>y \<in> Y. \<exists>x. y = f x)"
lemma surj_surj_on: "surj f \<Longrightarrow> surj_on f Y"
by (simp add: surjD surj_on_def)
lemma fun_surj_inj: "\<exists>g h. f = g \<circ> h \<and> surj_on h (range f) \<and> inj_on g (range f)"
proof-
obtain h where a: "\<forall>x. f x = h x"
by blast
then have "surj_on h (range f)"
by (metis (mono_tags, lifting) imageE surj_on_def)
then show ?thesis
unfolding inj_on_def surj_on_def fun_eq_iff using a by auto
qed
text \<open>Connections between downsets, upsets and closure operators are outlined next.\<close>
lemma preorder_clop: "clop (\<Down>::'a::preorder set \<Rightarrow> 'a set)"
by (simp add: clop_def downset_set_ext downset_set_iso)
lemma clop_preorder_aux: "clop f \<Longrightarrow> (x \<in> f {y} \<longleftrightarrow> f {x} \<subseteq> f {y})"
by (simp add: clop_alt)
lemma clop_preorder: "clop f \<Longrightarrow> class.preorder (\<lambda>x y. f {x} \<subseteq> f {y}) (\<lambda>x y. f {x} \<subset> f {y})"
unfolding clop_def mono_def le_fun_def id_def comp_def by standard (auto simp: subset_not_subset_eq)
lemma preorder_clop_dual: "clop (\<Up>::'a::preorder_with_dual set \<Rightarrow> 'a set)"
by (simp add: clop_def upset_set_anti upset_set_ext)
text \<open>The closed elements of any closure operator over a complete lattice form an Inf-closed set (a Moore family).\<close>
lemma clop_Inf_closed:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "clop f \<Longrightarrow> Inf_closed_set (Fix f)"
unfolding clop_def Inf_closed_set_def mono_def le_fun_def comp_def id_def Fix_def
by (smt Inf_greatest Inf_lower antisym mem_Collect_eq subsetCE)
lemma clop_top_Fix:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "clop f \<Longrightarrow> \<top> \<in> Fix f"
by (simp add: clop_Fix_range clop_closure clop_top)
text \<open>Conversely, every Inf-closed subset of a complete lattice is the set of fixpoints of some closure operator.\<close>
lemma Inf_closed_clop:
fixes X :: "'a::complete_lattice set"
shows "Inf_closed_set X \<Longrightarrow> clop (\<lambda>y. \<Sqinter>{x \<in> X. y \<le> x})"
by (smt Collect_mono_iff Inf_superset_mono clop_alt dual_order.trans le_Inf_iff mem_Collect_eq)
lemma Inf_closed_clop_var:
fixes X :: "'a::complete_lattice set"
shows "clop f \<Longrightarrow> \<forall>x \<in> X. x \<in> range f \<Longrightarrow> \<Sqinter>X \<in> range f"
by (metis Inf_closed_set_def clop_Fix_range clop_Inf_closed subsetI)
text \<open>It is well known that downsets and upsets over an ordering form subalgebras of the complete powerset lattice.\<close>
typedef (overloaded) 'a downsets = "range (\<Down>::'a::order set \<Rightarrow> 'a set)"
by fastforce
setup_lifting type_definition_downsets
typedef (overloaded) 'a upsets = "range (\<Up>::'a::order set \<Rightarrow> 'a set)"
by fastforce
setup_lifting type_definition_upsets
instantiation downsets :: (order) Inf_lattice
begin
lift_definition Inf_downsets :: "'a downsets set \<Rightarrow> 'a downsets" is "Abs_downsets \<circ> Inf \<circ> (`) Rep_downsets".
lift_definition less_eq_downsets :: "'a downsets \<Rightarrow> 'a downsets \<Rightarrow> bool" is "\<lambda>X Y. Rep_downsets X \<subseteq> Rep_downsets Y".
lift_definition less_downsets :: "'a downsets \<Rightarrow> 'a downsets \<Rightarrow> bool" is "\<lambda>X Y. Rep_downsets X \<subset> Rep_downsets Y".
instance
apply intro_classes
apply (transfer, simp)
apply (transfer, blast)
apply (simp add: Closure_Operators.less_eq_downsets.abs_eq Rep_downsets_inject)
apply (transfer, smt Abs_downsets_inverse INF_lower Inf_closed_clop_var Rep_downsets image_iff o_def preorder_clop)
by transfer (smt comp_def Abs_downsets_inverse Inf_closed_clop_var Rep_downsets image_iff le_INF_iff preorder_clop)
end
instantiation upsets :: (order_with_dual) Inf_lattice
begin
lift_definition Inf_upsets :: "'a upsets set \<Rightarrow> 'a upsets" is "Abs_upsets \<circ> Inf \<circ> (`) Rep_upsets".
lift_definition less_eq_upsets :: "'a upsets \<Rightarrow> 'a upsets \<Rightarrow> bool" is "\<lambda>X Y. Rep_upsets X \<subseteq> Rep_upsets Y".
lift_definition less_upsets :: "'a upsets \<Rightarrow> 'a upsets \<Rightarrow> bool" is "\<lambda>X Y. Rep_upsets X \<subset> Rep_upsets Y".
instance
apply intro_classes
apply (transfer, simp)
apply (transfer, blast)
apply (simp add: Closure_Operators.less_eq_upsets.abs_eq Rep_upsets_inject)
apply (transfer, smt Abs_upsets_inverse Inf_closed_clop_var Inf_lower Rep_upsets comp_apply image_iff preorder_clop_dual)
by transfer (smt comp_def Abs_upsets_inverse Inf_closed_clop_var Inter_iff Rep_upsets image_iff preorder_clop_dual subsetCE subsetI)
end
text \<open>It has already been shown in the section on representations that the map ds, which maps elements of the order to its downset, is an order
embedding. However, the duality between the underlying ordering and the lattices of up- and down-closed sets as categories can probably not be expressed,
as there is no easy access to contravariant functors. \<close>
subsection \<open>Co-Closure Operators\<close>
text \<open>Next, the co-closure (or kernel) operation satisfies dual laws.\<close>
definition coclop :: "('a::order \<Rightarrow> 'a::order) \<Rightarrow> bool" where
"coclop f = (f \<le> id \<and> mono f \<and> f \<le> f \<circ> f)"
lemma coclop_dual: "(coclop::('a::order_with_dual \<Rightarrow> 'a) \<Rightarrow> bool) = clop \<circ> \<partial>\<^sub>F"
unfolding coclop_def clop_def id_def mono_def map_dual_def comp_def fun_eq_iff le_fun_def
by (metis invol_dual_var ord_dual)
lemma coclop_dual_var:
fixes f :: "'a::order_with_dual \<Rightarrow> 'a"
shows "coclop f = clop (\<partial>\<^sub>F f)"
by (simp add: coclop_dual)
lemma clop_dual: "(clop::('a::order_with_dual \<Rightarrow> 'a) \<Rightarrow> bool) = coclop \<circ> \<partial>\<^sub>F"
by (simp add: coclop_dual comp_assoc map_dual_invol)
lemma clop_dual_var:
fixes f :: "'a::order_with_dual \<Rightarrow> 'a"
shows "clop f = coclop (\<partial>\<^sub>F f)"
by (simp add: clop_dual)
lemma coclop_coextensive: "coclop f \<Longrightarrow> f \<le> id"
by (simp add: coclop_def)
lemma coclop_coextensive_var: "coclop f \<Longrightarrow> f x \<le> x"
using coclop_def le_funD by fastforce
lemma coclop_iso: "coclop f \<Longrightarrow> mono f"
by (simp add: coclop_def)
lemma coclop_iso_var: "coclop f \<Longrightarrow> (x \<le> y \<longrightarrow> f x \<le> f y)"
by (simp add: coclop_iso monoD)
lemma coclop_idem: "coclop f \<Longrightarrow> f \<circ> f = f"
by (simp add: antisym coclop_def le_fun_def)
lemma coclop_closure: "coclop f \<Longrightarrow> (x \<in> range f) = (f x = x)"
by (simp add: coclop_idem retraction_prop)
lemma coclop_Fix_range: "coclop f \<Longrightarrow> (Fix f = range f)"
by (simp add: coclop_idem retraction_prop_fix)
lemma coclop_idem_var: "coclop f \<Longrightarrow> f (f x) = f x"
by (simp add: coclop_idem retraction_prop)
lemma coclop_Sup_closed_var:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'a"
shows "coclop f \<Longrightarrow> f \<circ> Sup \<circ> (`) f = Sup \<circ> (`) f"
unfolding coclop_def mono_def comp_def le_fun_def
by (metis (mono_tags, lifting) SUP_le_iff antisym id_apply order_refl)
lemma Sup_closed_coclop_var:
fixes X :: "'a::complete_lattice set"
shows "coclop f \<Longrightarrow> \<forall>x \<in> X. x \<in> range f \<Longrightarrow> \<Squnion>X \<in> range f"
by (smt Inf.INF_id_eq Sup.SUP_cong antisym coclop_closure coclop_coextensive_var coclop_iso id_apply mono_SUP)
lemma coclop_bot:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'a"
shows "coclop f \<Longrightarrow> f \<bottom> = \<bottom>"
by (simp add: bot.extremum_uniqueI coclop_coextensive_var)
lemma "coclop (f::'a::complete_lattice \<Rightarrow> 'a) \<Longrightarrow> f (\<Sqinter>x \<in> X. f x) = (\<Sqinter>x \<in> X. f x)" (*nitpick*)
oops
lemma "coclop (f::'a::complete_lattice \<Rightarrow> 'a) \<Longrightarrow> f (f x \<sqinter> f y) = f x \<sqinter> f y" (*nitpick*)
oops
lemma "coclop (f::'a::complete_lattice \<Rightarrow> 'a) \<Longrightarrow> f \<top> = \<top>" (*nitpick*)
oops
lemma "coclop (f::'a set \<Rightarrow> 'a set) \<Longrightarrow> f (\<Sqinter>x \<in> X. f x) = (\<Sqinter>x \<in> X. f x)" (*nitpick*)
oops
lemma "coclop (f::'a set \<Rightarrow> 'a set) \<Longrightarrow> f (f x \<sqinter> f y) = f x \<sqinter> f y" (*nitpick*)
oops
lemma "coclop (f::'a set \<Rightarrow> 'a set) \<Longrightarrow> f \<top> = \<top>" (*nitpick *)
oops
lemma coclop_coclosure: "coclop f \<Longrightarrow> f x = x \<longleftrightarrow> x \<in> range f"
by (simp add: coclop_idem retraction_prop)
lemma coclop_coclosure_set: "coclop f \<Longrightarrow> range f = Fix f"
by (simp add: coclop_idem retraction_prop_fix)
lemma coclop_coclosure_prop: "(coclop::('a::complete_lattice \<Rightarrow> 'a) \<Rightarrow> bool) (Sup \<circ> \<down>)"
by (simp add: coclop_def mono_def)
lemma coclop_coclosure_prop_var: "coclop (\<lambda>x::'a::complete_lattice. \<Squnion>{y. y \<le> x})"
by (metis (mono_tags, lifting) Sup_atMost atMost_def coclop_def comp_apply eq_id_iff eq_refl mono_def)
lemma coclop_alt: "(coclop f) = (\<forall>x y. f x \<le> y \<longleftrightarrow> f x \<le> f y)"
unfolding coclop_def mono_def le_fun_def comp_def id_def
by (meson dual_order.refl order_trans)
lemma coclop_adj:
fixes f :: "'a::order \<Rightarrow> 'b::order"
shows "f \<stileturn> g \<Longrightarrow> coclop (f \<circ> g)"
by (simp add: adj_cancel1 adj_idem1 adj_iso3 coclop_def)
text \<open>Finally, a subset of a complete lattice is Sup-closed if and only if it is the set of fixpoints
of some co-closure operator.\<close>
lemma coclop_Sup_closed:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "coclop f \<Longrightarrow> Sup_closed_set (Fix f)"
unfolding coclop_def Sup_closed_set_def mono_def le_fun_def comp_def id_def Fix_def
by (smt Sup_least Sup_upper antisym_conv mem_Collect_eq subsetCE)
lemma Sup_closed_coclop:
fixes X :: "'a::complete_lattice set"
shows "Sup_closed_set X \<Longrightarrow> coclop (\<lambda>y. \<Squnion>{x \<in> X. x \<le> y})"
unfolding Sup_closed_set_def coclop_def mono_def le_fun_def comp_def
apply safe
apply (metis (no_types, lifting) Sup_least eq_id_iff mem_Collect_eq)
apply (smt Collect_mono_iff Sup_subset_mono dual_order.trans)
by (simp add: Collect_mono_iff Sup_subset_mono Sup_upper)
subsection \<open>Complete Lattices of Closed Elements\<close>
text \<open>The machinery developed allows showing that the closed elements in a complete
lattice (with respect to some closure operation) form themselves a complete lattice.\<close>
class cl_op = ord +
fixes cl_op :: "'a \<Rightarrow> 'a"
assumes clop_ext: "x \<le> cl_op x"
and clop_iso: "x \<le> y \<Longrightarrow> cl_op x \<le> cl_op y"
and clop_wtrans: "cl_op (cl_op x) \<le> cl_op x"
class clattice_with_clop = complete_lattice + cl_op
begin
lemma clop_cl_op: "clop cl_op"
unfolding clop_def le_fun_def comp_def
by (simp add: cl_op_class.clop_ext cl_op_class.clop_iso cl_op_class.clop_wtrans order_class.mono_def)
lemma clop_idem [simp]: "cl_op \<circ> cl_op = cl_op"
using clop_ext clop_wtrans order.antisym by auto
lemma clop_idem_var [simp]: "cl_op (cl_op x) = cl_op x"
- by (simp add: antisym clop_ext clop_wtrans)
+ by (simp add: order.antisym clop_ext clop_wtrans)
lemma clop_range_Fix: "range cl_op = Fix cl_op"
by (simp add: retraction_prop_fix)
lemma Inf_closed_cl_op_var:
fixes X :: "'a set"
shows "\<forall>x \<in> X. x \<in> range cl_op \<Longrightarrow> \<Sqinter>X \<in> range cl_op"
proof-
assume h: "\<forall>x \<in> X. x \<in> range cl_op"
hence "\<forall>x \<in> X. cl_op x = x"
by (simp add: retraction_prop)
hence "cl_op (\<Sqinter>X) = \<Sqinter>X"
by (metis Inf_lower clop_ext clop_iso dual_order.antisym le_Inf_iff)
thus ?thesis
by (metis rangeI)
qed
lemma inf_closed_cl_op_var: "x \<in> range cl_op \<Longrightarrow> y \<in> range cl_op \<Longrightarrow> x \<sqinter> y \<in> range cl_op"
by (smt Inf_closed_cl_op_var UnI1 insert_iff insert_is_Un inf_Inf)
end
typedef (overloaded) 'a::clattice_with_clop cl_op_im = "range (cl_op::'a \<Rightarrow> 'a)"
by force
setup_lifting type_definition_cl_op_im
lemma cl_op_prop [iff]: "(cl_op (x \<squnion> y) = cl_op y) = (cl_op (x::'a::clattice_with_clop) \<le> cl_op y)"
by (smt cl_op_class.clop_iso clop_ext clop_wtrans inf_sup_ord(4) le_iff_sup sup.absorb_iff1 sup_left_commute)
lemma cl_op_prop_var [iff]: "(cl_op (x \<squnion> cl_op y) = cl_op y) = (cl_op (x::'a::clattice_with_clop) \<le> cl_op y)"
by (metis cl_op_prop clattice_with_clop_class.clop_idem_var)
instantiation cl_op_im :: (clattice_with_clop) complete_lattice
begin
lift_definition Inf_cl_op_im :: "'a cl_op_im set \<Rightarrow> 'a cl_op_im" is Inf
by (simp add: Inf_closed_cl_op_var)
lift_definition Sup_cl_op_im :: "'a cl_op_im set \<Rightarrow> 'a cl_op_im" is "\<lambda>X. cl_op (\<Squnion>X)"
by simp
lift_definition inf_cl_op_im :: "'a cl_op_im \<Rightarrow> 'a cl_op_im \<Rightarrow> 'a cl_op_im" is inf
by (simp add: inf_closed_cl_op_var)
lift_definition sup_cl_op_im :: "'a cl_op_im \<Rightarrow> 'a cl_op_im \<Rightarrow> 'a cl_op_im" is "\<lambda>x y. cl_op (x \<squnion> y)"
by simp
lift_definition less_eq_cl_op_im :: "'a cl_op_im \<Rightarrow> 'a cl_op_im \<Rightarrow> bool" is "(\<le>)".
lift_definition less_cl_op_im :: "'a cl_op_im \<Rightarrow> 'a cl_op_im \<Rightarrow> bool" is "(<)".
lift_definition bot_cl_op_im :: "'a cl_op_im" is "cl_op \<bottom>"
by simp
lift_definition top_cl_op_im :: "'a cl_op_im" is "\<top>"
by (simp add: clop_cl_op clop_closure clop_top)
instance
apply (intro_classes; transfer)
apply (simp_all add: less_le_not_le Inf_lower Inf_greatest)
apply (meson clop_cl_op clop_extensive_var dual_order.trans inf_sup_ord(3))
apply (meson clop_cl_op clop_extensive_var dual_order.trans sup_ge2)
apply (metis cl_op_class.clop_iso clop_cl_op clop_closure le_sup_iff)
apply (meson Sup_upper clop_cl_op clop_extensive_var dual_order.trans)
by (metis Sup_le_iff cl_op_class.clop_iso clop_cl_op clop_closure)
end
text \<open>This statement is perhaps less useful as it might seem, because it is difficult to make it cooperate with concrete closure operators,
which one would not generally like to define within a type class. Alternatively, a sublocale statement could perhaps be given. It would also
have been nice to prove this statement for Sup-lattices---this would have cut down the number of proof obligations significantly.
But this would require a tighter integration of these structures. A similar statement could have been proved for co-closure operators. But this would
not lead to new insights.\<close>
text \<open>Next I show that for every surjective Sup-preserving function between complete lattices there is a closure operator
such that the set of closed elements is isomorphic to the range of the surjection.\<close>
lemma surj_Sup_pres_id:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual"
assumes "surj f"
and "Sup_pres f"
shows "f \<circ> (radj f) = id"
proof-
have "f \<stileturn> (radj f)"
using Sup_pres_ladj assms(2) radj_adj by auto
thus ?thesis
using adj_sur_inv assms(1) by blast
qed
lemma surj_Sup_pres_inj:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual"
assumes "surj f"
and "Sup_pres f"
shows "inj (radj f)"
by (metis assms comp_eq_dest_lhs id_apply injI surj_Sup_pres_id)
lemma surj_Sup_pres_inj_on:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual"
assumes "surj f"
and "Sup_pres f"
shows "inj_on f (range (radj f \<circ> f))"
by (smt Sup_pres_ladj_aux adj_idem2 assms(2) comp_apply inj_on_def retraction_prop)
lemma surj_Sup_pres_bij_on:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual"
assumes "surj f"
and "Sup_pres f"
shows "bij_betw f (range (radj f \<circ> f)) UNIV"
unfolding bij_betw_def
apply safe
apply (simp add: assms(1) assms(2) surj_Sup_pres_inj_on cong del: image_cong_simp)
apply auto
apply (metis (mono_tags) UNIV_I assms(1) assms(2) comp_apply id_apply image_image surj_Sup_pres_id surj_def)
done
text \<open>Thus the restriction of $f$ to the set of closed elements is indeed a bijection. The final fact
shows that it preserves Sups of closed elements, and hence is an isomorphism of complete lattices.\<close>
lemma surj_Sup_pres_iso:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual"
assumes "surj f"
and "Sup_pres f"
shows "f ((radj f \<circ> f) (\<Squnion>X)) = (\<Squnion>x \<in> X. f x)"
by (metis assms(1) assms(2) comp_def pointfree_idE surj_Sup_pres_id)
subsection \<open>A Quick Example: Dedekind-MacNeille Completions\<close>
text \<open>I only outline the basic construction. Additional facts about join density, and that the completion yields
the least complete lattice that contains all Sups and Infs of the underlying posets, are left for future consideration.\<close>
abbreviation "dm \<equiv> lb_set \<circ> ub_set"
lemma up_set_prop: "(X::'a::preorder set) \<noteq> {} \<Longrightarrow> ub_set X = \<Inter>{\<up>x |x. x \<in> X}"
unfolding ub_set_def upset_def upset_set_def by (safe, simp_all, blast)
lemma lb_set_prop: "(X::'a::preorder set) \<noteq> {} \<Longrightarrow> lb_set X = \<Inter>{\<down>x |x. x \<in> X}"
unfolding lb_set_def downset_def downset_set_def by (safe, simp_all, blast)
lemma dm_downset_var: "dm {x} = \<down>(x::'a::preorder)"
unfolding lb_set_def ub_set_def downset_def downset_set_def
by (clarsimp, meson order_refl order_trans)
lemma dm_downset: "dm \<circ> \<eta> = (\<down>::'a::preorder \<Rightarrow> 'a set)"
using dm_downset_var fun.map_cong by fastforce
lemma dm_inj: "inj ((dm::'a::order set \<Rightarrow> 'a set) \<circ> \<eta>)"
by (simp add: dm_downset downset_inj)
lemma "clop (lb_set \<circ> ub_set)"
unfolding clop_def lb_set_def ub_set_def
apply safe
unfolding le_fun_def comp_def id_def mono_def
by auto
end
diff --git a/thys/Order_Lattice_Props/Order_Lattice_Props.thy b/thys/Order_Lattice_Props/Order_Lattice_Props.thy
--- a/thys/Order_Lattice_Props/Order_Lattice_Props.thy
+++ b/thys/Order_Lattice_Props/Order_Lattice_Props.thy
@@ -1,1273 +1,1269 @@
(*
Title: Properties of Orderings and Lattices
Author: Georg Struth
Maintainer: Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Properties of Orderings and Lattices\<close>
theory Order_Lattice_Props
imports Order_Duality
begin
subsection \<open>Basic Definitions for Orderings and Lattices\<close>
text \<open>The first definition is for order morphisms --- isotone (order-preserving, monotone) functions.
An order isomorphism is an order-preserving bijection. This should be defined in the class ord, but mono requires order.\<close>
definition ord_homset :: "('a::order \<Rightarrow> 'b::order) set" where
"ord_homset = {f::'a::order \<Rightarrow> 'b::order. mono f}"
definition ord_embed :: "('a::order \<Rightarrow> 'b::order) \<Rightarrow> bool" where
"ord_embed f = (\<forall>x y. f x \<le> f y \<longleftrightarrow> x \<le> y)"
definition ord_iso :: "('a::order \<Rightarrow> 'b::order) \<Rightarrow> bool" where
"ord_iso = bij \<sqinter> mono \<sqinter> (mono \<circ> the_inv)"
lemma ord_embed_alt: "ord_embed f = (mono f \<and> (\<forall>x y. f x \<le> f y \<longrightarrow> x \<le> y))"
using mono_def ord_embed_def by auto
lemma ord_embed_homset: "ord_embed f \<Longrightarrow> f \<in> ord_homset"
by (simp add: mono_def ord_embed_def ord_homset_def)
lemma ord_embed_inj: "ord_embed f \<Longrightarrow> inj f"
unfolding ord_embed_def inj_def by (simp add: eq_iff)
lemma ord_iso_ord_embed: "ord_iso f \<Longrightarrow> ord_embed f"
unfolding ord_iso_def ord_embed_def bij_def inj_def mono_def
by (clarsimp, metis inj_def the_inv_f_f)
lemma ord_iso_alt: "ord_iso f = (ord_embed f \<and> surj f)"
unfolding ord_iso_def ord_embed_def surj_def bij_def inj_def mono_def
apply safe
by simp_all (metis eq_iff inj_def the_inv_f_f)+
lemma ord_iso_the_inv: "ord_iso f \<Longrightarrow> mono (the_inv f)"
by (simp add: ord_iso_def)
lemma ord_iso_inv1: "ord_iso f \<Longrightarrow> (the_inv f) \<circ> f = id"
using ord_embed_inj ord_iso_ord_embed the_inv_into_f_f by fastforce
lemma ord_iso_inv2: "ord_iso f \<Longrightarrow> f \<circ> (the_inv f) = id"
using f_the_inv_into_f ord_embed_inj ord_iso_alt by fastforce
typedef (overloaded) ('a,'b) ord_homset = "ord_homset::('a::order \<Rightarrow> 'b::order) set"
by (force simp: ord_homset_def mono_def)
setup_lifting type_definition_ord_homset
text \<open>The next definition is for the set of fixpoints of a given function. It is important in the context of orders,
for instance for proving Tarski's fixpoint theorem, but does not really belong here.\<close>
definition Fix :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set" where
"Fix f = {x. f x = x}"
lemma retraction_prop: "f \<circ> f = f \<Longrightarrow> f x = x \<longleftrightarrow> x \<in> range f"
by (metis comp_apply f_inv_into_f rangeI)
lemma retraction_prop_fix: "f \<circ> f = f \<Longrightarrow> range f = Fix f"
unfolding Fix_def using retraction_prop by fastforce
lemma Fix_map_dual: "Fix \<circ> \<partial>\<^sub>F = (`) \<partial> \<circ> Fix"
unfolding Fix_def map_dual_def comp_def fun_eq_iff
by (smt Collect_cong invol_dual pointfree_idE setcompr_eq_image)
lemma Fix_map_dual_var: "Fix (\<partial>\<^sub>F f) = \<partial> ` (Fix f)"
by (metis Fix_map_dual o_def)
lemma gfp_dual: "(\<partial>::'a::complete_lattice_with_dual \<Rightarrow> 'a) \<circ> gfp = lfp \<circ> \<partial>\<^sub>F"
proof-
{fix f:: "'a \<Rightarrow> 'a"
have "\<partial> (gfp f) = \<partial> (\<Squnion>{u. u \<le> f u})"
by (simp add: gfp_def)
also have "... = \<Sqinter>(\<partial> ` {u. u \<le> f u})"
by (simp add: Sup_dual_def_var)
also have "... = \<Sqinter>{\<partial> u |u. u \<le> f u}"
by (simp add: setcompr_eq_image)
also have "... = \<Sqinter>{u |u. (\<partial>\<^sub>F f) u \<le> u}"
by (metis (no_types, hide_lams) dual_dual_ord dual_iff map_dual_def o_def)
finally have "\<partial> (gfp f) = lfp (\<partial>\<^sub>F f)"
by (metis lfp_def)}
thus ?thesis
by auto
qed
lemma gfp_dual_var:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'a"
shows "\<partial> (gfp f) = lfp (\<partial>\<^sub>F f)"
using comp_eq_elim gfp_dual by blast
lemma gfp_to_lfp: "gfp = (\<partial>::'a::complete_lattice_with_dual \<Rightarrow> 'a) \<circ> lfp \<circ> \<partial>\<^sub>F"
by (simp add: comp_assoc fun_dual2 gfp_dual)
lemma gfp_to_lfp_var:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'a"
shows "gfp f = \<partial> (lfp (\<partial>\<^sub>F f))"
by (metis gfp_dual_var invol_dual_var)
lemma lfp_dual: "(\<partial>::'a::complete_lattice_with_dual \<Rightarrow> 'a) \<circ> lfp = gfp \<circ> \<partial>\<^sub>F"
by (simp add: comp_assoc gfp_to_lfp map_dual_invol)
lemma lfp_dual_var:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'a"
shows "\<partial> (lfp f) = gfp (map_dual f)"
using comp_eq_dest_lhs lfp_dual by fastforce
lemma lfp_to_gfp: "lfp = (\<partial>::'a::complete_lattice_with_dual \<Rightarrow> 'a) \<circ> gfp \<circ> \<partial>\<^sub>F"
by (simp add: comp_assoc gfp_dual map_dual_invol)
lemma lfp_to_gfp_var:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'a"
shows "lfp f = \<partial> (gfp (\<partial>\<^sub>F f))"
by (metis invol_dual_var lfp_dual_var)
lemma lfp_in_Fix:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "mono f \<Longrightarrow> lfp f \<in> Fix f"
by (metis (mono_tags, lifting) Fix_def lfp_unfold mem_Collect_eq)
lemma gfp_in_Fix:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "mono f \<Longrightarrow> gfp f \<in> Fix f"
by (metis (mono_tags, lifting) Fix_def gfp_unfold mem_Collect_eq)
lemma nonempty_Fix:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "mono f \<Longrightarrow> Fix f \<noteq> {}"
using lfp_in_Fix by fastforce
text \<open>Next the minimal and maximal elements of an ordering are defined.\<close>
context ord
begin
definition min_set :: "'a set \<Rightarrow> 'a set" where
"min_set X = {y \<in> X. \<forall>x \<in> X. x \<le> y \<longrightarrow> x = y}"
definition max_set :: "'a set \<Rightarrow> 'a set" where
"max_set X = {x \<in> X. \<forall>y \<in> X. x \<le> y \<longrightarrow> x = y}"
end
context ord_with_dual
begin
lemma min_max_set_dual: "(`) \<partial> \<circ> min_set = max_set \<circ> (`) \<partial>"
unfolding max_set_def min_set_def fun_eq_iff comp_def
apply safe
using dual_dual_ord inj_dual_iff by auto
lemma min_max_set_dual_var: "\<partial> ` (min_set X) = max_set (\<partial> ` X)"
using comp_eq_dest min_max_set_dual by fastforce
lemma max_min_set_dual: "(`) \<partial> \<circ> max_set = min_set \<circ> (`) \<partial>"
by (metis (no_types, hide_lams) comp_id fun.map_comp id_comp image_dual min_max_set_dual)
lemma min_to_max_set: "min_set = (`) \<partial> \<circ> max_set \<circ> (`) \<partial>"
by (metis comp_id image_dual max_min_set_dual o_assoc)
lemma max_min_set_dual_var: "\<partial> ` (max_set X) = min_set (\<partial> ` X)"
using comp_eq_dest max_min_set_dual by fastforce
lemma min_to_max_set_var: "min_set X = \<partial> ` (max_set (\<partial> ` X))"
by (simp add: max_min_set_dual_var pointfree_idE)
end
text \<open>Next, directed and filtered sets, upsets, downsets, filters and ideals in posets are defined.\<close>
context ord
begin
definition directed :: "'a set \<Rightarrow> bool" where
"directed X = (\<forall>Y. finite Y \<and> Y \<subseteq> X \<longrightarrow> (\<exists>x \<in> X. \<forall>y \<in> Y. y \<le> x))"
definition filtered :: "'a set \<Rightarrow> bool" where
"filtered X = (\<forall>Y. finite Y \<and> Y \<subseteq> X \<longrightarrow> (\<exists>x \<in> X. \<forall>y \<in> Y. x \<le> y))"
definition downset_set :: "'a set \<Rightarrow> 'a set" ("\<Down>") where
"\<Down>X = {y. \<exists>x \<in> X. y \<le> x}"
definition upset_set :: "'a set \<Rightarrow> 'a set" ("\<Up>") where
"\<Up>X = {y. \<exists>x \<in> X. x \<le> y}"
definition downset :: "'a \<Rightarrow> 'a set" ("\<down>") where
"\<down> = \<Down> \<circ> \<eta>"
definition upset :: "'a \<Rightarrow> 'a set" ("\<up>") where
"\<up> = \<Up> \<circ> \<eta>"
definition downsets :: "'a set set" where
"downsets = Fix \<Down>"
definition upsets :: "'a set set" where
"upsets = Fix \<Up>"
definition "downclosed_set X = (X \<in> downsets)"
definition "upclosed_set X = (X \<in> upsets)"
definition ideals :: "'a set set" where
"ideals = {X. X \<noteq> {} \<and> downclosed_set X \<and> directed X}"
definition filters :: "'a set set" where
"filters = {X. X \<noteq> {} \<and> upclosed_set X \<and> filtered X}"
abbreviation "idealp X \<equiv> X \<in> ideals"
abbreviation "filterp X \<equiv> X \<in> filters"
end
text \<open>These notions are pair-wise dual.\<close>
text \<open>Filtered and directed sets are dual.\<close>
context ord_with_dual
begin
lemma filtered_directed_dual: "filtered \<circ> (`) \<partial> = directed"
unfolding filtered_def directed_def fun_eq_iff comp_def
apply clarsimp
apply safe
apply (meson finite_imageI imageI image_mono dual_dual_ord)
by (smt finite_subset_image imageE ord_dual)
lemma directed_filtered_dual: "directed \<circ> (`) \<partial> = filtered"
using filtered_directed_dual by (metis comp_id image_dual o_assoc)
lemma filtered_to_directed: "filtered X = directed (\<partial> ` X)"
by (metis comp_apply directed_filtered_dual)
text \<open>Upsets and downsets are dual.\<close>
lemma downset_set_upset_set_dual: "(`) \<partial> \<circ> \<Down> = \<Up> \<circ> (`) \<partial>"
unfolding downset_set_def upset_set_def fun_eq_iff comp_def
apply safe
apply (meson image_eqI ord_dual)
by (clarsimp, metis (mono_tags, lifting) dual_iff image_iff mem_Collect_eq ord_dual)
lemma upset_set_downset_set_dual: "(`) \<partial> \<circ> \<Up> = \<Down> \<circ> (`) \<partial>"
using downset_set_upset_set_dual by (metis (no_types, hide_lams) comp_id id_comp image_dual o_assoc)
lemma upset_set_to_downset_set: "\<Up> = (`) \<partial> \<circ> \<Down> \<circ> (`) \<partial>"
by (simp add: comp_assoc downset_set_upset_set_dual)
lemma upset_set_to_downset_set2: "\<Up> X = \<partial> ` (\<Down> (\<partial> ` X))"
by (simp add: upset_set_to_downset_set)
lemma downset_upset_dual: "(`) \<partial> \<circ> \<down> = \<up> \<circ> \<partial>"
using downset_def upset_def upset_set_to_downset_set by fastforce
lemma upset_to_downset: "(`) \<partial> \<circ> \<up> = \<down> \<circ> \<partial>"
by (metis comp_assoc id_apply ord.downset_def ord.upset_def power_set_func_nat_trans upset_set_downset_set_dual)
lemma upset_to_downset2: "\<up> = (`) \<partial> \<circ> \<down> \<circ> \<partial>"
by (simp add: comp_assoc downset_upset_dual)
lemma upset_to_downset3: "\<up> x = \<partial> ` (\<down> (\<partial> x))"
by (simp add: upset_to_downset2)
lemma downsets_upsets_dual: "(X \<in> downsets) = (\<partial> ` X \<in> upsets)"
unfolding downsets_def upsets_def Fix_def
by (smt comp_eq_dest downset_set_upset_set_dual image_inv_f_f inj_dual mem_Collect_eq)
lemma downset_setp_upset_setp_dual: "upclosed_set \<circ> (`) \<partial> = downclosed_set"
unfolding downclosed_set_def upclosed_set_def using downsets_upsets_dual by fastforce
lemma upsets_to_downsets: "(X \<in> upsets) = (\<partial> ` X \<in> downsets)"
by (simp add: downsets_upsets_dual image_comp)
lemma upset_setp_downset_setp_dual: "downclosed_set \<circ> (`) \<partial> = upclosed_set"
by (metis comp_id downset_setp_upset_setp_dual image_dual o_assoc)
text \<open>Filters and ideals are dual.\<close>
lemma ideals_filters_dual: "(X \<in> ideals) = ((\<partial> ` X) \<in> filters)"
by (smt comp_eq_dest_lhs directed_filtered_dual image_inv_f_f image_is_empty inv_unique_comp filters_def ideals_def inj_dual invol_dual mem_Collect_eq upset_setp_downset_setp_dual)
lemma idealp_filterp_dual: "idealp = filterp \<circ> (`) \<partial>"
unfolding fun_eq_iff by (simp add: ideals_filters_dual)
lemma filters_to_ideals: "(X \<in> filters) = ((\<partial> ` X) \<in> ideals)"
by (simp add: ideals_filters_dual image_comp)
lemma filterp_idealp_dual: "filterp = idealp \<circ> (`) \<partial>"
unfolding fun_eq_iff by (simp add: filters_to_ideals)
end
subsection \<open>Properties of Orderings\<close>
context ord
begin
lemma directed_nonempty: "directed X \<Longrightarrow> X \<noteq> {}"
unfolding directed_def by fastforce
lemma directed_ub: "directed X \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z)"
by (meson empty_subsetI directed_def finite.emptyI finite_insert insert_subset order_refl)
lemma downset_set_prop: "\<Down> = Union \<circ> (`) \<down>"
unfolding downset_set_def downset_def fun_eq_iff by fastforce
lemma downset_set_prop_var: "\<Down>X = (\<Union>x \<in> X. \<down>x)"
by (simp add: downset_set_prop)
lemma downset_prop: "\<down>x = {y. y \<le> x}"
unfolding downset_def downset_set_def fun_eq_iff by fastforce
lemma downset_prop2: "y \<le> x \<Longrightarrow> y \<in> \<down>x"
by (simp add: downset_prop)
lemma ideals_downsets: "X \<in> ideals \<Longrightarrow> X \<in> downsets"
by (simp add: downclosed_set_def ideals_def)
lemma ideals_directed: "X \<in> ideals \<Longrightarrow> directed X"
by (simp add: ideals_def)
end
context preorder
begin
lemma directed_prop: "X \<noteq> {} \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z) \<Longrightarrow> directed X"
proof-
assume h1: "X \<noteq> {}"
and h2: "\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z"
{fix Y
have "finite Y \<Longrightarrow> Y \<subseteq> X \<Longrightarrow> (\<exists>x \<in> X. \<forall>y \<in> Y. y \<le> x)"
proof (induct rule: finite_induct)
case empty
then show ?case
using h1 by blast
next
case (insert x F)
then show ?case
by (metis h2 insert_iff insert_subset order_trans)
qed}
thus ?thesis
by (simp add: directed_def)
qed
lemma directed_alt: "directed X = (X \<noteq> {} \<and> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z))"
by (metis directed_prop directed_nonempty directed_ub)
lemma downset_set_prop_var2: "x \<in> \<Down>X \<Longrightarrow> y \<le> x \<Longrightarrow> y \<in> \<Down>X"
unfolding downset_set_def using order_trans by blast
lemma downclosed_set_iff: "downclosed_set X = (\<forall>x \<in> X. \<forall>y. y \<le> x \<longrightarrow> y \<in> X)"
unfolding downclosed_set_def downsets_def Fix_def downset_set_def by auto
lemma downclosed_downset_set: "downclosed_set (\<Down>X)"
by (simp add: downclosed_set_iff downset_set_prop_var2 downset_def)
lemma downclosed_downset: "downclosed_set (\<down>x)"
by (simp add: downclosed_downset_set downset_def)
lemma downset_set_ext: "id \<le> \<Down>"
unfolding le_fun_def id_def downset_set_def by auto
lemma downset_set_iso: "mono \<Down>"
unfolding mono_def downset_set_def by blast
lemma downset_set_idem [simp]: "\<Down> \<circ> \<Down> = \<Down>"
unfolding fun_eq_iff downset_set_def using order_trans by auto
lemma downset_faithful: "\<down>x \<subseteq> \<down>y \<Longrightarrow> x \<le> y"
by (simp add: downset_prop subset_eq)
lemma downset_iso_iff: "(\<down>x \<subseteq> \<down>y) = (x \<le> y)"
using atMost_iff downset_prop order_trans by blast
text \<open>The following proof uses the Axiom of Choice.\<close>
lemma downset_directed_downset_var [simp]: "directed (\<Down>X) = directed X"
proof
assume h1: "directed X"
{fix Y
assume h2: "finite Y" and h3: "Y \<subseteq> \<Down>X"
hence "\<forall>y. \<exists>x. y \<in> Y \<longrightarrow> x \<in> X \<and> y \<le> x"
by (force simp: downset_set_def)
hence "\<exists>f. \<forall>y. y \<in> Y \<longrightarrow> f y \<in> X \<and> y \<le> f y"
by (rule choice)
hence "\<exists>f. finite (f ` Y) \<and> f ` Y \<subseteq> X \<and> (\<forall>y \<in> Y. y \<le> f y)"
by (metis finite_imageI h2 image_subsetI)
hence "\<exists>Z. finite Z \<and> Z \<subseteq> X \<and> (\<forall>y \<in> Y. \<exists> z \<in> Z. y \<le> z)"
by fastforce
hence "\<exists>Z. finite Z \<and> Z \<subseteq> X \<and> (\<forall>y \<in> Y. \<exists> z \<in> Z. y \<le> z) \<and> (\<exists>x \<in> X. \<forall> z \<in> Z. z \<le> x)"
by (metis directed_def h1)
hence "\<exists>x \<in> X. \<forall>y \<in> Y. y \<le> x"
by (meson order_trans)}
thus "directed (\<Down>X)"
unfolding directed_def downset_set_def by fastforce
next
assume "directed (\<Down>X)"
thus "directed X"
unfolding directed_def downset_set_def
apply clarsimp
by (smt Ball_Collect order_refl order_trans subsetCE)
qed
lemma downset_directed_downset [simp]: "directed \<circ> \<Down> = directed"
unfolding fun_eq_iff by simp
lemma directed_downset_ideals: "directed (\<Down>X) = (\<Down>X \<in> ideals)"
by (metis (mono_tags, lifting) CollectI Fix_def directed_alt downset_set_idem downclosed_set_def downsets_def ideals_def o_def ord.ideals_directed)
lemma downclosed_Fix: "downclosed_set X = (\<Down>X = X)"
by (metis (mono_tags, lifting) CollectD Fix_def downclosed_downset_set downclosed_set_def downsets_def)
end
lemma downset_iso: "mono (\<down>::'a::order \<Rightarrow> 'a set)"
by (simp add: downset_iso_iff mono_def)
lemma mono_downclosed:
fixes f :: "'a::order \<Rightarrow> 'b::order"
assumes "mono f"
shows "\<forall>Y. downclosed_set Y \<longrightarrow> downclosed_set (f -` Y)"
by (simp add: assms downclosed_set_iff monoD)
lemma
fixes f :: "'a::order \<Rightarrow> 'b::order"
assumes "mono f"
shows "\<forall>Y. downclosed_set X \<longrightarrow> downclosed_set (f ` X)" (*nitpick*)
oops
lemma downclosed_mono:
fixes f :: "'a::order \<Rightarrow> 'b::order"
assumes "\<forall>Y. downclosed_set Y \<longrightarrow> downclosed_set (f -` Y)"
shows "mono f"
proof-
{fix x y :: "'a::order"
assume h: "x \<le> y"
have "downclosed_set (\<down> (f y))"
unfolding downclosed_set_def downsets_def Fix_def downset_set_def downset_def by auto
hence "downclosed_set (f -` (\<down> (f y)))"
by (simp add: assms)
hence "downclosed_set {z. f z \<le> f y}"
unfolding vimage_def downset_def downset_set_def by auto
hence "\<forall>z w. (f z \<le> f y \<and> w \<le> z) \<longrightarrow> f w \<le> f y"
unfolding downclosed_set_def downclosed_set_def downsets_def Fix_def downset_set_def by force
hence "f x \<le> f y"
using h by blast}
thus ?thesis..
qed
lemma mono_downclosed_iff: "mono f = (\<forall>Y. downclosed_set Y \<longrightarrow> downclosed_set (f -` Y))"
using mono_downclosed downclosed_mono by auto
context order
begin
lemma downset_inj: "inj \<down>"
- by (metis injI downset_iso_iff eq_iff)
+ by (metis injI downset_iso_iff order.eq_iff)
lemma "(X \<subseteq> Y) = (\<Down>X \<subseteq> \<Down>Y)" (*nitpick*)
oops
end
context lattice
begin
lemma lat_ideals: "X \<in> ideals = (X \<noteq> {} \<and> X \<in> downsets \<and> (\<forall>x \<in> X. \<forall> y \<in> X. x \<squnion> y \<in> X))"
unfolding ideals_def directed_alt downsets_def Fix_def downset_set_def downclosed_set_def
by (clarsimp, smt sup.cobounded1 sup.orderE sup.orderI sup_absorb2 sup_left_commute mem_Collect_eq)
end
context bounded_lattice
begin
lemma bot_ideal: "X \<in> ideals \<Longrightarrow> \<bottom> \<in> X"
unfolding ideals_def downclosed_set_def downsets_def Fix_def downset_set_def by fastforce
end
context complete_lattice
begin
lemma Sup_downset_id [simp]: "Sup \<circ> \<down> = id"
using Sup_atMost atMost_def downset_prop by fastforce
lemma downset_Sup_id: "id \<le> \<down> \<circ> Sup"
by (simp add: Sup_upper downset_prop le_funI subsetI)
lemma Inf_Sup_var: "\<Squnion>(\<Inter>x \<in> X. \<down>x) = \<Sqinter>X"
unfolding downset_prop by (simp add: Collect_ball_eq Inf_eq_Sup)
lemma Inf_pres_downset_var: "(\<Inter>x \<in> X. \<down>x) = \<down>(\<Sqinter>X)"
unfolding downset_prop by (safe, simp_all add: le_Inf_iff)
end
subsection \<open>Dual Properties of Orderings\<close>
context ord_with_dual
begin
lemma filtered_nonempty: "filtered X \<Longrightarrow> X \<noteq> {}"
using filtered_to_directed ord.directed_nonempty by auto
lemma filtered_lb: "filtered X \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. z \<le> x \<and> z \<le> y)"
using filtered_to_directed directed_ub dual_dual_ord by fastforce
lemma upset_set_prop_var: "\<Up>X = (\<Union>x \<in> X. \<up>x)"
by (simp add: image_Union downset_set_prop_var upset_set_to_downset_set2 upset_to_downset2)
lemma upset_set_prop: "\<Up> = Union \<circ> (`) \<up>"
unfolding fun_eq_iff by (simp add: upset_set_prop_var)
lemma upset_prop: "\<up>x = {y. x \<le> y}"
unfolding upset_to_downset3 downset_prop image_def using dual_dual_ord by fastforce
lemma upset_prop2: "x \<le> y \<Longrightarrow> y \<in> \<up>x"
by (simp add: upset_prop)
lemma filters_upsets: "X \<in> filters \<Longrightarrow> X \<in> upsets"
by (simp add: upclosed_set_def filters_def)
lemma filters_filtered: "X \<in> filters \<Longrightarrow> filtered X"
by (simp add: filters_def)
end
context preorder_with_dual
begin
lemma filtered_prop: "X \<noteq> {} \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. z \<le> x \<and> z \<le> y) \<Longrightarrow> filtered X"
unfolding filtered_to_directed
by (rule directed_prop, blast, metis (full_types) image_iff ord_dual)
lemma filtered_alt: "filtered X = (X \<noteq> {} \<and> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. z \<le> x \<and> z \<le> y))"
by (metis image_empty directed_alt filtered_to_directed filtered_lb filtered_prop)
lemma up_set_prop_var2: "x \<in> \<Up>X \<Longrightarrow> x \<le> y \<Longrightarrow> y \<in> \<Up>X"
using downset_set_prop_var2 dual_iff ord_dual upset_set_to_downset_set2 by fastforce
lemma upclosed_set_iff: "upclosed_set X = (\<forall>x \<in> X. \<forall>y. x \<le> y \<longrightarrow> y \<in> X)"
unfolding upclosed_set_def upsets_def Fix_def upset_set_def by auto
lemma upclosed_upset_set: "upclosed_set (\<Up>X)"
using up_set_prop_var2 upclosed_set_iff by blast
lemma upclosed_upset: "upclosed_set (\<up>x)"
by (simp add: upset_def upclosed_upset_set)
lemma upset_set_ext: "id \<le> \<Up>"
by (smt comp_def comp_id image_mono le_fun_def downset_set_ext image_dual upset_set_to_downset_set2)
lemma upset_set_anti: "mono \<Up>"
by (metis image_mono downset_set_iso upset_set_to_downset_set2 mono_def)
lemma up_set_idem [simp]: "\<Up> \<circ> \<Up> = \<Up>"
by (metis comp_assoc downset_set_idem upset_set_downset_set_dual upset_set_to_downset_set)
lemma upset_faithful: "\<up>x \<subseteq> \<up>y \<Longrightarrow> y \<le> x"
by (metis inj_image_subset_iff downset_faithful dual_dual_ord inj_dual upset_to_downset3)
lemma upset_anti_iff: "(\<up>y \<subseteq> \<up>x) = (x \<le> y)"
by (metis downset_iso_iff ord_dual upset_to_downset3 subset_image_iff upset_faithful)
lemma upset_filtered_upset [simp]: "filtered \<circ> \<Up> = filtered"
by (metis comp_assoc directed_filtered_dual downset_directed_downset upset_set_downset_set_dual)
lemma filtered_upset_filters: "filtered (\<Up>X) = (\<Up>X \<in> filters)"
by (metis comp_apply directed_downset_ideals filtered_to_directed filterp_idealp_dual upset_set_downset_set_dual)
lemma upclosed_Fix: "upclosed_set X = (\<Up>X = X)"
by (simp add: Fix_def upclosed_set_def upsets_def)
end
lemma upset_anti: "antimono (\<up>::'a::order_with_dual \<Rightarrow> 'a set)"
by (simp add: antimono_def upset_anti_iff)
lemma mono_upclosed:
fixes f :: "'a::order_with_dual \<Rightarrow> 'b::order_with_dual"
assumes "mono f"
shows "\<forall>Y. upclosed_set Y \<longrightarrow> upclosed_set (f -` Y)"
by (simp add: assms monoD upclosed_set_iff)
lemma mono_upclosed:
fixes f :: "'a::order_with_dual \<Rightarrow> 'b::order_with_dual"
assumes "mono f"
shows "\<forall>Y. upclosed_set X \<longrightarrow> upclosed_set (f ` X)" (*nitpick*)
oops
lemma upclosed_mono:
fixes f :: "'a::order_with_dual \<Rightarrow> 'b::order_with_dual"
assumes "\<forall>Y. upclosed_set Y \<longrightarrow> upclosed_set (f -` Y)"
shows "mono f"
by (metis (mono_tags, lifting) assms dual_order.refl mem_Collect_eq monoI order.trans upclosed_set_iff vimageE vimageI2)
lemma mono_upclosed_iff:
fixes f :: "'a::order_with_dual \<Rightarrow> 'b::order_with_dual"
shows "mono f = (\<forall>Y. upclosed_set Y \<longrightarrow> upclosed_set (f -` Y))"
using mono_upclosed upclosed_mono by auto
context order_with_dual
begin
lemma upset_inj: "inj \<up>"
by (metis inj_compose inj_on_imageI2 downset_inj inj_dual upset_to_downset)
lemma "(X \<subseteq> Y) = (\<Up>Y \<subseteq> \<Up>X)" (*nitpick*)
oops
end
context lattice_with_dual
begin
lemma lat_filters: "X \<in> filters = (X \<noteq> {} \<and> X \<in> upsets \<and> (\<forall>x \<in> X. \<forall> y \<in> X. x \<sqinter> y \<in> X))"
unfolding filters_to_ideals upsets_to_downsets inf_to_sup lat_ideals
by (smt image_iff image_inv_f_f image_is_empty inj_image_mem_iff inv_unique_comp inj_dual invol_dual)
end
context bounded_lattice_with_dual
begin
lemma top_filter: "X \<in> filters \<Longrightarrow> \<top> \<in> X"
using bot_ideal inj_image_mem_iff inj_dual filters_to_ideals top_dual by fastforce
end
context complete_lattice_with_dual
begin
lemma Inf_upset_id [simp]: "Inf \<circ> \<up> = id"
by (metis comp_assoc comp_id Sup_downset_id Sups_dual_def downset_upset_dual invol_dual)
lemma upset_Inf_id: "id \<le> \<up> \<circ> Inf"
by (simp add: Inf_lower le_funI subsetI upset_prop)
lemma Sup_Inf_var: " \<Sqinter>(\<Inter>x \<in> X. \<up>x) = \<Squnion>X"
unfolding upset_prop by (simp add: Collect_ball_eq Sup_eq_Inf)
lemma Sup_dual_upset_var: "(\<Inter>x \<in> X. \<up>x) = \<up>(\<Squnion>X)"
unfolding upset_prop by (safe, simp_all add: Sup_le_iff)
end
subsection \<open>Shunting Laws\<close>
text \<open>The first set of laws supplies so-called shunting laws for boolean algebras.
Such laws rather belong into Isabelle Main.\<close>
context boolean_algebra
begin
lemma shunt1: "(x \<sqinter> y \<le> z) = (x \<le> -y \<squnion> z)"
proof standard
assume "x \<sqinter> y \<le> z"
hence "-y \<squnion> (x \<sqinter> y) \<le> -y \<squnion> z"
using sup.mono by blast
hence "-y \<squnion> x \<le> -y \<squnion> z"
by (simp add: sup_inf_distrib1)
thus "x \<le> -y \<squnion> z"
by simp
next
assume "x \<le> -y \<squnion> z"
hence "x \<sqinter> y \<le> (-y \<squnion> z) \<sqinter> y"
using inf_mono by auto
thus "x \<sqinter> y \<le> z"
using inf.boundedE inf_sup_distrib2 by auto
qed
lemma shunt2: "(x \<sqinter> -y \<le> z) = (x \<le> y \<squnion> z)"
by (simp add: shunt1)
lemma meet_shunt: "(x \<sqinter> y = \<bottom>) = (x \<le> -y)"
- by (simp add: eq_iff shunt1)
+ by (simp add: order.eq_iff shunt1)
lemma join_shunt: "(x \<squnion> y = \<top>) = (-x \<le> y)"
by (metis compl_sup compl_top_eq double_compl meet_shunt)
lemma meet_shunt_var: "(x - y = \<bottom>) = (x \<le> y)"
by (simp add: diff_eq meet_shunt)
lemma join_shunt_var: "(x \<longrightarrow> y = \<top>) = (x \<le> y)"
by simp
end
subsection \<open>Properties of Complete Lattices\<close>
definition "Inf_closed_set X = (\<forall>Y \<subseteq> X. \<Sqinter>Y \<in> X)"
definition "Sup_closed_set X = (\<forall>Y \<subseteq> X. \<Squnion>Y \<in> X)"
definition "inf_closed_set X = (\<forall>x \<in> X. \<forall>y \<in> X. x \<sqinter> y \<in> X)"
definition "sup_closed_set X = (\<forall>x \<in> X. \<forall>y \<in> X. x \<squnion> y \<in> X)"
text \<open>The following facts about complete lattices add to those in the Isabelle libraries.\<close>
context complete_lattice
begin
text \<open>The translation between sup and Sup could be improved. The sup-theorems should be direct
consequences of Sup-ones. In addition, duality between sup and inf is currently not exploited.\<close>
lemma sup_Sup: "x \<squnion> y = \<Squnion>{x,y}"
by simp
lemma inf_Inf: "x \<sqinter> y = \<Sqinter>{x,y}"
by simp
text \<open>The next two lemmas are about Sups and Infs of indexed families. These are interesting for
iterations and fixpoints.\<close>
lemma fSup_unfold: "(f::nat \<Rightarrow> 'a) 0 \<squnion> (\<Squnion>n. f (Suc n)) = (\<Squnion>n. f n)"
- apply (intro antisym sup_least)
+ apply (intro order.antisym sup_least)
apply (rule Sup_upper, force)
apply (rule Sup_mono, force)
apply (safe intro!: Sup_least)
by (case_tac n, simp_all add: Sup_upper le_supI2)
lemma fInf_unfold: "(f::nat \<Rightarrow> 'a) 0 \<sqinter> (\<Sqinter>n. f (Suc n)) = (\<Sqinter>n. f n)"
- apply (intro antisym inf_greatest)
+ apply (intro order.antisym inf_greatest)
apply (rule Inf_greatest, safe)
apply (case_tac n)
apply simp_all
using Inf_lower inf.coboundedI2 apply force
apply (simp add: Inf_lower)
by (auto intro: Inf_mono)
end
lemma Sup_sup_closed: "Sup_closed_set (X::'a::complete_lattice set) \<Longrightarrow> sup_closed_set X"
by (metis Sup_closed_set_def empty_subsetI insert_subsetI sup_Sup sup_closed_set_def)
lemma Inf_inf_closed: "Inf_closed_set (X::'a::complete_lattice set) \<Longrightarrow> inf_closed_set X"
by (metis Inf_closed_set_def empty_subsetI inf_Inf inf_closed_set_def insert_subset)
subsection \<open>Sup- and Inf-Preservation\<close>
text \<open>Next, important notation for morphism between posets and lattices is introduced:
sup-preservation, inf-preservation and related properties.\<close>
abbreviation Sup_pres :: "('a::Sup \<Rightarrow> 'b::Sup) \<Rightarrow> bool" where
"Sup_pres f \<equiv> f \<circ> Sup = Sup \<circ> (`) f"
abbreviation Inf_pres :: "('a::Inf \<Rightarrow> 'b::Inf) \<Rightarrow> bool" where
"Inf_pres f \<equiv> f \<circ> Inf = Inf \<circ> (`) f"
abbreviation sup_pres :: "('a::sup \<Rightarrow> 'b::sup) \<Rightarrow> bool" where
"sup_pres f \<equiv> (\<forall>x y. f (x \<squnion> y) = f x \<squnion> f y)"
abbreviation inf_pres :: "('a::inf \<Rightarrow> 'b::inf) \<Rightarrow> bool" where
"inf_pres f \<equiv> (\<forall>x y. f (x \<sqinter> y) = f x \<sqinter> f y)"
abbreviation bot_pres :: "('a::bot \<Rightarrow> 'b::bot) \<Rightarrow> bool" where
"bot_pres f \<equiv> f \<bottom> = \<bottom>"
abbreviation top_pres :: "('a::top \<Rightarrow> 'b::top) \<Rightarrow> bool" where
"top_pres f \<equiv> f \<top> = \<top>"
abbreviation Sup_dual :: "('a::Sup \<Rightarrow> 'b::Inf) \<Rightarrow> bool" where
"Sup_dual f \<equiv> f \<circ> Sup = Inf \<circ> (`) f"
abbreviation Inf_dual :: "('a::Inf \<Rightarrow> 'b::Sup) \<Rightarrow> bool" where
"Inf_dual f \<equiv> f \<circ> Inf = Sup \<circ> (`) f"
abbreviation sup_dual :: "('a::sup \<Rightarrow> 'b::inf) \<Rightarrow> bool" where
"sup_dual f \<equiv> (\<forall>x y. f (x \<squnion> y) = f x \<sqinter> f y)"
abbreviation inf_dual :: "('a::inf \<Rightarrow> 'b::sup) \<Rightarrow> bool" where
"inf_dual f \<equiv> (\<forall>x y. f (x \<sqinter> y) = f x \<squnion> f y)"
abbreviation bot_dual :: "('a::bot \<Rightarrow> 'b::top) \<Rightarrow> bool" where
"bot_dual f \<equiv> f \<bottom> = \<top>"
abbreviation top_dual :: "('a::top \<Rightarrow> 'b::bot) \<Rightarrow> bool" where
"top_dual f \<equiv> f \<top> = \<bottom>"
text \<open>Inf-preservation and sup-preservation relate with duality.\<close>
lemma Inf_pres_map_dual_var:
"Inf_pres f = Sup_pres (\<partial>\<^sub>F f)"
for f :: "'a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual"
proof -
{ fix x :: "'a set"
assume "\<partial> (f (\<Sqinter> (\<partial> ` x))) = (\<Squnion>y\<in>x. \<partial> (f (\<partial> y)))" for x
then have "\<Sqinter> (f ` \<partial> ` A) = f (\<partial> (\<Squnion> A))" for A
by (metis (no_types) Sup_dual_def_var image_image invol_dual_var subset_dual)
then have "\<Sqinter> (f ` x) = f (\<Sqinter> x)"
by (metis Sup_dual_def_var subset_dual) }
then show ?thesis
by (auto simp add: map_dual_def fun_eq_iff Inf_dual_var Sup_dual_def_var image_comp)
qed
lemma Inf_pres_map_dual: "Inf_pres = Sup_pres \<circ> (\<partial>\<^sub>F::('a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual) \<Rightarrow> 'a \<Rightarrow> 'b)"
proof-
{fix f::"'a \<Rightarrow> 'b"
have "Inf_pres f = (Sup_pres \<circ> \<partial>\<^sub>F) f"
by (simp add: Inf_pres_map_dual_var)}
thus ?thesis
by force
qed
lemma Sup_pres_map_dual_var:
fixes f :: "'a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual"
shows "Sup_pres f = Inf_pres (\<partial>\<^sub>F f)"
by (metis Inf_pres_map_dual_var fun_dual5 map_dual_def)
lemma Sup_pres_map_dual: "Sup_pres = Inf_pres \<circ> (\<partial>\<^sub>F::('a::complete_lattice_with_dual \<Rightarrow> 'b::complete_lattice_with_dual) \<Rightarrow> 'a \<Rightarrow> 'b)"
by (simp add: Inf_pres_map_dual comp_assoc map_dual_invol)
text \<open>The following lemmas relate isotonicity of functions between complete lattices
with weak (left) preservation properties of sups and infs.\<close>
lemma fun_isol: "mono f \<Longrightarrow> mono ((\<circ>) f)"
by (simp add: le_fun_def mono_def)
lemma fun_isor: "mono f \<Longrightarrow> mono (\<lambda>x. x \<circ> f)"
by (simp add: le_fun_def mono_def)
lemma Sup_sup_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Sup_pres f \<Longrightarrow> sup_pres f"
by (metis (no_types, hide_lams) Sup_empty Sup_insert comp_apply image_insert sup_bot.right_neutral)
lemma Inf_inf_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows"Inf_pres f \<Longrightarrow> inf_pres f"
by (smt INF_insert Inf_empty Inf_insert comp_eq_elim inf_top.right_neutral)
lemma Sup_bot_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Sup_pres f \<Longrightarrow> bot_pres f"
by (metis SUP_empty Sup_empty comp_eq_elim)
lemma Inf_top_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Inf_pres f \<Longrightarrow> top_pres f"
by (metis INF_empty Inf_empty comp_eq_elim)
lemma Sup_sup_dual:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Sup_dual f \<Longrightarrow> sup_dual f"
by (smt comp_eq_elim image_empty image_insert inf_Inf sup_Sup)
lemma Inf_inf_dual:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Inf_dual f \<Longrightarrow> inf_dual f"
by (smt comp_eq_elim image_empty image_insert inf_Inf sup_Sup)
lemma Sup_bot_dual:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Sup_dual f \<Longrightarrow> bot_dual f"
by (metis INF_empty Sup_empty comp_eq_elim)
lemma Inf_top_dual:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Inf_dual f \<Longrightarrow> top_dual f"
by (metis Inf_empty SUP_empty comp_eq_elim)
text \<open>However, Inf-preservation does not imply top-preservation and
Sup-preservation does not imply bottom-preservation.\<close>
lemma
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Sup_pres f \<Longrightarrow> top_pres f" (*nitpick*)
oops
lemma
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Inf_pres f \<Longrightarrow> bot_pres f" (*nitpick*)
oops
context complete_lattice
begin
lemma iso_Inf_subdistl:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow>f \<circ> Inf \<le> Inf \<circ> (`) f"
by (simp add: complete_lattice_class.le_Inf_iff le_funI Inf_lower monoD)
lemma iso_Sup_supdistl:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> Sup \<circ> (`) f \<le> f \<circ> Sup"
by (simp add: complete_lattice_class.Sup_le_iff le_funI Sup_upper monoD)
lemma Inf_subdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "f \<circ> Inf \<le> Inf \<circ> (`) f \<Longrightarrow> mono f"
unfolding mono_def le_fun_def comp_def by (metis complete_lattice_class.le_INF_iff Inf_atLeast atLeast_iff)
lemma Sup_supdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "Sup \<circ> (`) f \<le> f \<circ> Sup \<Longrightarrow> mono f"
unfolding mono_def le_fun_def comp_def by (metis complete_lattice_class.SUP_le_iff Sup_atMost atMost_iff)
lemma supdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "(Sup \<circ> (`) f \<le> f \<circ> Sup) = mono f"
using Sup_supdistl_iso iso_Sup_supdistl by force
lemma subdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "(f \<circ> Inf \<le> Inf \<circ> (`) f) = mono f"
using Inf_subdistl_iso iso_Inf_subdistl by force
end
lemma ord_iso_Inf_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "ord_iso f \<Longrightarrow> Inf \<circ> (`) f = f \<circ> Inf"
proof-
let ?g = "the_inv f"
assume h: "ord_iso f"
hence a: "mono ?g"
by (simp add: ord_iso_the_inv)
{fix X :: "'a::complete_lattice set"
{fix y :: "'b::complete_lattice"
have "(y \<le> f (\<Sqinter>X)) = (?g y \<le> \<Sqinter>X)"
by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt)
also have "... = (\<forall>x \<in> X. ?g y \<le> x)"
by (simp add: le_Inf_iff)
also have "... = (\<forall>x \<in> X. y \<le> f x)"
by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt)
also have "... = (y \<le> \<Sqinter> (f ` X))"
by (simp add: le_INF_iff)
finally have "(y \<le> f (\<Sqinter>X)) = (y \<le> \<Sqinter> (f ` X))".}
hence "f (\<Sqinter>X) = \<Sqinter> (f ` X)"
by (meson dual_order.antisym order_refl)}
thus ?thesis
unfolding fun_eq_iff by simp
qed
lemma ord_iso_Sup_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "ord_iso f \<Longrightarrow> Sup \<circ> (`) f = f \<circ> Sup"
proof-
let ?g = "the_inv f"
assume h: "ord_iso f"
hence a: "mono ?g"
by (simp add: ord_iso_the_inv)
{fix X :: "'a::complete_lattice set"
{fix y :: "'b::complete_lattice"
have "(f (\<Squnion>X) \<le> y) = (\<Squnion>X \<le> ?g y)"
by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt)
also have "... = (\<forall>x \<in> X. x \<le> ?g y)"
by (simp add: Sup_le_iff)
also have "... = (\<forall>x \<in> X. f x \<le> y)"
by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt)
also have "... = (\<Squnion> (f ` X) \<le> y)"
by (simp add: SUP_le_iff)
finally have "(f (\<Squnion>X) \<le> y) = (\<Squnion> (f ` X) \<le> y)".}
hence "f (\<Squnion>X) = \<Squnion> (f ` X)"
by (meson dual_order.antisym order_refl)}
thus ?thesis
unfolding fun_eq_iff by simp
qed
text \<open>Right preservation of sups and infs is trivial.\<close>
lemma fSup_distr: "Sup_pres (\<lambda>x. x \<circ> f)"
unfolding fun_eq_iff by (simp add: image_comp)
lemma fSup_distr_var: "\<Squnion>F \<circ> g = (\<Squnion>f \<in> F. f \<circ> g)"
unfolding fun_eq_iff by (simp add: image_comp)
lemma fInf_distr: "Inf_pres (\<lambda>x. x \<circ> f)"
unfolding fun_eq_iff comp_def
by (smt INF_apply Inf_fun_def Sup.SUP_cong)
lemma fInf_distr_var: "\<Sqinter>F \<circ> g = (\<Sqinter>f \<in> F. f \<circ> g)"
unfolding fun_eq_iff comp_def
by (smt INF_apply INF_cong INF_image Inf_apply image_comp image_def image_image)
text \<open>The next set of lemma revisits the preservation properties in the function space.\<close>
lemma fSup_subdistl:
assumes "mono (f::'a::complete_lattice \<Rightarrow> 'b::complete_lattice)"
shows "Sup \<circ> (`) ((\<circ>) f) \<le> (\<circ>) f \<circ> Sup"
using assms by (simp add: fun_isol supdistl_iso)
lemma fSup_subdistl_var:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> (\<Squnion>g \<in> G. f \<circ> g) \<le> f \<circ> \<Squnion>G"
by (simp add: fun_isol mono_Sup)
lemma fInf_subdistl:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> (\<circ>) f \<circ> Inf \<le> Inf \<circ> (`) ((\<circ>) f)"
by (simp add: fun_isol subdistl_iso)
lemma fInf_subdistl_var:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> f \<circ> \<Sqinter>G \<le> (\<Sqinter>g \<in> G. f \<circ> g)"
by (simp add: fun_isol mono_Inf)
lemma fSup_distl: "Sup_pres f \<Longrightarrow> Sup_pres ((\<circ>) f)"
unfolding fun_eq_iff by (simp add: image_comp)
lemma fSup_distl_var: "Sup_pres f \<Longrightarrow> f \<circ> \<Squnion>G = (\<Squnion>g \<in> G. f \<circ> g)"
unfolding fun_eq_iff by (simp add: image_comp)
lemma fInf_distl: "Inf_pres f \<Longrightarrow> Inf_pres ((\<circ>) f)"
unfolding fun_eq_iff by (simp add: image_comp)
lemma fInf_distl_var: "Inf_pres f \<Longrightarrow> f \<circ> \<Sqinter>G = (\<Sqinter>g \<in> G. f \<circ> g)"
unfolding fun_eq_iff by (simp add: image_comp)
text \<open>Downsets preserve infs whereas upsets preserve sups.\<close>
lemma Inf_pres_downset: "Inf_pres (\<down>::'a::complete_lattice_with_dual \<Rightarrow> 'a set)"
unfolding downset_prop fun_eq_iff
by (safe, simp_all add: le_Inf_iff)
lemma Sup_dual_upset: "Sup_dual (\<up>::'a::complete_lattice_with_dual \<Rightarrow> 'a set)"
unfolding upset_prop fun_eq_iff
by (safe, simp_all add: Sup_le_iff)
text \<open>Images of Sup-morphisms are closed under Sups and images of Inf-morphisms are closed under Infs.\<close>
lemma Sup_pres_Sup_closed: "Sup_pres f \<Longrightarrow> Sup_closed_set (range f)"
by (metis (mono_tags, lifting) Sup_closed_set_def comp_eq_elim range_eqI subset_image_iff)
lemma Inf_pres_Inf_closed: "Inf_pres f \<Longrightarrow> Inf_closed_set (range f)"
by (metis (mono_tags, lifting) Inf_closed_set_def comp_eq_elim range_eqI subset_image_iff)
text \<open>It is well known that functions into complete lattices form complete lattices. Here, such results are shown for
the subclasses of isotone functions, where additional closure conditions must be respected.\<close>
typedef (overloaded) 'a iso = "{f::'a::order \<Rightarrow> 'a::order. mono f}"
by (metis Abs_ord_homset_cases ord_homset_def)
setup_lifting type_definition_iso
instantiation iso :: (complete_lattice) complete_lattice
begin
lift_definition Inf_iso :: "'a::complete_lattice iso set \<Rightarrow> 'a iso" is Sup
by (metis (mono_tags, lifting) SUP_subset_mono Sup_apply mono_def subsetI)
lift_definition Sup_iso :: "'a::complete_lattice iso set \<Rightarrow> 'a iso" is Inf
by (smt INF_lower2 Inf_apply le_INF_iff mono_def)
lift_definition bot_iso :: "'a::complete_lattice iso" is "\<top>"
by (simp add: monoI)
lift_definition sup_iso :: "'a::complete_lattice iso \<Rightarrow> 'a iso \<Rightarrow> 'a iso" is inf
by (smt inf_apply inf_mono monoD monoI)
lift_definition top_iso :: "'a::complete_lattice iso" is "\<bottom>"
by (simp add: mono_def)
lift_definition inf_iso :: "'a::complete_lattice iso \<Rightarrow> 'a iso \<Rightarrow> 'a iso" is sup
by (smt mono_def sup.mono sup_apply)
lift_definition less_eq_iso :: "'a::complete_lattice iso \<Rightarrow> 'a iso \<Rightarrow> bool" is "(\<ge>)".
lift_definition less_iso :: "'a::complete_lattice iso \<Rightarrow> 'a iso \<Rightarrow> bool" is "(>)".
instance
by (intro_classes; transfer, simp_all add: less_fun_def Sup_upper Sup_least Inf_lower Inf_greatest)
end
text \<open>Duality has been baked into this result because of its relevance for predicate transformers. A proof
where Sups are mapped to Sups and Infs to Infs is certainly possible, but two instantiation of the same type
and the same classes are unfortunately impossible. Interpretations could be used instead.
A corresponding result for Inf-preseving functions and Sup-lattices, is proved in components on transformers,
as more advanced properties about Inf-preserving functions are needed.\<close>
subsection \<open>Alternative Definitions for Complete Boolean Algebras\<close>
text \<open>The current definitions of complete boolean algebras deviates from that in most textbooks in that
a distributive law with infinite sups and infinite infs is used. There are interesting applications, for instance
in topology, where weaker laws are needed --- for instance for frames and locales.\<close>
class complete_heyting_algebra = complete_lattice +
assumes ch_dist: "x \<sqinter> \<Squnion>Y = (\<Squnion>y \<in> Y. x \<sqinter> y)"
text \<open>Complete Heyting algebras are also known as frames or locales (they differ with respect to their morphisms).\<close>
class complete_co_heyting_algebra = complete_lattice +
assumes co_ch_dist: "x \<squnion> \<Sqinter>Y = (\<Sqinter>y \<in> Y. x \<squnion> y)"
class complete_boolean_algebra_alt = complete_lattice + boolean_algebra
instance set :: (type) complete_boolean_algebra_alt..
context complete_boolean_algebra_alt
begin
subclass complete_heyting_algebra
proof
fix x Y
{fix t
have "(x \<sqinter> \<Squnion>Y \<le> t) = (\<Squnion>Y \<le> -x \<squnion> t)"
by (simp add: inf.commute shunt1[symmetric])
also have "... = (\<forall>y \<in> Y. y \<le> -x \<squnion> t)"
using Sup_le_iff by blast
also have "... = (\<forall>y \<in> Y. x \<sqinter> y \<le> t)"
by (simp add: inf.commute shunt1)
finally have "(x \<sqinter> \<Squnion>Y \<le> t) = ((\<Squnion>y\<in>Y. x \<sqinter> y) \<le> t)"
by (simp add: local.SUP_le_iff)}
thus "x \<sqinter> \<Squnion>Y = (\<Squnion>y\<in>Y. x \<sqinter> y)"
- using eq_iff by blast
+ using order.eq_iff by blast
qed
subclass complete_co_heyting_algebra
apply unfold_locales
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: INF_greatest Inf_lower2)
by (meson eq_refl le_INF_iff le_Inf_iff shunt2)
lemma de_morgan1: "-(\<Squnion>X) = (\<Sqinter>x \<in> X. -x)"
proof-
{fix y
have "(y \<le> -(\<Squnion>X)) = (\<Squnion>X \<le> -y)"
using compl_le_swap1 by blast
also have "... = (\<forall>x \<in> X. x \<le> -y)"
by (simp add: Sup_le_iff)
also have "... = (\<forall>x \<in> X. y \<le> -x)"
using compl_le_swap1 by blast
also have "... = (y \<le> (\<Sqinter>x \<in> X. -x))"
using le_INF_iff by force
finally have "(y \<le> -(\<Squnion>X)) = (y \<le>(\<Sqinter>x \<in> X. -x))".}
thus ?thesis
- using antisym by blast
+ using order.antisym by blast
qed
lemma de_morgan2: "-(\<Sqinter>X) = (\<Squnion>x \<in> X. -x)"
by (metis de_morgan1 ba_dual.dual_iff ba_dual.image_dual pointfree_idE)
end
class complete_boolean_algebra_alt_with_dual = complete_lattice_with_dual + complete_boolean_algebra_alt
instantiation set :: (type) complete_boolean_algebra_alt_with_dual
begin
definition dual_set :: "'a set \<Rightarrow> 'a set" where
"dual_set = uminus"
instance
by intro_classes (simp_all add: ba_dual.inj_dual dual_set_def comp_def uminus_Sup id_def)
end
context complete_boolean_algebra_alt
begin
sublocale cba_dual: complete_boolean_algebra_alt_with_dual _ _ _ _ _ _ _ _ uminus _ _
by unfold_locales (auto simp: de_morgan2 de_morgan1)
end
subsection \<open>Atomic Boolean Algebras\<close>
text \<open>Next, atomic boolean algebras are defined.\<close>
context bounded_lattice
begin
text \<open>Atoms are covers of bottom.\<close>
definition "atom x = (x \<noteq> \<bottom> \<and> \<not>(\<exists>y. \<bottom> < y \<and> y < x))"
definition "atom_map x = {y. atom y \<and> y \<le> x}"
lemma atom_map_def_var: "atom_map x = \<down>x \<inter> Collect atom"
unfolding atom_map_def downset_def downset_set_def comp_def atom_def by fastforce
lemma atom_map_atoms: "\<Union>(range atom_map) = Collect atom"
unfolding atom_map_def atom_def by auto
end
typedef (overloaded) 'a atoms = "range (atom_map::'a::bounded_lattice \<Rightarrow> 'a set)"
by blast
setup_lifting type_definition_atoms
definition at_map :: "'a::bounded_lattice \<Rightarrow> 'a atoms" where
"at_map = Abs_atoms \<circ> atom_map"
class atomic_boolean_algebra = boolean_algebra +
assumes atomicity: "x \<noteq> \<bottom> \<Longrightarrow> (\<exists>y. atom y \<and> y \<le> x)"
class complete_atomic_boolean_algebra = complete_lattice + atomic_boolean_algebra
begin
subclass complete_boolean_algebra_alt..
end
text \<open>Here are two equivalent definitions for atoms; first in boolean algebras, and then in complete
boolean algebras.\<close>
context boolean_algebra
begin
text \<open>The following two conditions are taken from Koppelberg's book~\cite{Koppelberg89}.\<close>
lemma atom_neg: "atom x \<Longrightarrow> x \<noteq> \<bottom> \<and> (\<forall>y z. x \<le> y \<or> x \<le> -y)"
by (metis atom_def dual_order.order_iff_strict inf.cobounded1 inf.commute meet_shunt)
lemma atom_sup: "(\<forall>y. x \<le> y \<or> x \<le> -y) \<Longrightarrow> (\<forall>y z. (x \<le> y \<or> x \<le> z) = (x \<le> y \<squnion> z))"
by (metis inf.orderE le_supI1 shunt2)
lemma sup_atom: "x \<noteq> \<bottom> \<Longrightarrow> (\<forall>y z. (x \<le> y \<or> x \<le> z) = (x \<le> y \<squnion> z)) \<Longrightarrow> atom x"
unfolding atom_def apply clarsimp by (metis bot_less inf.absorb2 less_le_not_le meet_shunt sup_compl_top)
lemma atom_sup_iff: "atom x = (x \<noteq> \<bottom> \<and> (\<forall>y z. (x \<le> y \<or> x \<le> z) = (x \<le> y \<squnion> z)))"
by (standard, auto simp add: atom_neg atom_sup sup_atom)
lemma atom_neg_iff: "atom x = (x \<noteq> \<bottom> \<and> (\<forall>y z. x \<le> y \<or> x \<le> -y))"
by (standard, auto simp add: atom_neg atom_sup sup_atom)
lemma atom_map_bot_pres: "atom_map \<bottom> = {}"
using atom_def atom_map_def le_bot by auto
lemma atom_map_top_pres: "atom_map \<top> = Collect atom"
using atom_map_def by auto
end
context complete_boolean_algebra_alt
begin
lemma atom_Sup: "\<And>Y. x \<noteq> \<bottom> \<Longrightarrow> (\<forall>y. x \<le> y \<or> x \<le> -y) \<Longrightarrow> ((\<exists>y \<in> Y. x \<le> y) = (x \<le> \<Squnion>Y))"
by (metis Sup_least Sup_upper2 compl_le_swap1 le_iff_inf meet_shunt)
lemma Sup_atom: "x \<noteq> \<bottom> \<Longrightarrow> (\<forall>Y. (\<exists>y \<in> Y. x \<le> y) = (x \<le> \<Squnion>Y)) \<Longrightarrow> atom x"
proof-
assume h1: "x \<noteq> \<bottom>"
and h2: "\<forall>Y. (\<exists>y \<in> Y. x \<le> y) = (x \<le> \<Squnion>Y)"
hence "\<forall>y z. (x \<le> y \<or> x \<le> z) = (x \<le> y \<squnion> z)"
by (smt insert_iff sup_Sup sup_bot.right_neutral)
thus "atom x"
by (simp add: h1 sup_atom)
qed
lemma atom_Sup_iff: "atom x = (x \<noteq> \<bottom> \<and> (\<forall>Y. (\<exists>y \<in> Y. x \<le> y) = (x \<le> \<Squnion>Y)))"
by standard (auto simp: atom_neg atom_Sup Sup_atom)
end
end
-
-
-
-
diff --git a/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy b/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy
--- a/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy
+++ b/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy
@@ -1,575 +1,575 @@
(*
Title: Locale-Based Duality
Author: Georg Struth
Maintainer:Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Locale-Based Duality\<close>
theory Order_Lattice_Props_Loc
imports Main
"HOL-Library.Lattice_Syntax"
begin
text \<open>This section explores order and lattice duality based on locales. Used within the context of a class or locale,
this is very effective, though more opaque than the previous approach. Outside of such a context, however, it apparently
cannot be used for dualising theorems. Examples are properties of functions between orderings or lattices.\<close>
definition Fix :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set" where
"Fix f = {x. f x = x}"
context ord
begin
definition min_set :: "'a set \<Rightarrow> 'a set" where
"min_set X = {y \<in> X. \<forall>x \<in> X. x \<le> y \<longrightarrow> x = y}"
definition max_set :: "'a set \<Rightarrow> 'a set" where
"max_set X = {x \<in> X. \<forall>y \<in> X. x \<le> y \<longrightarrow> x = y}"
definition directed :: "'a set \<Rightarrow> bool" where
"directed X = (\<forall>Y. finite Y \<and> Y \<subseteq> X \<longrightarrow> (\<exists>x \<in> X. \<forall>y \<in> Y. y \<le> x))"
definition filtered :: "'a set \<Rightarrow> bool" where
"filtered X = (\<forall>Y. finite Y \<and> Y \<subseteq> X \<longrightarrow> (\<exists>x \<in> X. \<forall>y \<in> Y. x \<le> y))"
definition downset_set :: "'a set \<Rightarrow> 'a set" ("\<Down>") where
"\<Down>X = {y. \<exists>x \<in> X. y \<le> x}"
definition upset_set :: "'a set \<Rightarrow> 'a set" ("\<Up>") where
"\<Up>X = {y. \<exists>x \<in> X. x \<le> y}"
definition downset :: "'a \<Rightarrow> 'a set" ("\<down>") where
"\<down> = \<Down> \<circ> (\<lambda>x. {x})"
definition upset :: "'a \<Rightarrow> 'a set" ("\<up>") where
"\<up> = \<Up> \<circ> (\<lambda>x. {x})"
definition downsets :: "'a set set" where
"downsets = Fix \<Down>"
definition upsets :: "'a set set" where
"upsets = Fix \<Up>"
abbreviation "downset_setp X \<equiv> X \<in> downsets"
abbreviation "upset_setp X \<equiv> X \<in> upsets"
definition ideals :: "'a set set" where
"ideals = {X. X \<noteq> {} \<and> downset_setp X \<and> directed X}"
definition filters :: "'a set set" where
"filters = {X. X \<noteq> {} \<and> upset_setp X \<and> filtered X}"
abbreviation "idealp X \<equiv> X \<in> ideals"
abbreviation "filterp X \<equiv> X \<in> filters"
end
abbreviation Sup_pres :: "('a::Sup \<Rightarrow> 'b::Sup) \<Rightarrow> bool" where
"Sup_pres f \<equiv> f \<circ> Sup = Sup \<circ> (`) f"
abbreviation Inf_pres :: "('a::Inf \<Rightarrow> 'b::Inf) \<Rightarrow> bool" where
"Inf_pres f \<equiv> f \<circ> Inf = Inf \<circ> (`) f"
abbreviation sup_pres :: "('a::sup \<Rightarrow> 'b::sup) \<Rightarrow> bool" where
"sup_pres f \<equiv> (\<forall>x y. f (x \<squnion> y) = f x \<squnion> f y)"
abbreviation inf_pres :: "('a::inf \<Rightarrow> 'b::inf) \<Rightarrow> bool" where
"inf_pres f \<equiv> (\<forall>x y. f (x \<sqinter> y) = f x \<sqinter> f y)"
abbreviation bot_pres :: "('a::bot \<Rightarrow> 'b::bot) \<Rightarrow> bool" where
"bot_pres f \<equiv> f \<bottom> = \<bottom>"
abbreviation top_pres :: "('a::top \<Rightarrow> 'b::top) \<Rightarrow> bool" where
"top_pres f \<equiv> f \<top> = \<top>"
abbreviation Sup_dual :: "('a::Sup \<Rightarrow> 'b::Inf) \<Rightarrow> bool" where
"Sup_dual f \<equiv> f \<circ> Sup = Inf \<circ> (`) f"
abbreviation Inf_dual :: "('a::Inf \<Rightarrow> 'b::Sup) \<Rightarrow> bool" where
"Inf_dual f \<equiv> f \<circ> Inf = Sup \<circ> (`) f"
abbreviation sup_dual :: "('a::sup \<Rightarrow> 'b::inf) \<Rightarrow> bool" where
"sup_dual f \<equiv> (\<forall>x y. f (x \<squnion> y) = f x \<sqinter> f y)"
abbreviation inf_dual :: "('a::inf \<Rightarrow> 'b::sup) \<Rightarrow> bool" where
"inf_dual f \<equiv> (\<forall>x y. f (x \<sqinter> y) = f x \<squnion> f y)"
abbreviation bot_dual :: "('a::bot \<Rightarrow> 'b::top) \<Rightarrow> bool" where
"bot_dual f \<equiv> f \<bottom> = \<top>"
abbreviation top_dual :: "('a::top \<Rightarrow> 'b::bot) \<Rightarrow> bool" where
"top_dual f \<equiv> f \<top> = \<bottom>"
subsection \<open>Duality via Locales\<close>
sublocale ord \<subseteq> dual_ord: ord "(\<ge>)" "(>)"
rewrites dual_max_set: "max_set = dual_ord.min_set"
and dual_filtered: "filtered = dual_ord.directed"
and dual_upset_set: "upset_set = dual_ord.downset_set"
and dual_upset: "upset = dual_ord.downset"
and dual_upsets: "upsets = dual_ord.downsets"
and dual_filters: "filters = dual_ord.ideals"
apply unfold_locales
unfolding max_set_def ord.min_set_def fun_eq_iff apply blast
unfolding filtered_def ord.directed_def apply simp
unfolding upset_set_def ord.downset_set_def apply simp
apply (simp add: ord.downset_def ord.downset_set_def ord.upset_def ord.upset_set_def)
unfolding upsets_def ord.downsets_def apply (metis ord.downset_set_def upset_set_def)
unfolding filters_def ord.ideals_def Fix_def ord.downsets_def upsets_def ord.downset_set_def upset_set_def ord.directed_def filtered_def
by simp
sublocale preorder \<subseteq> dual_preorder: preorder "(\<ge>)" "(>)"
apply unfold_locales
apply (simp add: less_le_not_le)
apply simp
using order_trans by blast
sublocale order \<subseteq> dual_order: order "(\<ge>)" "(>)"
by (unfold_locales, simp)
sublocale lattice \<subseteq> dual_lattice: lattice sup "(\<ge>)" "(>)" inf
by (unfold_locales, simp_all)
sublocale bounded_lattice \<subseteq> dual_bounded_lattice: bounded_lattice sup "(\<ge>)" "(>)" inf \<top> \<bottom>
by (unfold_locales, simp_all)
sublocale boolean_algebra \<subseteq> dual_boolean_algebra: boolean_algebra "\<lambda>x y. x \<squnion> -y" uminus sup "(\<ge>)" "(>)" inf \<top> \<bottom>
by (unfold_locales, simp_all add: inf_sup_distrib1)
sublocale complete_lattice \<subseteq> dual_complete_lattice: complete_lattice Sup Inf sup "(\<ge>)" "(>)" inf \<top> \<bottom>
rewrites dual_gfp: "gfp = dual_complete_lattice.lfp"
proof-
show "class.complete_lattice Sup Inf sup (\<ge>) (>) inf \<top> \<bottom>"
by (unfold_locales, simp_all add: Sup_upper Sup_least Inf_lower Inf_greatest)
then interpret dual_complete_lattice: complete_lattice Sup Inf sup "(\<ge>)" "(>)" inf \<top> \<bottom>.
show "gfp = dual_complete_lattice.lfp"
unfolding gfp_def dual_complete_lattice.lfp_def fun_eq_iff by simp
qed
context ord
begin
lemma dual_min_set: "min_set = dual_ord.max_set"
by (simp add: dual_ord.dual_max_set)
lemma dual_directed: "directed = dual_ord.filtered"
by (simp add:dual_ord.dual_filtered)
lemma dual_downset: "downset = dual_ord.upset"
by (simp add: dual_ord.dual_upset)
lemma dual_downset_set: "downset_set = dual_ord.upset_set"
by (simp add: dual_ord.dual_upset_set)
lemma dual_downsets: "downsets = dual_ord.upsets"
by (simp add: dual_ord.dual_upsets)
lemma dual_ideals: "ideals = dual_ord.filters"
by (simp add: dual_ord.dual_filters)
end
context complete_lattice
begin
lemma dual_lfp: "lfp = dual_complete_lattice.gfp"
by (simp add: dual_complete_lattice.dual_gfp)
end
subsection \<open>Properties of Orderings, Again\<close>
context ord
begin
lemma directed_nonempty: "directed X \<Longrightarrow> X \<noteq> {}"
unfolding directed_def by fastforce
lemma directed_ub: "directed X \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z)"
by (meson empty_subsetI directed_def finite.emptyI finite_insert insert_subset order_refl)
lemma downset_set_prop: "\<Down> = Union \<circ> (`) \<down>"
unfolding downset_set_def downset_def fun_eq_iff by fastforce
lemma downset_set_prop_var: "\<Down>X = (\<Union>x \<in> X. \<down>x)"
by (simp add: downset_set_prop)
lemma downset_prop: "\<down>x = {y. y \<le> x}"
unfolding downset_def downset_set_def fun_eq_iff comp_def by fastforce
end
context preorder
begin
lemma directed_prop: "X \<noteq> {} \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z) \<Longrightarrow> directed X"
proof-
assume h1: "X \<noteq> {}"
and h2: "\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z"
{fix Y
have "finite Y \<Longrightarrow> Y \<subseteq> X \<Longrightarrow> (\<exists>x \<in> X. \<forall>y \<in> Y. y \<le> x)"
proof (induct rule: finite_induct)
case empty
then show ?case
using h1 by blast
next
case (insert x F)
then show ?case
by (metis h2 insert_iff insert_subset order_trans)
qed}
thus ?thesis
by (simp add: directed_def)
qed
lemma directed_alt: "directed X = (X \<noteq> {} \<and> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. x \<le> z \<and> y \<le> z))"
by (metis directed_prop directed_nonempty directed_ub)
lemma downset_set_ext: "id \<le> \<Down>"
unfolding le_fun_def id_def downset_set_def by auto
lemma downset_set_iso: "mono \<Down>"
unfolding mono_def downset_set_def by blast
lemma downset_set_idem [simp]: "\<Down> \<circ> \<Down> = \<Down>"
unfolding fun_eq_iff downset_set_def comp_def using order_trans by auto
lemma downset_faithful: "\<down>x \<subseteq> \<down>y \<Longrightarrow> x \<le> y"
by (simp add: downset_prop subset_eq)
lemma downset_iso_iff: "(\<down>x \<subseteq> \<down>y) = (x \<le> y)"
using atMost_iff downset_prop order_trans by blast
lemma downset_directed_downset_var [simp]: "directed (\<Down>X) = directed X"
proof
assume h1: "directed X"
{fix Y
assume h2: "finite Y" and h3: "Y \<subseteq> \<Down>X"
hence "\<forall>y. \<exists>x. y \<in> Y \<longrightarrow> x \<in> X \<and> y \<le> x"
by (force simp: downset_set_def)
hence "\<exists>f. \<forall>y. y \<in> Y \<longrightarrow> f y \<in> X \<and> y \<le> f y"
by (rule choice)
hence "\<exists>f. finite (f ` Y) \<and> f ` Y \<subseteq> X \<and> (\<forall>y \<in> Y. y \<le> f y)"
by (metis finite_imageI h2 image_subsetI)
hence "\<exists>Z. finite Z \<and> Z \<subseteq> X \<and> (\<forall>y \<in> Y. \<exists> z \<in> Z. y \<le> z)"
by fastforce
hence "\<exists>Z. finite Z \<and> Z \<subseteq> X \<and> (\<forall>y \<in> Y. \<exists> z \<in> Z. y \<le> z) \<and> (\<exists>x \<in> X. \<forall> z \<in> Z. z \<le> x)"
by (metis directed_def h1)
hence "\<exists>x \<in> X. \<forall>y \<in> Y. y \<le> x"
by (meson order_trans)}
thus "directed (\<Down>X)"
unfolding directed_def downset_set_def by fastforce
next
assume "directed (\<Down>X)"
thus "directed X"
unfolding directed_def downset_set_def
apply clarsimp
by (smt Ball_Collect order_refl order_trans subsetCE)
qed
lemma downset_directed_downset [simp]: "directed \<circ> \<Down> = directed"
unfolding fun_eq_iff comp_def by simp
lemma directed_downset_ideals: "directed (\<Down>X) = (\<Down>X \<in> ideals)"
by (metis (mono_tags, lifting) Fix_def comp_apply directed_alt downset_set_idem downsets_def ideals_def mem_Collect_eq)
end
lemma downset_iso: "mono (\<down>::'a::order \<Rightarrow> 'a set)"
by (simp add: downset_iso_iff mono_def)
context order
begin
lemma downset_inj: "inj \<down>"
- by (metis injI downset_iso_iff eq_iff)
+ by (metis injI downset_iso_iff order.eq_iff)
end
context lattice
begin
lemma lat_ideals: "X \<in> ideals = (X \<noteq> {} \<and> X \<in> downsets \<and> (\<forall>x \<in> X. \<forall> y \<in> X. x \<squnion> y \<in> X))"
unfolding ideals_def directed_alt downsets_def Fix_def downset_set_def
by (clarsimp, smt sup.cobounded1 sup.orderE sup.orderI sup_absorb2 sup_left_commute mem_Collect_eq)
end
context bounded_lattice
begin
lemma bot_ideal: "X \<in> ideals \<Longrightarrow> \<bottom> \<in> X"
unfolding ideals_def downsets_def Fix_def downset_set_def by fastforce
end
context complete_lattice
begin
lemma Sup_downset_id [simp]: "Sup \<circ> \<down> = id"
using Sup_atMost atMost_def downset_prop by fastforce
lemma downset_Sup_id: "id \<le> \<down> \<circ> Sup"
by (simp add: Sup_upper downset_prop le_funI subsetI)
lemma Inf_Sup_var: "\<Squnion>(\<Inter>x \<in> X. \<down>x) = \<Sqinter>X"
unfolding downset_prop by (simp add: Collect_ball_eq Inf_eq_Sup)
lemma Inf_pres_downset_var: "(\<Inter>x \<in> X. \<down>x) = \<down>(\<Sqinter>X)"
unfolding downset_prop by (safe, simp_all add: le_Inf_iff)
end
lemma lfp_in_Fix:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "mono f \<Longrightarrow> lfp f \<in> Fix f"
using Fix_def lfp_unfold by fastforce
lemma gfp_in_Fix:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "mono f \<Longrightarrow> gfp f \<in> Fix f"
using Fix_def gfp_unfold by fastforce
lemma nonempty_Fix:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
shows "mono f \<Longrightarrow> Fix f \<noteq> {}"
using lfp_in_Fix by fastforce
subsection \<open>Dual Properties of Orderings from Locales\<close>
text \<open>These properties can be proved very smoothly overall. But only within the context of a class
or locale!\<close>
context ord
begin
lemma filtered_nonempty: "filtered X \<Longrightarrow> X \<noteq> {}"
by (simp add: dual_filtered dual_ord.directed_nonempty)
lemma filtered_lb: "filtered X \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. z \<le> x \<and> z \<le> y)"
by (simp add: dual_filtered dual_ord.directed_ub)
lemma upset_set_prop: "\<Up> = Union \<circ> (`) \<up>"
by (simp add: dual_ord.downset_set_prop dual_upset dual_upset_set)
lemma upset_set_prop_var: "\<Up>X = (\<Union>x \<in> X. \<up>x)"
by (simp add: dual_ord.downset_set_prop_var dual_upset dual_upset_set)
lemma upset_prop: "\<up>x = {y. x \<le> y}"
by (simp add: dual_ord.downset_prop dual_upset)
end
context preorder
begin
lemma filtered_prop: "X \<noteq> {} \<Longrightarrow> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. z \<le> x \<and> z \<le> y) \<Longrightarrow> filtered X"
by (simp add: dual_filtered dual_preorder.directed_prop)
lemma filtered_alt: "filtered X = (X \<noteq> {} \<and> (\<forall>x \<in> X. \<forall>y \<in> X. \<exists>z \<in> X. z \<le> x \<and> z \<le> y))"
by (simp add: dual_filtered dual_preorder.directed_alt)
lemma upset_set_ext: "id \<le> \<Up>"
by (simp add: dual_preorder.downset_set_ext dual_upset_set)
lemma upset_set_anti: "mono \<Up>"
by (simp add: dual_preorder.downset_set_iso dual_upset_set)
lemma up_set_idem [simp]: "\<Up> \<circ> \<Up> = \<Up>"
by (simp add: dual_upset_set)
lemma upset_faithful: "\<up>x \<subseteq> \<up>y \<Longrightarrow> y \<le> x"
by (metis dual_preorder.downset_faithful dual_upset)
lemma upset_anti_iff: "(\<up>y \<subseteq> \<up>x) = (x \<le> y)"
by (simp add: dual_preorder.downset_iso_iff dual_upset)
lemma upset_filtered_upset [simp]: "filtered \<circ> \<Up> = filtered"
by (simp add: dual_filtered dual_upset_set)
lemma filtered_upset_filters: "filtered (\<Up>X) = (\<Up>X \<in> filters)"
using dual_filtered dual_preorder.directed_downset_ideals dual_upset_set ord.dual_filters by fastforce
end
context order
begin
lemma upset_inj: "inj \<up>"
by (simp add: dual_order.downset_inj dual_upset)
end
context lattice
begin
lemma lat_filters: "X \<in> filters = (X \<noteq> {} \<and> X \<in> upsets \<and> (\<forall>x \<in> X. \<forall> y \<in> X. x \<sqinter> y \<in> X))"
by (simp add: dual_filters dual_lattice.lat_ideals dual_ord.downsets_def dual_upset_set upsets_def)
end
context bounded_lattice
begin
lemma top_filter: "X \<in> filters \<Longrightarrow> \<top> \<in> X"
by (simp add: dual_bounded_lattice.bot_ideal dual_filters)
end
context complete_lattice
begin
lemma Inf_upset_id [simp]: "Inf \<circ> \<up> = id"
by (simp add: dual_upset)
lemma upset_Inf_id: "id \<le> \<up> \<circ> Inf"
by (simp add: dual_complete_lattice.downset_Sup_id dual_upset)
lemma Sup_Inf_var: " \<Sqinter>(\<Inter>x \<in> X. \<up>x) = \<Squnion>X"
by (simp add: dual_complete_lattice.Inf_Sup_var dual_upset)
lemma Sup_dual_upset_var: "(\<Inter>x \<in> X. \<up>x) = \<up>(\<Squnion>X)"
by (simp add: dual_complete_lattice.Inf_pres_downset_var dual_upset)
end
subsection \<open>Examples that Do Not Dualise\<close>
lemma upset_anti: "antimono (\<up>::'a::order \<Rightarrow> 'a set)"
by (simp add: antimono_def upset_anti_iff)
context complete_lattice
begin
lemma fSup_unfold: "(f::nat \<Rightarrow> 'a) 0 \<squnion> (\<Squnion>n. f (Suc n)) = (\<Squnion>n. f n)"
- apply (intro antisym sup_least)
+ apply (intro order.antisym sup_least)
apply (rule Sup_upper, force)
apply (rule Sup_mono, force)
apply (safe intro!: Sup_least)
by (case_tac n, simp_all add: Sup_upper le_supI2)
lemma fInf_unfold: "(f::nat \<Rightarrow> 'a) 0 \<sqinter> (\<Sqinter>n. f (Suc n)) = (\<Sqinter>n. f n)"
- apply (intro antisym inf_greatest)
+ apply (intro order.antisym inf_greatest)
apply (rule Inf_greatest, safe)
apply (case_tac n)
apply simp_all
using Inf_lower inf.coboundedI2 apply force
apply (simp add: Inf_lower)
by (auto intro: Inf_mono)
end
lemma fun_isol: "mono f \<Longrightarrow> mono ((\<circ>) f)"
by (simp add: le_fun_def mono_def)
lemma fun_isor: "mono f \<Longrightarrow> mono (\<lambda>x. x \<circ> f)"
by (simp add: le_fun_def mono_def)
lemma Sup_sup_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Sup_pres f \<Longrightarrow> sup_pres f"
by (metis (no_types, hide_lams) Sup_empty Sup_insert comp_apply image_insert sup_bot.right_neutral)
lemma Inf_inf_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows"Inf_pres f \<Longrightarrow> inf_pres f"
by (smt INF_insert comp_eq_elim dual_complete_lattice.Sup_empty dual_complete_lattice.Sup_insert inf_top.right_neutral)
lemma Sup_bot_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Sup_pres f \<Longrightarrow> bot_pres f"
by (metis SUP_empty Sup_empty comp_eq_elim)
lemma Inf_top_pres:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "Inf_pres f \<Longrightarrow> top_pres f"
by (metis INF_empty comp_eq_elim dual_complete_lattice.Sup_empty)
context complete_lattice
begin
lemma iso_Inf_subdistl:
assumes "mono (f::'a \<Rightarrow> 'b::complete_lattice)"
shows "f \<circ> Inf \<le> Inf \<circ> (`) f"
by (simp add: assms complete_lattice_class.le_Inf_iff le_funI Inf_lower monoD)
lemma iso_Sup_supdistl:
assumes "mono (f::'a \<Rightarrow> 'b::complete_lattice)"
shows "Sup \<circ> (`) f \<le> f \<circ> Sup"
by (simp add: assms complete_lattice_class.SUP_le_iff le_funI dual_complete_lattice.Inf_lower monoD)
lemma Inf_subdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "f \<circ> Inf \<le> Inf \<circ> (`) f \<Longrightarrow> mono f"
unfolding mono_def le_fun_def comp_def by (metis complete_lattice_class.le_INF_iff Inf_atLeast atLeast_iff)
lemma Sup_supdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "Sup \<circ> (`) f \<le> f \<circ> Sup \<Longrightarrow> mono f"
unfolding mono_def le_fun_def comp_def by (metis complete_lattice_class.SUP_le_iff Sup_atMost atMost_iff)
lemma supdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "(Sup \<circ> (`) f \<le> f \<circ> Sup) = mono f"
using Sup_supdistl_iso iso_Sup_supdistl by force
lemma subdistl_iso:
fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
shows "(f \<circ> Inf \<le> Inf \<circ> (`) f) = mono f"
using Inf_subdistl_iso iso_Inf_subdistl by force
end
lemma fSup_distr: "Sup_pres (\<lambda>x. x \<circ> f)"
unfolding fun_eq_iff comp_def
by (smt Inf.INF_cong SUP_apply Sup_apply)
lemma fSup_distr_var: "\<Squnion>F \<circ> g = (\<Squnion>f \<in> F. f \<circ> g)"
unfolding fun_eq_iff comp_def
by (smt Inf.INF_cong SUP_apply Sup_apply)
lemma fInf_distr: "Inf_pres (\<lambda>x. x \<circ> f)"
unfolding fun_eq_iff comp_def
by (smt INF_apply Inf.INF_cong Inf_apply)
lemma fInf_distr_var: "\<Sqinter>F \<circ> g = (\<Sqinter>f \<in> F. f \<circ> g)"
unfolding fun_eq_iff comp_def
by (smt INF_apply Inf.INF_cong Inf_apply)
lemma fSup_subdistl:
assumes "mono (f::'a::complete_lattice \<Rightarrow> 'b::complete_lattice)"
shows "Sup \<circ> (`) ((\<circ>) f) \<le> (\<circ>) f \<circ> Sup"
using assms by (simp add: SUP_least Sup_upper le_fun_def monoD image_comp)
lemma fSup_subdistl_var:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> (\<Squnion>g \<in> G. f \<circ> g) \<le> f \<circ> \<Squnion>G"
by (simp add: SUP_least Sup_upper le_fun_def monoD image_comp)
lemma fInf_subdistl:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> (\<circ>) f \<circ> Inf \<le> Inf \<circ> (`) ((\<circ>) f)"
by (simp add: INF_greatest Inf_lower le_fun_def monoD image_comp)
lemma fInf_subdistl_var:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> f \<circ> \<Sqinter>G \<le> (\<Sqinter>g \<in> G. f \<circ> g)"
by (simp add: INF_greatest Inf_lower le_fun_def monoD image_comp)
lemma Inf_pres_downset: "Inf_pres (\<down>::'a::complete_lattice \<Rightarrow> 'a set)"
unfolding downset_prop fun_eq_iff comp_def
by (safe, simp_all add: le_Inf_iff)
lemma Sup_dual_upset: "Sup_dual (\<up>::'a::complete_lattice \<Rightarrow> 'a set)"
unfolding upset_prop fun_eq_iff comp_def
by (safe, simp_all add: Sup_le_iff)
text \<open>This approach could probably be combined with the explicit functor-based one. This may be good for proofs, but seems conceptually rather ugly.\<close>
end
\ No newline at end of file
diff --git a/thys/Ordered_Resolution_Prover/Abstract_Substitution.thy b/thys/Ordered_Resolution_Prover/Abstract_Substitution.thy
--- a/thys/Ordered_Resolution_Prover/Abstract_Substitution.thy
+++ b/thys/Ordered_Resolution_Prover/Abstract_Substitution.thy
@@ -1,1270 +1,1271 @@
(* Title: Abstract Substitutions
Author: Dmitriy Traytel <traytel at inf.ethz.ch>, 2014
Author: Jasmin Blanchette <j.c.blanchette at vu.nl>, 2014, 2017
Author: Anders Schlichtkrull <andschl at dtu.dk>, 2016, 2017
Maintainer: Anders Schlichtkrull <andschl at dtu.dk>
*)
section \<open>Abstract Substitutions\<close>
theory Abstract_Substitution
imports Clausal_Logic Map2
begin
text \<open>
Atoms and substitutions are abstracted away behind some locales, to avoid having a direct dependency
on the IsaFoR library.
Conventions: \<open>'s\<close> substitutions, \<open>'a\<close> atoms.
\<close>
subsection \<open>Library\<close>
lemma f_Suc_decr_eventually_const:
fixes f :: "nat \<Rightarrow> nat"
assumes leq: "\<forall>i. f (Suc i) \<le> f i"
shows "\<exists>l. \<forall>l' \<ge> l. f l' = f (Suc l')"
proof (rule ccontr)
assume a: "\<nexists>l. \<forall>l' \<ge> l. f l' = f (Suc l')"
have "\<forall>i. \<exists>i'. i' > i \<and> f i' < f i"
proof
fix i
from a have "\<exists>l' \<ge> i. f l' \<noteq> f (Suc l')"
by auto
then obtain l' where
l'_p: "l' \<ge> i \<and> f l' \<noteq> f (Suc l')"
by metis
then have "f l' > f (Suc l')"
using leq le_eq_less_or_eq by auto
moreover have "f i \<ge> f l'"
using leq l'_p by (induction l' arbitrary: i) (blast intro: lift_Suc_antimono_le)+
ultimately show "\<exists>i' > i. f i' < f i"
using l'_p less_le_trans by blast
qed
then obtain g_sm :: "nat \<Rightarrow> nat" where
g_sm_p: "\<forall>i. g_sm i > i \<and> f (g_sm i) < f i"
by metis
define c :: "nat \<Rightarrow> nat" where
"\<And>n. c n = (g_sm ^^ n) 0"
have "f (c i) > f (c (Suc i))" for i
by (induction i) (auto simp: c_def g_sm_p)
then have "\<forall>i. (f \<circ> c) i > (f \<circ> c) (Suc i)"
by auto
then have "\<exists>fc :: nat \<Rightarrow> nat. \<forall>i. fc i > fc (Suc i)"
by metis
then show False
using wf_less_than by (simp add: wf_iff_no_infinite_down_chain)
qed
subsection \<open>Substitution Operators\<close>
locale substitution_ops =
fixes
subst_atm :: "'a \<Rightarrow> 's \<Rightarrow> 'a" and
id_subst :: 's and
comp_subst :: "'s \<Rightarrow> 's \<Rightarrow> 's"
begin
abbreviation subst_atm_abbrev :: "'a \<Rightarrow> 's \<Rightarrow> 'a" (infixl "\<cdot>a" 67) where
"subst_atm_abbrev \<equiv> subst_atm"
abbreviation comp_subst_abbrev :: "'s \<Rightarrow> 's \<Rightarrow> 's" (infixl "\<odot>" 67) where
"comp_subst_abbrev \<equiv> comp_subst"
definition comp_substs :: "'s list \<Rightarrow> 's list \<Rightarrow> 's list" (infixl "\<odot>s" 67) where
"\<sigma>s \<odot>s \<tau>s = map2 comp_subst \<sigma>s \<tau>s"
definition subst_atms :: "'a set \<Rightarrow> 's \<Rightarrow> 'a set" (infixl "\<cdot>as" 67) where
"AA \<cdot>as \<sigma> = (\<lambda>A. A \<cdot>a \<sigma>) ` AA"
definition subst_atmss :: "'a set set \<Rightarrow> 's \<Rightarrow> 'a set set" (infixl "\<cdot>ass" 67) where
"AAA \<cdot>ass \<sigma> = (\<lambda>AA. AA \<cdot>as \<sigma>) ` AAA"
definition subst_atm_list :: "'a list \<Rightarrow> 's \<Rightarrow> 'a list" (infixl "\<cdot>al" 67) where
"As \<cdot>al \<sigma> = map (\<lambda>A. A \<cdot>a \<sigma>) As"
definition subst_atm_mset :: "'a multiset \<Rightarrow> 's \<Rightarrow> 'a multiset" (infixl "\<cdot>am" 67) where
"AA \<cdot>am \<sigma> = image_mset (\<lambda>A. A \<cdot>a \<sigma>) AA"
definition
subst_atm_mset_list :: "'a multiset list \<Rightarrow> 's \<Rightarrow> 'a multiset list" (infixl "\<cdot>aml" 67)
where
"AAA \<cdot>aml \<sigma> = map (\<lambda>AA. AA \<cdot>am \<sigma>) AAA"
definition
subst_atm_mset_lists :: "'a multiset list \<Rightarrow> 's list \<Rightarrow> 'a multiset list" (infixl "\<cdot>\<cdot>aml" 67)
where
"AAs \<cdot>\<cdot>aml \<sigma>s = map2 (\<cdot>am) AAs \<sigma>s"
definition subst_lit :: "'a literal \<Rightarrow> 's \<Rightarrow> 'a literal" (infixl "\<cdot>l" 67) where
"L \<cdot>l \<sigma> = map_literal (\<lambda>A. A \<cdot>a \<sigma>) L"
lemma atm_of_subst_lit[simp]: "atm_of (L \<cdot>l \<sigma>) = atm_of L \<cdot>a \<sigma>"
unfolding subst_lit_def by (cases L) simp+
definition subst_cls :: "'a clause \<Rightarrow> 's \<Rightarrow> 'a clause" (infixl "\<cdot>" 67) where
"AA \<cdot> \<sigma> = image_mset (\<lambda>A. A \<cdot>l \<sigma>) AA"
definition subst_clss :: "'a clause set \<Rightarrow> 's \<Rightarrow> 'a clause set" (infixl "\<cdot>cs" 67) where
"AA \<cdot>cs \<sigma> = (\<lambda>A. A \<cdot> \<sigma>) ` AA"
definition subst_cls_list :: "'a clause list \<Rightarrow> 's \<Rightarrow> 'a clause list" (infixl "\<cdot>cl" 67) where
"Cs \<cdot>cl \<sigma> = map (\<lambda>A. A \<cdot> \<sigma>) Cs"
definition subst_cls_lists :: "'a clause list \<Rightarrow> 's list \<Rightarrow> 'a clause list" (infixl "\<cdot>\<cdot>cl" 67) where
"Cs \<cdot>\<cdot>cl \<sigma>s = map2 (\<cdot>) Cs \<sigma>s"
definition subst_cls_mset :: "'a clause multiset \<Rightarrow> 's \<Rightarrow> 'a clause multiset" (infixl "\<cdot>cm" 67) where
"CC \<cdot>cm \<sigma> = image_mset (\<lambda>A. A \<cdot> \<sigma>) CC"
lemma subst_cls_add_mset[simp]: "add_mset L C \<cdot> \<sigma> = add_mset (L \<cdot>l \<sigma>) (C \<cdot> \<sigma>)"
unfolding subst_cls_def by simp
lemma subst_cls_mset_add_mset[simp]: "add_mset C CC \<cdot>cm \<sigma> = add_mset (C \<cdot> \<sigma>) (CC \<cdot>cm \<sigma>)"
unfolding subst_cls_mset_def by simp
definition generalizes_atm :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
"generalizes_atm A B \<longleftrightarrow> (\<exists>\<sigma>. A \<cdot>a \<sigma> = B)"
definition strictly_generalizes_atm :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
"strictly_generalizes_atm A B \<longleftrightarrow> generalizes_atm A B \<and> \<not> generalizes_atm B A"
definition generalizes_lit :: "'a literal \<Rightarrow> 'a literal \<Rightarrow> bool" where
"generalizes_lit L M \<longleftrightarrow> (\<exists>\<sigma>. L \<cdot>l \<sigma> = M)"
definition strictly_generalizes_lit :: "'a literal \<Rightarrow> 'a literal \<Rightarrow> bool" where
"strictly_generalizes_lit L M \<longleftrightarrow> generalizes_lit L M \<and> \<not> generalizes_lit M L"
definition generalizes :: "'a clause \<Rightarrow> 'a clause \<Rightarrow> bool" where
"generalizes C D \<longleftrightarrow> (\<exists>\<sigma>. C \<cdot> \<sigma> = D)"
definition strictly_generalizes :: "'a clause \<Rightarrow> 'a clause \<Rightarrow> bool" where
"strictly_generalizes C D \<longleftrightarrow> generalizes C D \<and> \<not> generalizes D C"
definition subsumes :: "'a clause \<Rightarrow> 'a clause \<Rightarrow> bool" where
"subsumes C D \<longleftrightarrow> (\<exists>\<sigma>. C \<cdot> \<sigma> \<subseteq># D)"
definition strictly_subsumes :: "'a clause \<Rightarrow> 'a clause \<Rightarrow> bool" where
"strictly_subsumes C D \<longleftrightarrow> subsumes C D \<and> \<not> subsumes D C"
definition variants :: "'a clause \<Rightarrow> 'a clause \<Rightarrow> bool" where
"variants C D \<longleftrightarrow> generalizes C D \<and> generalizes D C"
definition is_renaming :: "'s \<Rightarrow> bool" where
"is_renaming \<sigma> \<longleftrightarrow> (\<exists>\<tau>. \<sigma> \<odot> \<tau> = id_subst)"
definition is_renaming_list :: "'s list \<Rightarrow> bool" where
"is_renaming_list \<sigma>s \<longleftrightarrow> (\<forall>\<sigma> \<in> set \<sigma>s. is_renaming \<sigma>)"
definition inv_renaming :: "'s \<Rightarrow> 's" where
"inv_renaming \<sigma> = (SOME \<tau>. \<sigma> \<odot> \<tau> = id_subst)"
definition is_ground_atm :: "'a \<Rightarrow> bool" where
"is_ground_atm A \<longleftrightarrow> (\<forall>\<sigma>. A = A \<cdot>a \<sigma>)"
definition is_ground_atms :: "'a set \<Rightarrow> bool" where
"is_ground_atms AA = (\<forall>A \<in> AA. is_ground_atm A)"
definition is_ground_atm_list :: "'a list \<Rightarrow> bool" where
"is_ground_atm_list As \<longleftrightarrow> (\<forall>A \<in> set As. is_ground_atm A)"
definition is_ground_atm_mset :: "'a multiset \<Rightarrow> bool" where
"is_ground_atm_mset AA \<longleftrightarrow> (\<forall>A. A \<in># AA \<longrightarrow> is_ground_atm A)"
definition is_ground_lit :: "'a literal \<Rightarrow> bool" where
"is_ground_lit L \<longleftrightarrow> is_ground_atm (atm_of L)"
definition is_ground_cls :: "'a clause \<Rightarrow> bool" where
"is_ground_cls C \<longleftrightarrow> (\<forall>L. L \<in># C \<longrightarrow> is_ground_lit L)"
definition is_ground_clss :: "'a clause set \<Rightarrow> bool" where
"is_ground_clss CC \<longleftrightarrow> (\<forall>C \<in> CC. is_ground_cls C)"
definition is_ground_cls_list :: "'a clause list \<Rightarrow> bool" where
"is_ground_cls_list CC \<longleftrightarrow> (\<forall>C \<in> set CC. is_ground_cls C)"
definition is_ground_subst :: "'s \<Rightarrow> bool" where
"is_ground_subst \<sigma> \<longleftrightarrow> (\<forall>A. is_ground_atm (A \<cdot>a \<sigma>))"
definition is_ground_subst_list :: "'s list \<Rightarrow> bool" where
"is_ground_subst_list \<sigma>s \<longleftrightarrow> (\<forall>\<sigma> \<in> set \<sigma>s. is_ground_subst \<sigma>)"
definition grounding_of_cls :: "'a clause \<Rightarrow> 'a clause set" where
"grounding_of_cls C = {C \<cdot> \<sigma> |\<sigma>. is_ground_subst \<sigma>}"
definition grounding_of_clss :: "'a clause set \<Rightarrow> 'a clause set" where
"grounding_of_clss CC = (\<Union>C \<in> CC. grounding_of_cls C)"
definition is_unifier :: "'s \<Rightarrow> 'a set \<Rightarrow> bool" where
"is_unifier \<sigma> AA \<longleftrightarrow> card (AA \<cdot>as \<sigma>) \<le> 1"
definition is_unifiers :: "'s \<Rightarrow> 'a set set \<Rightarrow> bool" where
"is_unifiers \<sigma> AAA \<longleftrightarrow> (\<forall>AA \<in> AAA. is_unifier \<sigma> AA)"
definition is_mgu :: "'s \<Rightarrow> 'a set set \<Rightarrow> bool" where
"is_mgu \<sigma> AAA \<longleftrightarrow> is_unifiers \<sigma> AAA \<and> (\<forall>\<tau>. is_unifiers \<tau> AAA \<longrightarrow> (\<exists>\<gamma>. \<tau> = \<sigma> \<odot> \<gamma>))"
definition var_disjoint :: "'a clause list \<Rightarrow> bool" where
"var_disjoint Cs \<longleftrightarrow>
(\<forall>\<sigma>s. length \<sigma>s = length Cs \<longrightarrow> (\<exists>\<tau>. \<forall>i < length Cs. \<forall>S. S \<subseteq># Cs ! i \<longrightarrow> S \<cdot> \<sigma>s ! i = S \<cdot> \<tau>))"
end
subsection \<open>Substitution Lemmas\<close>
locale substitution = substitution_ops subst_atm id_subst comp_subst
for
subst_atm :: "'a \<Rightarrow> 's \<Rightarrow> 'a" and
id_subst :: 's and
comp_subst :: "'s \<Rightarrow> 's \<Rightarrow> 's" +
fixes
renamings_apart :: "'a clause list \<Rightarrow> 's list" and
atm_of_atms :: "'a list \<Rightarrow> 'a"
assumes
subst_atm_id_subst[simp]: "A \<cdot>a id_subst = A" and
subst_atm_comp_subst[simp]: "A \<cdot>a (\<sigma> \<odot> \<tau>) = (A \<cdot>a \<sigma>) \<cdot>a \<tau>" and
subst_ext: "(\<And>A. A \<cdot>a \<sigma> = A \<cdot>a \<tau>) \<Longrightarrow> \<sigma> = \<tau>" and
make_ground_subst: "is_ground_cls (C \<cdot> \<sigma>) \<Longrightarrow> \<exists>\<tau>. is_ground_subst \<tau> \<and>C \<cdot> \<tau> = C \<cdot> \<sigma>" and
wf_strictly_generalizes_atm: "wfP strictly_generalizes_atm" and
renamings_apart_length: "length (renamings_apart Cs) = length Cs" and
renamings_apart_renaming: "\<rho> \<in> set (renamings_apart Cs) \<Longrightarrow> is_renaming \<rho>" and
renamings_apart_var_disjoint: "var_disjoint (Cs \<cdot>\<cdot>cl (renamings_apart Cs))" and
atm_of_atms_subst:
"\<And>As Bs. atm_of_atms As \<cdot>a \<sigma> = atm_of_atms Bs \<longleftrightarrow> map (\<lambda>A. A \<cdot>a \<sigma>) As = Bs"
begin
lemma subst_ext_iff: "\<sigma> = \<tau> \<longleftrightarrow> (\<forall>A. A \<cdot>a \<sigma> = A \<cdot>a \<tau>)"
by (blast intro: subst_ext)
subsubsection \<open>Identity Substitution\<close>
lemma id_subst_comp_subst[simp]: "id_subst \<odot> \<sigma> = \<sigma>"
by (rule subst_ext) simp
lemma comp_subst_id_subst[simp]: "\<sigma> \<odot> id_subst = \<sigma>"
by (rule subst_ext) simp
lemma id_subst_comp_substs[simp]: "replicate (length \<sigma>s) id_subst \<odot>s \<sigma>s = \<sigma>s"
using comp_substs_def by (induction \<sigma>s) auto
lemma comp_substs_id_subst[simp]: "\<sigma>s \<odot>s replicate (length \<sigma>s) id_subst = \<sigma>s"
using comp_substs_def by (induction \<sigma>s) auto
lemma subst_atms_id_subst[simp]: "AA \<cdot>as id_subst = AA"
unfolding subst_atms_def by simp
lemma subst_atmss_id_subst[simp]: "AAA \<cdot>ass id_subst = AAA"
unfolding subst_atmss_def by simp
lemma subst_atm_list_id_subst[simp]: "As \<cdot>al id_subst = As"
unfolding subst_atm_list_def by auto
lemma subst_atm_mset_id_subst[simp]: "AA \<cdot>am id_subst = AA"
unfolding subst_atm_mset_def by simp
lemma subst_atm_mset_list_id_subst[simp]: "AAs \<cdot>aml id_subst = AAs"
unfolding subst_atm_mset_list_def by simp
lemma subst_atm_mset_lists_id_subst[simp]: "AAs \<cdot>\<cdot>aml replicate (length AAs) id_subst = AAs"
unfolding subst_atm_mset_lists_def by (induct AAs) auto
lemma subst_lit_id_subst[simp]: "L \<cdot>l id_subst = L"
unfolding subst_lit_def by (simp add: literal.map_ident)
lemma subst_cls_id_subst[simp]: "C \<cdot> id_subst = C"
unfolding subst_cls_def by simp
lemma subst_clss_id_subst[simp]: "CC \<cdot>cs id_subst = CC"
unfolding subst_clss_def by simp
lemma subst_cls_list_id_subst[simp]: "Cs \<cdot>cl id_subst = Cs"
unfolding subst_cls_list_def by simp
lemma subst_cls_lists_id_subst[simp]: "Cs \<cdot>\<cdot>cl replicate (length Cs) id_subst = Cs"
unfolding subst_cls_lists_def by (induct Cs) auto
lemma subst_cls_mset_id_subst[simp]: "CC \<cdot>cm id_subst = CC"
unfolding subst_cls_mset_def by simp
subsubsection \<open>Associativity of Composition\<close>
lemma comp_subst_assoc[simp]: "\<sigma> \<odot> (\<tau> \<odot> \<gamma>) = \<sigma> \<odot> \<tau> \<odot> \<gamma>"
by (rule subst_ext) simp
subsubsection \<open>Compatibility of Substitution and Composition\<close>
lemma subst_atms_comp_subst[simp]: "AA \<cdot>as (\<tau> \<odot> \<sigma>) = AA \<cdot>as \<tau> \<cdot>as \<sigma>"
unfolding subst_atms_def by auto
lemma subst_atmss_comp_subst[simp]: "AAA \<cdot>ass (\<tau> \<odot> \<sigma>) = AAA \<cdot>ass \<tau> \<cdot>ass \<sigma>"
unfolding subst_atmss_def by auto
lemma subst_atm_list_comp_subst[simp]: "As \<cdot>al (\<tau> \<odot> \<sigma>) = As \<cdot>al \<tau> \<cdot>al \<sigma>"
unfolding subst_atm_list_def by auto
lemma subst_atm_mset_comp_subst[simp]: "AA \<cdot>am (\<tau> \<odot> \<sigma>) = AA \<cdot>am \<tau> \<cdot>am \<sigma>"
unfolding subst_atm_mset_def by auto
lemma subst_atm_mset_list_comp_subst[simp]: "AAs \<cdot>aml (\<tau> \<odot> \<sigma>) = (AAs \<cdot>aml \<tau>) \<cdot>aml \<sigma>"
unfolding subst_atm_mset_list_def by auto
lemma subst_atm_mset_lists_comp_substs[simp]: "AAs \<cdot>\<cdot>aml (\<tau>s \<odot>s \<sigma>s) = AAs \<cdot>\<cdot>aml \<tau>s \<cdot>\<cdot>aml \<sigma>s"
unfolding subst_atm_mset_lists_def comp_substs_def map_zip_map map_zip_map2 map_zip_assoc
by (simp add: split_def)
lemma subst_lit_comp_subst[simp]: "L \<cdot>l (\<tau> \<odot> \<sigma>) = L \<cdot>l \<tau> \<cdot>l \<sigma>"
unfolding subst_lit_def by (auto simp: literal.map_comp o_def)
lemma subst_cls_comp_subst[simp]: "C \<cdot> (\<tau> \<odot> \<sigma>) = C \<cdot> \<tau> \<cdot> \<sigma>"
unfolding subst_cls_def by auto
lemma subst_clsscomp_subst[simp]: "CC \<cdot>cs (\<tau> \<odot> \<sigma>) = CC \<cdot>cs \<tau> \<cdot>cs \<sigma>"
unfolding subst_clss_def by auto
lemma subst_cls_list_comp_subst[simp]: "Cs \<cdot>cl (\<tau> \<odot> \<sigma>) = Cs \<cdot>cl \<tau> \<cdot>cl \<sigma>"
unfolding subst_cls_list_def by auto
lemma subst_cls_lists_comp_substs[simp]: "Cs \<cdot>\<cdot>cl (\<tau>s \<odot>s \<sigma>s) = Cs \<cdot>\<cdot>cl \<tau>s \<cdot>\<cdot>cl \<sigma>s"
unfolding subst_cls_lists_def comp_substs_def map_zip_map map_zip_map2 map_zip_assoc
by (simp add: split_def)
lemma subst_cls_mset_comp_subst[simp]: "CC \<cdot>cm (\<tau> \<odot> \<sigma>) = CC \<cdot>cm \<tau> \<cdot>cm \<sigma>"
unfolding subst_cls_mset_def by auto
subsubsection \<open>``Commutativity'' of Membership and Substitution\<close>
lemma Melem_subst_atm_mset[simp]: "A \<in># AA \<cdot>am \<sigma> \<longleftrightarrow> (\<exists>B. B \<in># AA \<and> A = B \<cdot>a \<sigma>)"
unfolding subst_atm_mset_def by auto
lemma Melem_subst_cls[simp]: "L \<in># C \<cdot> \<sigma> \<longleftrightarrow> (\<exists>M. M \<in># C \<and> L = M \<cdot>l \<sigma>)"
unfolding subst_cls_def by auto
lemma Melem_subst_cls_mset[simp]: "AA \<in># CC \<cdot>cm \<sigma> \<longleftrightarrow> (\<exists>BB. BB \<in># CC \<and> AA = BB \<cdot> \<sigma>)"
unfolding subst_cls_mset_def by auto
subsubsection \<open>Signs and Substitutions\<close>
lemma subst_lit_is_neg[simp]: "is_neg (L \<cdot>l \<sigma>) = is_neg L"
unfolding subst_lit_def by auto
lemma subst_lit_is_pos[simp]: "is_pos (L \<cdot>l \<sigma>) = is_pos L"
unfolding subst_lit_def by auto
lemma subst_minus[simp]: "(- L) \<cdot>l \<mu> = - (L \<cdot>l \<mu>)"
by (simp add: literal.map_sel subst_lit_def uminus_literal_def)
subsubsection \<open>Substitution on Literal(s)\<close>
lemma eql_neg_lit_eql_atm[simp]: "(Neg A' \<cdot>l \<eta>) = Neg A \<longleftrightarrow> A' \<cdot>a \<eta> = A"
by (simp add: subst_lit_def)
lemma eql_pos_lit_eql_atm[simp]: "(Pos A' \<cdot>l \<eta>) = Pos A \<longleftrightarrow> A' \<cdot>a \<eta> = A"
by (simp add: subst_lit_def)
lemma subst_cls_negs[simp]: "(negs AA) \<cdot> \<sigma> = negs (AA \<cdot>am \<sigma>)"
unfolding subst_cls_def subst_lit_def subst_atm_mset_def by auto
lemma subst_cls_poss[simp]: "(poss AA) \<cdot> \<sigma> = poss (AA \<cdot>am \<sigma>)"
unfolding subst_cls_def subst_lit_def subst_atm_mset_def by auto
lemma atms_of_subst_atms: "atms_of C \<cdot>as \<sigma> = atms_of (C \<cdot> \<sigma>)"
proof -
have "atms_of (C \<cdot> \<sigma>) = set_mset (image_mset atm_of (image_mset (map_literal (\<lambda>A. A \<cdot>a \<sigma>)) C))"
unfolding subst_cls_def subst_atms_def subst_lit_def atms_of_def by auto
also have "... = set_mset (image_mset (\<lambda>A. A \<cdot>a \<sigma>) (image_mset atm_of C))"
by simp (meson literal.map_sel)
finally show "atms_of C \<cdot>as \<sigma> = atms_of (C \<cdot> \<sigma>)"
unfolding subst_atms_def atms_of_def by auto
qed
lemma in_image_Neg_is_neg[simp]: "L \<cdot>l \<sigma> \<in> Neg ` AA \<Longrightarrow> is_neg L"
by (metis bex_imageD literal.disc(2) literal.map_disc_iff subst_lit_def)
lemma subst_lit_in_negs_subst_is_neg: "L \<cdot>l \<sigma> \<in># (negs AA) \<cdot> \<tau> \<Longrightarrow> is_neg L"
by simp
lemma subst_lit_in_negs_is_neg: "L \<cdot>l \<sigma> \<in># negs AA \<Longrightarrow> is_neg L"
by simp
subsubsection \<open>Substitution on Empty\<close>
lemma subst_atms_empty[simp]: "{} \<cdot>as \<sigma> = {}"
unfolding subst_atms_def by auto
lemma subst_atmss_empty[simp]: "{} \<cdot>ass \<sigma> = {}"
unfolding subst_atmss_def by auto
lemma comp_substs_empty_iff[simp]: "\<sigma>s \<odot>s \<eta>s = [] \<longleftrightarrow> \<sigma>s = [] \<or> \<eta>s = []"
using comp_substs_def map2_empty_iff by auto
lemma subst_atm_list_empty[simp]: "[] \<cdot>al \<sigma> = []"
unfolding subst_atm_list_def by auto
lemma subst_atm_mset_empty[simp]: "{#} \<cdot>am \<sigma> = {#}"
unfolding subst_atm_mset_def by auto
lemma subst_atm_mset_list_empty[simp]: "[] \<cdot>aml \<sigma> = []"
unfolding subst_atm_mset_list_def by auto
lemma subst_atm_mset_lists_empty[simp]: "[] \<cdot>\<cdot>aml \<sigma>s = []"
unfolding subst_atm_mset_lists_def by auto
lemma subst_cls_empty[simp]: "{#} \<cdot> \<sigma> = {#}"
unfolding subst_cls_def by auto
lemma subst_clss_empty[simp]: "{} \<cdot>cs \<sigma> = {}"
unfolding subst_clss_def by auto
lemma subst_cls_list_empty[simp]: "[] \<cdot>cl \<sigma> = []"
unfolding subst_cls_list_def by auto
lemma subst_cls_lists_empty[simp]: "[] \<cdot>\<cdot>cl \<sigma>s = []"
unfolding subst_cls_lists_def by auto
lemma subst_scls_mset_empty[simp]: "{#} \<cdot>cm \<sigma> = {#}"
unfolding subst_cls_mset_def by auto
lemma subst_atms_empty_iff[simp]: "AA \<cdot>as \<eta> = {} \<longleftrightarrow> AA = {}"
unfolding subst_atms_def by auto
lemma subst_atmss_empty_iff[simp]: "AAA \<cdot>ass \<eta> = {} \<longleftrightarrow> AAA = {}"
unfolding subst_atmss_def by auto
lemma subst_atm_list_empty_iff[simp]: "As \<cdot>al \<eta> = [] \<longleftrightarrow> As = []"
unfolding subst_atm_list_def by auto
lemma subst_atm_mset_empty_iff[simp]: "AA \<cdot>am \<eta> = {#} \<longleftrightarrow> AA = {#}"
unfolding subst_atm_mset_def by auto
lemma subst_atm_mset_list_empty_iff[simp]: "AAs \<cdot>aml \<eta> = [] \<longleftrightarrow> AAs = []"
unfolding subst_atm_mset_list_def by auto
lemma subst_atm_mset_lists_empty_iff[simp]: "AAs \<cdot>\<cdot>aml \<eta>s = [] \<longleftrightarrow> (AAs = [] \<or> \<eta>s = [])"
using map2_empty_iff subst_atm_mset_lists_def by auto
lemma subst_cls_empty_iff[simp]: "C \<cdot> \<eta> = {#} \<longleftrightarrow> C = {#}"
unfolding subst_cls_def by auto
lemma subst_clss_empty_iff[simp]: "CC \<cdot>cs \<eta> = {} \<longleftrightarrow> CC = {}"
unfolding subst_clss_def by auto
lemma subst_cls_list_empty_iff[simp]: "Cs \<cdot>cl \<eta> = [] \<longleftrightarrow> Cs = []"
unfolding subst_cls_list_def by auto
lemma subst_cls_lists_empty_iff[simp]: "Cs \<cdot>\<cdot>cl \<eta>s = [] \<longleftrightarrow> Cs = [] \<or> \<eta>s = []"
using map2_empty_iff subst_cls_lists_def by auto
lemma subst_cls_mset_empty_iff[simp]: "CC \<cdot>cm \<eta> = {#} \<longleftrightarrow> CC = {#}"
unfolding subst_cls_mset_def by auto
subsubsection \<open>Substitution on a Union\<close>
lemma subst_atms_union[simp]: "(AA \<union> BB) \<cdot>as \<sigma> = AA \<cdot>as \<sigma> \<union> BB \<cdot>as \<sigma>"
unfolding subst_atms_def by auto
lemma subst_atmss_union[simp]: "(AAA \<union> BBB) \<cdot>ass \<sigma> = AAA \<cdot>ass \<sigma> \<union> BBB \<cdot>ass \<sigma>"
unfolding subst_atmss_def by auto
lemma subst_atm_list_append[simp]: "(As @ Bs) \<cdot>al \<sigma> = As \<cdot>al \<sigma> @ Bs \<cdot>al \<sigma>"
unfolding subst_atm_list_def by auto
lemma subst_atm_mset_union[simp]: "(AA + BB) \<cdot>am \<sigma> = AA \<cdot>am \<sigma> + BB \<cdot>am \<sigma>"
unfolding subst_atm_mset_def by auto
lemma subst_atm_mset_list_append[simp]: "(AAs @ BBs) \<cdot>aml \<sigma> = AAs \<cdot>aml \<sigma> @ BBs \<cdot>aml \<sigma>"
unfolding subst_atm_mset_list_def by auto
lemma subst_cls_union[simp]: "(C + D) \<cdot> \<sigma> = C \<cdot> \<sigma> + D \<cdot> \<sigma>"
unfolding subst_cls_def by auto
lemma subst_clss_union[simp]: "(CC \<union> DD) \<cdot>cs \<sigma> = CC \<cdot>cs \<sigma> \<union> DD \<cdot>cs \<sigma>"
unfolding subst_clss_def by auto
lemma subst_cls_list_append[simp]: "(Cs @ Ds) \<cdot>cl \<sigma> = Cs \<cdot>cl \<sigma> @ Ds \<cdot>cl \<sigma>"
unfolding subst_cls_list_def by auto
lemma subst_cls_lists_append[simp]:
"length Cs = length \<sigma>s \<Longrightarrow> length Cs' = length \<sigma>s' \<Longrightarrow>
(Cs @ Cs') \<cdot>\<cdot>cl (\<sigma>s @ \<sigma>s') = Cs \<cdot>\<cdot>cl \<sigma>s @ Cs' \<cdot>\<cdot>cl \<sigma>s'"
unfolding subst_cls_lists_def by auto
lemma subst_cls_mset_union[simp]: "(CC + DD) \<cdot>cm \<sigma> = CC \<cdot>cm \<sigma> + DD \<cdot>cm \<sigma>"
unfolding subst_cls_mset_def by auto
subsubsection \<open>Substitution on a Singleton\<close>
lemma subst_atms_single[simp]: "{A} \<cdot>as \<sigma> = {A \<cdot>a \<sigma>}"
unfolding subst_atms_def by auto
lemma subst_atmss_single[simp]: "{AA} \<cdot>ass \<sigma> = {AA \<cdot>as \<sigma>}"
unfolding subst_atmss_def by auto
lemma subst_atm_list_single[simp]: "[A] \<cdot>al \<sigma> = [A \<cdot>a \<sigma>]"
unfolding subst_atm_list_def by auto
lemma subst_atm_mset_single[simp]: "{#A#} \<cdot>am \<sigma> = {#A \<cdot>a \<sigma>#}"
unfolding subst_atm_mset_def by auto
lemma subst_atm_mset_list[simp]: "[AA] \<cdot>aml \<sigma> = [AA \<cdot>am \<sigma>]"
unfolding subst_atm_mset_list_def by auto
lemma subst_cls_single[simp]: "{#L#} \<cdot> \<sigma> = {#L \<cdot>l \<sigma>#}"
by simp
lemma subst_clss_single[simp]: "{C} \<cdot>cs \<sigma> = {C \<cdot> \<sigma>}"
unfolding subst_clss_def by auto
lemma subst_cls_list_single[simp]: "[C] \<cdot>cl \<sigma> = [C \<cdot> \<sigma>]"
unfolding subst_cls_list_def by auto
lemma subst_cls_lists_single[simp]: "[C] \<cdot>\<cdot>cl [\<sigma>] = [C \<cdot> \<sigma>]"
unfolding subst_cls_lists_def by auto
lemma subst_cls_mset_single[simp]: "{#C#} \<cdot>cm \<sigma> = {#C \<cdot> \<sigma>#}"
by simp
subsubsection \<open>Substitution on @{term Cons}\<close>
lemma subst_atm_list_Cons[simp]: "(A # As) \<cdot>al \<sigma> = A \<cdot>a \<sigma> # As \<cdot>al \<sigma>"
unfolding subst_atm_list_def by auto
lemma subst_atm_mset_list_Cons[simp]: "(A # As) \<cdot>aml \<sigma> = A \<cdot>am \<sigma> # As \<cdot>aml \<sigma>"
unfolding subst_atm_mset_list_def by auto
lemma subst_atm_mset_lists_Cons[simp]: "(C # Cs) \<cdot>\<cdot>aml (\<sigma> # \<sigma>s) = C \<cdot>am \<sigma> # Cs \<cdot>\<cdot>aml \<sigma>s"
unfolding subst_atm_mset_lists_def by auto
lemma subst_cls_list_Cons[simp]: "(C # Cs) \<cdot>cl \<sigma> = C \<cdot> \<sigma> # Cs \<cdot>cl \<sigma>"
unfolding subst_cls_list_def by auto
lemma subst_cls_lists_Cons[simp]: "(C # Cs) \<cdot>\<cdot>cl (\<sigma> # \<sigma>s) = C \<cdot> \<sigma> # Cs \<cdot>\<cdot>cl \<sigma>s"
unfolding subst_cls_lists_def by auto
subsubsection \<open>Substitution on @{term tl}\<close>
lemma subst_atm_list_tl[simp]: "tl (As \<cdot>al \<sigma>) = tl As \<cdot>al \<sigma>"
by (cases As) auto
lemma subst_atm_mset_list_tl[simp]: "tl (AAs \<cdot>aml \<sigma>) = tl AAs \<cdot>aml \<sigma>"
by (cases AAs) auto
lemma subst_cls_list_tl[simp]: "tl (Cs \<cdot>cl \<sigma>) = tl Cs \<cdot>cl \<sigma>"
by (cases Cs) auto
lemma subst_cls_lists_tl[simp]: "length Cs = length \<sigma>s \<Longrightarrow> tl (Cs \<cdot>\<cdot>cl \<sigma>s) = tl Cs \<cdot>\<cdot>cl tl \<sigma>s"
by (cases Cs; cases \<sigma>s) auto
subsubsection \<open>Substitution on @{term nth}\<close>
lemma comp_substs_nth[simp]:
"length \<tau>s = length \<sigma>s \<Longrightarrow> i < length \<tau>s \<Longrightarrow> (\<tau>s \<odot>s \<sigma>s) ! i = (\<tau>s ! i) \<odot> (\<sigma>s ! i)"
by (simp add: comp_substs_def)
lemma subst_atm_list_nth[simp]: "i < length As \<Longrightarrow> (As \<cdot>al \<tau>) ! i = As ! i \<cdot>a \<tau>"
unfolding subst_atm_list_def using less_Suc_eq_0_disj nth_map by force
lemma subst_atm_mset_list_nth[simp]: "i < length AAs \<Longrightarrow> (AAs \<cdot>aml \<eta>) ! i = (AAs ! i) \<cdot>am \<eta>"
unfolding subst_atm_mset_list_def by auto
lemma subst_atm_mset_lists_nth[simp]:
"length AAs = length \<sigma>s \<Longrightarrow> i < length AAs \<Longrightarrow> (AAs \<cdot>\<cdot>aml \<sigma>s) ! i = (AAs ! i) \<cdot>am (\<sigma>s ! i)"
unfolding subst_atm_mset_lists_def by auto
lemma subst_cls_list_nth[simp]: "i < length Cs \<Longrightarrow> (Cs \<cdot>cl \<tau>) ! i = (Cs ! i) \<cdot> \<tau>"
unfolding subst_cls_list_def using less_Suc_eq_0_disj nth_map by (induction Cs) auto
lemma subst_cls_lists_nth[simp]:
"length Cs = length \<sigma>s \<Longrightarrow> i < length Cs \<Longrightarrow> (Cs \<cdot>\<cdot>cl \<sigma>s) ! i = (Cs ! i) \<cdot> (\<sigma>s ! i)"
unfolding subst_cls_lists_def by auto
subsubsection \<open>Substitution on Various Other Functions\<close>
lemma subst_clss_image[simp]: "image f X \<cdot>cs \<sigma> = {f x \<cdot> \<sigma> | x. x \<in> X}"
unfolding subst_clss_def by auto
lemma subst_cls_mset_image_mset[simp]: "image_mset f X \<cdot>cm \<sigma> = {# f x \<cdot> \<sigma>. x \<in># X #}"
unfolding subst_cls_mset_def by auto
lemma mset_subst_atm_list_subst_atm_mset[simp]: "mset (As \<cdot>al \<sigma>) = mset (As) \<cdot>am \<sigma>"
unfolding subst_atm_list_def subst_atm_mset_def by auto
lemma mset_subst_cls_list_subst_cls_mset: "mset (Cs \<cdot>cl \<sigma>) = (mset Cs) \<cdot>cm \<sigma>"
unfolding subst_cls_mset_def subst_cls_list_def by auto
lemma sum_list_subst_cls_list_subst_cls[simp]: "sum_list (Cs \<cdot>cl \<eta>) = sum_list Cs \<cdot> \<eta>"
unfolding subst_cls_list_def by (induction Cs) auto
lemma set_mset_subst_cls_mset_subst_clss: "set_mset (CC \<cdot>cm \<mu>) = (set_mset CC) \<cdot>cs \<mu>"
by (simp add: subst_cls_mset_def subst_clss_def)
lemma Neg_Melem_subst_atm_subst_cls[simp]: "Neg A \<in># C \<Longrightarrow> Neg (A \<cdot>a \<sigma>) \<in># C \<cdot> \<sigma> "
by (metis Melem_subst_cls eql_neg_lit_eql_atm)
lemma Pos_Melem_subst_atm_subst_cls[simp]: "Pos A \<in># C \<Longrightarrow> Pos (A \<cdot>a \<sigma>) \<in># C \<cdot> \<sigma> "
by (metis Melem_subst_cls eql_pos_lit_eql_atm)
lemma in_atms_of_subst[simp]: "B \<in> atms_of C \<Longrightarrow> B \<cdot>a \<sigma> \<in> atms_of (C \<cdot> \<sigma>)"
by (metis atms_of_subst_atms image_iff subst_atms_def)
subsubsection \<open>Renamings\<close>
lemma is_renaming_id_subst[simp]: "is_renaming id_subst"
unfolding is_renaming_def by simp
lemma is_renamingD: "is_renaming \<sigma> \<Longrightarrow> (\<forall>A1 A2. A1 \<cdot>a \<sigma> = A2 \<cdot>a \<sigma> \<longleftrightarrow> A1 = A2)"
by (metis is_renaming_def subst_atm_comp_subst subst_atm_id_subst)
lemma inv_renaming_cancel_r[simp]: "is_renaming r \<Longrightarrow> r \<odot> inv_renaming r = id_subst"
unfolding inv_renaming_def is_renaming_def by (metis (mono_tags) someI_ex)
lemma inv_renaming_cancel_r_list[simp]:
"is_renaming_list rs \<Longrightarrow> rs \<odot>s map inv_renaming rs = replicate (length rs) id_subst"
unfolding is_renaming_list_def by (induction rs) (auto simp add: comp_substs_def)
lemma Nil_comp_substs[simp]: "[] \<odot>s s = []"
unfolding comp_substs_def by auto
lemma comp_substs_Nil[simp]: "s \<odot>s [] = []"
unfolding comp_substs_def by auto
lemma is_renaming_idempotent_id_subst: "is_renaming r \<Longrightarrow> r \<odot> r = r \<Longrightarrow> r = id_subst"
by (metis comp_subst_assoc comp_subst_id_subst inv_renaming_cancel_r)
lemma is_renaming_left_id_subst_right_id_subst:
"is_renaming r \<Longrightarrow> s \<odot> r = id_subst \<Longrightarrow> r \<odot> s = id_subst"
by (metis comp_subst_assoc comp_subst_id_subst is_renaming_def)
lemma is_renaming_closure: "is_renaming r1 \<Longrightarrow> is_renaming r2 \<Longrightarrow> is_renaming (r1 \<odot> r2)"
unfolding is_renaming_def by (metis comp_subst_assoc comp_subst_id_subst)
lemma is_renaming_inv_renaming_cancel_atm[simp]: "is_renaming \<rho> \<Longrightarrow> A \<cdot>a \<rho> \<cdot>a inv_renaming \<rho> = A"
by (metis inv_renaming_cancel_r subst_atm_comp_subst subst_atm_id_subst)
lemma is_renaming_inv_renaming_cancel_atms[simp]: "is_renaming \<rho> \<Longrightarrow> AA \<cdot>as \<rho> \<cdot>as inv_renaming \<rho> = AA"
by (metis inv_renaming_cancel_r subst_atms_comp_subst subst_atms_id_subst)
lemma is_renaming_inv_renaming_cancel_atmss[simp]: "is_renaming \<rho> \<Longrightarrow> AAA \<cdot>ass \<rho> \<cdot>ass inv_renaming \<rho> = AAA"
by (metis inv_renaming_cancel_r subst_atmss_comp_subst subst_atmss_id_subst)
lemma is_renaming_inv_renaming_cancel_atm_list[simp]: "is_renaming \<rho> \<Longrightarrow> As \<cdot>al \<rho> \<cdot>al inv_renaming \<rho> = As"
by (metis inv_renaming_cancel_r subst_atm_list_comp_subst subst_atm_list_id_subst)
lemma is_renaming_inv_renaming_cancel_atm_mset[simp]: "is_renaming \<rho> \<Longrightarrow> AA \<cdot>am \<rho> \<cdot>am inv_renaming \<rho> = AA"
by (metis inv_renaming_cancel_r subst_atm_mset_comp_subst subst_atm_mset_id_subst)
lemma is_renaming_inv_renaming_cancel_atm_mset_list[simp]: "is_renaming \<rho> \<Longrightarrow> (AAs \<cdot>aml \<rho>) \<cdot>aml inv_renaming \<rho> = AAs"
by (metis inv_renaming_cancel_r subst_atm_mset_list_comp_subst subst_atm_mset_list_id_subst)
lemma is_renaming_list_inv_renaming_cancel_atm_mset_lists[simp]:
"length AAs = length \<rho>s \<Longrightarrow> is_renaming_list \<rho>s \<Longrightarrow> AAs \<cdot>\<cdot>aml \<rho>s \<cdot>\<cdot>aml map inv_renaming \<rho>s = AAs"
by (metis inv_renaming_cancel_r_list subst_atm_mset_lists_comp_substs
subst_atm_mset_lists_id_subst)
lemma is_renaming_inv_renaming_cancel_lit[simp]: "is_renaming \<rho> \<Longrightarrow> (L \<cdot>l \<rho>) \<cdot>l inv_renaming \<rho> = L"
by (metis inv_renaming_cancel_r subst_lit_comp_subst subst_lit_id_subst)
lemma is_renaming_inv_renaming_cancel_cls[simp]: "is_renaming \<rho> \<Longrightarrow> C \<cdot> \<rho> \<cdot> inv_renaming \<rho> = C"
by (metis inv_renaming_cancel_r subst_cls_comp_subst subst_cls_id_subst)
lemma is_renaming_inv_renaming_cancel_clss[simp]:
"is_renaming \<rho> \<Longrightarrow> CC \<cdot>cs \<rho> \<cdot>cs inv_renaming \<rho> = CC"
by (metis inv_renaming_cancel_r subst_clss_id_subst subst_clsscomp_subst)
lemma is_renaming_inv_renaming_cancel_cls_list[simp]:
"is_renaming \<rho> \<Longrightarrow> Cs \<cdot>cl \<rho> \<cdot>cl inv_renaming \<rho> = Cs"
by (metis inv_renaming_cancel_r subst_cls_list_comp_subst subst_cls_list_id_subst)
lemma is_renaming_list_inv_renaming_cancel_cls_list[simp]:
"length Cs = length \<rho>s \<Longrightarrow> is_renaming_list \<rho>s \<Longrightarrow> Cs \<cdot>\<cdot>cl \<rho>s \<cdot>\<cdot>cl map inv_renaming \<rho>s = Cs"
by (metis inv_renaming_cancel_r_list subst_cls_lists_comp_substs subst_cls_lists_id_subst)
lemma is_renaming_inv_renaming_cancel_cls_mset[simp]:
"is_renaming \<rho> \<Longrightarrow> CC \<cdot>cm \<rho> \<cdot>cm inv_renaming \<rho> = CC"
by (metis inv_renaming_cancel_r subst_cls_mset_comp_subst subst_cls_mset_id_subst)
subsubsection \<open>Monotonicity\<close>
lemma subst_cls_mono: "set_mset C \<subseteq> set_mset D \<Longrightarrow> set_mset (C \<cdot> \<sigma>) \<subseteq> set_mset (D \<cdot> \<sigma>)"
by force
lemma subst_cls_mono_mset: "C \<subseteq># D \<Longrightarrow> C \<cdot> \<sigma> \<subseteq># D \<cdot> \<sigma>"
unfolding subst_clss_def by (metis mset_subset_eq_exists_conv subst_cls_union)
lemma subst_subset_mono: "D \<subset># C \<Longrightarrow> D \<cdot> \<sigma> \<subset># C \<cdot> \<sigma>"
unfolding subst_cls_def by (simp add: image_mset_subset_mono)
subsubsection \<open>Size after Substitution\<close>
lemma size_subst[simp]: "size (D \<cdot> \<sigma>) = size D"
unfolding subst_cls_def by auto
lemma subst_atm_list_length[simp]: "length (As \<cdot>al \<sigma>) = length As"
unfolding subst_atm_list_def by auto
lemma length_subst_atm_mset_list[simp]: "length (AAs \<cdot>aml \<eta>) = length AAs"
unfolding subst_atm_mset_list_def by auto
lemma subst_atm_mset_lists_length[simp]: "length (AAs \<cdot>\<cdot>aml \<sigma>s) = min (length AAs) (length \<sigma>s)"
unfolding subst_atm_mset_lists_def by auto
lemma subst_cls_list_length[simp]: "length (Cs \<cdot>cl \<sigma>) = length Cs"
unfolding subst_cls_list_def by auto
lemma comp_substs_length[simp]: "length (\<tau>s \<odot>s \<sigma>s) = min (length \<tau>s) (length \<sigma>s)"
unfolding comp_substs_def by auto
lemma subst_cls_lists_length[simp]: "length (Cs \<cdot>\<cdot>cl \<sigma>s) = min (length Cs) (length \<sigma>s)"
unfolding subst_cls_lists_def by auto
subsubsection \<open>Variable Disjointness\<close>
lemma var_disjoint_clauses:
assumes "var_disjoint Cs"
shows "\<forall>\<sigma>s. length \<sigma>s = length Cs \<longrightarrow> (\<exists>\<tau>. Cs \<cdot>\<cdot>cl \<sigma>s = Cs \<cdot>cl \<tau>)"
proof clarify
fix \<sigma>s :: "'s list"
assume a: "length \<sigma>s = length Cs"
then obtain \<tau> where "\<forall>i < length Cs. \<forall>S. S \<subseteq># Cs ! i \<longrightarrow> S \<cdot> \<sigma>s ! i = S \<cdot> \<tau>"
using assms unfolding var_disjoint_def by blast
then have "\<forall>i < length Cs. (Cs ! i) \<cdot> \<sigma>s ! i = (Cs ! i) \<cdot> \<tau>"
by auto
then have "Cs \<cdot>\<cdot>cl \<sigma>s = Cs \<cdot>cl \<tau>"
using a by (auto intro: nth_equalityI)
then show "\<exists>\<tau>. Cs \<cdot>\<cdot>cl \<sigma>s = Cs \<cdot>cl \<tau>"
by auto
qed
subsubsection \<open>Ground Expressions and Substitutions\<close>
lemma ex_ground_subst: "\<exists>\<sigma>. is_ground_subst \<sigma>"
using make_ground_subst[of "{#}"]
by (simp add: is_ground_cls_def)
lemma is_ground_cls_list_Cons[simp]:
"is_ground_cls_list (C # Cs) = (is_ground_cls C \<and> is_ground_cls_list Cs)"
unfolding is_ground_cls_list_def by auto
paragraph \<open>Ground union\<close>
lemma is_ground_atms_union[simp]: "is_ground_atms (AA \<union> BB) \<longleftrightarrow> is_ground_atms AA \<and> is_ground_atms BB"
unfolding is_ground_atms_def by auto
lemma is_ground_atm_mset_union[simp]:
"is_ground_atm_mset (AA + BB) \<longleftrightarrow> is_ground_atm_mset AA \<and> is_ground_atm_mset BB"
unfolding is_ground_atm_mset_def by auto
lemma is_ground_cls_union[simp]: "is_ground_cls (C + D) \<longleftrightarrow> is_ground_cls C \<and> is_ground_cls D"
unfolding is_ground_cls_def by auto
lemma is_ground_clss_union[simp]:
"is_ground_clss (CC \<union> DD) \<longleftrightarrow> is_ground_clss CC \<and> is_ground_clss DD"
unfolding is_ground_clss_def by auto
lemma is_ground_cls_list_is_ground_cls_sum_list[simp]:
"is_ground_cls_list Cs \<Longrightarrow> is_ground_cls (sum_list Cs)"
by (meson in_mset_sum_list2 is_ground_cls_def is_ground_cls_list_def)
paragraph \<open>Grounding monotonicity\<close>
lemma is_ground_cls_mono: "C \<subseteq># D \<Longrightarrow> is_ground_cls D \<Longrightarrow> is_ground_cls C"
unfolding is_ground_cls_def by (metis set_mset_mono subsetD)
lemma is_ground_clss_mono: "CC \<subseteq> DD \<Longrightarrow> is_ground_clss DD \<Longrightarrow> is_ground_clss CC"
unfolding is_ground_clss_def by blast
lemma grounding_of_clss_mono: "CC \<subseteq> DD \<Longrightarrow> grounding_of_clss CC \<subseteq> grounding_of_clss DD"
using grounding_of_clss_def by auto
lemma sum_list_subseteq_mset_is_ground_cls_list[simp]:
"sum_list Cs \<subseteq># sum_list Ds \<Longrightarrow> is_ground_cls_list Ds \<Longrightarrow> is_ground_cls_list Cs"
by (meson in_mset_sum_list is_ground_cls_def is_ground_cls_list_is_ground_cls_sum_list
is_ground_cls_mono is_ground_cls_list_def)
paragraph \<open>Substituting on ground expression preserves ground\<close>
lemma is_ground_comp_subst[simp]: "is_ground_subst \<sigma> \<Longrightarrow> is_ground_subst (\<tau> \<odot> \<sigma>)"
unfolding is_ground_subst_def is_ground_atm_def by auto
lemma ground_subst_ground_atm[simp]: "is_ground_subst \<sigma> \<Longrightarrow> is_ground_atm (A \<cdot>a \<sigma>)"
by (simp add: is_ground_subst_def)
lemma ground_subst_ground_lit[simp]: "is_ground_subst \<sigma> \<Longrightarrow> is_ground_lit (L \<cdot>l \<sigma>)"
unfolding is_ground_lit_def subst_lit_def by (cases L) auto
lemma ground_subst_ground_cls[simp]: "is_ground_subst \<sigma> \<Longrightarrow> is_ground_cls (C \<cdot> \<sigma>)"
unfolding is_ground_cls_def by auto
lemma ground_subst_ground_clss[simp]: "is_ground_subst \<sigma> \<Longrightarrow> is_ground_clss (CC \<cdot>cs \<sigma>)"
unfolding is_ground_clss_def subst_clss_def by auto
lemma ground_subst_ground_cls_list[simp]: "is_ground_subst \<sigma> \<Longrightarrow> is_ground_cls_list (Cs \<cdot>cl \<sigma>)"
unfolding is_ground_cls_list_def subst_cls_list_def by auto
lemma ground_subst_ground_cls_lists[simp]:
"\<forall>\<sigma> \<in> set \<sigma>s. is_ground_subst \<sigma> \<Longrightarrow> is_ground_cls_list (Cs \<cdot>\<cdot>cl \<sigma>s)"
unfolding is_ground_cls_list_def subst_cls_lists_def by (auto simp: set_zip)
lemma subst_cls_eq_grounding_of_cls_subset_eq:
assumes "D \<cdot> \<sigma> = C"
shows "grounding_of_cls C \<subseteq> grounding_of_cls D"
proof
fix C\<sigma>'
assume "C\<sigma>' \<in> grounding_of_cls C"
then obtain \<sigma>' where
C\<sigma>': "C \<cdot> \<sigma>' = C\<sigma>'" "is_ground_subst \<sigma>'"
unfolding grounding_of_cls_def by auto
then have "C \<cdot> \<sigma>' = D \<cdot> \<sigma> \<cdot> \<sigma>' \<and> is_ground_subst (\<sigma> \<odot> \<sigma>')"
using assms by auto
then show "C\<sigma>' \<in> grounding_of_cls D"
unfolding grounding_of_cls_def using C\<sigma>'(1) by force
qed
paragraph \<open>Substituting on ground expression has no effect\<close>
lemma is_ground_subst_atm[simp]: "is_ground_atm A \<Longrightarrow> A \<cdot>a \<sigma> = A"
unfolding is_ground_atm_def by simp
lemma is_ground_subst_atms[simp]: "is_ground_atms AA \<Longrightarrow> AA \<cdot>as \<sigma> = AA"
unfolding is_ground_atms_def subst_atms_def image_def by auto
lemma is_ground_subst_atm_mset[simp]: "is_ground_atm_mset AA \<Longrightarrow> AA \<cdot>am \<sigma> = AA"
unfolding is_ground_atm_mset_def subst_atm_mset_def by auto
lemma is_ground_subst_atm_list[simp]: "is_ground_atm_list As \<Longrightarrow> As \<cdot>al \<sigma> = As"
unfolding is_ground_atm_list_def subst_atm_list_def by (auto intro: nth_equalityI)
lemma is_ground_subst_atm_list_member[simp]:
"is_ground_atm_list As \<Longrightarrow> i < length As \<Longrightarrow> As ! i \<cdot>a \<sigma> = As ! i"
unfolding is_ground_atm_list_def by auto
lemma is_ground_subst_lit[simp]: "is_ground_lit L \<Longrightarrow> L \<cdot>l \<sigma> = L"
unfolding is_ground_lit_def subst_lit_def by (cases L) simp_all
lemma is_ground_subst_cls[simp]: "is_ground_cls C \<Longrightarrow> C \<cdot> \<sigma> = C"
unfolding is_ground_cls_def subst_cls_def by simp
lemma is_ground_subst_clss[simp]: "is_ground_clss CC \<Longrightarrow> CC \<cdot>cs \<sigma> = CC"
unfolding is_ground_clss_def subst_clss_def image_def by auto
lemma is_ground_subst_cls_lists[simp]:
assumes "length P = length Cs" and "is_ground_cls_list Cs"
shows "Cs \<cdot>\<cdot>cl P = Cs"
using assms by (metis is_ground_cls_list_def is_ground_subst_cls min.idem nth_equalityI nth_mem
subst_cls_lists_nth subst_cls_lists_length)
lemma is_ground_subst_lit_iff: "is_ground_lit L \<longleftrightarrow> (\<forall>\<sigma>. L = L \<cdot>l \<sigma>)"
using is_ground_atm_def is_ground_lit_def subst_lit_def by (cases L) auto
lemma is_ground_subst_cls_iff: "is_ground_cls C \<longleftrightarrow> (\<forall>\<sigma>. C = C \<cdot> \<sigma>)"
by (metis ex_ground_subst ground_subst_ground_cls is_ground_subst_cls)
paragraph \<open>Members of ground expressions are ground\<close>
lemma is_ground_cls_as_atms: "is_ground_cls C \<longleftrightarrow> (\<forall>A \<in> atms_of C. is_ground_atm A)"
by (auto simp: atms_of_def is_ground_cls_def is_ground_lit_def)
lemma is_ground_cls_imp_is_ground_lit: "L \<in># C \<Longrightarrow> is_ground_cls C \<Longrightarrow> is_ground_lit L"
by (simp add: is_ground_cls_def)
lemma is_ground_cls_imp_is_ground_atm: "A \<in> atms_of C \<Longrightarrow> is_ground_cls C \<Longrightarrow> is_ground_atm A"
by (simp add: is_ground_cls_as_atms)
lemma is_ground_cls_is_ground_atms_atms_of[simp]: "is_ground_cls C \<Longrightarrow> is_ground_atms (atms_of C)"
by (simp add: is_ground_cls_imp_is_ground_atm is_ground_atms_def)
lemma grounding_ground: "C \<in> grounding_of_clss M \<Longrightarrow> is_ground_cls C"
unfolding grounding_of_clss_def grounding_of_cls_def by auto
lemma in_subset_eq_grounding_of_clss_is_ground_cls[simp]:
"C \<in> CC \<Longrightarrow> CC \<subseteq> grounding_of_clss DD \<Longrightarrow> is_ground_cls C"
unfolding grounding_of_clss_def grounding_of_cls_def by auto
lemma is_ground_cls_empty[simp]: "is_ground_cls {#}"
unfolding is_ground_cls_def by simp
lemma grounding_of_cls_ground: "is_ground_cls C \<Longrightarrow> grounding_of_cls C = {C}"
unfolding grounding_of_cls_def by (simp add: ex_ground_subst)
lemma grounding_of_cls_empty[simp]: "grounding_of_cls {#} = {{#}}"
by (simp add: grounding_of_cls_ground)
lemma union_grounding_of_cls_ground: "is_ground_clss (\<Union> (grounding_of_cls ` N))"
by (simp add: grounding_ground grounding_of_clss_def is_ground_clss_def)
paragraph \<open>Grounding idempotence\<close>
lemma grounding_of_grounding_of_cls: "E \<in> grounding_of_cls D \<Longrightarrow> D \<in> grounding_of_cls C \<Longrightarrow> E = D"
using grounding_of_cls_def by auto
subsubsection \<open>Subsumption\<close>
lemma subsumes_empty_left[simp]: "subsumes {#} C"
unfolding subsumes_def subst_cls_def by simp
lemma strictly_subsumes_empty_left[simp]: "strictly_subsumes {#} C \<longleftrightarrow> C \<noteq> {#}"
unfolding strictly_subsumes_def subsumes_def subst_cls_def by simp
subsubsection \<open>Unifiers\<close>
lemma card_le_one_alt: "finite X \<Longrightarrow> card X \<le> 1 \<longleftrightarrow> X = {} \<or> (\<exists>x. X = {x})"
by (induct rule: finite_induct) auto
lemma is_unifier_subst_atm_eqI:
assumes "finite AA"
shows "is_unifier \<sigma> AA \<Longrightarrow> A \<in> AA \<Longrightarrow> B \<in> AA \<Longrightarrow> A \<cdot>a \<sigma> = B \<cdot>a \<sigma>"
unfolding is_unifier_def subst_atms_def card_le_one_alt[OF finite_imageI[OF assms]]
by (metis equals0D imageI insert_iff)
lemma is_unifier_alt:
assumes "finite AA"
shows "is_unifier \<sigma> AA \<longleftrightarrow> (\<forall>A \<in> AA. \<forall>B \<in> AA. A \<cdot>a \<sigma> = B \<cdot>a \<sigma>)"
unfolding is_unifier_def subst_atms_def card_le_one_alt[OF finite_imageI[OF assms(1)]]
by (rule iffI, metis empty_iff insert_iff insert_image, blast)
lemma is_unifiers_subst_atm_eqI:
assumes "finite AA" "is_unifiers \<sigma> AAA" "AA \<in> AAA" "A \<in> AA" "B \<in> AA"
shows "A \<cdot>a \<sigma> = B \<cdot>a \<sigma>"
by (metis assms is_unifiers_def is_unifier_subst_atm_eqI)
theorem is_unifiers_comp:
"is_unifiers \<sigma> (set_mset ` set (map2 add_mset As Bs) \<cdot>ass \<eta>) \<longleftrightarrow>
is_unifiers (\<eta> \<odot> \<sigma>) (set_mset ` set (map2 add_mset As Bs))"
unfolding is_unifiers_def is_unifier_def subst_atmss_def by auto
subsubsection \<open>Most General Unifier\<close>
lemma is_mgu_is_unifiers: "is_mgu \<sigma> AAA \<Longrightarrow> is_unifiers \<sigma> AAA"
using is_mgu_def by blast
lemma is_mgu_is_most_general: "is_mgu \<sigma> AAA \<Longrightarrow> is_unifiers \<tau> AAA \<Longrightarrow> \<exists>\<gamma>. \<tau> = \<sigma> \<odot> \<gamma>"
using is_mgu_def by blast
lemma is_unifiers_is_unifier: "is_unifiers \<sigma> AAA \<Longrightarrow> AA \<in> AAA \<Longrightarrow> is_unifier \<sigma> AA"
using is_unifiers_def by simp
subsubsection \<open>Generalization and Subsumption\<close>
lemma variants_sym: "variants D D' \<longleftrightarrow> variants D' D"
unfolding variants_def by auto
lemma variants_iff_subsumes: "variants C D \<longleftrightarrow> subsumes C D \<and> subsumes D C"
proof
assume "variants C D"
then show "subsumes C D \<and> subsumes D C"
- unfolding variants_def generalizes_def subsumes_def by (metis subset_mset.order.refl)
+ unfolding variants_def generalizes_def subsumes_def
+ by (metis subset_mset.refl)
next
assume sub: "subsumes C D \<and> subsumes D C"
then have "size C = size D"
unfolding subsumes_def by (metis antisym size_mset_mono size_subst)
then show "variants C D"
using sub unfolding subsumes_def variants_def generalizes_def
by (metis leD mset_subset_size size_mset_mono size_subst
- subset_mset.order.not_eq_order_implies_strict)
+ subset_mset.not_eq_order_implies_strict)
qed
lemma wf_strictly_generalizes: "wfP strictly_generalizes"
proof -
{
assume "\<exists>C_at. \<forall>i. strictly_generalizes (C_at (Suc i)) (C_at i)"
then obtain C_at :: "nat \<Rightarrow> 'a clause" where
sg_C: "\<And>i. strictly_generalizes (C_at (Suc i)) (C_at i)"
by blast
define n :: nat where
"n = size (C_at 0)"
have sz_C: "size (C_at i) = n" for i
proof (induct i)
case (Suc i)
then show ?case
using sg_C[of i] unfolding strictly_generalizes_def generalizes_def subst_cls_def
by (metis size_image_mset)
qed (simp add: n_def)
obtain \<sigma>_at :: "nat \<Rightarrow> 's" where
C_\<sigma>: "\<And>i. image_mset (\<lambda>L. L \<cdot>l \<sigma>_at i) (C_at (Suc i)) = C_at i"
using sg_C[unfolded strictly_generalizes_def generalizes_def subst_cls_def] by metis
define Ls_at :: "nat \<Rightarrow> 'a literal list" where
"Ls_at = rec_nat (SOME Ls. mset Ls = C_at 0)
(\<lambda>i Lsi. SOME Ls. mset Ls = C_at (Suc i) \<and> map (\<lambda>L. L \<cdot>l \<sigma>_at i) Ls = Lsi)"
have
Ls_at_0: "Ls_at 0 = (SOME Ls. mset Ls = C_at 0)" and
Ls_at_Suc: "\<And>i. Ls_at (Suc i) =
(SOME Ls. mset Ls = C_at (Suc i) \<and> map (\<lambda>L. L \<cdot>l \<sigma>_at i) Ls = Ls_at i)"
unfolding Ls_at_def by simp+
have mset_Lt_at_0: "mset (Ls_at 0) = C_at 0"
unfolding Ls_at_0 by (rule someI_ex) (metis list_of_mset_exi)
have "mset (Ls_at (Suc i)) = C_at (Suc i) \<and> map (\<lambda>L. L \<cdot>l \<sigma>_at i) (Ls_at (Suc i)) = Ls_at i"
for i
proof (induct i)
case 0
then show ?case
by (simp add: Ls_at_Suc, rule someI_ex,
metis C_\<sigma> image_mset_of_subset_list mset_Lt_at_0)
next
case Suc
then show ?case
by (subst (1 2) Ls_at_Suc) (rule someI_ex, metis C_\<sigma> image_mset_of_subset_list)
qed
note mset_Ls = this[THEN conjunct1] and Ls_\<sigma> = this[THEN conjunct2]
have len_Ls: "\<And>i. length (Ls_at i) = n"
by (metis mset_Ls mset_Lt_at_0 not0_implies_Suc size_mset sz_C)
have is_pos_Ls: "\<And>i j. j < n \<Longrightarrow> is_pos (Ls_at (Suc i) ! j) \<longleftrightarrow> is_pos (Ls_at i ! j)"
using Ls_\<sigma> len_Ls by (metis literal.map_disc_iff nth_map subst_lit_def)
have Ls_\<tau>_strict_lit: "\<And>i \<tau>. map (\<lambda>L. L \<cdot>l \<tau>) (Ls_at i) \<noteq> Ls_at (Suc i)"
by (metis C_\<sigma> mset_Ls Ls_\<sigma> mset_map sg_C generalizes_def strictly_generalizes_def
subst_cls_def)
have Ls_\<tau>_strict_tm:
"map ((\<lambda>t. t \<cdot>a \<tau>) \<circ> atm_of) (Ls_at i) \<noteq> map atm_of (Ls_at (Suc i))" for i \<tau>
proof -
obtain j :: nat where
j_lt: "j < n" and
j_\<tau>: "Ls_at i ! j \<cdot>l \<tau> \<noteq> Ls_at (Suc i) ! j"
using Ls_\<tau>_strict_lit[of \<tau> i] len_Ls
by (metis (no_types, lifting) length_map list_eq_iff_nth_eq nth_map)
have "atm_of (Ls_at i ! j) \<cdot>a \<tau> \<noteq> atm_of (Ls_at (Suc i) ! j)"
using j_\<tau> is_pos_Ls[OF j_lt]
by (metis (mono_guards) literal.expand literal.map_disc_iff literal.map_sel subst_lit_def)
then show ?thesis
using j_lt len_Ls by (metis nth_map o_apply)
qed
define tm_at :: "nat \<Rightarrow> 'a" where
"\<And>i. tm_at i = atm_of_atms (map atm_of (Ls_at i))"
have "\<And>i. generalizes_atm (tm_at (Suc i)) (tm_at i)"
unfolding tm_at_def generalizes_atm_def atm_of_atms_subst
using Ls_\<sigma>[THEN arg_cong, of "map atm_of"] by (auto simp: comp_def)
moreover have "\<And>i. \<not> generalizes_atm (tm_at i) (tm_at (Suc i))"
unfolding tm_at_def generalizes_atm_def atm_of_atms_subst by (simp add: Ls_\<tau>_strict_tm)
ultimately have "\<And>i. strictly_generalizes_atm (tm_at (Suc i)) (tm_at i)"
unfolding strictly_generalizes_atm_def by blast
then have False
using wf_strictly_generalizes_atm[unfolded wfP_def wf_iff_no_infinite_down_chain] by blast
}
then show "wfP (strictly_generalizes :: 'a clause \<Rightarrow> _ \<Rightarrow> _)"
unfolding wfP_def by (blast intro: wf_iff_no_infinite_down_chain[THEN iffD2])
qed
lemma strict_subset_subst_strictly_subsumes: "C \<cdot> \<eta> \<subset># D \<Longrightarrow> strictly_subsumes C D"
by (metis leD mset_subset_size size_mset_mono size_subst strictly_subsumes_def
subset_mset.dual_order.strict_implies_order substitution_ops.subsumes_def)
lemma generalizes_refl: "generalizes C C"
unfolding generalizes_def by (rule exI[of _ id_subst]) auto
lemma generalizes_trans: "generalizes C D \<Longrightarrow> generalizes D E \<Longrightarrow> generalizes C E"
unfolding generalizes_def using subst_cls_comp_subst by blast
lemma subsumes_refl: "subsumes C C"
unfolding subsumes_def by (rule exI[of _ id_subst]) auto
lemma subsumes_trans: "subsumes C D \<Longrightarrow> subsumes D E \<Longrightarrow> subsumes C E"
unfolding subsumes_def
- by (metis (no_types) subset_mset.order.trans subst_cls_comp_subst subst_cls_mono_mset)
+ by (metis (no_types) subset_mset.trans subst_cls_comp_subst subst_cls_mono_mset)
lemma strictly_generalizes_irrefl: "\<not> strictly_generalizes C C"
unfolding strictly_generalizes_def by blast
lemma strictly_generalizes_antisym: "strictly_generalizes C D \<Longrightarrow> \<not> strictly_generalizes D C"
unfolding strictly_generalizes_def by blast
lemma strictly_generalizes_trans:
"strictly_generalizes C D \<Longrightarrow> strictly_generalizes D E \<Longrightarrow> strictly_generalizes C E"
unfolding strictly_generalizes_def using generalizes_trans by blast
lemma strictly_subsumes_irrefl: "\<not> strictly_subsumes C C"
unfolding strictly_subsumes_def by blast
lemma strictly_subsumes_antisym: "strictly_subsumes C D \<Longrightarrow> \<not> strictly_subsumes D C"
unfolding strictly_subsumes_def by blast
lemma strictly_subsumes_trans:
"strictly_subsumes C D \<Longrightarrow> strictly_subsumes D E \<Longrightarrow> strictly_subsumes C E"
unfolding strictly_subsumes_def using subsumes_trans by blast
lemma subset_strictly_subsumes: "C \<subset># D \<Longrightarrow> strictly_subsumes C D"
using strict_subset_subst_strictly_subsumes[of C id_subst] by auto
lemma strictly_generalizes_neq: "strictly_generalizes D' D \<Longrightarrow> D' \<noteq> D \<cdot> \<sigma>"
unfolding strictly_generalizes_def generalizes_def by blast
lemma strictly_subsumes_neq: "strictly_subsumes D' D \<Longrightarrow> D' \<noteq> D \<cdot> \<sigma>"
unfolding strictly_subsumes_def subsumes_def by blast
lemma strictly_subsumes_has_minimum:
assumes "CC \<noteq> {}"
shows "\<exists>C \<in> CC. \<forall>D \<in> CC. \<not> strictly_subsumes D C"
proof (rule ccontr)
assume "\<not> (\<exists>C \<in> CC. \<forall>D\<in>CC. \<not> strictly_subsumes D C)"
then have "\<forall>C \<in> CC. \<exists>D \<in> CC. strictly_subsumes D C"
by blast
then obtain f where
f_p: "\<forall>C \<in> CC. f C \<in> CC \<and> strictly_subsumes (f C) C"
by metis
from assms obtain C where
C_p: "C \<in> CC"
by auto
define c :: "nat \<Rightarrow> 'a clause" where
"\<And>n. c n = (f ^^ n) C"
have incc: "c i \<in> CC" for i
by (induction i) (auto simp: c_def f_p C_p)
have ps: "\<forall>i. strictly_subsumes (c (Suc i)) (c i)"
using incc f_p unfolding c_def by auto
have "\<forall>i. size (c i) \<ge> size (c (Suc i))"
using ps unfolding strictly_subsumes_def subsumes_def by (metis size_mset_mono size_subst)
then have lte: "\<forall>i. (size \<circ> c) i \<ge> (size \<circ> c) (Suc i)"
unfolding comp_def .
then have "\<exists>l. \<forall>l' \<ge> l. size (c l') = size (c (Suc l'))"
using f_Suc_decr_eventually_const comp_def by auto
then obtain l where
l_p: "\<forall>l' \<ge> l. size (c l') = size (c (Suc l'))"
by metis
then have "\<forall>l' \<ge> l. strictly_generalizes (c (Suc l')) (c l')"
using ps unfolding strictly_generalizes_def generalizes_def
by (metis size_subst less_irrefl strictly_subsumes_def mset_subset_size subset_mset_def
subsumes_def strictly_subsumes_neq)
then have "\<forall>i. strictly_generalizes (c (Suc i + l)) (c (i + l))"
unfolding strictly_generalizes_def generalizes_def by auto
then have "\<exists>f. \<forall>i. strictly_generalizes (f (Suc i)) (f i)"
by (rule exI[of _ "\<lambda>x. c (x + l)"])
then show False
using wf_strictly_generalizes
wf_iff_no_infinite_down_chain[of "{(x, y). strictly_generalizes x y}"]
unfolding wfP_def by auto
qed
lemma wf_strictly_subsumes: "wfP strictly_subsumes"
using strictly_subsumes_has_minimum by (metis equals0D wfP_eq_minimal)
lemma variants_imp_exists_substitution: "variants D D' \<Longrightarrow> \<exists>\<sigma>. D \<cdot> \<sigma> = D'"
unfolding variants_iff_subsumes subsumes_def
by (meson strictly_subsumes_def subset_mset_def strict_subset_subst_strictly_subsumes subsumes_def)
lemma strictly_subsumes_variants:
assumes "strictly_subsumes E D" and "variants D D'"
shows "strictly_subsumes E D'"
proof -
from assms obtain \<sigma> \<sigma>' where
\<sigma>_\<sigma>'_p: "D \<cdot> \<sigma> = D' \<and> D' \<cdot> \<sigma>' = D"
using variants_imp_exists_substitution variants_sym by metis
from assms obtain \<sigma>'' where
"E \<cdot> \<sigma>'' \<subseteq># D"
unfolding strictly_subsumes_def subsumes_def by auto
then have "E \<cdot> \<sigma>'' \<cdot> \<sigma> \<subseteq># D \<cdot> \<sigma>"
using subst_cls_mono_mset by blast
then have "E \<cdot> (\<sigma>'' \<odot> \<sigma>) \<subseteq># D'"
using \<sigma>_\<sigma>'_p by auto
moreover from assms have n: "\<nexists>\<sigma>. D \<cdot> \<sigma> \<subseteq># E"
unfolding strictly_subsumes_def subsumes_def by auto
have "\<nexists>\<sigma>. D' \<cdot> \<sigma> \<subseteq># E"
proof
assume "\<exists>\<sigma>'''. D' \<cdot> \<sigma>''' \<subseteq># E"
then obtain \<sigma>''' where
"D' \<cdot> \<sigma>''' \<subseteq># E"
by auto
then have "D \<cdot> (\<sigma> \<odot> \<sigma>''') \<subseteq># E"
using \<sigma>_\<sigma>'_p by auto
then show False
using n by metis
qed
ultimately show ?thesis
unfolding strictly_subsumes_def subsumes_def by metis
qed
lemma neg_strictly_subsumes_variants:
assumes "\<not> strictly_subsumes E D" and "variants D D'"
shows "\<not> strictly_subsumes E D'"
using assms strictly_subsumes_variants variants_sym by auto
end
subsection \<open>Most General Unifiers\<close>
locale mgu = substitution subst_atm id_subst comp_subst renamings_apart atm_of_atms
for
subst_atm :: "'a \<Rightarrow> 's \<Rightarrow> 'a" and
id_subst :: 's and
comp_subst :: "'s \<Rightarrow> 's \<Rightarrow> 's" and
atm_of_atms :: "'a list \<Rightarrow> 'a" and
renamings_apart :: "'a literal multiset list \<Rightarrow> 's list" +
fixes
mgu :: "'a set set \<Rightarrow> 's option"
assumes
mgu_sound: "finite AAA \<Longrightarrow> (\<forall>AA \<in> AAA. finite AA) \<Longrightarrow> mgu AAA = Some \<sigma> \<Longrightarrow> is_mgu \<sigma> AAA" and
mgu_complete:
"finite AAA \<Longrightarrow> (\<forall>AA \<in> AAA. finite AA) \<Longrightarrow> is_unifiers \<sigma> AAA \<Longrightarrow> \<exists>\<tau>. mgu AAA = Some \<tau>"
begin
lemmas is_unifiers_mgu = mgu_sound[unfolded is_mgu_def, THEN conjunct1]
lemmas is_mgu_most_general = mgu_sound[unfolded is_mgu_def, THEN conjunct2]
lemma mgu_unifier:
assumes
aslen: "length As = n" and
aaslen: "length AAs = n" and
mgu: "Some \<sigma> = mgu (set_mset ` set (map2 add_mset As AAs))" and
i_lt: "i < n" and
a_in: "A \<in># AAs ! i"
shows "A \<cdot>a \<sigma> = As ! i \<cdot>a \<sigma>"
proof -
from mgu have "is_mgu \<sigma> (set_mset ` set (map2 add_mset As AAs))"
using mgu_sound by auto
then have "is_unifiers \<sigma> (set_mset ` set (map2 add_mset As AAs))"
using is_mgu_is_unifiers by auto
then have "is_unifier \<sigma> (set_mset (add_mset (As ! i) (AAs ! i)))"
using i_lt aslen aaslen unfolding is_unifiers_def is_unifier_def
by simp (metis length_zip min.idem nth_mem nth_zip prod.case set_mset_add_mset_insert)
then show ?thesis
using aslen aaslen a_in is_unifier_subst_atm_eqI
by (metis finite_set_mset insertCI set_mset_add_mset_insert)
qed
end
end
diff --git a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy
--- a/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy
+++ b/thys/Ordered_Resolution_Prover/FO_Ordered_Resolution.thy
@@ -1,1412 +1,1412 @@
(* Title: First-Order Ordered Resolution Calculus with Selection
Author: Anders Schlichtkrull <andschl at dtu.dk>, 2016, 2017
Author: Jasmin Blanchette <j.c.blanchette at vu.nl>, 2014, 2017
Author: Dmitriy Traytel <traytel at inf.ethz.ch>, 2014
Author: Sophie Tourret <stourret at mpi-inf.mpg.de>, 2020
Maintainer: Anders Schlichtkrull <andschl at dtu.dk>
*)
section \<open>First-Order Ordered Resolution Calculus with Selection\<close>
theory FO_Ordered_Resolution
imports Abstract_Substitution Ordered_Ground_Resolution Standard_Redundancy
begin
text \<open>
This material is based on Section 4.3 (``A Simple Resolution Prover for First-Order Clauses'') of
Bachmair and Ganzinger's chapter. Specifically, it formalizes the ordered resolution calculus for
first-order standard clauses presented in Figure 4 and its related lemmas and theorems, including
soundness and Lemma 4.12 (the lifting lemma).
The following corresponds to pages 41--42 of Section 4.3, until Figure 5 and its explanation.
\<close>
locale FO_resolution = mgu subst_atm id_subst comp_subst atm_of_atms renamings_apart mgu
for
subst_atm :: "'a :: wellorder \<Rightarrow> 's \<Rightarrow> 'a" and
id_subst :: "'s" and
comp_subst :: "'s \<Rightarrow> 's \<Rightarrow> 's" and
renamings_apart :: "'a literal multiset list \<Rightarrow> 's list" and
atm_of_atms :: "'a list \<Rightarrow> 'a" and
mgu :: "'a set set \<Rightarrow> 's option" +
fixes
less_atm :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
assumes
less_atm_stable: "less_atm A B \<Longrightarrow> less_atm (A \<cdot>a \<sigma>) (B \<cdot>a \<sigma>)" and
less_atm_ground: "is_ground_atm A \<Longrightarrow> is_ground_atm B \<Longrightarrow> less_atm A B \<Longrightarrow> A < B"
begin
subsection \<open>Library\<close>
lemma Bex_cartesian_product: "(\<exists>xy \<in> A \<times> B. P xy) \<equiv> (\<exists>x \<in> A. \<exists>y \<in> B. P (x, y))"
by simp
lemma eql_map_neg_lit_eql_atm:
assumes "map (\<lambda>L. L \<cdot>l \<eta>) (map Neg As') = map Neg As"
shows "As' \<cdot>al \<eta> = As"
using assms by (induction As' arbitrary: As) auto
lemma instance_list:
assumes "negs (mset As) = SDA' \<cdot> \<eta>"
shows "\<exists>As'. negs (mset As') = SDA' \<and> As' \<cdot>al \<eta> = As"
proof -
from assms have negL: "\<forall>L \<in># SDA'. is_neg L"
using Melem_subst_cls subst_lit_in_negs_is_neg by metis
from assms have "{#L \<cdot>l \<eta>. L \<in># SDA'#} = mset (map Neg As)"
using subst_cls_def by auto
then have "\<exists>NAs'. map (\<lambda>L. L \<cdot>l \<eta>) NAs' = map Neg As \<and> mset NAs' = SDA'"
using image_mset_of_subset_list[of "\<lambda>L. L \<cdot>l \<eta>" SDA' "map Neg As"] by auto
then obtain As' where As'_p:
"map (\<lambda>L. L \<cdot>l \<eta>) (map Neg As') = map Neg As \<and> mset (map Neg As') = SDA'"
by (metis (no_types, lifting) Neg_atm_of_iff negL ex_map_conv set_mset_mset)
have "negs (mset As') = SDA'"
using As'_p by auto
moreover have "map (\<lambda>L. L \<cdot>l \<eta>) (map Neg As') = map Neg As"
using As'_p by auto
then have "As' \<cdot>al \<eta> = As"
using eql_map_neg_lit_eql_atm by auto
ultimately show ?thesis
by blast
qed
lemma map2_add_mset_map:
assumes "length AAs' = n" and "length As' = n"
shows "map2 add_mset (As' \<cdot>al \<eta>) (AAs' \<cdot>aml \<eta>) = map2 add_mset As' AAs' \<cdot>aml \<eta>"
using assms
proof (induction n arbitrary: AAs' As')
case (Suc n)
then have "map2 add_mset (tl (As' \<cdot>al \<eta>)) (tl (AAs' \<cdot>aml \<eta>)) = map2 add_mset (tl As') (tl AAs') \<cdot>aml \<eta>"
by simp
moreover have Succ: "length (As' \<cdot>al \<eta>) = Suc n" "length (AAs' \<cdot>aml \<eta>) = Suc n"
using Suc(3) Suc(2) by auto
then have "length (tl (As' \<cdot>al \<eta>)) = n" "length (tl (AAs' \<cdot>aml \<eta>)) = n"
by auto
then have "length (map2 add_mset (tl (As' \<cdot>al \<eta>)) (tl (AAs' \<cdot>aml \<eta>))) = n"
"length (map2 add_mset (tl As') (tl AAs') \<cdot>aml \<eta>) = n"
using Suc(2,3) by auto
ultimately have "\<forall>i < n. tl (map2 add_mset ( (As' \<cdot>al \<eta>)) ((AAs' \<cdot>aml \<eta>))) ! i =
tl (map2 add_mset (As') (AAs') \<cdot>aml \<eta>) ! i"
using Suc(2,3) Succ by (simp add: map2_tl map_tl subst_atm_mset_list_def del: subst_atm_list_tl)
moreover have nn: "length (map2 add_mset ((As' \<cdot>al \<eta>)) ((AAs' \<cdot>aml \<eta>))) = Suc n"
"length (map2 add_mset (As') (AAs') \<cdot>aml \<eta>) = Suc n"
using Succ Suc by auto
ultimately have "\<forall>i. i < Suc n \<longrightarrow> i > 0 \<longrightarrow>
map2 add_mset (As' \<cdot>al \<eta>) (AAs' \<cdot>aml \<eta>) ! i = (map2 add_mset As' AAs' \<cdot>aml \<eta>) ! i"
by (auto simp: subst_atm_mset_list_def gr0_conv_Suc subst_atm_mset_def)
moreover have "add_mset (hd As' \<cdot>a \<eta>) (hd AAs' \<cdot>am \<eta>) = add_mset (hd As') (hd AAs') \<cdot>am \<eta>"
unfolding subst_atm_mset_def by auto
then have "(map2 add_mset (As' \<cdot>al \<eta>) (AAs' \<cdot>aml \<eta>)) ! 0 = (map2 add_mset (As') (AAs') \<cdot>aml \<eta>) ! 0"
using Suc by (simp add: Succ(2) subst_atm_mset_def)
ultimately have "\<forall>i < Suc n. (map2 add_mset (As' \<cdot>al \<eta>) (AAs' \<cdot>aml \<eta>)) ! i =
(map2 add_mset (As') (AAs') \<cdot>aml \<eta>) ! i"
using Suc by auto
then show ?case
using nn list_eq_iff_nth_eq by metis
qed auto
context
fixes S :: "'a clause \<Rightarrow> 'a clause"
begin
subsection \<open>Calculus\<close>
text \<open>
The following corresponds to Figure 4.
\<close>
definition maximal_wrt :: "'a \<Rightarrow> 'a literal multiset \<Rightarrow> bool" where
"maximal_wrt A C \<longleftrightarrow> (\<forall>B \<in> atms_of C. \<not> less_atm A B)"
definition strictly_maximal_wrt :: "'a \<Rightarrow> 'a literal multiset \<Rightarrow> bool" where
"strictly_maximal_wrt A C \<equiv> \<forall>B \<in> atms_of C. A \<noteq> B \<and> \<not> less_atm A B"
lemma strictly_maximal_wrt_maximal_wrt: "strictly_maximal_wrt A C \<Longrightarrow> maximal_wrt A C"
unfolding maximal_wrt_def strictly_maximal_wrt_def by auto
lemma maximal_wrt_subst: "maximal_wrt (A \<cdot>a \<sigma>) (C \<cdot> \<sigma>) \<Longrightarrow> maximal_wrt A C"
unfolding maximal_wrt_def using in_atms_of_subst less_atm_stable by blast
lemma strictly_maximal_wrt_subst:
"strictly_maximal_wrt (A \<cdot>a \<sigma>) (C \<cdot> \<sigma>) \<Longrightarrow> strictly_maximal_wrt A C"
unfolding strictly_maximal_wrt_def using in_atms_of_subst less_atm_stable by blast
inductive eligible :: "'s \<Rightarrow> 'a list \<Rightarrow> 'a clause \<Rightarrow> bool" where
eligible:
"S DA = negs (mset As) \<or> S DA = {#} \<and> length As = 1 \<and> maximal_wrt (As ! 0 \<cdot>a \<sigma>) (DA \<cdot> \<sigma>) \<Longrightarrow>
eligible \<sigma> As DA"
inductive
ord_resolve
:: "'a clause list \<Rightarrow> 'a clause \<Rightarrow> 'a multiset list \<Rightarrow> 'a list \<Rightarrow> 's \<Rightarrow> 'a clause \<Rightarrow> bool"
where
ord_resolve:
"length CAs = n \<Longrightarrow>
length Cs = n \<Longrightarrow>
length AAs = n \<Longrightarrow>
length As = n \<Longrightarrow>
n \<noteq> 0 \<Longrightarrow>
(\<forall>i < n. CAs ! i = Cs ! i + poss (AAs ! i)) \<Longrightarrow>
(\<forall>i < n. AAs ! i \<noteq> {#}) \<Longrightarrow>
Some \<sigma> = mgu (set_mset ` set (map2 add_mset As AAs)) \<Longrightarrow>
eligible \<sigma> As (D + negs (mset As)) \<Longrightarrow>
(\<forall>i < n. strictly_maximal_wrt (As ! i \<cdot>a \<sigma>) (Cs ! i \<cdot> \<sigma>)) \<Longrightarrow>
(\<forall>i < n. S (CAs ! i) = {#}) \<Longrightarrow>
ord_resolve CAs (D + negs (mset As)) AAs As \<sigma> ((\<Sum>\<^sub># (mset Cs) + D) \<cdot> \<sigma>)"
inductive
ord_resolve_rename
:: "'a clause list \<Rightarrow> 'a clause \<Rightarrow> 'a multiset list \<Rightarrow> 'a list \<Rightarrow> 's \<Rightarrow> 'a clause \<Rightarrow> bool"
where
ord_resolve_rename:
"length CAs = n \<Longrightarrow>
length AAs = n \<Longrightarrow>
length As = n \<Longrightarrow>
(\<forall>i < n. poss (AAs ! i) \<subseteq># CAs ! i) \<Longrightarrow>
negs (mset As) \<subseteq># DA \<Longrightarrow>
\<rho> = hd (renamings_apart (DA # CAs)) \<Longrightarrow>
\<rho>s = tl (renamings_apart (DA # CAs)) \<Longrightarrow>
ord_resolve (CAs \<cdot>\<cdot>cl \<rho>s) (DA \<cdot> \<rho>) (AAs \<cdot>\<cdot>aml \<rho>s) (As \<cdot>al \<rho>) \<sigma> E \<Longrightarrow>
ord_resolve_rename CAs DA AAs As \<sigma> E"
lemma ord_resolve_empty_main_prem: "\<not> ord_resolve Cs {#} AAs As \<sigma> E"
by (simp add: ord_resolve.simps)
lemma ord_resolve_rename_empty_main_prem: "\<not> ord_resolve_rename Cs {#} AAs As \<sigma> E"
by (simp add: ord_resolve_empty_main_prem ord_resolve_rename.simps)
subsection \<open>Soundness\<close>
text \<open>
Soundness is not discussed in the chapter, but it is an important property.
\<close>
lemma ord_resolve_ground_inst_sound:
assumes
res_e: "ord_resolve CAs DA AAs As \<sigma> E" and
cc_inst_true: "I \<Turnstile>m mset CAs \<cdot>cm \<sigma> \<cdot>cm \<eta>" and
d_inst_true: "I \<Turnstile> DA \<cdot> \<sigma> \<cdot> \<eta>" and
ground_subst_\<eta>: "is_ground_subst \<eta>"
shows "I \<Turnstile> E \<cdot> \<eta>"
using res_e
proof (cases rule: ord_resolve.cases)
case (ord_resolve n Cs D)
note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and
aas_len = this(5) and as_len = this(6) and cas = this(8) and mgu = this(10) and
len = this(1)
have len: "length CAs = length As"
using as_len cas_len by auto
have "is_ground_subst (\<sigma> \<odot> \<eta>)"
using ground_subst_\<eta> by (rule is_ground_comp_subst)
then have cc_true: "I \<Turnstile>m mset CAs \<cdot>cm \<sigma> \<cdot>cm \<eta>" and d_true: "I \<Turnstile> DA \<cdot> \<sigma> \<cdot> \<eta>"
using cc_inst_true d_inst_true by auto
from mgu have unif: "\<forall>i < n. \<forall>A\<in>#AAs ! i. A \<cdot>a \<sigma> = As ! i \<cdot>a \<sigma>"
using mgu_unifier as_len aas_len by blast
show "I \<Turnstile> E \<cdot> \<eta>"
proof (cases "\<forall>A \<in> set As. A \<cdot>a \<sigma> \<cdot>a \<eta> \<in> I")
case True
then have "\<not> I \<Turnstile> negs (mset As) \<cdot> \<sigma> \<cdot> \<eta>"
unfolding true_cls_def[of I] by auto
then have "I \<Turnstile> D \<cdot> \<sigma> \<cdot> \<eta>"
using d_true da by auto
then show ?thesis
unfolding e by auto
next
case False
then obtain i where a_in_aa: "i < length CAs" and a_false: "(As ! i) \<cdot>a \<sigma> \<cdot>a \<eta> \<notin> I"
using da len by (metis in_set_conv_nth)
define C where "C \<equiv> Cs ! i"
define BB where "BB \<equiv> AAs ! i"
have c_cf': "C \<subseteq># \<Sum>\<^sub># (mset CAs)"
unfolding C_def using a_in_aa cas cas_len
- by (metis less_subset_eq_Union_mset mset_subset_eq_add_left subset_mset.order.trans)
+ by (metis less_subset_eq_Union_mset mset_subset_eq_add_left subset_mset.trans)
have c_in_cc: "C + poss BB \<in># mset CAs"
using C_def BB_def a_in_aa cas_len in_set_conv_nth cas by fastforce
{
fix B
assume "B \<in># BB"
then have "B \<cdot>a \<sigma> = (As ! i) \<cdot>a \<sigma>"
using unif a_in_aa cas_len unfolding BB_def by auto
}
then have "\<not> I \<Turnstile> poss BB \<cdot> \<sigma> \<cdot> \<eta>"
using a_false by (auto simp: true_cls_def)
moreover have "I \<Turnstile> (C + poss BB) \<cdot> \<sigma> \<cdot> \<eta>"
using c_in_cc cc_true true_cls_mset_true_cls[of I "mset CAs \<cdot>cm \<sigma> \<cdot>cm \<eta>"] by force
ultimately have "I \<Turnstile> C \<cdot> \<sigma> \<cdot> \<eta>"
by simp
then show ?thesis
unfolding e subst_cls_union using c_cf' C_def a_in_aa cas_len cs_len
by (metis (no_types, lifting) mset_subset_eq_add_left nth_mem_mset set_mset_mono sum_mset.remove true_cls_mono subst_cls_mono)
qed
qed
text \<open>
The previous lemma is not only used to prove soundness, but also the following lemma which is
used to prove Lemma 4.10.
\<close>
lemma ord_resolve_rename_ground_inst_sound:
assumes
"ord_resolve_rename CAs DA AAs As \<sigma> E" and
"\<rho>s = tl (renamings_apart (DA # CAs))" and
"\<rho> = hd (renamings_apart (DA # CAs))" and
"I \<Turnstile>m (mset (CAs \<cdot>\<cdot>cl \<rho>s)) \<cdot>cm \<sigma> \<cdot>cm \<eta>" and
"I \<Turnstile> DA \<cdot> \<rho> \<cdot> \<sigma> \<cdot> \<eta>" and
"is_ground_subst \<eta>"
shows "I \<Turnstile> E \<cdot> \<eta>"
using assms by (cases rule: ord_resolve_rename.cases) (fast intro: ord_resolve_ground_inst_sound)
text \<open>
Here follows the soundness theorem for the resolution rule.
\<close>
theorem ord_resolve_sound:
assumes
res_e: "ord_resolve CAs DA AAs As \<sigma> E" and
cc_d_true: "\<And>\<sigma>. is_ground_subst \<sigma> \<Longrightarrow> I \<Turnstile>m (mset CAs + {#DA#}) \<cdot>cm \<sigma>" and
ground_subst_\<eta>: "is_ground_subst \<eta>"
shows "I \<Turnstile> E \<cdot> \<eta>"
proof (use res_e in \<open>cases rule: ord_resolve.cases\<close>)
case (ord_resolve n Cs D)
note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4)
and aas_len = this(5) and as_len = this(6) and cas = this(8) and mgu = this(10)
have ground_subst_\<sigma>_\<eta>: "is_ground_subst (\<sigma> \<odot> \<eta>)"
using ground_subst_\<eta> by (rule is_ground_comp_subst)
have cas_true: "I \<Turnstile>m mset CAs \<cdot>cm \<sigma> \<cdot>cm \<eta>"
using cc_d_true ground_subst_\<sigma>_\<eta> by fastforce
have da_true: "I \<Turnstile> DA \<cdot> \<sigma> \<cdot> \<eta>"
using cc_d_true ground_subst_\<sigma>_\<eta> by fastforce
show "I \<Turnstile> E \<cdot> \<eta>"
using ord_resolve_ground_inst_sound[OF res_e cas_true da_true] ground_subst_\<eta> by auto
qed
lemma subst_sound:
assumes
"\<And>\<sigma>. is_ground_subst \<sigma> \<Longrightarrow> I \<Turnstile> C \<cdot> \<sigma>" and
"is_ground_subst \<eta>"
shows "I \<Turnstile> C \<cdot> \<rho> \<cdot> \<eta>"
using assms is_ground_comp_subst subst_cls_comp_subst by metis
lemma subst_sound_scl:
assumes
len: "length P = length CAs" and
true_cas: "\<And>\<sigma>. is_ground_subst \<sigma> \<Longrightarrow> I \<Turnstile>m mset CAs \<cdot>cm \<sigma>" and
ground_subst_\<eta>: "is_ground_subst \<eta>"
shows "I \<Turnstile>m mset (CAs \<cdot>\<cdot>cl P) \<cdot>cm \<eta>"
proof -
from true_cas have "\<And>CA. CA\<in># mset CAs \<Longrightarrow> (\<And>\<sigma>. is_ground_subst \<sigma> \<Longrightarrow> I \<Turnstile> CA \<cdot> \<sigma>)"
unfolding true_cls_mset_def by force
then have "\<forall>i < length CAs. \<forall>\<sigma>. is_ground_subst \<sigma> \<longrightarrow> (I \<Turnstile> CAs ! i \<cdot> \<sigma>)"
using in_set_conv_nth by auto
then have true_cp: "\<forall>i < length CAs. \<forall>\<sigma>. is_ground_subst \<sigma> \<longrightarrow> I \<Turnstile> CAs ! i \<cdot> P ! i \<cdot> \<sigma>"
using subst_sound len by auto
{
fix CA
assume "CA \<in># mset (CAs \<cdot>\<cdot>cl P)"
then obtain i where
i_x: "i < length (CAs \<cdot>\<cdot>cl P)" "CA = (CAs \<cdot>\<cdot>cl P) ! i"
by (metis in_mset_conv_nth)
then have "\<forall>\<sigma>. is_ground_subst \<sigma> \<longrightarrow> I \<Turnstile> CA \<cdot> \<sigma>"
using true_cp unfolding subst_cls_lists_def by (simp add: len)
}
then show ?thesis
using assms unfolding true_cls_mset_def by auto
qed
text \<open>
Here follows the soundness theorem for the resolution rule with renaming.
\<close>
lemma ord_resolve_rename_sound:
assumes
res_e: "ord_resolve_rename CAs DA AAs As \<sigma> E" and
cc_d_true: "\<And>\<sigma>. is_ground_subst \<sigma> \<Longrightarrow> I \<Turnstile>m ((mset CAs) + {#DA#}) \<cdot>cm \<sigma>" and
ground_subst_\<eta>: "is_ground_subst \<eta>"
shows "I \<Turnstile> E \<cdot> \<eta>"
using res_e
proof (cases rule: ord_resolve_rename.cases)
case (ord_resolve_rename n \<rho> \<rho>s)
note \<rho>s = this(7) and res = this(8)
have len: "length \<rho>s = length CAs"
using \<rho>s renamings_apart_length by auto
have "\<And>\<sigma>. is_ground_subst \<sigma> \<Longrightarrow> I \<Turnstile>m (mset (CAs \<cdot>\<cdot>cl \<rho>s) + {#DA \<cdot> \<rho>#}) \<cdot>cm \<sigma>"
using subst_sound_scl[OF len, of I] subst_sound cc_d_true by auto
then show "I \<Turnstile> E \<cdot> \<eta>"
using ground_subst_\<eta> ord_resolve_sound[OF res] by simp
qed
subsection \<open>Other Basic Properties\<close>
lemma ord_resolve_unique:
assumes
"ord_resolve CAs DA AAs As \<sigma> E" and
"ord_resolve CAs DA AAs As \<sigma>' E'"
shows "\<sigma> = \<sigma>' \<and> E = E'"
using assms
proof (cases rule: ord_resolve.cases[case_product ord_resolve.cases], intro conjI)
case (ord_resolve_ord_resolve CAs n Cs AAs As \<sigma>'' DA CAs' n' Cs' AAs' As' \<sigma>''' DA')
note res = this(1-17) and res' = this(18-34)
show \<sigma>: "\<sigma> = \<sigma>'"
using res(3-5,14) res'(3-5,14) by (metis option.inject)
have "Cs = Cs'"
using res(1,3,7,8,12) res'(1,3,7,8,12) by (metis add_right_imp_eq nth_equalityI)
moreover have "DA = DA'"
using res(2,4) res'(2,4) by fastforce
ultimately show "E = E'"
using res(5,6) res'(5,6) \<sigma> by blast
qed
lemma ord_resolve_rename_unique:
assumes
"ord_resolve_rename CAs DA AAs As \<sigma> E" and
"ord_resolve_rename CAs DA AAs As \<sigma>' E'"
shows "\<sigma> = \<sigma>' \<and> E = E'"
using assms unfolding ord_resolve_rename.simps using ord_resolve_unique by meson
lemma ord_resolve_max_side_prems: "ord_resolve CAs DA AAs As \<sigma> E \<Longrightarrow> length CAs \<le> size DA"
by (auto elim!: ord_resolve.cases)
lemma ord_resolve_rename_max_side_prems:
"ord_resolve_rename CAs DA AAs As \<sigma> E \<Longrightarrow> length CAs \<le> size DA"
by (elim ord_resolve_rename.cases, drule ord_resolve_max_side_prems, simp add: renamings_apart_length)
subsection \<open>Inference System\<close>
definition ord_FO_\<Gamma> :: "'a inference set" where
"ord_FO_\<Gamma> = {Infer (mset CAs) DA E | CAs DA AAs As \<sigma> E. ord_resolve_rename CAs DA AAs As \<sigma> E}"
interpretation ord_FO_resolution: inference_system ord_FO_\<Gamma> .
lemma finite_ord_FO_resolution_inferences_between:
assumes fin_cc: "finite CC"
shows "finite (ord_FO_resolution.inferences_between CC C)"
proof -
let ?CCC = "CC \<union> {C}"
define all_AA where "all_AA = (\<Union>D \<in> ?CCC. atms_of D)"
define max_ary where "max_ary = Max (size ` ?CCC)"
define CAS where "CAS = {CAs. CAs \<in> lists ?CCC \<and> length CAs \<le> max_ary}"
define AS where "AS = {As. As \<in> lists all_AA \<and> length As \<le> max_ary}"
define AAS where "AAS = {AAs. AAs \<in> lists (mset ` AS) \<and> length AAs \<le> max_ary}"
note defs = all_AA_def max_ary_def CAS_def AS_def AAS_def
let ?infer_of =
"\<lambda>CAs DA AAs As. Infer (mset CAs) DA (THE E. \<exists>\<sigma>. ord_resolve_rename CAs DA AAs As \<sigma> E)"
let ?Z = "{\<gamma> | CAs DA AAs As \<sigma> E \<gamma>. \<gamma> = Infer (mset CAs) DA E
\<and> ord_resolve_rename CAs DA AAs As \<sigma> E \<and> infer_from ?CCC \<gamma> \<and> C \<in># prems_of \<gamma>}"
let ?Y = "{Infer (mset CAs) DA E | CAs DA AAs As \<sigma> E.
ord_resolve_rename CAs DA AAs As \<sigma> E \<and> set CAs \<union> {DA} \<subseteq> ?CCC}"
let ?X = "{?infer_of CAs DA AAs As | CAs DA AAs As. CAs \<in> CAS \<and> DA \<in> ?CCC \<and> AAs \<in> AAS \<and> As \<in> AS}"
let ?W = "CAS \<times> ?CCC \<times> AAS \<times> AS"
have fin_w: "finite ?W"
unfolding defs using fin_cc by (simp add: finite_lists_length_le lists_eq_set)
have "?Z \<subseteq> ?Y"
by (force simp: infer_from_def)
also have "\<dots> \<subseteq> ?X"
proof -
{
fix CAs DA AAs As \<sigma> E
assume
res_e: "ord_resolve_rename CAs DA AAs As \<sigma> E" and
da_in: "DA \<in> ?CCC" and
cas_sub: "set CAs \<subseteq> ?CCC"
have "E = (THE E. \<exists>\<sigma>. ord_resolve_rename CAs DA AAs As \<sigma> E)
\<and> CAs \<in> CAS \<and> AAs \<in> AAS \<and> As \<in> AS" (is "?e \<and> ?cas \<and> ?aas \<and> ?as")
proof (intro conjI)
show ?e
using res_e ord_resolve_rename_unique by (blast intro: the_equality[symmetric])
next
show ?cas
unfolding CAS_def max_ary_def using cas_sub
ord_resolve_rename_max_side_prems[OF res_e] da_in fin_cc
by (auto simp add: Max_ge_iff)
next
show ?aas
using res_e
proof (cases rule: ord_resolve_rename.cases)
case (ord_resolve_rename n \<rho> \<rho>s)
note len_cas = this(1) and len_aas = this(2) and len_as = this(3) and
aas_sub = this(4) and as_sub = this(5) and res_e' = this(8)
show ?thesis
unfolding AAS_def
proof (clarify, intro conjI)
show "AAs \<in> lists (mset ` AS)"
unfolding AS_def image_def
proof clarsimp
fix AA
assume "AA \<in> set AAs"
then obtain i where
i_lt: "i < n" and
aa: "AA = AAs ! i"
by (metis in_set_conv_nth len_aas)
have casi_in: "CAs ! i \<in> ?CCC"
using i_lt len_cas cas_sub nth_mem by blast
have pos_aa_sub: "poss AA \<subseteq># CAs ! i"
using aa aas_sub i_lt by blast
then have "set_mset AA \<subseteq> atms_of (CAs ! i)"
by (metis atms_of_poss lits_subseteq_imp_atms_subseteq set_mset_mono)
also have aa_sub: "\<dots> \<subseteq> all_AA"
unfolding all_AA_def using casi_in by force
finally have aa_sub: "set_mset AA \<subseteq> all_AA"
.
have "size AA = size (poss AA)"
by simp
also have "\<dots> \<le> size (CAs ! i)"
by (rule size_mset_mono[OF pos_aa_sub])
also have "\<dots> \<le> max_ary"
unfolding max_ary_def using fin_cc casi_in by auto
finally have sz_aa: "size AA \<le> max_ary"
.
let ?As' = "sorted_list_of_multiset AA"
have "?As' \<in> lists all_AA"
using aa_sub by auto
moreover have "length ?As' \<le> max_ary"
using sz_aa by simp
moreover have "AA = mset ?As'"
by simp
ultimately show "\<exists>xa. xa \<in> lists all_AA \<and> length xa \<le> max_ary \<and> AA = mset xa"
by blast
qed
next
have "length AAs = length As"
unfolding len_aas len_as ..
also have "\<dots> \<le> size DA"
using as_sub size_mset_mono by fastforce
also have "\<dots> \<le> max_ary"
unfolding max_ary_def using fin_cc da_in by auto
finally show "length AAs \<le> max_ary"
.
qed
qed
next
show ?as
unfolding AS_def
proof (clarify, intro conjI)
have "set As \<subseteq> atms_of DA"
using res_e[simplified ord_resolve_rename.simps]
by (metis atms_of_negs lits_subseteq_imp_atms_subseteq set_mset_mono set_mset_mset)
also have "\<dots> \<subseteq> all_AA"
unfolding all_AA_def using da_in by blast
finally show "As \<in> lists all_AA"
unfolding lists_eq_set by simp
next
have "length As \<le> size DA"
using res_e[simplified ord_resolve_rename.simps]
ord_resolve_rename_max_side_prems[OF res_e] by auto
also have "size DA \<le> max_ary"
unfolding max_ary_def using fin_cc da_in by auto
finally show "length As \<le> max_ary"
.
qed
qed
}
then show ?thesis
by simp fast
qed
also have "\<dots> \<subseteq> (\<lambda>(CAs, DA, AAs, As). ?infer_of CAs DA AAs As) ` ?W"
unfolding image_def Bex_cartesian_product by fast
finally show ?thesis
unfolding inference_system.inferences_between_def ord_FO_\<Gamma>_def mem_Collect_eq
by (fast intro: rev_finite_subset[OF finite_imageI[OF fin_w]])
qed
lemma ord_FO_resolution_inferences_between_empty_empty:
"ord_FO_resolution.inferences_between {} {#} = {}"
unfolding ord_FO_resolution.inferences_between_def inference_system.inferences_between_def
infer_from_def ord_FO_\<Gamma>_def
using ord_resolve_rename_empty_main_prem by auto
subsection \<open>Lifting\<close>
text \<open>
The following corresponds to the passage between Lemmas 4.11 and 4.12.
\<close>
context
fixes M :: "'a clause set"
assumes select: "selection S"
begin
interpretation selection
by (rule select)
definition S_M :: "'a literal multiset \<Rightarrow> 'a literal multiset" where
"S_M C =
(if C \<in> grounding_of_clss M then
(SOME C'. \<exists>D \<sigma>. D \<in> M \<and> C = D \<cdot> \<sigma> \<and> C' = S D \<cdot> \<sigma> \<and> is_ground_subst \<sigma>)
else
S C)"
lemma S_M_grounding_of_clss:
assumes "C \<in> grounding_of_clss M"
obtains D \<sigma> where
"D \<in> M \<and> C = D \<cdot> \<sigma> \<and> S_M C = S D \<cdot> \<sigma> \<and> is_ground_subst \<sigma>"
proof (atomize_elim, unfold S_M_def eqTrueI[OF assms] if_True, rule someI_ex)
from assms show "\<exists>C' D \<sigma>. D \<in> M \<and> C = D \<cdot> \<sigma> \<and> C' = S D \<cdot> \<sigma> \<and> is_ground_subst \<sigma>"
by (auto simp: grounding_of_clss_def grounding_of_cls_def)
qed
lemma S_M_not_grounding_of_clss: "C \<notin> grounding_of_clss M \<Longrightarrow> S_M C = S C"
unfolding S_M_def by simp
lemma S_M_selects_subseteq: "S_M C \<subseteq># C"
by (metis S_M_grounding_of_clss S_M_not_grounding_of_clss S_selects_subseteq subst_cls_mono_mset)
lemma S_M_selects_neg_lits: "L \<in># S_M C \<Longrightarrow> is_neg L"
by (metis Melem_subst_cls S_M_grounding_of_clss S_M_not_grounding_of_clss S_selects_neg_lits
subst_lit_is_neg)
end
end
text \<open>
The following corresponds to Lemma 4.12:
\<close>
lemma ground_resolvent_subset:
assumes
gr_cas: "is_ground_cls_list CAs" and
gr_da: "is_ground_cls DA" and
res_e: "ord_resolve S CAs DA AAs As \<sigma> E"
shows "E \<subseteq># \<Sum>\<^sub># (mset CAs) + DA"
using res_e
proof (cases rule: ord_resolve.cases)
case (ord_resolve n Cs D)
note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4)
and aas_len = this(5) and as_len = this(6) and cas = this(8) and mgu = this(10)
then have cs_sub_cas: "\<Sum>\<^sub># (mset Cs) \<subseteq># \<Sum>\<^sub># (mset CAs)"
using subseteq_list_Union_mset cas_len cs_len by force
then have cs_sub_cas: "\<Sum>\<^sub># (mset Cs) \<subseteq># \<Sum>\<^sub># (mset CAs)"
using subseteq_list_Union_mset cas_len cs_len by force
then have gr_cs: "is_ground_cls_list Cs"
using gr_cas by simp
have d_sub_da: "D \<subseteq># DA"
by (simp add: da)
then have gr_d: "is_ground_cls D"
using gr_da is_ground_cls_mono by auto
have "is_ground_cls (\<Sum>\<^sub># (mset Cs) + D)"
using gr_cs gr_d by auto
with e have "E = \<Sum>\<^sub># (mset Cs) + D"
by auto
then show ?thesis
using cs_sub_cas d_sub_da by (auto simp: subset_mset.add_mono)
qed
lemma ord_resolve_obtain_clauses:
assumes
res_e: "ord_resolve (S_M S M) CAs DA AAs As \<sigma> E" and
select: "selection S" and
grounding: "{DA} \<union> set CAs \<subseteq> grounding_of_clss M" and
n: "length CAs = n" and
d: "DA = D + negs (mset As)" and
c: "(\<forall>i < n. CAs ! i = Cs ! i + poss (AAs ! i))" "length Cs = n" "length AAs = n"
obtains DA0 \<eta>0 CAs0 \<eta>s0 As0 AAs0 D0 Cs0 where
"length CAs0 = n"
"length \<eta>s0 = n"
"DA0 \<in> M"
"DA0 \<cdot> \<eta>0 = DA"
"S DA0 \<cdot> \<eta>0 = S_M S M DA"
"\<forall>CA0 \<in> set CAs0. CA0 \<in> M"
"CAs0 \<cdot>\<cdot>cl \<eta>s0 = CAs"
"map S CAs0 \<cdot>\<cdot>cl \<eta>s0 = map (S_M S M) CAs"
"is_ground_subst \<eta>0"
"is_ground_subst_list \<eta>s0"
"As0 \<cdot>al \<eta>0 = As"
"AAs0 \<cdot>\<cdot>aml \<eta>s0 = AAs"
"length As0 = n"
"D0 \<cdot> \<eta>0 = D"
"DA0 = D0 + (negs (mset As0))"
"S_M S M (D + negs (mset As)) \<noteq> {#} \<Longrightarrow> negs (mset As0) = S DA0"
"length Cs0 = n"
"Cs0 \<cdot>\<cdot>cl \<eta>s0 = Cs"
"\<forall>i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)"
"length AAs0 = n"
using res_e
proof (cases rule: ord_resolve.cases)
case (ord_resolve n_twin Cs_twins D_twin)
note da = this(1) and e = this(2) and cas = this(8) and mgu = this(10) and eligible = this(11)
from ord_resolve have "n_twin = n" "D_twin = D"
using n d by auto
moreover have "Cs_twins = Cs"
using c cas n calculation(1) \<open>length Cs_twins = n_twin\<close> by (auto simp add: nth_equalityI)
ultimately
have nz: "n \<noteq> 0" and cs_len: "length Cs = n" and aas_len: "length AAs = n" and as_len: "length As = n"
and da: "DA = D + negs (mset As)" and eligible: "eligible (S_M S M) \<sigma> As (D + negs (mset As))"
and cas: "\<forall>i<n. CAs ! i = Cs ! i + poss (AAs ! i)"
using ord_resolve by force+
note n = \<open>n \<noteq> 0\<close> \<open>length CAs = n\<close> \<open>length Cs = n\<close> \<open>length AAs = n\<close> \<open>length As = n\<close>
interpret S: selection S by (rule select)
\<comment> \<open>Obtain FO side premises\<close>
have "\<forall>CA \<in> set CAs. \<exists>CA0 \<eta>c0. CA0 \<in> M \<and> CA0 \<cdot> \<eta>c0 = CA \<and> S CA0 \<cdot> \<eta>c0 = S_M S M CA \<and> is_ground_subst \<eta>c0"
using grounding S_M_grounding_of_clss select by (metis (no_types) le_supE subset_iff)
then have "\<forall>i < n. \<exists>CA0 \<eta>c0. CA0 \<in> M \<and> CA0 \<cdot> \<eta>c0 = (CAs ! i) \<and> S CA0 \<cdot> \<eta>c0 = S_M S M (CAs ! i) \<and> is_ground_subst \<eta>c0"
using n by force
then obtain \<eta>s0f CAs0f where f_p:
"\<forall>i < n. CAs0f i \<in> M"
"\<forall>i < n. (CAs0f i) \<cdot> (\<eta>s0f i) = (CAs ! i)"
"\<forall>i < n. S (CAs0f i) \<cdot> (\<eta>s0f i) = S_M S M (CAs ! i)"
"\<forall>i < n. is_ground_subst (\<eta>s0f i)"
using n by (metis (no_types))
define \<eta>s0 where
"\<eta>s0 = map \<eta>s0f [0 ..<n]"
define CAs0 where
"CAs0 = map CAs0f [0 ..<n]"
have "length \<eta>s0 = n" "length CAs0 = n"
unfolding \<eta>s0_def CAs0_def by auto
note n = \<open>length \<eta>s0 = n\<close> \<open>length CAs0 = n\<close> n
\<comment> \<open>The properties we need of the FO side premises\<close>
have CAs0_in_M: "\<forall>CA0 \<in> set CAs0. CA0 \<in> M"
unfolding CAs0_def using f_p(1) by auto
have CAs0_to_CAs: "CAs0 \<cdot>\<cdot>cl \<eta>s0 = CAs"
unfolding CAs0_def \<eta>s0_def using f_p(2) by (auto simp: n intro: nth_equalityI)
have SCAs0_to_SMCAs: "(map S CAs0) \<cdot>\<cdot>cl \<eta>s0 = map (S_M S M) CAs"
unfolding CAs0_def \<eta>s0_def using f_p(3) n by (force intro: nth_equalityI)
have sub_ground: "\<forall>\<eta>c0 \<in> set \<eta>s0. is_ground_subst \<eta>c0"
unfolding \<eta>s0_def using f_p n by force
then have "is_ground_subst_list \<eta>s0"
using n unfolding is_ground_subst_list_def by auto
\<comment> \<open>Split side premises CAs0 into Cs0 and AAs0\<close>
obtain AAs0 Cs0 where AAs0_Cs0_p:
"AAs0 \<cdot>\<cdot>aml \<eta>s0 = AAs" "length Cs0 = n" "Cs0 \<cdot>\<cdot>cl \<eta>s0 = Cs"
"\<forall>i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)" "length AAs0 = n"
proof -
have "\<forall>i < n. \<exists>AA0. AA0 \<cdot>am \<eta>s0 ! i = AAs ! i \<and> poss AA0 \<subseteq># CAs0 ! i"
proof (rule, rule)
fix i
assume "i < n"
have "CAs0 ! i \<cdot> \<eta>s0 ! i = CAs ! i"
using \<open>i < n\<close> \<open>CAs0 \<cdot>\<cdot>cl \<eta>s0 = CAs\<close> n by force
moreover have "poss (AAs ! i) \<subseteq># CAs !i"
using \<open>i < n\<close> cas by auto
ultimately obtain poss_AA0 where
nn: "poss_AA0 \<cdot> \<eta>s0 ! i = poss (AAs ! i) \<and> poss_AA0 \<subseteq># CAs0 ! i"
using cas image_mset_of_subset unfolding subst_cls_def by metis
then have l: "\<forall>L \<in># poss_AA0. is_pos L"
unfolding subst_cls_def by (metis Melem_subst_cls imageE literal.disc(1)
literal.map_disc_iff set_image_mset subst_cls_def subst_lit_def)
define AA0 where
"AA0 = image_mset atm_of poss_AA0"
have na: "poss AA0 = poss_AA0"
using l unfolding AA0_def by auto
then have "AA0 \<cdot>am \<eta>s0 ! i = AAs ! i"
using nn by (metis (mono_tags) literal.inject(1) multiset.inj_map_strong subst_cls_poss)
moreover have "poss AA0 \<subseteq># CAs0 ! i"
using na nn by auto
ultimately show "\<exists>AA0. AA0 \<cdot>am \<eta>s0 ! i = AAs ! i \<and> poss AA0 \<subseteq># CAs0 ! i"
by blast
qed
then obtain AAs0f where
AAs0f_p: "\<forall>i < n. AAs0f i \<cdot>am \<eta>s0 ! i = AAs ! i \<and> (poss (AAs0f i)) \<subseteq># CAs0 ! i"
by metis
define AAs0 where "AAs0 = map AAs0f [0 ..<n]"
then have "length AAs0 = n"
by auto
note n = n \<open>length AAs0 = n\<close>
from AAs0_def have "\<forall>i < n. AAs0 ! i \<cdot>am \<eta>s0 ! i = AAs ! i"
using AAs0f_p by auto
then have AAs0_AAs: "AAs0 \<cdot>\<cdot>aml \<eta>s0 = AAs"
using n by (auto intro: nth_equalityI)
from AAs0_def have AAs0_in_CAs0: "\<forall>i < n. poss (AAs0 ! i) \<subseteq># CAs0 ! i"
using AAs0f_p by auto
define Cs0 where
"Cs0 = map2 (-) CAs0 (map poss AAs0)"
have "length Cs0 = n"
using Cs0_def n by auto
note n = n \<open>length Cs0 = n\<close>
have "\<forall>i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)"
using AAs0_in_CAs0 Cs0_def n by auto
then have "Cs0 \<cdot>\<cdot>cl \<eta>s0 = Cs"
using \<open>CAs0 \<cdot>\<cdot>cl \<eta>s0 = CAs\<close> AAs0_AAs cas n by (auto intro: nth_equalityI)
show ?thesis
using that
\<open>AAs0 \<cdot>\<cdot>aml \<eta>s0 = AAs\<close> \<open>Cs0 \<cdot>\<cdot>cl \<eta>s0 = Cs\<close> \<open>\<forall>i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)\<close>
\<open>length AAs0 = n\<close> \<open>length Cs0 = n\<close>
by blast
qed
\<comment> \<open>Obtain FO main premise\<close>
have "\<exists>DA0 \<eta>0. DA0 \<in> M \<and> DA = DA0 \<cdot> \<eta>0 \<and> S DA0 \<cdot> \<eta>0 = S_M S M DA \<and> is_ground_subst \<eta>0"
using grounding S_M_grounding_of_clss select by (metis le_supE singletonI subsetCE)
then obtain DA0 \<eta>0 where
DA0_\<eta>0_p: "DA0 \<in> M \<and> DA = DA0 \<cdot> \<eta>0 \<and> S DA0 \<cdot> \<eta>0 = S_M S M DA \<and> is_ground_subst \<eta>0"
by auto
\<comment> \<open>The properties we need of the FO main premise\<close>
have DA0_in_M: "DA0 \<in> M"
using DA0_\<eta>0_p by auto
have DA0_to_DA: "DA0 \<cdot> \<eta>0 = DA"
using DA0_\<eta>0_p by auto
have SDA0_to_SMDA: "S DA0 \<cdot> \<eta>0 = S_M S M DA"
using DA0_\<eta>0_p by auto
have "is_ground_subst \<eta>0"
using DA0_\<eta>0_p by auto
\<comment> \<open>Split main premise DA0 into D0 and As0\<close>
obtain D0 As0 where D0As0_p:
"As0 \<cdot>al \<eta>0 = As" "length As0 = n" "D0 \<cdot> \<eta>0 = D" "DA0 = D0 + (negs (mset As0))"
"S_M S M (D + negs (mset As)) \<noteq> {#} \<Longrightarrow> negs (mset As0) = S DA0"
proof -
{
assume a: "S_M S M (D + negs (mset As)) = {#} \<and> length As = (Suc 0)
\<and> maximal_wrt (As ! 0 \<cdot>a \<sigma>) ((D + negs (mset As)) \<cdot> \<sigma>)"
then have as: "mset As = {#As ! 0#}"
by (auto intro: nth_equalityI)
then have "negs (mset As) = {#Neg (As ! 0)#}"
by (simp add: \<open>mset As = {#As ! 0#}\<close>)
then have "DA = D + {#Neg (As ! 0)#}"
using da by auto
then obtain L where "L \<in># DA0 \<and> L \<cdot>l \<eta>0 = Neg (As ! 0)"
using DA0_to_DA by (metis Melem_subst_cls mset_subset_eq_add_right single_subset_iff)
then have "Neg (atm_of L) \<in># DA0 \<and> Neg (atm_of L) \<cdot>l \<eta>0 = Neg (As ! 0)"
by (metis Neg_atm_of_iff literal.sel(2) subst_lit_is_pos)
then have "[atm_of L] \<cdot>al \<eta>0 = As \<and> negs (mset [atm_of L]) \<subseteq># DA0"
using as subst_lit_def by auto
then have "\<exists>As0. As0 \<cdot>al \<eta>0 = As \<and> negs (mset As0) \<subseteq># DA0
\<and> (S_M S M (D + negs (mset As)) \<noteq> {#} \<longrightarrow> negs (mset As0) = S DA0)"
using a by blast
}
moreover
{
assume "S_M S M (D + negs (mset As)) = negs (mset As)"
then have "negs (mset As) = S DA0 \<cdot> \<eta>0"
using da \<open>S DA0 \<cdot> \<eta>0 = S_M S M DA\<close> by auto
then have "\<exists>As0. negs (mset As0) = S DA0 \<and> As0 \<cdot>al \<eta>0 = As"
using instance_list[of As "S DA0" \<eta>0] S.S_selects_neg_lits by auto
then have "\<exists>As0. As0 \<cdot>al \<eta>0 = As \<and> negs (mset As0) \<subseteq># DA0
\<and> (S_M S M (D + negs (mset As)) \<noteq> {#} \<longrightarrow> negs (mset As0) = S DA0)"
using S.S_selects_subseteq by auto
}
ultimately have "\<exists>As0. As0 \<cdot>al \<eta>0 = As \<and> (negs (mset As0)) \<subseteq># DA0
\<and> (S_M S M (D + negs (mset As)) \<noteq> {#} \<longrightarrow> negs (mset As0) = S DA0)"
using eligible unfolding eligible.simps by auto
then obtain As0 where
As0_p: "As0 \<cdot>al \<eta>0 = As \<and> negs (mset As0) \<subseteq># DA0
\<and> (S_M S M (D + negs (mset As)) \<noteq> {#} \<longrightarrow> negs (mset As0) = S DA0)"
by blast
then have "length As0 = n"
using as_len by auto
note n = n this
have "As0 \<cdot>al \<eta>0 = As"
using As0_p by auto
define D0 where
"D0 = DA0 - negs (mset As0)"
then have "DA0 = D0 + negs (mset As0)"
using As0_p by auto
then have "D0 \<cdot> \<eta>0 = D"
using DA0_to_DA da As0_p by auto
have "S_M S M (D + negs (mset As)) \<noteq> {#} \<Longrightarrow> negs (mset As0) = S DA0"
using As0_p by blast
then show ?thesis
using that \<open>As0 \<cdot>al \<eta>0 = As\<close> \<open>D0 \<cdot> \<eta>0= D\<close> \<open>DA0 = D0 + (negs (mset As0))\<close> \<open>length As0 = n\<close>
by metis
qed
show ?thesis
using that[OF n(2,1) DA0_in_M DA0_to_DA SDA0_to_SMDA CAs0_in_M CAs0_to_CAs SCAs0_to_SMCAs
\<open>is_ground_subst \<eta>0\<close> \<open>is_ground_subst_list \<eta>s0\<close> \<open>As0 \<cdot>al \<eta>0 = As\<close>
\<open>AAs0 \<cdot>\<cdot>aml \<eta>s0 = AAs\<close>
\<open>length As0 = n\<close>
\<open>D0 \<cdot> \<eta>0 = D\<close>
\<open>DA0 = D0 + (negs (mset As0))\<close>
\<open>S_M S M (D + negs (mset As)) \<noteq> {#} \<Longrightarrow> negs (mset As0) = S DA0\<close>
\<open>length Cs0 = n\<close>
\<open>Cs0 \<cdot>\<cdot>cl \<eta>s0 = Cs\<close>
\<open>\<forall>i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)\<close>
\<open>length AAs0 = n\<close>]
by auto
qed
lemma ord_resolve_rename_lifting:
assumes
sel_stable: "\<And>\<rho> C. is_renaming \<rho> \<Longrightarrow> S (C \<cdot> \<rho>) = S C \<cdot> \<rho>" and
res_e: "ord_resolve (S_M S M) CAs DA AAs As \<sigma> E" and
select: "selection S" and
grounding: "{DA} \<union> set CAs \<subseteq> grounding_of_clss M"
obtains \<eta>s \<eta> \<eta>2 CAs0 DA0 AAs0 As0 E0 \<tau> where
"is_ground_subst \<eta>"
"is_ground_subst_list \<eta>s"
"is_ground_subst \<eta>2"
"ord_resolve_rename S CAs0 DA0 AAs0 As0 \<tau> E0"
"CAs0 \<cdot>\<cdot>cl \<eta>s = CAs" "DA0 \<cdot> \<eta> = DA" "E0 \<cdot> \<eta>2 = E"
"{DA0} \<union> set CAs0 \<subseteq> M"
"length CAs0 = length CAs"
"length \<eta>s = length CAs"
using res_e
proof (cases rule: ord_resolve.cases)
case (ord_resolve n Cs D)
note da = this(1) and e = this(2) and cas_len = this(3) and cs_len = this(4) and
aas_len = this(5) and as_len = this(6) and nz = this(7) and cas = this(8) and
aas_not_empt = this(9) and mgu = this(10) and eligible = this(11) and str_max = this(12) and
sel_empt = this(13)
have sel_ren_list_inv:
"\<And>\<rho>s Cs. length \<rho>s = length Cs \<Longrightarrow> is_renaming_list \<rho>s \<Longrightarrow> map S (Cs \<cdot>\<cdot>cl \<rho>s) = map S Cs \<cdot>\<cdot>cl \<rho>s"
using sel_stable unfolding is_renaming_list_def by (auto intro: nth_equalityI)
note n = \<open>n \<noteq> 0\<close> \<open>length CAs = n\<close> \<open>length Cs = n\<close> \<open>length AAs = n\<close> \<open>length As = n\<close>
interpret S: selection S by (rule select)
obtain DA0 \<eta>0 CAs0 \<eta>s0 As0 AAs0 D0 Cs0 where as0:
"length CAs0 = n"
"length \<eta>s0 = n"
"DA0 \<in> M"
"DA0 \<cdot> \<eta>0 = DA"
"S DA0 \<cdot> \<eta>0 = S_M S M DA"
"\<forall>CA0 \<in> set CAs0. CA0 \<in> M"
"CAs0 \<cdot>\<cdot>cl \<eta>s0 = CAs"
"map S CAs0 \<cdot>\<cdot>cl \<eta>s0 = map (S_M S M) CAs"
"is_ground_subst \<eta>0"
"is_ground_subst_list \<eta>s0"
"As0 \<cdot>al \<eta>0 = As"
"AAs0 \<cdot>\<cdot>aml \<eta>s0 = AAs"
"length As0 = n"
"D0 \<cdot> \<eta>0 = D"
"DA0 = D0 + (negs (mset As0))"
"S_M S M (D + negs (mset As)) \<noteq> {#} \<Longrightarrow> negs (mset As0) = S DA0"
"length Cs0 = n"
"Cs0 \<cdot>\<cdot>cl \<eta>s0 = Cs"
"\<forall>i < n. CAs0 ! i = Cs0 ! i + poss (AAs0 ! i)"
"length AAs0 = n"
using ord_resolve_obtain_clauses[of S M CAs DA, OF res_e select grounding n(2) \<open>DA = D + negs (mset As)\<close>
\<open>\<forall>i<n. CAs ! i = Cs ! i + poss (AAs ! i)\<close> \<open>length Cs = n\<close> \<open>length AAs = n\<close>, of thesis] by blast
note n = \<open>length CAs0 = n\<close> \<open>length \<eta>s0 = n\<close> \<open>length As0 = n\<close> \<open>length AAs0 = n\<close> \<open>length Cs0 = n\<close> n
have "length (renamings_apart (DA0 # CAs0)) = Suc n"
using n renamings_apart_length by auto
note n = this n
define \<rho> where
"\<rho> = hd (renamings_apart (DA0 # CAs0))"
define \<rho>s where
"\<rho>s = tl (renamings_apart (DA0 # CAs0))"
define DA0' where
"DA0' = DA0 \<cdot> \<rho>"
define D0' where
"D0' = D0 \<cdot> \<rho>"
define As0' where
"As0' = As0 \<cdot>al \<rho>"
define CAs0' where
"CAs0' = CAs0 \<cdot>\<cdot>cl \<rho>s"
define Cs0' where
"Cs0' = Cs0 \<cdot>\<cdot>cl \<rho>s"
define AAs0' where
"AAs0' = AAs0 \<cdot>\<cdot>aml \<rho>s"
define \<eta>0' where
"\<eta>0' = inv_renaming \<rho> \<odot> \<eta>0"
define \<eta>s0' where
"\<eta>s0' = map inv_renaming \<rho>s \<odot>s \<eta>s0"
have renames_DA0: "is_renaming \<rho>"
using renamings_apart_length renamings_apart_renaming unfolding \<rho>_def
by (metis length_greater_0_conv list.exhaust_sel list.set_intros(1) list.simps(3))
have renames_CAs0: "is_renaming_list \<rho>s"
using renamings_apart_length renamings_apart_renaming unfolding \<rho>s_def
by (metis is_renaming_list_def length_greater_0_conv list.set_sel(2) list.simps(3))
have "length \<rho>s = n"
unfolding \<rho>s_def using n by auto
note n = n \<open>length \<rho>s = n\<close>
have "length As0' = n"
unfolding As0'_def using n by auto
have "length CAs0' = n"
using as0(1) n unfolding CAs0'_def by auto
have "length Cs0' = n"
unfolding Cs0'_def using n by auto
have "length AAs0' = n"
unfolding AAs0'_def using n by auto
have "length \<eta>s0' = n"
using as0(2) n unfolding \<eta>s0'_def by auto
note n = \<open>length CAs0' = n\<close> \<open>length \<eta>s0' = n\<close> \<open>length As0' = n\<close> \<open>length AAs0' = n\<close> \<open>length Cs0' = n\<close> n
have DA0'_DA: "DA0' \<cdot> \<eta>0' = DA"
using as0(4) unfolding \<eta>0'_def DA0'_def using renames_DA0 by simp
have D0'_D: "D0' \<cdot> \<eta>0' = D"
using as0(14) unfolding \<eta>0'_def D0'_def using renames_DA0 by simp
have As0'_As: "As0' \<cdot>al \<eta>0' = As"
using as0(11) unfolding \<eta>0'_def As0'_def using renames_DA0 by auto
have "S DA0' \<cdot> \<eta>0' = S_M S M DA"
using as0(5) unfolding \<eta>0'_def DA0'_def using renames_DA0 sel_stable by auto
have CAs0'_CAs: "CAs0' \<cdot>\<cdot>cl \<eta>s0' = CAs"
using as0(7) unfolding CAs0'_def \<eta>s0'_def using renames_CAs0 n by auto
have Cs0'_Cs: "Cs0' \<cdot>\<cdot>cl \<eta>s0' = Cs"
using as0(18) unfolding Cs0'_def \<eta>s0'_def using renames_CAs0 n by auto
have AAs0'_AAs: "AAs0' \<cdot>\<cdot>aml \<eta>s0' = AAs"
using as0(12) unfolding \<eta>s0'_def AAs0'_def using renames_CAs0 using n by auto
have "map S CAs0' \<cdot>\<cdot>cl \<eta>s0' = map (S_M S M) CAs"
unfolding CAs0'_def \<eta>s0'_def using as0(8) n renames_CAs0 sel_ren_list_inv by auto
have DA0'_split: "DA0' = D0' + negs (mset As0')"
using as0(15) DA0'_def D0'_def As0'_def by auto
then have D0'_subset_DA0': "D0' \<subseteq># DA0'"
by auto
from DA0'_split have negs_As0'_subset_DA0': "negs (mset As0') \<subseteq># DA0'"
by auto
have CAs0'_split: "\<forall>i<n. CAs0' ! i = Cs0' ! i + poss (AAs0' ! i)"
using as0(19) CAs0'_def Cs0'_def AAs0'_def n by auto
then have "\<forall>i<n. Cs0' ! i \<subseteq># CAs0' ! i"
by auto
from CAs0'_split have poss_AAs0'_subset_CAs0': "\<forall>i<n. poss (AAs0' ! i) \<subseteq># CAs0' ! i"
by auto
then have AAs0'_in_atms_of_CAs0': "\<forall>i < n. \<forall>A\<in>#AAs0' ! i. A \<in> atms_of (CAs0' ! i)"
by (auto simp add: atm_iff_pos_or_neg_lit)
have as0':
"S_M S M (D + negs (mset As)) \<noteq> {#} \<Longrightarrow> negs (mset As0') = S DA0'"
proof -
assume a: "S_M S M (D + negs (mset As)) \<noteq> {#}"
then have "negs (mset As0) \<cdot> \<rho> = S DA0 \<cdot> \<rho>"
using as0(16) unfolding \<rho>_def by metis
then show "negs (mset As0') = S DA0'"
using As0'_def DA0'_def using sel_stable[of \<rho> DA0] renames_DA0 by auto
qed
have vd: "var_disjoint (DA0' # CAs0')"
unfolding DA0'_def CAs0'_def using renamings_apart_var_disjoint
unfolding \<rho>_def \<rho>s_def
by (metis length_greater_0_conv list.exhaust_sel n(6) substitution.subst_cls_lists_Cons
substitution_axioms zero_less_Suc)
\<comment> \<open>Introduce ground substitution\<close>
from vd DA0'_DA CAs0'_CAs have "\<exists>\<eta>. \<forall>i < Suc n. \<forall>S. S \<subseteq># (DA0' # CAs0') ! i \<longrightarrow> S \<cdot> (\<eta>0'#\<eta>s0') ! i = S \<cdot> \<eta>"
unfolding var_disjoint_def using n by auto
then obtain \<eta> where \<eta>_p: "\<forall>i < Suc n. \<forall>S. S \<subseteq># (DA0' # CAs0') ! i \<longrightarrow> S \<cdot> (\<eta>0'#\<eta>s0') ! i = S \<cdot> \<eta>"
by auto
have \<eta>_p_lit: "\<forall>i < Suc n. \<forall>L. L \<in># (DA0' # CAs0') ! i \<longrightarrow> L \<cdot>l (\<eta>0'#\<eta>s0') ! i = L \<cdot>l \<eta>"
proof (rule, rule, rule, rule)
fix i :: "nat" and L :: "'a literal"
assume a:
"i < Suc n"
"L \<in># (DA0' # CAs0') ! i"
then have "\<forall>S. S \<subseteq># (DA0' # CAs0') ! i \<longrightarrow> S \<cdot> (\<eta>0' # \<eta>s0') ! i = S \<cdot> \<eta>"
using \<eta>_p by auto
then have "{# L #} \<cdot> (\<eta>0' # \<eta>s0') ! i = {# L #} \<cdot> \<eta>"
using a by (meson single_subset_iff)
then show "L \<cdot>l (\<eta>0' # \<eta>s0') ! i = L \<cdot>l \<eta>" by auto
qed
have \<eta>_p_atm: "\<forall>i < Suc n. \<forall>A. A \<in> atms_of ((DA0' # CAs0') ! i) \<longrightarrow> A \<cdot>a (\<eta>0'#\<eta>s0') ! i = A \<cdot>a \<eta>"
proof (rule, rule, rule, rule)
fix i :: "nat" and A :: "'a"
assume a:
"i < Suc n"
"A \<in> atms_of ((DA0' # CAs0') ! i)"
then obtain L where L_p: "atm_of L = A \<and> L \<in># (DA0' # CAs0') ! i"
unfolding atms_of_def by auto
then have "L \<cdot>l (\<eta>0'#\<eta>s0') ! i = L \<cdot>l \<eta>"
using \<eta>_p_lit a by auto
then show "A \<cdot>a (\<eta>0' # \<eta>s0') ! i = A \<cdot>a \<eta>"
using L_p unfolding subst_lit_def by (cases L) auto
qed
have DA0'_DA: "DA0' \<cdot> \<eta> = DA"
using DA0'_DA \<eta>_p by auto
have "D0' \<cdot> \<eta> = D" using \<eta>_p D0'_D n D0'_subset_DA0' by auto
have "As0' \<cdot>al \<eta> = As"
proof (rule nth_equalityI)
show "length (As0' \<cdot>al \<eta>) = length As"
using n by auto
next
fix i
show "i<length (As0' \<cdot>al \<eta>) \<Longrightarrow> (As0' \<cdot>al \<eta>) ! i = As ! i"
proof -
assume a: "i < length (As0' \<cdot>al \<eta>)"
have A_eq: "\<forall>A. A \<in> atms_of DA0' \<longrightarrow> A \<cdot>a \<eta>0' = A \<cdot>a \<eta>"
using \<eta>_p_atm n by force
have "As0' ! i \<in> atms_of DA0'"
using negs_As0'_subset_DA0' unfolding atms_of_def
using a n by force
then have "As0' ! i \<cdot>a \<eta>0' = As0' ! i \<cdot>a \<eta>"
using A_eq by simp
then show "(As0' \<cdot>al \<eta>) ! i = As ! i"
using As0'_As \<open>length As0' = n\<close> a by auto
qed
qed
interpret selection
by (rule select)
have "S DA0' \<cdot> \<eta> = S_M S M DA"
using \<open>S DA0' \<cdot> \<eta>0' = S_M S M DA\<close> \<eta>_p S.S_selects_subseteq by auto
from \<eta>_p have \<eta>_p_CAs0': "\<forall>i < n. (CAs0' ! i) \<cdot> (\<eta>s0' ! i) = (CAs0'! i) \<cdot> \<eta>"
using n by auto
then have "CAs0' \<cdot>\<cdot>cl \<eta>s0' = CAs0' \<cdot>cl \<eta>"
using n by (auto intro: nth_equalityI)
then have CAs0'_\<eta>_fo_CAs: "CAs0' \<cdot>cl \<eta> = CAs"
using CAs0'_CAs \<eta>_p n by auto
from \<eta>_p have "\<forall>i < n. S (CAs0' ! i) \<cdot> \<eta>s0' ! i = S (CAs0' ! i) \<cdot> \<eta>"
using S.S_selects_subseteq n by auto
then have "map S CAs0' \<cdot>\<cdot>cl \<eta>s0' = map S CAs0' \<cdot>cl \<eta>"
using n by (auto intro: nth_equalityI)
then have SCAs0'_\<eta>_fo_SMCAs: "map S CAs0' \<cdot>cl \<eta> = map (S_M S M) CAs"
using \<open>map S CAs0' \<cdot>\<cdot>cl \<eta>s0' = map (S_M S M) CAs\<close> by auto
have "Cs0' \<cdot>cl \<eta> = Cs"
proof (rule nth_equalityI)
show "length (Cs0' \<cdot>cl \<eta>) = length Cs"
using n by auto
next
fix i
show "i<length (Cs0' \<cdot>cl \<eta>) \<Longrightarrow> (Cs0' \<cdot>cl \<eta>) ! i = Cs ! i"
proof -
assume "i < length (Cs0' \<cdot>cl \<eta>)"
then have a: "i < n"
using n by force
have "(Cs0' \<cdot>\<cdot>cl \<eta>s0') ! i = Cs ! i"
using Cs0'_Cs a n by force
moreover
have \<eta>_p_CAs0': "\<forall>S. S \<subseteq># CAs0' ! i \<longrightarrow> S \<cdot> \<eta>s0' ! i = S \<cdot> \<eta>"
using \<eta>_p a by force
have "Cs0' ! i \<cdot> \<eta>s0' ! i = (Cs0' \<cdot>cl \<eta>) ! i"
using \<eta>_p_CAs0' \<open>\<forall>i<n. Cs0' ! i \<subseteq># CAs0' ! i\<close> a n by force
then have "(Cs0' \<cdot>\<cdot>cl \<eta>s0') ! i = (Cs0' \<cdot>cl \<eta>) ! i "
using a n by force
ultimately show "(Cs0' \<cdot>cl \<eta>) ! i = Cs ! i"
by auto
qed
qed
have AAs0'_AAs: "AAs0' \<cdot>aml \<eta> = AAs"
proof (rule nth_equalityI)
show "length (AAs0' \<cdot>aml \<eta>) = length AAs"
using n by auto
next
fix i
show "i<length (AAs0' \<cdot>aml \<eta>) \<Longrightarrow> (AAs0' \<cdot>aml \<eta>) ! i = AAs ! i"
proof -
assume a: "i < length (AAs0' \<cdot>aml \<eta>)"
then have "i < n"
using n by force
then have "\<forall>A. A \<in> atms_of ((DA0' # CAs0') ! Suc i) \<longrightarrow> A \<cdot>a (\<eta>0' # \<eta>s0') ! Suc i = A \<cdot>a \<eta>"
using \<eta>_p_atm n by force
then have A_eq: "\<forall>A. A \<in> atms_of (CAs0' ! i) \<longrightarrow> A \<cdot>a \<eta>s0' ! i = A \<cdot>a \<eta>"
by auto
have AAs_CAs0': "\<forall>A \<in># AAs0' ! i. A \<in> atms_of (CAs0' ! i)"
using AAs0'_in_atms_of_CAs0' unfolding atms_of_def
using a n by force
then have "AAs0' ! i \<cdot>am \<eta>s0' ! i = AAs0' ! i \<cdot>am \<eta>"
unfolding subst_atm_mset_def using A_eq unfolding subst_atm_mset_def by auto
then show "(AAs0' \<cdot>aml \<eta>) ! i = AAs ! i"
using AAs0'_AAs \<open>length AAs0' = n\<close> \<open>length \<eta>s0' = n\<close> a by auto
qed
qed
\<comment> \<open>Obtain MGU and substitution\<close>
obtain \<tau> \<phi> where \<tau>\<phi>:
"Some \<tau> = mgu (set_mset ` set (map2 add_mset As0' AAs0'))"
"\<tau> \<odot> \<phi> = \<eta> \<odot> \<sigma>"
proof -
have uu: "is_unifiers \<sigma> (set_mset ` set (map2 add_mset (As0' \<cdot>al \<eta>) (AAs0' \<cdot>aml \<eta>)))"
using mgu mgu_sound is_mgu_def unfolding \<open>AAs0' \<cdot>aml \<eta> = AAs\<close> using \<open>As0' \<cdot>al \<eta> = As\<close> by auto
have \<eta>\<sigma>uni: "is_unifiers (\<eta> \<odot> \<sigma>) (set_mset ` set (map2 add_mset As0' AAs0'))"
proof -
have "set_mset ` set (map2 add_mset As0' AAs0' \<cdot>aml \<eta>) =
set_mset ` set (map2 add_mset As0' AAs0') \<cdot>ass \<eta>"
unfolding subst_atmss_def subst_atm_mset_list_def using subst_atm_mset_def subst_atms_def
by (simp add: image_image subst_atm_mset_def subst_atms_def)
then have "is_unifiers \<sigma> (set_mset ` set (map2 add_mset As0' AAs0') \<cdot>ass \<eta>)"
using uu by (auto simp: n map2_add_mset_map)
then show ?thesis
using is_unifiers_comp by auto
qed
then obtain \<tau> where
\<tau>_p: "Some \<tau> = mgu (set_mset ` set (map2 add_mset As0' AAs0'))"
using mgu_complete
by (metis (mono_tags, hide_lams) List.finite_set finite_imageI finite_set_mset image_iff)
moreover then obtain \<phi> where \<phi>_p: "\<tau> \<odot> \<phi> = \<eta> \<odot> \<sigma>"
by (metis (mono_tags, hide_lams) finite_set \<eta>\<sigma>uni finite_imageI finite_set_mset image_iff
mgu_sound set_mset_mset substitution_ops.is_mgu_def) (* should be simpler *)
ultimately show thesis
using that by auto
qed
\<comment> \<open>Lifting eligibility\<close>
have eligible0': "eligible S \<tau> As0' (D0' + negs (mset As0'))"
proof -
have "S_M S M (D + negs (mset As)) = negs (mset As) \<or> S_M S M (D + negs (mset As)) = {#} \<and>
length As = 1 \<and> maximal_wrt (As ! 0 \<cdot>a \<sigma>) ((D + negs (mset As)) \<cdot> \<sigma>)"
using eligible unfolding eligible.simps by auto
then show ?thesis
proof
assume "S_M S M (D + negs (mset As)) = negs (mset As)"
then have "S_M S M (D + negs (mset As)) \<noteq> {#}"
using n by force
then have "S (D0' + negs (mset As0')) = negs (mset As0')"
using as0' DA0'_split by auto
then show ?thesis
unfolding eligible.simps[simplified] by auto
next
assume asm: "S_M S M (D + negs (mset As)) = {#} \<and> length As = 1 \<and>
maximal_wrt (As ! 0 \<cdot>a \<sigma>) ((D + negs (mset As)) \<cdot> \<sigma>)"
then have "S (D0' + negs (mset As0')) = {#}"
using \<open>D0' \<cdot> \<eta> = D\<close>[symmetric] \<open>As0' \<cdot>al \<eta> = As\<close>[symmetric] \<open>S (DA0') \<cdot> \<eta> = S_M S M (DA)\<close>
da DA0'_split subst_cls_empty_iff by metis
moreover from asm have l: "length As0' = 1"
using \<open>As0' \<cdot>al \<eta> = As\<close> by auto
moreover from asm have "maximal_wrt (As0' ! 0 \<cdot>a (\<tau> \<odot> \<phi>)) ((D0' + negs (mset As0')) \<cdot> (\<tau> \<odot> \<phi>))"
using \<open>As0' \<cdot>al \<eta> = As\<close> \<open>D0' \<cdot> \<eta> = D\<close> using l \<tau>\<phi> by auto
then have "maximal_wrt (As0' ! 0 \<cdot>a \<tau> \<cdot>a \<phi>) ((D0' + negs (mset As0')) \<cdot> \<tau> \<cdot> \<phi>)"
by auto
then have "maximal_wrt (As0' ! 0 \<cdot>a \<tau>) ((D0' + negs (mset As0')) \<cdot> \<tau>)"
using maximal_wrt_subst by blast
ultimately show ?thesis
unfolding eligible.simps[simplified] by auto
qed
qed
\<comment> \<open>Lifting maximality\<close>
have maximality: "\<forall>i < n. strictly_maximal_wrt (As0' ! i \<cdot>a \<tau>) (Cs0' ! i \<cdot> \<tau>)"
(* Reformulate in list notation? *)
proof -
from str_max have "\<forall>i < n. strictly_maximal_wrt ((As0' \<cdot>al \<eta>) ! i \<cdot>a \<sigma>) ((Cs0' \<cdot>cl \<eta>) ! i \<cdot> \<sigma>)"
using \<open>As0' \<cdot>al \<eta> = As\<close> \<open>Cs0' \<cdot>cl \<eta> = Cs\<close> by simp
then have "\<forall>i < n. strictly_maximal_wrt (As0' ! i \<cdot>a (\<tau> \<odot> \<phi>)) (Cs0' ! i \<cdot> (\<tau> \<odot> \<phi>))"
using n \<tau>\<phi> by simp
then have "\<forall>i < n. strictly_maximal_wrt (As0' ! i \<cdot>a \<tau> \<cdot>a \<phi>) (Cs0' ! i \<cdot> \<tau> \<cdot> \<phi>)"
by auto
then show "\<forall>i < n. strictly_maximal_wrt (As0' ! i \<cdot>a \<tau>) (Cs0' ! i \<cdot> \<tau>)"
using strictly_maximal_wrt_subst \<tau>\<phi> by blast
qed
\<comment> \<open>Lifting nothing being selected\<close>
have nothing_selected: "\<forall>i < n. S (CAs0' ! i) = {#}"
proof -
have "\<forall>i < n. (map S CAs0' \<cdot>cl \<eta>) ! i = map (S_M S M) CAs ! i"
by (simp add: \<open>map S CAs0' \<cdot>cl \<eta> = map (S_M S M) CAs\<close>)
then have "\<forall>i < n. S (CAs0' ! i) \<cdot> \<eta> = S_M S M (CAs ! i)"
using n by auto
then have "\<forall>i < n. S (CAs0' ! i) \<cdot> \<eta> = {#}"
using sel_empt \<open>\<forall>i < n. S (CAs0' ! i) \<cdot> \<eta> = S_M S M (CAs ! i)\<close> by auto
then show "\<forall>i < n. S (CAs0' ! i) = {#}"
using subst_cls_empty_iff by blast
qed
\<comment> \<open>Lifting AAs0's non-emptiness\<close>
have "\<forall>i < n. AAs0' ! i \<noteq> {#}"
using n aas_not_empt \<open>AAs0' \<cdot>aml \<eta> = AAs\<close> by auto
\<comment> \<open>Resolve the lifted clauses\<close>
define E0' where
"E0' = ((\<Sum>\<^sub># (mset Cs0')) + D0') \<cdot> \<tau>"
have res_e0': "ord_resolve S CAs0' DA0' AAs0' As0' \<tau> E0'"
using ord_resolve.intros[of CAs0' n Cs0' AAs0' As0' \<tau> S D0',
OF _ _ _ _ _ _ \<open>\<forall>i < n. AAs0' ! i \<noteq> {#}\<close> \<tau>\<phi>(1) eligible0'
\<open>\<forall>i < n. strictly_maximal_wrt (As0' ! i \<cdot>a \<tau>) (Cs0' ! i \<cdot> \<tau>)\<close> \<open>\<forall>i < n. S (CAs0' ! i) = {#}\<close>]
unfolding E0'_def using DA0'_split n \<open>\<forall>i<n. CAs0' ! i = Cs0' ! i + poss (AAs0' ! i)\<close> by blast
\<comment> \<open>Prove resolvent instantiates to ground resolvent\<close>
have e0'\<phi>e: "E0' \<cdot> \<phi> = E"
proof -
have "E0' \<cdot> \<phi> = ((\<Sum>\<^sub># (mset Cs0')) + D0') \<cdot> (\<tau> \<odot> \<phi>)"
unfolding E0'_def by auto
also have "\<dots> = (\<Sum>\<^sub># (mset Cs0') + D0') \<cdot> (\<eta> \<odot> \<sigma>)"
using \<tau>\<phi> by auto
also have "\<dots> = (\<Sum>\<^sub># (mset Cs) + D) \<cdot> \<sigma>"
using \<open>Cs0' \<cdot>cl \<eta> = Cs\<close> \<open>D0' \<cdot> \<eta> = D\<close> by auto
also have "\<dots> = E"
using e by auto
finally show e0'\<phi>e: "E0' \<cdot> \<phi> = E"
.
qed
\<comment> \<open>Replace @{term \<phi>} with a true ground substitution\<close>
obtain \<eta>2 where
ground_\<eta>2: "is_ground_subst \<eta>2" "E0' \<cdot> \<eta>2 = E"
proof -
have "is_ground_cls_list CAs" "is_ground_cls DA"
using grounding grounding_ground unfolding is_ground_cls_list_def by auto
then have "is_ground_cls E"
using res_e ground_resolvent_subset by (force intro: is_ground_cls_mono)
then show thesis
using that e0'\<phi>e make_ground_subst by auto
qed
have \<open>length CAs0 = length CAs\<close>
using n by simp
have \<open>length \<eta>s0 = length CAs\<close>
using n by simp
\<comment> \<open>Wrap up the proof\<close>
have "ord_resolve S (CAs0 \<cdot>\<cdot>cl \<rho>s) (DA0 \<cdot> \<rho>) (AAs0 \<cdot>\<cdot>aml \<rho>s) (As0 \<cdot>al \<rho>) \<tau> E0'"
using res_e0' As0'_def \<rho>_def AAs0'_def \<rho>s_def DA0'_def \<rho>_def CAs0'_def \<rho>s_def by simp
moreover have "\<forall>i<n. poss (AAs0 ! i) \<subseteq># CAs0 ! i"
using as0(19) by auto
moreover have "negs (mset As0) \<subseteq># DA0"
using local.as0(15) by auto
ultimately have "ord_resolve_rename S CAs0 DA0 AAs0 As0 \<tau> E0'"
using ord_resolve_rename[of CAs0 n AAs0 As0 DA0 \<rho> \<rho>s S \<tau> E0'] \<rho>_def \<rho>s_def n by auto
then show thesis
using that[of \<eta>0 \<eta>s0 \<eta>2 CAs0 DA0] \<open>is_ground_subst \<eta>0\<close> \<open>is_ground_subst_list \<eta>s0\<close>
\<open>is_ground_subst \<eta>2\<close> \<open>CAs0 \<cdot>\<cdot>cl \<eta>s0 = CAs\<close> \<open>DA0 \<cdot> \<eta>0 = DA\<close> \<open>E0' \<cdot> \<eta>2 = E\<close> \<open>DA0 \<in> M\<close>
\<open>\<forall>CA \<in> set CAs0. CA \<in> M\<close> \<open>length CAs0 = length CAs\<close> \<open>length \<eta>s0 = length CAs\<close>
by blast
qed
lemma ground_ord_resolve_ground:
assumes
select: "selection S" and
CAs_p: "ground_resolution_with_selection.ord_resolve S CAs DA AAs As E" and
ground_cas: "is_ground_cls_list CAs" and
ground_da: "is_ground_cls DA"
shows "is_ground_cls E"
proof -
have a1: "atms_of E \<subseteq> (\<Union>CA \<in> set CAs. atms_of CA) \<union> atms_of DA"
using ground_resolution_with_selection.ord_resolve_atms_of_concl_subset[OF _ CAs_p]
ground_resolution_with_selection.intro[OF select] by blast
{
fix L :: "'a literal"
assume "L \<in># E"
then have "atm_of L \<in> atms_of E"
by (meson atm_of_lit_in_atms_of)
then have "is_ground_atm (atm_of L)"
using a1 ground_cas ground_da is_ground_cls_imp_is_ground_atm is_ground_cls_list_def
by auto
}
then show ?thesis
unfolding is_ground_cls_def is_ground_lit_def by simp
qed
lemma ground_ord_resolve_imp_ord_resolve:
assumes
ground_da: \<open>is_ground_cls DA\<close> and
ground_cas: \<open>is_ground_cls_list CAs\<close> and
gr: "ground_resolution_with_selection S_G" and
gr_res: \<open>ground_resolution_with_selection.ord_resolve S_G CAs DA AAs As E\<close>
shows \<open>\<exists>\<sigma>. ord_resolve S_G CAs DA AAs As \<sigma> E\<close>
proof (cases rule: ground_resolution_with_selection.ord_resolve.cases[OF gr gr_res])
case (1 CAs n Cs AAs As D)
note cas = this(1) and da = this(2) and aas = this(3) and as = this(4) and e = this(5) and
cas_len = this(6) and cs_len = this(7) and aas_len = this(8) and as_len = this(9) and
nz = this(10) and casi = this(11) and aas_not_empt = this(12) and as_aas = this(13) and
eligibility = this(14) and str_max = this(15) and sel_empt = this(16)
have len_aas_len_as: "length AAs = length As"
using aas_len as_len by auto
from as_aas have "\<forall>i < n. \<forall>A \<in># add_mset (As ! i) (AAs ! i). A = As ! i"
by simp
then have "\<forall>i < n. card (set_mset (add_mset (As ! i) (AAs ! i))) \<le> Suc 0"
using all_the_same by metis
then have "\<forall>i < length AAs. card (set_mset (add_mset (As ! i) (AAs ! i))) \<le> Suc 0"
using aas_len by auto
then have "\<forall>AA \<in> set (map2 add_mset As AAs). card (set_mset AA) \<le> Suc 0"
using set_map2_ex[of AAs As add_mset, OF len_aas_len_as] by auto
then have "is_unifiers id_subst (set_mset ` set (map2 add_mset As AAs))"
unfolding is_unifiers_def is_unifier_def by auto
moreover have "finite (set_mset ` set (map2 add_mset As AAs))"
by auto
moreover have "\<forall>AA \<in> set_mset ` set (map2 add_mset As AAs). finite AA"
by auto
ultimately obtain \<sigma> where
\<sigma>_p: "Some \<sigma> = mgu (set_mset ` set (map2 add_mset As AAs))"
using mgu_complete by metis
have ground_elig: "ground_resolution_with_selection.eligible S_G As (D + negs (mset As))"
using eligibility by simp
have ground_cs: "\<forall>i < n. is_ground_cls (Cs ! i)"
using cas cas_len cs_len casi ground_cas nth_mem unfolding is_ground_cls_list_def by force
have ground_set_as: "is_ground_atms (set As)"
using da ground_da by (metis atms_of_negs is_ground_cls_is_ground_atms_atms_of
is_ground_cls_union set_mset_mset)
then have ground_mset_as: "is_ground_atm_mset (mset As)"
unfolding is_ground_atm_mset_def is_ground_atms_def by auto
have ground_as: "is_ground_atm_list As"
using ground_set_as is_ground_atm_list_def is_ground_atms_def by auto
have ground_d: "is_ground_cls D"
using ground_da da by simp
from as_len nz have atms:
"atms_of D \<union> set As \<noteq> {}"
"finite (atms_of D \<union> set As)"
by auto
then have "Max (atms_of D \<union> set As) \<in> atms_of D \<union> set As"
using Max_in by metis
then have is_ground_Max: "is_ground_atm (Max (atms_of D \<union> set As))"
using ground_d ground_mset_as is_ground_cls_imp_is_ground_atm
unfolding is_ground_atm_mset_def by auto
have "maximal_wrt (Max (atms_of D \<union> set As)) (D + negs (mset As))"
unfolding maximal_wrt_def
by clarsimp (metis atms Max_less_iff UnCI ground_d ground_set_as infinite_growing
is_ground_Max is_ground_atms_def is_ground_cls_imp_is_ground_atm less_atm_ground)
moreover have
"Max (atms_of D \<union> set As) \<cdot>a \<sigma> = Max (atms_of D \<union> set As)" and
"D \<cdot> \<sigma> + negs (mset As \<cdot>am \<sigma>) = D + negs (mset As)"
using ground_elig is_ground_Max ground_mset_as ground_d by auto
ultimately have fo_elig: "eligible S_G \<sigma> As (D + negs (mset As))"
using ground_elig unfolding ground_resolution_with_selection.eligible.simps[OF gr]
ground_resolution_with_selection.maximal_wrt_def[OF gr] eligible.simps
by auto
have "\<forall>i < n. strictly_maximal_wrt (As ! i) (Cs ! i)"
using str_max[unfolded ground_resolution_with_selection.strictly_maximal_wrt_def[OF gr]]
ground_as[unfolded is_ground_atm_list_def] ground_cs as_len less_atm_ground
unfolding strictly_maximal_wrt_def by clarsimp (fastforce simp: is_ground_cls_as_atms)+
then have ll: "\<forall>i < n. strictly_maximal_wrt (As ! i \<cdot>a \<sigma>) (Cs ! i \<cdot> \<sigma>)"
by (simp add: ground_as ground_cs as_len)
have ground_e: "is_ground_cls E"
using ground_d ground_cs cs_len unfolding e is_ground_cls_def
by simp (metis in_mset_sum_list2 in_set_conv_nth)
show ?thesis
using cas da aas as e ground_e ord_resolve.intros[OF cas_len cs_len aas_len as_len nz casi
aas_not_empt \<sigma>_p fo_elig ll sel_empt]
by auto
qed
end
end
diff --git a/thys/PAC_Checker/More_Loops.thy b/thys/PAC_Checker/More_Loops.thy
--- a/thys/PAC_Checker/More_Loops.thy
+++ b/thys/PAC_Checker/More_Loops.thy
@@ -1,128 +1,128 @@
(*
File: More_Loops.thy
Author: Mathias Fleury, Daniela Kaufmann, JKU
Maintainer: Mathias Fleury, JKU
*)
theory More_Loops
imports
"Refine_Monadic.Refine_While"
"Refine_Monadic.Refine_Foreach"
"HOL-Library.Rewrite"
begin
subsection \<open>More Theorem about Loops\<close>
text \<open>Most theorem below have a counterpart in the Refinement Framework that is weaker (by missing
assertions for example that are critical for code generation).
\<close>
lemma Down_id_eq:
\<open>\<Down>Id x = x\<close>
by auto
lemma while_upt_while_direct1:
"b \<ge> a \<Longrightarrow>
do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (FOREACH_cond c) (\<lambda>x. do {ASSERT (FOREACH_cond c x); FOREACH_body f x})
([a..<b],\<sigma>);
RETURN \<sigma>
} \<le> do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (\<lambda>(i, x). i < b \<and> c x) (\<lambda>(i, x). do {ASSERT (i < b); \<sigma>'\<leftarrow>f i x; RETURN (i+1,\<sigma>')
}) (a,\<sigma>);
RETURN \<sigma>
}"
apply (rewrite at \<open>_ \<le> \<hole>\<close> Down_id_eq[symmetric])
apply (refine_vcg WHILET_refine[where R = \<open>{((l, x'), (i::nat, x::'a)). x= x' \<and> i \<le> b \<and> i \<ge> a \<and>
l = drop (i-a) [a..<b]}\<close>])
subgoal by auto
subgoal by (auto simp: FOREACH_cond_def)
subgoal by (auto simp: FOREACH_body_def intro!: bind_refine[OF Id_refine])
subgoal by auto
done
lemma while_upt_while_direct2:
"b \<ge> a \<Longrightarrow>
do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (FOREACH_cond c) (\<lambda>x. do {ASSERT (FOREACH_cond c x); FOREACH_body f x})
([a..<b],\<sigma>);
RETURN \<sigma>
} \<ge> do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (\<lambda>(i, x). i < b \<and> c x) (\<lambda>(i, x). do {ASSERT (i < b); \<sigma>'\<leftarrow>f i x; RETURN (i+1,\<sigma>')
}) (a,\<sigma>);
RETURN \<sigma>
}"
apply (rewrite at \<open>_ \<le> \<hole>\<close> Down_id_eq[symmetric])
apply (refine_vcg WHILET_refine[where R = \<open>{((i::nat, x::'a), (l, x')). x= x' \<and> i \<le> b \<and> i \<ge> a \<and>
l = drop (i-a) [a..<b]}\<close>])
subgoal by auto
subgoal by (auto simp: FOREACH_cond_def)
subgoal by (auto simp: FOREACH_body_def intro!: bind_refine[OF Id_refine])
subgoal by (auto simp: FOREACH_body_def intro!: bind_refine[OF Id_refine])
subgoal by auto
done
lemma while_upt_while_direct:
"b \<ge> a \<Longrightarrow>
do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (FOREACH_cond c) (\<lambda>x. do {ASSERT (FOREACH_cond c x); FOREACH_body f x})
([a..<b],\<sigma>);
RETURN \<sigma>
} = do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (\<lambda>(i, x). i < b \<and> c x) (\<lambda>(i, x). do {ASSERT (i < b); \<sigma>'\<leftarrow>f i x; RETURN (i+1,\<sigma>')
}) (a,\<sigma>);
RETURN \<sigma>
}"
using while_upt_while_direct1[of a b] while_upt_while_direct2[of a b]
- unfolding order_class.eq_iff by fast
+ unfolding order_eq_iff by fast
lemma while_nfoldli:
"do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (FOREACH_cond c) (\<lambda>x. do {ASSERT (FOREACH_cond c x); FOREACH_body f x}) (l,\<sigma>);
RETURN \<sigma>
} \<le> nfoldli l c f \<sigma>"
apply (induct l arbitrary: \<sigma>)
apply (subst WHILET_unfold)
apply (simp add: FOREACH_cond_def)
apply (subst WHILET_unfold)
apply (auto
simp: FOREACH_cond_def FOREACH_body_def
intro: bind_mono Refine_Basic.bind_mono(1))
done
lemma nfoldli_while: "nfoldli l c f \<sigma>
\<le>
(WHILE\<^sub>T\<^bsup>I\<^esup>
(FOREACH_cond c) (\<lambda>x. do {ASSERT (FOREACH_cond c x); FOREACH_body f x}) (l, \<sigma>) \<bind>
(\<lambda>(_, \<sigma>). RETURN \<sigma>))"
proof (induct l arbitrary: \<sigma>)
case Nil thus ?case by (subst WHILEIT_unfold) (auto simp: FOREACH_cond_def)
next
case (Cons x ls)
show ?case
proof (cases "c \<sigma>")
case False thus ?thesis
apply (subst WHILEIT_unfold)
unfolding FOREACH_cond_def
by simp
next
case [simp]: True
from Cons show ?thesis
apply (subst WHILEIT_unfold)
unfolding FOREACH_cond_def FOREACH_body_def
apply clarsimp
apply (rule Refine_Basic.bind_mono)
apply simp_all
done
qed
qed
lemma while_eq_nfoldli: "do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (FOREACH_cond c) (\<lambda>x. do {ASSERT (FOREACH_cond c x); FOREACH_body f x}) (l,\<sigma>);
RETURN \<sigma>
} = nfoldli l c f \<sigma>"
apply (rule antisym)
apply (rule while_nfoldli)
apply (rule order_trans[OF nfoldli_while[where I="\<lambda>_. True"]])
apply (simp add: WHILET_def)
done
end
\ No newline at end of file
diff --git a/thys/PSemigroupsConvolution/Quantales.thy b/thys/PSemigroupsConvolution/Quantales.thy
--- a/thys/PSemigroupsConvolution/Quantales.thy
+++ b/thys/PSemigroupsConvolution/Quantales.thy
@@ -1,553 +1,553 @@
(* Title: Quantales
Author: Brijesh Dongol, Victor Gomes, Ian J Hayes, Georg Struth
Maintainer: Victor Gomes <victor.gomes@cl.cam.ac.uk>
Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Quantales\<close>
text \<open>This entry will be merged eventually with other quantale entries and become a standalone one.\<close>
theory Quantales
imports Main
begin
notation sup (infixl "\<squnion>" 60)
and inf (infixl "\<sqinter>" 55)
and top ("\<top>")
and bot ("\<bottom>")
and relcomp (infixl ";" 70)
and times (infixl "\<cdot>" 70)
and Sup ("\<Squnion>_" [900] 900)
and Inf ("\<Sqinter>_" [900] 900)
subsection \<open>Properties of Complete Lattices\<close>
lemma (in complete_lattice) Sup_sup_pred: "x \<squnion> \<Squnion>{y. P y} = \<Squnion>{y. y = x \<or> P y}"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: Collect_mono Sup_subset_mono Sup_upper)
using Sup_least Sup_upper sup.coboundedI2 by force
lemma (in complete_lattice) sup_Sup: "x \<squnion> y = \<Squnion>{x,y}"
by simp
lemma (in complete_lattice) sup_Sup_var: "x \<squnion> y = \<Squnion>{z. z \<in> {x,y}}"
by (metis Collect_mem_eq sup_Sup)
lemma (in complete_boolean_algebra) shunt1: "x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> -y \<squnion> z"
proof standard
assume "x \<sqinter> y \<le> z"
hence "-y \<squnion> (x \<sqinter> y) \<le> -y \<squnion> z"
using sup.mono by blast
hence "-y \<squnion> x \<le> -y \<squnion> z"
by (simp add: sup_inf_distrib1)
thus "x \<le> -y \<squnion> z"
by simp
next
assume "x \<le> - y \<squnion> z"
hence "x \<sqinter> y \<le> (-y \<squnion> z) \<sqinter> y"
using inf_mono by auto
thus "x \<sqinter> y \<le> z"
using inf.boundedE inf_sup_distrib2 by auto
qed
lemma (in complete_boolean_algebra) meet_shunt: "x \<sqinter> y = \<bottom> \<longleftrightarrow> x \<le> -y"
by (metis bot_least inf_absorb2 inf_compl_bot_left2 shunt1 sup_absorb1)
lemma (in complete_boolean_algebra) join_shunt: "x \<squnion> y = \<top> \<longleftrightarrow> -x \<le> y"
by (metis compl_sup compl_top_eq double_compl meet_shunt)
subsection \<open> Familes of Proto-Quantales\<close>
text \<open>Proto-Quanales are complete lattices equipped with an operation of composition or multiplication
that need not be associative.\<close>
class proto_near_quantale = complete_lattice + times +
assumes Sup_distr: "\<Squnion>X \<cdot> y = \<Squnion>{x \<cdot> y |x. x \<in> X}"
begin
lemma mult_botl [simp]: "\<bottom> \<cdot> x = \<bottom>"
using Sup_distr[where X="{}"] by auto
lemma sup_distr: "(x \<squnion> y) \<cdot> z = (x \<cdot> z) \<squnion> (y \<cdot> z)"
using Sup_distr[where X="{x, y}"] by (fastforce intro!: Sup_eqI)
lemma mult_isor: "x \<le> y \<Longrightarrow> x \<cdot> z \<le> y \<cdot> z"
by (metis sup.absorb_iff1 sup_distr)
definition bres :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<rightarrow>" 60) where
"x \<rightarrow> z = \<Squnion>{y. x \<cdot> y \<le> z}"
definition fres :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<leftarrow>" 60) where
"z \<leftarrow> y = \<Squnion>{x. x \<cdot> y \<le> z}"
lemma bres_galois_imp: "x \<cdot> y \<le> z \<longrightarrow> y \<le> x \<rightarrow> z"
by (simp add: Sup_upper bres_def)
lemma fres_galois: "x \<cdot> y \<le> z \<longleftrightarrow> x \<le> z \<leftarrow> y"
proof
show "x \<cdot> y \<le> z \<Longrightarrow> x \<le> z \<leftarrow> y"
by (simp add: Sup_upper fres_def)
next
assume "x \<le> z \<leftarrow> y"
hence "x \<cdot> y \<le> \<Squnion>{x. x \<cdot> y \<le> z} \<cdot> y"
by (simp add: fres_def mult_isor)
also have "... = \<Squnion>{x \<cdot> y |x. x \<cdot> y \<le> z}"
by (simp add: Sup_distr)
also have "... \<le> z"
by (rule Sup_least, auto)
finally show "x \<cdot> y \<le> z" .
qed
end
class proto_pre_quantale = proto_near_quantale +
assumes Sup_subdistl: "\<Squnion>{x \<cdot> y | y . y \<in> Y} \<le> x \<cdot> \<Squnion>Y"
begin
lemma sup_subdistl: "(x \<cdot> y) \<squnion> (x \<cdot> z) \<le> x \<cdot> (y \<squnion> z)"
using Sup_subdistl[where Y="{y, z}"] Sup_le_iff by auto
lemma mult_isol: "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
by (metis le_iff_sup le_sup_iff sup_subdistl)
end
class weak_proto_quantale = proto_near_quantale +
assumes weak_Sup_distl: "Y \<noteq> {} \<Longrightarrow> x \<cdot> \<Squnion>Y = \<Squnion>{x \<cdot> y |y. y \<in> Y}"
begin
subclass proto_pre_quantale
proof standard
have a: "\<And>x Y. Y = {} \<Longrightarrow> \<Squnion>{x \<cdot> y |y. y \<in> Y} \<le> x \<cdot> \<Squnion>Y"
by simp
have b: "\<And>x Y. Y \<noteq> {} \<Longrightarrow> \<Squnion>{x \<cdot> y |y. y \<in> Y} \<le> x \<cdot> \<Squnion>Y"
by (simp add: weak_Sup_distl)
show "\<And>x Y. \<Squnion>{x \<cdot> y |y. y \<in> Y} \<le> x \<cdot> \<Squnion>Y"
using a b by blast
qed
lemma sup_distl: "x \<cdot> (y \<squnion> z) = (x \<cdot> y) \<squnion> (x \<cdot> z)"
using weak_Sup_distl[where Y="{y, z}"] by (fastforce intro!: Sup_eqI)
lemma "y \<le> x \<rightarrow> z \<longrightarrow> x \<cdot> y \<le> z" (* nitpick [expect = genuine] *)
oops
end
class proto_quantale = proto_near_quantale +
assumes Sup_distl: "x \<cdot> \<Squnion>Y = \<Squnion>{x \<cdot> y |y. y \<in> Y}"
begin
subclass weak_proto_quantale
by standard (simp add: Sup_distl)
lemma bres_galois: "x \<cdot> y \<le> z \<longleftrightarrow> y \<le> x \<rightarrow> z"
proof
show "x \<cdot> y \<le> z \<Longrightarrow> y \<le> x \<rightarrow> z"
by (simp add: Sup_upper bres_def)
next
assume "y \<le> x \<rightarrow> z"
hence "x \<cdot> y \<le> x \<cdot> \<Squnion>{y. x \<cdot> y \<le> z}"
by (simp add: bres_def mult_isol)
also have "... = \<Squnion>{x \<cdot> y |y. x \<cdot> y \<le> z}"
by (simp add: Sup_distl)
also have "... \<le> z"
by (rule Sup_least, auto)
finally show "x \<cdot> y \<le> z" .
qed
end
subsection \<open>Families of Quantales\<close>
class near_quantale = proto_near_quantale + semigroup_mult
class unital_near_quantale = near_quantale + monoid_mult
begin
definition iter :: "'a \<Rightarrow> 'a" where
"iter x \<equiv> \<Sqinter>{y. \<exists>i. y = x ^ i}"
lemma iter_ref [simp]: "iter x \<le> 1"
apply (simp add: iter_def)
by (metis (mono_tags, lifting) Inf_lower local.power.power_0 mem_Collect_eq)
lemma le_top: "x \<le> \<top> \<cdot> x"
by (metis mult_isor mult.monoid_axioms top_greatest monoid.left_neutral)
end
class pre_quantale = proto_pre_quantale + semigroup_mult
subclass (in pre_quantale) near_quantale ..
class unital_pre_quantale = pre_quantale + monoid_mult
subclass (in unital_pre_quantale) unital_near_quantale ..
class weak_quantale = weak_proto_quantale + semigroup_mult
subclass (in weak_quantale) pre_quantale ..
text \<open>The following counterexample shows an important consequence of weakness:
the absence of right annihilation.\<close>
lemma (in weak_quantale) "x \<cdot> \<bottom> = \<bottom>" (*nitpick[expect=genuine]*)
oops
class unital_weak_quantale = weak_quantale + monoid_mult
lemma (in unital_weak_quantale) "x \<cdot> \<bottom> = \<bottom>" (*nitpick[expect=genuine]*)
oops
subclass (in unital_weak_quantale) unital_pre_quantale ..
class quantale = proto_quantale + semigroup_mult
begin
subclass weak_quantale ..
lemma mult_botr [simp]: "x \<cdot> \<bottom> = \<bottom>"
using Sup_distl[where Y="{}"] by auto
end
class unital_quantale = quantale + monoid_mult
subclass (in unital_quantale) unital_weak_quantale ..
class ab_quantale = quantale + ab_semigroup_mult
begin
lemma bres_fres_eq: "x \<rightarrow> y = y \<leftarrow> x"
by (simp add: fres_def bres_def mult_commute)
end
class ab_unital_quantale = ab_quantale + unital_quantale
class distrib_quantale = quantale + complete_distrib_lattice
class bool_quantale = quantale + complete_boolean_algebra
class distrib_unital_quantale = unital_quantale + complete_distrib_lattice
class bool_unital_quantale = unital_quantale + complete_boolean_algebra
class distrib_ab_quantale = distrib_quantale + ab_quantale
class bool_ab_quantale = bool_quantale + ab_quantale
class distrib_ab_unital_quantale = distrib_quantale + unital_quantale
class bool_ab_unital_quantale = bool_ab_quantale + unital_quantale
subsection \<open>Quantales of Booleans and Complete Boolean Algebras\<close>
instantiation bool :: bool_ab_unital_quantale
begin
definition "one_bool = True"
definition "times_bool = (\<lambda>x y. x \<and> y)"
instance
by standard (auto simp: times_bool_def one_bool_def)
end
context complete_distrib_lattice
begin
interpretation cdl_quantale: distrib_quantale _ _ _ _ _ _ _ _ inf
by standard (simp_all add: inf.assoc Setcompr_eq_image Sup_inf inf_Sup)
end
context complete_boolean_algebra
begin
interpretation cba_quantale: bool_ab_unital_quantale inf _ _ _ _ _ _ _ _ _ _ top
apply standard
apply (simp add: inf.assoc)
apply (simp add: inf.commute)
apply (simp add: Setcompr_eq_image Sup_inf)
apply (simp add: Setcompr_eq_image inf_Sup)
by simp_all
text \<open>In this setting, residuation can be translated like classical implication.\<close>
lemma cba_bres1: "x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> cba_quantale.bres y z"
using cba_quantale.bres_galois inf.commute by fastforce
lemma cba_bres2: "x \<le> -y \<squnion> z \<longleftrightarrow> x \<le> cba_quantale.bres y z"
using cba_bres1 shunt1 by auto
lemma cba_bres_prop: "cba_quantale.bres x y = -x \<squnion> y"
- using cba_bres2 eq_iff by blast
+ using cba_bres2 order.eq_iff by blast
end
text \<open>Other models will follow.\<close>
subsection \<open>Products of Quantales\<close>
definition "Inf_prod X = (\<Sqinter>{fst x |x. x \<in> X}, \<Sqinter>{snd x |x. x \<in> X})"
definition "inf_prod x y = (fst x \<sqinter> fst y, snd x \<sqinter> snd y)"
definition "bot_prod = (bot,bot)"
definition "Sup_prod X = (\<Squnion>{fst x |x. x \<in> X}, \<Squnion>{snd x |x. x \<in> X})"
definition "sup_prod x y = (fst x \<squnion> fst y, snd x \<squnion> snd y)"
definition "top_prod = (top,top)"
definition "less_eq_prod x y \<equiv> less_eq (fst x) (fst y) \<and> less_eq (snd x) (snd y)"
definition "less_prod x y \<equiv> less_eq (fst x) (fst y) \<and> less_eq (snd x) (snd y) \<and> x \<noteq> y"
definition "times_prod' x y = (fst x \<cdot> fst y, snd x \<cdot> snd y)"
definition "one_prod = (1,1)"
interpretation prod: complete_lattice Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::complete_lattice \<times> 'b::complete_lattice)"
apply standard
apply (simp_all add: Inf_prod_def Sup_prod_def inf_prod_def sup_prod_def bot_prod_def top_prod_def less_eq_prod_def less_prod_def Sup_distl Sup_distr)
apply force+
apply (rule conjI, (rule Inf_lower, force)+)
apply (rule conjI, (rule Inf_greatest, force)+)
apply (rule conjI, (rule Sup_upper, force)+)
by (rule conjI, (rule Sup_least, force)+)
interpretation prod: proto_near_quantale Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::proto_near_quantale \<times> 'b::proto_near_quantale)" times_prod'
apply (standard, simp add: times_prod'_def Sup_prod_def)
by (rule conjI, (simp add: Sup_distr, clarify, metis)+)
interpretation prod: proto_quantale Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::proto_quantale \<times> 'b::proto_quantale)" times_prod'
apply (standard, simp add: times_prod'_def Sup_prod_def less_eq_prod_def)
by (rule conjI, (simp add: Sup_distl, metis)+)
interpretation prod: unital_quantale one_prod times_prod' Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::unital_quantale \<times> 'b::unital_quantale)"
by standard (simp_all add: one_prod_def times_prod'_def mult.assoc)
subsection \<open>Quantale Modules and Semidirect Products\<close>
text \<open>Quantale modules are extensions of semigroup actions in that a quantale acts on a complete lattice.\<close>
locale unital_quantale_module =
fixes act :: "'a::unital_quantale \<Rightarrow> 'b::complete_lattice \<Rightarrow> 'b" ("\<alpha>")
assumes act1: "\<alpha> (x \<cdot> y) p = \<alpha> x (\<alpha> y p)"
and act2 [simp]: "\<alpha> 1 p = p"
and act3: "\<alpha> (\<Squnion>X) p = \<Squnion>{\<alpha> x p |x. x \<in> X}"
and act4: "\<alpha> x (\<Squnion>P) = \<Squnion>{\<alpha> x p |p. p \<in> P}"
context unital_quantale_module
begin
text \<open>Actions are morphisms. The curried notation is particularly convenient for this.\<close>
lemma act_morph1: "\<alpha> (x \<cdot> y) = (\<alpha> x) \<circ> (\<alpha> y)"
by (simp add: fun_eq_iff act1)
lemma act_morph2: "\<alpha> 1 = id"
by (simp add: fun_eq_iff)
lemma emp_act: "\<alpha> (\<Squnion>{}) p = \<bottom>"
by (simp only: act3, force)
lemma emp_act_var: "\<alpha> \<bottom> p = \<bottom>"
using emp_act by auto
lemma act_emp: "\<alpha> x (\<Squnion>{}) = \<bottom>"
by (simp only: act4, force)
lemma act_emp_var: "\<alpha> x \<bottom> = \<bottom>"
using act_emp by auto
lemma act_sup_distl: "\<alpha> x (p \<squnion> q) = (\<alpha> x p) \<squnion> (\<alpha> x q)"
proof-
have "\<alpha> x (p \<squnion> q) = \<alpha> x (\<Squnion>{p,q})"
by simp
also have "... = \<Squnion>{\<alpha> x y |y. y \<in> {p,q}}"
by (rule act4)
also have "... = \<Squnion>{v. v = \<alpha> x p \<or> v = \<alpha> x q}"
by (metis empty_iff insert_iff)
finally show ?thesis
by (metis Collect_disj_eq Sup_insert ccpo_Sup_singleton insert_def singleton_conv)
qed
lemma act_sup_distr: "\<alpha> (x \<squnion> y) p = (\<alpha> x p) \<squnion> (\<alpha> y p)"
proof-
have "\<alpha> (x \<squnion> y) p = \<alpha> (\<Squnion>{x,y}) p"
by simp
also have "... = \<Squnion>{\<alpha> v p |v. v \<in> {x,y}}"
by (rule act3)
also have "... = \<Squnion>{v. v = \<alpha> x p \<or> v = \<alpha> y p}"
by (metis empty_iff insert_iff)
finally show ?thesis
by (metis Collect_disj_eq Sup_insert ccpo_Sup_singleton insert_def singleton_conv)
qed
lemma act_sup_distr_var: "\<alpha> (x \<squnion> y) = (\<alpha> x) \<squnion> (\<alpha> y)"
by (simp add: fun_eq_iff act_sup_distr)
text \<open>Next we define the semidirect product of a unital quantale and a complete lattice. \<close>
definition "sd_prod x y = (fst x \<cdot> fst y, snd x \<squnion> \<alpha> (fst x) (snd y))"
lemma sd_distr_aux:
"\<Squnion>{snd x |x. x \<in> X} \<squnion> \<Squnion>{\<alpha> (fst x) p |x. x \<in> X} = \<Squnion>{snd x \<squnion> \<alpha> (fst x) p |x. x \<in> X}"
proof (rule antisym, rule sup_least)
show "\<Squnion>{snd x |x. x \<in> X} \<le> \<Squnion>{snd x \<squnion> \<alpha> (fst x) p |x. x \<in> X}"
proof (rule Sup_least)
fix x :: 'b
assume "x \<in> {snd x |x. x \<in> X}"
hence "\<exists>b pa. x \<squnion> b = snd pa \<squnion> \<alpha> (fst pa) p \<and> pa \<in> X"
by blast
hence "\<exists>b. x \<squnion> b \<in> {snd pa \<squnion> \<alpha> (fst pa) p |pa. pa \<in> X}"
by blast
thus "x \<le> \<Squnion>{snd pa \<squnion> \<alpha> (fst pa) p |pa. pa \<in> X}"
by (meson Sup_upper sup.bounded_iff)
qed
next
show "\<Squnion>{\<alpha> (fst x) p |x. x \<in> X} \<le> \<Squnion>{snd x \<squnion> \<alpha> (fst x) p |x. x \<in> X}"
proof (rule Sup_least)
fix x :: 'b
assume "x \<in> {\<alpha> (fst x) p |x. x \<in> X}"
then have "\<exists>b pa. b \<squnion> x = snd pa \<squnion> \<alpha> (fst pa) p \<and> pa \<in> X"
by blast
then have "\<exists>b. b \<squnion> x \<in> {snd pa \<squnion> \<alpha> (fst pa) p |pa. pa \<in> X}"
by blast
then show "x \<le> \<Squnion>{snd pa \<squnion> \<alpha> (fst pa) p |pa. pa \<in> X}"
by (meson Sup_upper le_supE)
qed
next
show "\<Squnion>{snd x \<squnion> \<alpha> (fst x) p |x. x \<in> X} \<le> \<Squnion>{snd x |x. x \<in> X} \<squnion> \<Squnion>{\<alpha> (fst x) p |x. x \<in> X}"
apply (rule Sup_least)
apply safe
apply (rule sup_least)
apply (metis (mono_tags, lifting) Sup_upper mem_Collect_eq sup.coboundedI1)
by (metis (mono_tags, lifting) Sup_upper mem_Collect_eq sup.coboundedI2)
qed
lemma sd_distr: "sd_prod (Sup_prod X) y = Sup_prod {sd_prod x y |x. x \<in> X}"
proof -
have "sd_prod (Sup_prod X) y = sd_prod (\<Squnion>{fst x |x. x \<in> X}, \<Squnion>{snd x |x. x \<in> X}) y"
by (simp add: Sup_prod_def)
also have
"... = ((\<Squnion>{fst x |x. x \<in> X}) \<cdot> fst y, (\<Squnion>{snd x |x. x \<in> X}) \<squnion> (\<alpha> (\<Squnion>{fst x | x. x \<in> X}) (snd y)))"
by (simp add: sd_prod_def)
also have
"... = (\<Squnion>{fst x \<cdot> fst y|x. x \<in> X}, (\<Squnion>{snd x |x. x \<in> X}) \<squnion> (\<alpha> (\<Squnion>{fst x | x. x \<in> X}) (snd y)))"
by (simp add: Sup_distr)
also have
"... = (\<Squnion>{fst x \<cdot> fst y|x. x \<in> X}, (\<Squnion> {snd x |x. x \<in> X}) \<squnion> (\<Squnion>{\<alpha> (fst x) (snd y) | x. x \<in> X}))"
by (simp add: act3)
also have "... = (\<Squnion>{fst x \<cdot> fst y|x. x \<in> X}, \<Squnion>{snd x \<squnion> (\<alpha> (fst x) (snd y)) | x. x \<in> X})"
by (simp only: sd_distr_aux)
also have "... = Sup_prod {(fst x \<cdot> fst y, snd x \<squnion> (\<alpha> (fst x) (snd y))) |x. x \<in> X}"
by (simp add: Sup_prod_def, metis)
finally show ?thesis
by (simp add: sd_prod_def)
qed
lemma sd_distl_aux: "Y \<noteq> {} \<Longrightarrow> p \<squnion> (\<Squnion>{\<alpha> x (snd y) |y. y \<in> Y}) = \<Squnion>{p \<squnion> \<alpha> x (snd y) |y. y \<in> Y}"
proof (rule antisym, rule sup_least)
show "Y \<noteq> {} \<Longrightarrow> p \<le> \<Squnion>{p \<squnion> \<alpha> x (snd y) |y. y \<in> Y}"
proof -
assume "Y \<noteq> {}"
hence "\<exists>b. b \<in> {p \<squnion> \<alpha> x (snd pa) |pa. pa \<in> Y} \<and> p \<le> b"
by fastforce
thus "p \<le> \<Squnion>{p \<squnion> \<alpha> x (snd pa) |pa. pa \<in> Y}"
by (meson Sup_upper2)
qed
next
show "Y \<noteq> {} \<Longrightarrow> \<Squnion>{\<alpha> x (snd y) |y. y \<in> Y} \<le> \<Squnion>{p \<squnion> \<alpha> x (snd y) |y. y \<in> Y}"
apply (rule Sup_least)
proof -
fix xa :: 'b
assume "xa \<in> {\<alpha> x (snd y) |y. y \<in> Y}"
then have "\<exists>b. (\<exists>pa. b = p \<squnion> \<alpha> x (snd pa) \<and> pa \<in> Y) \<and> xa \<le> b"
by fastforce
then have "\<exists>b. b \<in> {p \<squnion> \<alpha> x (snd pa) |pa. pa \<in> Y} \<and> xa \<le> b"
by blast
then show "xa \<le> \<Squnion>{p \<squnion> \<alpha> x (snd pa) |pa. pa \<in> Y}"
by (meson Sup_upper2)
qed
next
show "Y \<noteq> {} \<Longrightarrow> \<Squnion>{p \<squnion> \<alpha> x (snd y) |y. y \<in> Y} \<le> p \<squnion> \<Squnion>{\<alpha> x (snd y) |y. y \<in> Y}"
apply (rule Sup_least)
apply safe
by (metis (mono_tags, lifting) Sup_le_iff le_sup_iff mem_Collect_eq sup_ge1 sup_ge2)
qed
lemma sd_distl: "Y \<noteq> {} \<Longrightarrow> sd_prod x (Sup_prod Y) = Sup_prod {sd_prod x y |y. y \<in> Y}"
proof -
assume a: "Y \<noteq> {}"
have "sd_prod x (Sup_prod Y) = sd_prod x (\<Squnion>{fst y |y. y \<in> Y}, \<Squnion>{snd y |y. y \<in> Y})"
by (simp add: Sup_prod_def)
also have "... = ((fst x) \<cdot> (\<Squnion>{fst y |y. y \<in> Y}), (snd x \<squnion> (\<alpha> (fst x) (\<Squnion>{snd y |y. y \<in> Y}))))"
by (simp add: sd_prod_def)
also have "... = (\<Squnion>{fst x \<cdot> fst y |y. y \<in> Y}, (snd x \<squnion> (\<alpha> (fst x) (\<Squnion>{snd y |y. y \<in> Y}))))"
by (simp add: Sup_distl)
also have "... = (\<Squnion>{fst x \<cdot> fst y |y. y \<in> Y}, (snd x \<squnion> (\<Squnion>{\<alpha> (fst x) (snd y) |y. y \<in> Y})))"
by (simp add: act4, meson)
also have "... = (\<Squnion>{fst x \<cdot> fst y |y. y \<in> Y}, \<Squnion>{snd x \<squnion> (\<alpha> (fst x) (snd y)) |y. y \<in> Y})"
using a sd_distl_aux by blast
also have "... = Sup_prod {(fst x \<cdot> fst y, snd x \<squnion> (\<alpha> (fst x) (snd y))) |y. y \<in> Y}"
by (simp add: Sup_prod_def, metis)
finally show ?thesis
by (simp add: sd_prod_def)
qed
definition "sd_unit = (1,\<bottom>)"
lemma sd_unitl [simp]: "sd_prod sd_unit x = x"
by (simp add: sd_prod_def sd_unit_def)
lemma sd_unitr [simp]: "sd_prod x sd_unit = x"
apply (simp add: sd_prod_def sd_unit_def)
using act_emp by force
text \<open>The following counterexamples rule out that semidirect products of quantales and complete lattices form quantales.
The reason is that the right annihilation law fails.\<close>
lemma "sd_prod x (Sup_prod Y) = Sup_prod {sd_prod x y |y. y \<in> Y}" (*nitpick[show_all,expect=genuine]*)
oops
lemma "sd_prod x bot_prod = bot_prod" (*nitpick[show_all,expect=genuine]*)
oops
text \<open>However we can show that semidirect products of (unital) quantales with complete lattices form weak (unital) quantales. \<close>
interpretation dp_quantale: weak_quantale sd_prod Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod top_prod
apply standard
apply (simp_all add: sd_distl sd_distr)
apply (simp_all add: sd_prod_def Inf_prod_def Sup_prod_def bot_prod_def sup_prod_def top_prod_def inf_prod_def less_eq_prod_def less_prod_def)
by (rule conjI, simp add: mult.assoc, simp add: act1 act_sup_distl sup_assoc)
interpretation dpu_quantale: unital_weak_quantale sd_unit sd_prod Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod top_prod
by (standard; simp_all)
end
end
diff --git a/thys/Partial_Order_Reduction/Extensions/CCPO_Extensions.thy b/thys/Partial_Order_Reduction/Extensions/CCPO_Extensions.thy
--- a/thys/Partial_Order_Reduction/Extensions/CCPO_Extensions.thy
+++ b/thys/Partial_Order_Reduction/Extensions/CCPO_Extensions.thy
@@ -1,162 +1,162 @@
section \<open>Chain-Complete Partial Orders\<close>
theory CCPO_Extensions
imports
"HOL-Library.Complete_Partial_Order2"
ENat_Extensions
Set_Extensions
begin
lemma chain_split[dest]:
assumes "Complete_Partial_Order.chain ord C" "x \<in> C"
shows "C = {y \<in> C. ord x y} \<union> {y \<in> C. ord y x}"
proof -
have 1: "\<And> y. y \<in> C \<Longrightarrow> ord x y \<or> ord y x" using chainD assms by this
show ?thesis using 1 by blast
qed
lemma infinite_chain_below[dest]:
assumes "Complete_Partial_Order.chain ord C" "infinite C" "x \<in> C"
assumes "finite {y \<in> C. ord x y}"
shows "infinite {y \<in> C. ord y x}"
proof -
have 1: "C = {y \<in> C. ord x y} \<union> {y \<in> C. ord y x}" using assms(1, 3) by rule
show ?thesis using finite_Un assms(2, 4) 1 by (metis (poly_guards_query))
qed
lemma infinite_chain_above[dest]:
assumes "Complete_Partial_Order.chain ord C" "infinite C" "x \<in> C"
assumes "finite {y \<in> C. ord y x}"
shows "infinite {y \<in> C. ord x y}"
proof -
have 1: "C = {y \<in> C. ord x y} \<union> {y \<in> C. ord y x}" using assms(1, 3) by rule
show ?thesis using finite_Un assms(2, 4) 1 by (metis (poly_guards_query))
qed
lemma (in ccpo) ccpo_Sup_upper_inv:
assumes "Complete_Partial_Order.chain less_eq C" "x > \<Squnion> C"
shows "x \<notin> C"
using assms ccpo_Sup_upper by fastforce
lemma (in ccpo) ccpo_Sup_least_inv:
assumes "Complete_Partial_Order.chain less_eq C" "\<Squnion> C > x"
obtains y
where "y \<in> C" "\<not> y \<le> x"
using assms ccpo_Sup_least that by fastforce
lemma ccpo_Sup_least_inv':
fixes C :: "'a :: {ccpo, linorder} set"
assumes "Complete_Partial_Order.chain less_eq C" "\<Squnion> C > x"
obtains y
where "y \<in> C" "y > x"
proof -
obtain y where 1: "y \<in> C" "\<not> y \<le> x" using ccpo_Sup_least_inv assms by this
show ?thesis using that 1 by simp
qed
lemma mcont2mcont_lessThan[THEN lfp.mcont2mcont, simp, cont_intro]:
shows mcont_lessThan: "mcont Sup less_eq Sup less_eq
(lessThan :: 'a :: {ccpo, linorder} \<Rightarrow> 'a set)"
proof
show "monotone less_eq less_eq (lessThan :: 'a \<Rightarrow> 'a set)" by (rule, auto)
show "cont Sup less_eq Sup less_eq (lessThan :: 'a \<Rightarrow> 'a set)"
proof
fix C :: "'a set"
assume 1: "Complete_Partial_Order.chain less_eq C"
show "{..< \<Squnion> C} = \<Union> (lessThan ` C)"
proof (intro equalityI subsetI)
fix A
assume 2: "A \<in> {..< \<Squnion> C}"
obtain B where 3: "B \<in> C" "B > A" using ccpo_Sup_least_inv' 1 2 by blast
show "A \<in> \<Union> (lessThan ` C)" using 3 by auto
next
fix A
assume 2: "A \<in> \<Union> (lessThan ` C)"
show "A \<in> {..< \<Squnion> C}" using ccpo_Sup_upper 2 by force
qed
qed
qed
class esize =
fixes esize :: "'a \<Rightarrow> enat"
class esize_order = esize + order +
assumes esize_finite[dest]: "esize x \<noteq> \<infinity> \<Longrightarrow> finite {y. y \<le> x}"
assumes esize_mono[intro]: "x \<le> y \<Longrightarrow> esize x \<le> esize y"
assumes esize_strict_mono[intro]: "esize x \<noteq> \<infinity> \<Longrightarrow> x < y \<Longrightarrow> esize x < esize y"
begin
lemma infinite_chain_eSuc_esize[dest]:
assumes "Complete_Partial_Order.chain less_eq C" "infinite C" "x \<in> C"
obtains y
where "y \<in> C" "esize y \<ge> eSuc (esize x)"
proof (cases "esize x")
case (enat k)
have 1: "finite {y \<in> C. y \<le> x}" using esize_finite enat by simp
have 2: "infinite {y \<in> C. y \<ge> x}" using assms 1 by rule
have 3: "{y \<in> C. y > x} = {y \<in> C. y \<ge> x} - {x}" by auto
have 4: "infinite {y \<in> C. y > x}" using 2 unfolding 3 by simp
obtain y where 5: "y \<in> C" "y > x" using 4 by auto
have 6: "esize y > esize x" using esize_strict_mono enat 5(2) by blast
show ?thesis using that 5(1) 6 ileI1 by simp
next
case (infinity)
show ?thesis using that infinity assms(3) by simp
qed
lemma infinite_chain_arbitrary_esize[dest]:
assumes "Complete_Partial_Order.chain less_eq C" "infinite C"
obtains x
where "x \<in> C" "esize x \<ge> enat n"
proof (induct n arbitrary: thesis)
case 0
show ?case using assms(2) 0 by force
next
case (Suc n)
obtain x where 1: "x \<in> C" "esize x \<ge> enat n" using Suc(1) by blast
obtain y where 2: "y \<in> C" "esize y \<ge> eSuc (esize x)" using assms 1(1) by rule
show ?case using gfp.leq_trans Suc(2) 1(2) 2 by fastforce
qed
end
class esize_ccpo = esize_order + ccpo
begin
lemma esize_cont[dest]:
assumes "Complete_Partial_Order.chain less_eq C" "C \<noteq> {}"
shows "esize (\<Squnion> C) = \<Squnion> (esize ` C)"
proof (cases "finite C")
case False
have 1: "esize (\<Squnion> C) = \<infinity>"
proof
fix n
obtain A where 1: "A \<in> C" "esize A \<ge> enat n" using assms(1) False by rule
have 2: "A \<le> \<Squnion> C" using ccpo_Sup_upper assms(1) 1(1) by this
have "enat n \<le> esize A" using 1(2) by this
also have "\<dots> \<le> esize (\<Squnion> C)" using 2 by rule
finally show "enat n \<le> esize (\<Squnion> C)" by this
qed
have 2: "(\<Squnion> A \<in> C. esize A) = \<infinity>"
proof
fix n
obtain A where 1: "A \<in> C" "esize A \<ge> enat n" using assms(1) False by rule
show "enat n \<le> (\<Squnion> A \<in> C. esize A)" using SUP_upper2 1 by this
qed
show ?thesis using 1 2 by simp
next
case True
have 1: "esize (\<Squnion> C) = (\<Squnion> x \<in> C. esize x)"
- proof (intro order_class.antisym SUP_upper SUP_least esize_mono)
+ proof (intro order_class.order.antisym SUP_upper SUP_least esize_mono)
show "\<Squnion> C \<in> C" using in_chain_finite assms(1) True assms(2) by this
show "\<And> x. x \<in> C \<Longrightarrow> x \<le> \<Squnion> C" using ccpo_Sup_upper assms(1) by this
qed
show ?thesis using 1 by simp
qed
lemma esize_mcont: "mcont Sup less_eq Sup less_eq esize"
by (blast intro: mcontI monotoneI contI)
lemmas mcont2mcont_esize = esize_mcont[THEN lfp.mcont2mcont, simp, cont_intro]
end
end
diff --git a/thys/Polynomials/MPoly_Type_Class_OAlist.thy b/thys/Polynomials/MPoly_Type_Class_OAlist.thy
--- a/thys/Polynomials/MPoly_Type_Class_OAlist.thy
+++ b/thys/Polynomials/MPoly_Type_Class_OAlist.thy
@@ -1,962 +1,963 @@
(* Author: Fabian Immler, TU Muenchen *)
(* Author: Florian Haftmann, TU Muenchen *)
(* Author: Andreas Lochbihler, ETH Zurich *)
(* Author: Alexander Maletzky, RISC Linz *)
section \<open>Executable Representation of Polynomial Mappings as Association Lists\<close>
theory MPoly_Type_Class_OAlist
imports Term_Order
begin
instantiation pp :: (type, "{equal, zero}") equal
begin
definition equal_pp :: "('a, 'b) pp \<Rightarrow> ('a, 'b) pp \<Rightarrow> bool" where
"equal_pp p q \<equiv> (\<forall>t. lookup_pp p t = lookup_pp q t)"
instance by standard (auto simp: equal_pp_def intro: pp_eqI)
end
instantiation poly_mapping :: (type, "{equal, zero}") equal
begin
definition equal_poly_mapping :: "('a, 'b) poly_mapping \<Rightarrow> ('a, 'b) poly_mapping \<Rightarrow> bool" where
equal_poly_mapping_def [code del]: "equal_poly_mapping p q \<equiv> (\<forall>t. lookup p t = lookup q t)"
instance by standard (auto simp: equal_poly_mapping_def intro: poly_mapping_eqI)
end
subsection \<open>Power-Products Represented by @{type oalist_tc}\<close>
definition PP_oalist :: "('a::linorder, 'b::zero) oalist_tc \<Rightarrow> ('a, 'b) pp"
where "PP_oalist xs = pp_of_fun (OAlist_tc_lookup xs)"
code_datatype PP_oalist
lemma lookup_PP_oalist [simp, code]: "lookup_pp (PP_oalist xs) = OAlist_tc_lookup xs"
unfolding PP_oalist_def
proof (rule lookup_pp_of_fun)
have "{x. OAlist_tc_lookup xs x \<noteq> 0} \<subseteq> fst ` set (list_of_oalist_tc xs)"
proof (rule, simp)
fix x
assume "OAlist_tc_lookup xs x \<noteq> 0"
thus "x \<in> fst ` set (list_of_oalist_tc xs)"
using in_OAlist_tc_sorted_domain_iff_lookup set_OAlist_tc_sorted_domain by blast
qed
also have "finite ..." by simp
finally (finite_subset) show "finite {x. OAlist_tc_lookup xs x \<noteq> 0}" .
qed
lemma keys_PP_oalist [code]: "keys_pp (PP_oalist xs) = set (OAlist_tc_sorted_domain xs)"
by (rule set_eqI, simp add: keys_pp_iff in_OAlist_tc_sorted_domain_iff_lookup)
lemma lex_comp_PP_oalist [code]:
"lex_comp' (PP_oalist xs) (PP_oalist ys) =
the (OAlist_tc_lex_ord (\<lambda>_ x y. Some (comparator_of x y)) xs ys)"
for xs ys::"('a::nat, 'b::nat) oalist_tc"
proof (cases "lex_comp' (PP_oalist xs) (PP_oalist ys) = Eq")
case True
hence "PP_oalist xs = PP_oalist ys" by (rule lex_comp'_EqD)
hence eq: "OAlist_tc_lookup xs = OAlist_tc_lookup ys" by (simp add: pp_eq_iff)
have "OAlist_tc_lex_ord (\<lambda>_ x y. Some (comparator_of x y)) xs ys = Some Eq"
by (rule OAlist_tc_lex_ord_EqI, simp add: eq)
thus ?thesis by (simp add: True)
next
case False
then obtain x where 1: "x \<in> keys_pp (rep_nat_pp (PP_oalist xs)) \<union> keys_pp (rep_nat_pp (PP_oalist ys))"
and 2: "comparator_of (lookup_pp (rep_nat_pp (PP_oalist xs)) x) (lookup_pp (rep_nat_pp (PP_oalist ys)) x) =
lex_comp' (PP_oalist xs) (PP_oalist ys)"
and 3: "\<And>y. y < x \<Longrightarrow> lookup_pp (rep_nat_pp (PP_oalist xs)) y = lookup_pp (rep_nat_pp (PP_oalist ys)) y"
by (rule lex_comp'_valE, blast)
have "OAlist_tc_lex_ord (\<lambda>_ x y. Some (comparator_of x y)) xs ys = Some (lex_comp' (PP_oalist xs) (PP_oalist ys))"
proof (rule OAlist_tc_lex_ord_valI)
from False show "Some (lex_comp' (PP_oalist xs) (PP_oalist ys)) \<noteq> Some Eq" by simp
next
from 1 have "abs_nat x \<in> abs_nat ` (keys_pp (rep_nat_pp (PP_oalist xs)) \<union> keys_pp (rep_nat_pp (PP_oalist ys)))"
by (rule imageI)
also have "... = fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys)"
by (simp add: keys_rep_nat_pp_pp keys_PP_oalist OAlist_tc_sorted_domain_def image_Un image_image)
finally show "abs_nat x \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys)" .
next
show "Some (lex_comp' (PP_oalist xs) (PP_oalist ys)) =
Some (comparator_of (OAlist_tc_lookup xs (abs_nat x)) (OAlist_tc_lookup ys (abs_nat x)))"
by (simp add: 2[symmetric] lookup_rep_nat_pp_pp)
next
fix y::'a
assume "y < abs_nat x"
hence "rep_nat y < x" by (metis abs_inverse ord_iff(2))
hence "lookup_pp (rep_nat_pp (PP_oalist xs)) (rep_nat y) = lookup_pp (rep_nat_pp (PP_oalist ys)) (rep_nat y)"
by (rule 3)
hence "OAlist_tc_lookup xs y = OAlist_tc_lookup ys y" by (auto simp: lookup_rep_nat_pp_pp elim: rep_inj)
thus "Some (comparator_of (OAlist_tc_lookup xs y) (OAlist_tc_lookup ys y)) = Some Eq" by simp
qed
thus ?thesis by simp
qed
lemma zero_PP_oalist [code]: "(0::('a::linorder, 'b::zero) pp) = PP_oalist OAlist_tc_empty"
by (rule pp_eqI, simp add: lookup_OAlist_tc_empty)
lemma plus_PP_oalist [code]:
"PP_oalist xs + PP_oalist ys = PP_oalist (OAlist_tc_map2_val_neutr (\<lambda>_. (+)) xs ys)"
by (rule pp_eqI, simp add: lookup_plus_pp, rule lookup_OAlist_tc_map2_val_neutr[symmetric], simp_all)
lemma minus_PP_oalist [code]:
"PP_oalist xs - PP_oalist ys = PP_oalist (OAlist_tc_map2_val_rneutr (\<lambda>_. (-)) xs ys)"
by (rule pp_eqI, simp add: lookup_minus_pp, rule lookup_OAlist_tc_map2_val_rneutr[symmetric], simp)
lemma equal_PP_oalist [code]: "equal_class.equal (PP_oalist xs) (PP_oalist ys) = (xs = ys)"
by (simp add: equal_eq pp_eq_iff, auto elim: OAlist_tc_lookup_inj)
lemma lcs_PP_oalist [code]:
"lcs (PP_oalist xs) (PP_oalist ys) = PP_oalist (OAlist_tc_map2_val_neutr (\<lambda>_. max) xs ys)"
for xs ys :: "('a::linorder, 'b::add_linorder_min) oalist_tc"
by (rule pp_eqI, simp add: lookup_lcs_pp, rule lookup_OAlist_tc_map2_val_neutr[symmetric], simp_all add: max_def)
lemma deg_pp_PP_oalist [code]: "deg_pp (PP_oalist xs) = sum_list (map snd (list_of_oalist_tc xs))"
proof -
have "irreflp ((<)::_::linorder \<Rightarrow> _)" by (rule irreflpI, simp)
have "deg_pp (PP_oalist xs) = sum (OAlist_tc_lookup xs) (set (OAlist_tc_sorted_domain xs))"
by (simp add: deg_pp_alt keys_PP_oalist)
also have "... = sum_list (map (OAlist_tc_lookup xs) (OAlist_tc_sorted_domain xs))"
by (rule sum.distinct_set_conv_list, rule distinct_sorted_wrt_irrefl,
fact, fact transp_less, fact sorted_OAlist_tc_sorted_domain)
also have "... = sum_list (map snd (list_of_oalist_tc xs))"
by (rule arg_cong[where f=sum_list], simp add: OAlist_tc_sorted_domain_def OAlist_tc_lookup_eq_valueI)
finally show ?thesis .
qed
lemma single_PP_oalist [code]: "single_pp x e = PP_oalist (oalist_tc_of_list [(x, e)])"
by (rule pp_eqI, simp add: lookup_single_pp OAlist_tc_lookup_single)
definition adds_pp_add_linorder :: "('b, 'a::add_linorder) pp \<Rightarrow> _ \<Rightarrow> bool"
where [code_abbrev]: "adds_pp_add_linorder = (adds)"
lemma adds_pp_PP_oalist [code]:
"adds_pp_add_linorder (PP_oalist xs) (PP_oalist ys) = OAlist_tc_prod_ord (\<lambda>_. less_eq) xs ys"
for xs ys::"('a::linorder, 'b::add_linorder_min) oalist_tc"
proof (simp add: adds_pp_add_linorder_def adds_pp_iff adds_poly_mapping lookup_pp.rep_eq[symmetric] OAlist_tc_prod_ord_alt le_fun_def,
intro iffI allI ballI)
fix k
assume "\<forall>x. OAlist_tc_lookup xs x \<le> OAlist_tc_lookup ys x"
thus "OAlist_tc_lookup xs k \<le> OAlist_tc_lookup ys k" by blast
next
fix x
assume *: "\<forall>k\<in>fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys).
OAlist_tc_lookup xs k \<le> OAlist_tc_lookup ys k"
show "OAlist_tc_lookup xs x \<le> OAlist_tc_lookup ys x"
proof (cases "x \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys)")
case True
with * show ?thesis ..
next
case False
hence "x \<notin> set (OAlist_tc_sorted_domain xs)" and "x \<notin> set (OAlist_tc_sorted_domain ys)"
by (simp_all add: set_OAlist_tc_sorted_domain)
thus ?thesis by (simp add: in_OAlist_tc_sorted_domain_iff_lookup)
qed
qed
subsubsection \<open>Constructor\<close>
definition "sparse\<^sub>0 xs = PP_oalist (oalist_tc_of_list xs)" \<comment>\<open>sparse representation\<close>
subsubsection \<open>Computations\<close>
experiment begin
abbreviation "X \<equiv> 0::nat"
abbreviation "Y \<equiv> 1::nat"
abbreviation "Z \<equiv> 2::nat"
value [code] "sparse\<^sub>0 [(X, 2::nat), (Z, 7)]"
lemma
"sparse\<^sub>0 [(X, 2::nat), (Z, 7)] - sparse\<^sub>0 [(X, 2), (Z, 2)] = sparse\<^sub>0 [(Z, 5)]"
by eval
lemma
"lcs (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 7)]) (sparse\<^sub>0 [(Y, 3), (Z, 2)]) = sparse\<^sub>0 [(X, 2), (Y, 3), (Z, 7)]"
by eval
lemma
"(sparse\<^sub>0 [(X, 2::nat), (Z, 1)]) adds (sparse\<^sub>0 [(X, 3), (Y, 2), (Z, 1)])"
by eval
lemma
"lookup_pp (sparse\<^sub>0 [(X, 2::nat), (Z, 3)]) X = 2"
by eval
lemma
"deg_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3), (X, 1)]) = 6"
by eval
lemma
"lex_comp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)]) = Lt"
by eval
lemma
"lex_comp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)], 3::nat) (sparse\<^sub>0 [(X, 4)], 2) = Lt"
by eval
lemma
"lex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)])"
by eval
lemma
"lex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)])"
by eval
lemma
"\<not> dlex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 3)]) (sparse\<^sub>0 [(X, 4)])"
by eval
lemma
"dlex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 2)]) (sparse\<^sub>0 [(X, 5)])"
by eval
lemma
"\<not> drlex_pp (sparse\<^sub>0 [(X, 2::nat), (Y, 1), (Z, 2)]) (sparse\<^sub>0 [(X, 5)])"
by eval
end
subsection \<open>\<open>MP_oalist\<close>\<close>
lift_definition MP_oalist :: "('a::nat_term, 'b::zero) oalist_ntm \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b"
is OAlist_lookup_ntm
proof -
fix xs :: "('a, 'b) oalist_ntm"
have "{x. OAlist_lookup_ntm xs x \<noteq> 0} \<subseteq> fst ` set (fst (list_of_oalist_ntm xs))"
proof (rule, simp)
fix x
assume "OAlist_lookup_ntm xs x \<noteq> 0"
thus "x \<in> fst ` set (fst (list_of_oalist_ntm xs))"
using oa_ntm.in_sorted_domain_iff_lookup oa_ntm.set_sorted_domain by blast
qed
also have "finite ..." by simp
finally (finite_subset) show "finite {x. OAlist_lookup_ntm xs x \<noteq> 0}" .
qed
lemmas [simp, code] = MP_oalist.rep_eq
code_datatype MP_oalist
lemma keys_MP_oalist [code]: "keys (MP_oalist xs) = set (map fst (fst (list_of_oalist_ntm xs)))"
by (rule set_eqI, simp add: in_keys_iff oa_ntm.in_sorted_domain_iff_lookup[simplified oa_ntm.set_sorted_domain])
lemma MP_oalist_empty [simp]: "MP_oalist (OAlist_empty_ntm ko) = 0"
by (rule poly_mapping_eqI, simp add: oa_ntm.lookup_empty)
lemma zero_MP_oalist [code]: "(0::('a::{linorder,nat_term} \<Rightarrow>\<^sub>0 'b::zero)) = MP_oalist (OAlist_empty_ntm nat_term_order_of_le)"
by simp
definition is_zero :: "('a \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> bool"
where [code_abbrev]: "is_zero p \<longleftrightarrow> (p = 0)"
lemma is_zero_MP_oalist [code]: "is_zero (MP_oalist xs) = List.null (fst (list_of_oalist_ntm xs))"
unfolding is_zero_def List.null_def
proof
assume "MP_oalist xs = 0"
hence "OAlist_lookup_ntm xs k = 0" for k by (simp add: poly_mapping_eq_iff)
thus "fst (list_of_oalist_ntm xs) = []"
by (metis image_eqI ko_ntm.min_key_val_raw_in oa_ntm.in_sorted_domain_iff_lookup oa_ntm.set_sorted_domain)
next
assume "fst (list_of_oalist_ntm xs) = []"
hence "OAlist_lookup_ntm xs k = 0" for k
by (metis oa_ntm.list_of_oalist_empty oa_ntm.lookup_empty oalist_ntm_eqI surjective_pairing)
thus "MP_oalist xs = 0" by (simp add: poly_mapping_eq_iff ext)
qed
lemma plus_MP_oalist [code]: "MP_oalist xs + MP_oalist ys = MP_oalist (OAlist_map2_val_neutr_ntm (\<lambda>_. (+)) xs ys)"
by (rule poly_mapping_eqI, simp add: lookup_plus_fun, rule oa_ntm.lookup_map2_val_neutr[symmetric], simp_all)
lemma minus_MP_oalist [code]: "MP_oalist xs - MP_oalist ys = MP_oalist (OAlist_map2_val_rneutr_ntm (\<lambda>_. (-)) xs ys)"
by (rule poly_mapping_eqI, simp add: lookup_minus_fun, rule oa_ntm.lookup_map2_val_rneutr[symmetric], simp)
lemma uminus_MP_oalist [code]: "- MP_oalist xs = MP_oalist (OAlist_map_val_ntm (\<lambda>_. uminus) xs)"
by (rule poly_mapping_eqI, simp, rule oa_ntm.lookup_map_val[symmetric], simp)
lemma equal_MP_oalist [code]: "equal_class.equal (MP_oalist xs) (MP_oalist ys) = (OAlist_eq_ntm xs ys)"
by (simp add: oa_ntm.oalist_eq_alt equal_eq poly_mapping_eq_iff)
lemma map_MP_oalist [code]: "Poly_Mapping.map f (MP_oalist xs) = MP_oalist (OAlist_map_val_ntm (\<lambda>_. f) xs)"
proof -
have eq: "OAlist_map_val_ntm (\<lambda>_. f) xs = OAlist_map_val_ntm (\<lambda>_ c. f c when c \<noteq> 0) xs"
proof (rule oa_ntm.map_val_cong)
fix t c
assume *: "(t, c) \<in> set (fst (list_of_oalist_ntm xs))"
hence "fst (t, c) \<in> fst ` set (fst (list_of_oalist_ntm xs))" by (rule imageI)
hence "OAlist_lookup_ntm xs t \<noteq> 0"
by (simp add: oa_ntm.in_sorted_domain_iff_lookup[simplified oa_ntm.set_sorted_domain])
moreover from * have "OAlist_lookup_ntm xs t = c" by (rule oa_ntm.lookup_eq_valueI)
ultimately have "c \<noteq> 0" by simp
thus "f c = (f c when c \<noteq> 0)" by simp
qed
show ?thesis
by (rule poly_mapping_eqI, simp add: Poly_Mapping.map.rep_eq eq, rule oa_ntm.lookup_map_val[symmetric], simp)
qed
lemma range_MP_oalist [code]: "Poly_Mapping.range (MP_oalist xs) = set (map snd (fst (list_of_oalist_ntm xs)))"
proof (simp add: Poly_Mapping.range.rep_eq, intro set_eqI iffI)
fix c
assume "c \<in> range (OAlist_lookup_ntm xs) - {0}"
hence "c \<in> range (OAlist_lookup_ntm xs)" and "c \<noteq> 0" by simp_all
from this(1) obtain t where "OAlist_lookup_ntm xs t = c" by fastforce
with \<open>c \<noteq> 0\<close> have "(t, c) \<in> set (fst (list_of_oalist_ntm xs))" by (simp add: oa_ntm.lookup_eq_value)
hence "snd (t, c) \<in> snd ` set (fst (list_of_oalist_ntm xs))" by (rule imageI)
thus "c \<in> snd ` set (fst (list_of_oalist_ntm xs))" by simp
next
fix c
assume "c \<in> snd ` set (fst (list_of_oalist_ntm xs))"
then obtain t where *: "(t, c) \<in> set (fst (list_of_oalist_ntm xs))" by fastforce
hence "fst (t, c) \<in> fst ` set (fst (list_of_oalist_ntm xs))" by (rule imageI)
hence "OAlist_lookup_ntm xs t \<noteq> 0"
by (simp add: oa_ntm.in_sorted_domain_iff_lookup[simplified oa_ntm.set_sorted_domain])
moreover from * have "OAlist_lookup_ntm xs t = c" by (rule oa_ntm.lookup_eq_valueI)
ultimately show "c \<in> range (OAlist_lookup_ntm xs) - {0}" by fastforce
qed
lemma if_poly_mapping_eq_iff:
"(if x = y then a else b) = (if (\<forall>i\<in>keys x \<union> keys y. lookup x i = lookup y i) then a else b)"
by simp (metis UnI1 UnI2 in_keys_iff poly_mapping_eqI)
lemma keys_add_eq: "keys (a + b) = keys a \<union> keys b - {x \<in> keys a \<inter> keys b. lookup a x + lookup b x = 0}"
by (auto simp: in_keys_iff lookup_add add_eq_0_iff
simp del: lookup_not_eq_zero_eq_in_keys)
locale gd_nat_term =
gd_term pair_of_term term_of_pair
"\<lambda>s t. le_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))"
"\<lambda>s t. lt_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))"
"le_of_nat_term_order cmp_term"
"lt_of_nat_term_order cmp_term"
for pair_of_term::"'t::nat_term \<Rightarrow> ('a::{nat_term,graded_dickson_powerprod} \<times> 'k::{countable,the_min,wellorder})"
and term_of_pair::"('a \<times> 'k) \<Rightarrow> 't"
and cmp_term +
assumes splus_eq_splus: "t \<oplus> u = nat_term_class.splus (term_of_pair (t, the_min)) u"
begin
definition shift_map_keys :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('t, 'b) oalist_ntm \<Rightarrow> ('t, 'b::semiring_0) oalist_ntm"
where "shift_map_keys t f xs = OAlist_ntm (map_raw (\<lambda>kv. (t \<oplus> fst kv, f (snd kv))) (list_of_oalist_ntm xs))"
lemma list_of_oalist_shift_keys:
"list_of_oalist_ntm (shift_map_keys t f xs) = (map_raw (\<lambda>kv. (t \<oplus> fst kv, f (snd kv))) (list_of_oalist_ntm xs))"
unfolding shift_map_keys_def
by (rule oa_ntm.list_of_oalist_of_list_id, rule ko_ntm.oalist_inv_map_raw, fact oalist_inv_list_of_oalist_ntm,
simp add: nat_term_compare_inv_conv[symmetric] nat_term_compare_inv_def splus_eq_splus nat_term_compare_splus)
lemma lookup_shift_map_keys_plus:
"lookup (MP_oalist (shift_map_keys t ((*) c) xs)) (t \<oplus> u) = c * lookup (MP_oalist xs) u" (is "?l = ?r")
proof -
let ?f = "\<lambda>kv. (t \<oplus> fst kv, c * snd kv)"
have "?l = lookup_ko_ntm (map_raw ?f (list_of_oalist_ntm xs)) (fst (?f (u, c)))"
by (simp add: oa_ntm.lookup_def list_of_oalist_shift_keys)
also have "... = snd (?f (u, lookup_ko_ntm (list_of_oalist_ntm xs) u))"
by (rule ko_ntm.lookup_raw_map_raw, fact oalist_inv_list_of_oalist_ntm, simp,
simp add: nat_term_compare_inv_conv[symmetric] nat_term_compare_inv_def splus_eq_splus nat_term_compare_splus)
also have "... = ?r" by (simp add: oa_ntm.lookup_def)
finally show ?thesis .
qed
lemma keys_shift_map_keys_subset:
"keys (MP_oalist (shift_map_keys t ((*) c) xs)) \<subseteq> ((\<oplus>) t) ` keys (MP_oalist xs)" (is "?l \<subseteq> ?r")
proof -
let ?f = "\<lambda>kv. (t \<oplus> fst kv, c * snd kv)"
have "?l = fst ` set (fst (map_raw ?f (list_of_oalist_ntm xs)))"
by (simp add: keys_MP_oalist list_of_oalist_shift_keys)
also from ko_ntm.map_raw_subset have "... \<subseteq> fst ` ?f ` set (fst (list_of_oalist_ntm xs))"
by (rule image_mono)
also have "... \<subseteq> ?r" by (simp add: keys_MP_oalist image_image)
finally show ?thesis .
qed
lemma monom_mult_MP_oalist [code]:
"monom_mult c t (MP_oalist xs) =
MP_oalist (if c = 0 then OAlist_empty_ntm (snd (list_of_oalist_ntm xs)) else shift_map_keys t ((*) c) xs)"
proof (cases "c = 0")
case True
hence "monom_mult c t (MP_oalist xs) = 0" using monom_mult_zero_left by simp
thus ?thesis using True by simp
next
case False
have "monom_mult c t (MP_oalist xs) = MP_oalist (shift_map_keys t ((*) c) xs)"
proof (rule poly_mapping_eqI, simp add: lookup_monom_mult del: MP_oalist.rep_eq, intro conjI impI)
fix u
assume "t adds\<^sub>p u"
then obtain v where "u = t \<oplus> v" by (rule adds_ppE)
thus "c * lookup (MP_oalist xs) (u \<ominus> t) = lookup (MP_oalist (shift_map_keys t ((*) c) xs)) u"
by (simp add: splus_sminus lookup_shift_map_keys_plus del: MP_oalist.rep_eq)
next
fix u
assume "\<not> t adds\<^sub>p u"
have "u \<notin> keys (MP_oalist (shift_map_keys t ((*) c) xs))"
proof
assume "u \<in> keys (MP_oalist (shift_map_keys t ((*) c) xs))"
also have "... \<subseteq> ((\<oplus>) t) ` keys (MP_oalist xs)" by (fact keys_shift_map_keys_subset)
finally obtain v where "u = t \<oplus> v" ..
hence "t adds\<^sub>p u" by (rule adds_ppI)
with \<open>\<not> t adds\<^sub>p u\<close> show False ..
qed
thus "lookup (MP_oalist (shift_map_keys t ((*) c) xs)) u = 0" by (simp add: in_keys_iff)
qed
thus ?thesis by (simp add: False)
qed
lemma mult_scalar_MP_oalist [code]:
"(MP_oalist xs) \<odot> (MP_oalist ys) =
(if is_zero (MP_oalist xs) then
MP_oalist (OAlist_empty_ntm (snd (list_of_oalist_ntm ys)))
else
let ct = OAlist_hd_ntm xs in
monom_mult (snd ct) (fst ct) (MP_oalist ys) + (MP_oalist (OAlist_tl_ntm xs)) \<odot> (MP_oalist ys))"
proof (split if_split, intro conjI impI)
assume "is_zero (MP_oalist xs)"
thus "MP_oalist xs \<odot> MP_oalist ys = MP_oalist (OAlist_empty_ntm (snd (list_of_oalist_ntm ys)))"
by (simp add: is_zero_def)
next
assume "\<not> is_zero (MP_oalist xs)"
hence *: "fst (list_of_oalist_ntm xs) \<noteq> []" by (simp add: is_zero_MP_oalist List.null_def)
define ct where "ct = OAlist_hd_ntm xs"
have eq: "except (MP_oalist xs) {fst ct} = MP_oalist (OAlist_tl_ntm xs)"
by (rule poly_mapping_eqI, simp add: lookup_except ct_def oa_ntm.lookup_tl')
have "MP_oalist xs \<odot> MP_oalist ys =
monom_mult (lookup (MP_oalist xs) (fst ct)) (fst ct) (MP_oalist ys) +
except (MP_oalist xs) {fst ct} \<odot> MP_oalist ys" by (fact mult_scalar_rec_left)
also have "... = monom_mult (snd ct) (fst ct) (MP_oalist ys) + except (MP_oalist xs) {fst ct} \<odot> MP_oalist ys"
using * by (simp add: ct_def oa_ntm.snd_hd)
also have "... = monom_mult (snd ct) (fst ct) (MP_oalist ys) + MP_oalist (OAlist_tl_ntm xs) \<odot> MP_oalist ys"
by (simp only: eq)
finally show "MP_oalist xs \<odot> MP_oalist ys =
(let ct = OAlist_hd_ntm xs in
monom_mult (snd ct) (fst ct) (MP_oalist ys) + MP_oalist (OAlist_tl_ntm xs) \<odot> MP_oalist ys)"
by (simp add: ct_def Let_def)
qed
end (* ordered_nat_term *)
subsubsection \<open>Special case of addition: adding monomials\<close>
definition plus_monomial_less :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b::monoid_add)"
where "plus_monomial_less p c u = p + monomial c u"
text \<open>@{const plus_monomial_less} is useful when adding a monomial to a polynomial, where the term
of the monomial is known to be smaller than all terms in the polynomial, because it can be
implemented more efficiently than general addition.\<close>
lemma plus_monomial_less_MP_oalist [code]:
"plus_monomial_less (MP_oalist xs) c u = MP_oalist (OAlist_update_by_fun_gr_ntm u (\<lambda>c0. c0 + c) xs)"
unfolding plus_monomial_less_def oa_ntm.update_by_fun_gr_eq_update_by_fun
by (rule poly_mapping_eqI, simp add: lookup_plus_fun oa_ntm.lookup_update_by_fun lookup_single)
text \<open>@{const plus_monomial_less} is computed by @{const OAlist_update_by_fun_gr_ntm}, because greater
terms come @{emph \<open>before\<close>} smaller ones in @{type oalist_ntm}.\<close>
subsubsection \<open>Constructors\<close>
definition "distr\<^sub>0 ko xs = MP_oalist (oalist_of_list_ntm (xs, ko))" \<comment>\<open>sparse representation\<close>
definition V\<^sub>0 :: "'a \<Rightarrow> ('a, nat) pp \<Rightarrow>\<^sub>0 'b::{one,zero}" where
"V\<^sub>0 n \<equiv> monomial 1 (single_pp n 1)"
definition C\<^sub>0 :: "'b \<Rightarrow> ('a, nat) pp \<Rightarrow>\<^sub>0 'b::zero" where "C\<^sub>0 c \<equiv> monomial c 0"
lemma C\<^sub>0_one: "C\<^sub>0 1 = 1"
by (simp add: C\<^sub>0_def)
lemma C\<^sub>0_numeral: "C\<^sub>0 (numeral x) = numeral x"
by (auto intro!: poly_mapping_eqI simp: C\<^sub>0_def lookup_numeral)
lemma C\<^sub>0_minus: "C\<^sub>0 (- x) = - C\<^sub>0 x"
by (simp add: C\<^sub>0_def single_uminus)
lemma C\<^sub>0_zero: "C\<^sub>0 0 = 0"
by (auto intro!: poly_mapping_eqI simp: C\<^sub>0_def)
lemma V\<^sub>0_power: "V\<^sub>0 v ^ n = monomial 1 (single_pp v n)"
by (induction n) (auto simp: V\<^sub>0_def mult_single single_pp_plus)
lemma single_MP_oalist [code]: "Poly_Mapping.single k v = distr\<^sub>0 nat_term_order_of_le [(k, v)]"
unfolding distr\<^sub>0_def by (rule poly_mapping_eqI, simp add: lookup_single OAlist_lookup_ntm_single)
lemma one_MP_oalist [code]: "1 = distr\<^sub>0 nat_term_order_of_le [(0, 1)]"
by (metis single_MP_oalist single_one)
lemma except_MP_oalist [code]: "except (MP_oalist xs) S = MP_oalist (OAlist_filter_ntm (\<lambda>kv. fst kv \<notin> S) xs)"
by (rule poly_mapping_eqI, simp add: lookup_except oa_ntm.lookup_filter)
subsubsection \<open>Changing the Internal Order\<close>
definition change_ord :: "'a::nat_term_compare nat_term_order \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b)"
where "change_ord to = (\<lambda>x. x)"
lemma change_ord_MP_oalist [code]: "change_ord to (MP_oalist xs) = MP_oalist (OAlist_reorder_ntm to xs)"
by (rule poly_mapping_eqI, simp add: change_ord_def oa_ntm.lookup_reorder)
subsubsection \<open>Ordered Power-Products\<close>
lemma foldl_assoc:
assumes "\<And>x y z. f (f x y) z = f x (f y z)"
shows "foldl f (f a b) xs = f a (foldl f b xs)"
proof (induct xs arbitrary: a b)
fix a b
show "foldl f (f a b) [] = f a (foldl f b [])" by simp
next
fix a b x xs
assume "\<And>a b. foldl f (f a b) xs = f a (foldl f b xs)"
from assms[of a b x] this[of a "f b x"]
show "foldl f (f a b) (x # xs) = f a (foldl f b (x # xs))" unfolding foldl_Cons by simp
qed
context gd_nat_term
begin
definition ord_pp :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
where "ord_pp s t = le_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))"
definition ord_pp_strict :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
where "ord_pp_strict s t = lt_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min))"
lemma lt_MP_oalist [code]:
"lt (MP_oalist xs) = (if is_zero (MP_oalist xs) then min_term else fst (OAlist_min_key_val_ntm cmp_term xs))"
proof (split if_split, intro conjI impI)
assume "is_zero (MP_oalist xs)"
thus "lt (MP_oalist xs) = min_term" by (simp add: is_zero_def)
next
assume "\<not> is_zero (MP_oalist xs)"
hence "fst (list_of_oalist_ntm xs) \<noteq> []" by (simp add: is_zero_MP_oalist List.null_def)
show "lt (MP_oalist xs) = fst (OAlist_min_key_val_ntm cmp_term xs)"
proof (rule lt_eqI_keys)
show "fst (OAlist_min_key_val_ntm cmp_term xs) \<in> keys (MP_oalist xs)"
by (simp add: keys_MP_oalist, rule imageI, rule oa_ntm.min_key_val_in, fact)
next
fix u
assume "u \<in> keys (MP_oalist xs)"
also have "... = fst ` set (fst (list_of_oalist_ntm xs))" by (simp add: keys_MP_oalist)
finally obtain z where "z \<in> set (fst (list_of_oalist_ntm xs))" and "u = fst z" ..
from this(1) have "ko.le (key_order_of_nat_term_order_inv cmp_term) (fst (OAlist_min_key_val_ntm cmp_term xs)) u"
unfolding \<open>u = fst z\<close> by (rule oa_ntm.min_key_val_minimal)
thus "le_of_nat_term_order cmp_term u (fst (OAlist_min_key_val_ntm cmp_term xs))"
by (simp add: le_of_nat_term_order_alt)
qed
qed
lemma lc_MP_oalist [code]:
"lc (MP_oalist xs) = (if is_zero (MP_oalist xs) then 0 else snd (OAlist_min_key_val_ntm cmp_term xs))"
proof (split if_split, intro conjI impI)
assume "is_zero (MP_oalist xs)"
thus "lc (MP_oalist xs) = 0" by (simp add: is_zero_def)
next
assume "\<not> is_zero (MP_oalist xs)"
moreover from this have "fst (list_of_oalist_ntm xs) \<noteq> []" by (simp add: is_zero_MP_oalist List.null_def)
ultimately show "lc (MP_oalist xs) = snd (OAlist_min_key_val_ntm cmp_term xs)"
by (simp add: lc_def lt_MP_oalist oa_ntm.snd_min_key_val)
qed
lemma tail_MP_oalist [code]: "tail (MP_oalist xs) = MP_oalist (OAlist_except_min_ntm cmp_term xs)"
proof (cases "is_zero (MP_oalist xs)")
case True
hence "fst (list_of_oalist_ntm xs) = []" by (simp add: is_zero_MP_oalist List.null_def)
hence "fst (list_of_oalist_ntm (OAlist_except_min_ntm cmp_term xs)) = []"
by (rule oa_ntm.except_min_Nil)
hence "is_zero (MP_oalist (OAlist_except_min_ntm cmp_term xs))"
by (simp add: is_zero_MP_oalist List.null_def)
with True show ?thesis by (simp add: is_zero_def)
next
case False
show ?thesis by (rule poly_mapping_eqI, simp add: lookup_tail_2 oa_ntm.lookup_except_min' lt_MP_oalist False)
qed
definition comp_opt_p :: "('t \<Rightarrow>\<^sub>0 'c::zero, 't \<Rightarrow>\<^sub>0 'c) comp_opt"
where "comp_opt_p p q =
(if p = q then Some Eq else if ord_strict_p p q then Some Lt else if ord_strict_p q p then Some Gt else None)"
lemma comp_opt_p_MP_oalist [code]:
"comp_opt_p (MP_oalist xs) (MP_oalist ys) =
OAlist_lex_ord_ntm cmp_term (\<lambda>_ x y. if x = y then Some Eq else if x = 0 then Some Lt else if y = 0 then Some Gt else None) xs ys"
proof -
let ?f = "\<lambda>_ x y. if x = y then Some Eq else if x = 0 then Some Lt else if y = 0 then Some Gt else None"
show ?thesis
proof (cases "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Eq")
case True
hence "MP_oalist xs = MP_oalist ys" by (simp add: comp_opt_p_def split: if_splits)
hence "lookup (MP_oalist xs) = lookup (MP_oalist ys)" by (rule arg_cong)
hence eq: "OAlist_lookup_ntm xs = OAlist_lookup_ntm ys" by simp
have "OAlist_lex_ord_ntm cmp_term ?f xs ys = Some Eq"
by (rule oa_ntm.lex_ord_EqI, simp add: eq)
with True show ?thesis by simp
next
case False
hence neq: "MP_oalist xs \<noteq> MP_oalist ys" by (simp add: comp_opt_p_def split: if_splits)
then obtain v where 1: "v \<in> keys (MP_oalist xs) \<union> keys (MP_oalist ys)"
and 2: "lookup (MP_oalist xs) v \<noteq> lookup (MP_oalist ys) v"
and 3: "\<And>u. lt_of_nat_term_order cmp_term v u \<Longrightarrow> lookup (MP_oalist xs) u = lookup (MP_oalist ys) u"
by (rule poly_mapping_neqE, blast)
show ?thesis
proof (rule HOL.sym, rule oa_ntm.lex_ord_valI)
from 1 show "v \<in> fst ` set (fst (list_of_oalist_ntm xs)) \<union> fst ` set (fst (list_of_oalist_ntm ys))"
by (simp add: keys_MP_oalist)
next
from 2 have 4: "OAlist_lookup_ntm xs v \<noteq> OAlist_lookup_ntm ys v" by simp
show "comp_opt_p (MP_oalist xs) (MP_oalist ys) =
(if OAlist_lookup_ntm xs v = OAlist_lookup_ntm ys v then Some Eq
else if OAlist_lookup_ntm xs v = 0 then Some Lt
else if OAlist_lookup_ntm ys v = 0 then Some Gt else None)"
proof (simp add: 4, intro conjI impI)
assume "OAlist_lookup_ntm ys v = 0" and "OAlist_lookup_ntm xs v = 0"
with 4 show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Lt" by simp
next
assume "OAlist_lookup_ntm xs v \<noteq> 0" and "OAlist_lookup_ntm ys v = 0"
hence "lookup (MP_oalist ys) v = 0" and "lookup (MP_oalist xs) v \<noteq> 0" by simp_all
hence "ord_strict_p (MP_oalist ys) (MP_oalist xs)" using 3[symmetric]
by (rule ord_strict_pI)
with neq show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Gt" by (auto simp: comp_opt_p_def)
next
assume "OAlist_lookup_ntm ys v \<noteq> 0" and "OAlist_lookup_ntm xs v = 0"
hence "lookup (MP_oalist xs) v = 0" and "lookup (MP_oalist ys) v \<noteq> 0" by simp_all
hence "ord_strict_p (MP_oalist xs) (MP_oalist ys)" using 3 by (rule ord_strict_pI)
with neq show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = Some Lt" by (auto simp: comp_opt_p_def)
next
assume "OAlist_lookup_ntm xs v \<noteq> 0"
hence "lookup (MP_oalist xs) v \<noteq> 0" by simp
with 2 have a: "\<not> ord_strict_p (MP_oalist xs) (MP_oalist ys)" using 3 by (rule not_ord_strict_pI)
assume "OAlist_lookup_ntm ys v \<noteq> 0"
hence "lookup (MP_oalist ys) v \<noteq> 0" by simp
with 2[symmetric] have "\<not> ord_strict_p (MP_oalist ys) (MP_oalist xs)"
using 3[symmetric] by (rule not_ord_strict_pI)
with neq a show "comp_opt_p (MP_oalist xs) (MP_oalist ys) = None" by (auto simp: comp_opt_p_def)
qed
next
fix u
assume "ko.lt (key_order_of_nat_term_order_inv cmp_term) u v"
hence "lt_of_nat_term_order cmp_term v u" by (simp only: lt_of_nat_term_order_alt)
hence "lookup (MP_oalist xs) u = lookup (MP_oalist ys) u" by (rule 3)
thus "(if OAlist_lookup_ntm xs u = OAlist_lookup_ntm ys u then Some Eq
else if OAlist_lookup_ntm xs u = 0 then Some Lt
else if OAlist_lookup_ntm ys u = 0 then Some Gt else None) = Some Eq" by simp
qed fact
qed
qed
lemma compute_ord_p [code]: "ord_p p q = (let aux = comp_opt_p p q in aux = Some Lt \<or> aux = Some Eq)"
by (auto simp: ord_p_def comp_opt_p_def)
lemma compute_ord_p_strict [code]: "ord_strict_p p q = (comp_opt_p p q = Some Lt)"
by (auto simp: comp_opt_p_def)
lemma keys_to_list_MP_oalist [code]: "keys_to_list (MP_oalist xs) = OAlist_sorted_domain_ntm cmp_term xs"
proof -
have eq: "ko.lt (key_order_of_nat_term_order_inv cmp_term) = ord_term_strict_conv"
by (intro ext, simp add: lt_of_nat_term_order_alt)
have 1: "irreflp ord_term_strict_conv" by (rule irreflpI, simp)
have 2: "transp ord_term_strict_conv" by (rule transpI, simp)
have "antisymp ord_term_strict_conv" by (rule antisympI, simp)
moreover have 3: "sorted_wrt ord_term_strict_conv (keys_to_list (MP_oalist xs))"
unfolding keys_to_list_def by (fact pps_to_list_sorted_wrt)
moreover note _
moreover have 4: "sorted_wrt ord_term_strict_conv (OAlist_sorted_domain_ntm cmp_term xs)"
unfolding eq[symmetric] by (fact oa_ntm.sorted_sorted_domain)
ultimately show ?thesis
proof (rule sorted_wrt_distinct_set_unique)
from 1 2 3 show "distinct (keys_to_list (MP_oalist xs))" by (rule distinct_sorted_wrt_irrefl)
next
from 1 2 4 show "distinct (OAlist_sorted_domain_ntm cmp_term xs)" by (rule distinct_sorted_wrt_irrefl)
next
show "set (keys_to_list (MP_oalist xs)) = set (OAlist_sorted_domain_ntm cmp_term xs)"
by (simp add: set_keys_to_list keys_MP_oalist oa_ntm.set_sorted_domain)
qed
qed
end (* ordered_nat_term *)
lifting_update poly_mapping.lifting
lifting_forget poly_mapping.lifting
subsection \<open>Interpretations\<close>
lemma term_powerprod_gd_term:
fixes pair_of_term :: "'t::nat_term \<Rightarrow> ('a::{graded_dickson_powerprod,nat_pp_compare} \<times> 'k::{the_min,wellorder})"
assumes "term_powerprod pair_of_term term_of_pair"
and "\<And>v. fst (rep_nat_term v) = rep_nat_pp (fst (pair_of_term v))"
and "\<And>t. snd (rep_nat_term (term_of_pair (t, the_min))) = 0"
and "\<And>v w. snd (pair_of_term v) \<le> snd (pair_of_term w) \<Longrightarrow> snd (rep_nat_term v) \<le> snd (rep_nat_term w)"
and "\<And>s t k. term_of_pair (s + t, k) = splus (term_of_pair (s, k)) (term_of_pair (t, k))"
and "\<And>t v. term_powerprod.splus pair_of_term term_of_pair t v = splus (term_of_pair (t, the_min)) v"
shows "gd_term pair_of_term term_of_pair
(\<lambda>s t. le_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min)))
(\<lambda>s t. lt_of_nat_term_order cmp_term (term_of_pair (s, the_min)) (term_of_pair (t, the_min)))
(le_of_nat_term_order cmp_term)
(lt_of_nat_term_order cmp_term)"
proof -
from assms(1) interpret tp: term_powerprod pair_of_term term_of_pair .
let ?f = "\<lambda>x. term_of_pair (x, the_min)"
show ?thesis
proof (intro gd_term.intro ordered_term.intro)
from assms(1) show "term_powerprod pair_of_term term_of_pair" .
next
show "ordered_powerprod (\<lambda>s t. le_of_nat_term_order cmp_term (?f s) (?f t))
(\<lambda>s t. lt_of_nat_term_order cmp_term (?f s) (?f t))"
proof (intro ordered_powerprod.intro ordered_powerprod_axioms.intro)
show "class.linorder (\<lambda>s t. le_of_nat_term_order cmp_term (?f s) (?f t))
(\<lambda>s t. lt_of_nat_term_order cmp_term (?f s) (?f t))"
proof (unfold_locales, simp_all add: lt_of_nat_term_order_alt le_of_nat_term_order_alt ko.linear ko.less_le_not_le)
fix x y
assume "ko.le (key_order_of_nat_term_order_inv cmp_term) (term_of_pair (x, the_min)) (term_of_pair (y, the_min))"
and "ko.le (key_order_of_nat_term_order_inv cmp_term) (term_of_pair (y, the_min)) (term_of_pair (x, the_min))"
- hence "term_of_pair (x, the_min) = term_of_pair (y, the_min)" by (rule ko.antisym)
+ hence "term_of_pair (x, the_min) = term_of_pair (y, the_min)"
+ by (rule ko.antisym)
hence "(x, the_min) = (y, the_min::'k)" by (rule tp.term_of_pair_injective)
thus "x = y" by simp
qed
next
fix t
show "le_of_nat_term_order cmp_term (?f 0) (?f t)"
unfolding le_of_nat_term_order
by (rule nat_term_compD1', fact comparator_nat_term_compare, fact nat_term_comp_nat_term_compare,
simp add: assms(3), simp add: assms(2) zero_pp tp.pair_term)
next
fix s t u
assume "le_of_nat_term_order cmp_term (?f s) (?f t)"
hence "le_of_nat_term_order cmp_term (?f (u + s)) (?f (u + t))"
by (simp add: le_of_nat_term_order assms(5) nat_term_compare_splus)
thus "le_of_nat_term_order cmp_term (?f (s + u)) (?f (t + u))" by (simp only: ac_simps)
qed
next
show "class.linorder (le_of_nat_term_order cmp_term) (lt_of_nat_term_order cmp_term)"
by (fact linorder_le_of_nat_term_order)
next
show "ordered_term_axioms pair_of_term term_of_pair (\<lambda>s t. le_of_nat_term_order cmp_term (?f s) (?f t))
(le_of_nat_term_order cmp_term)"
proof
fix v w t
assume "le_of_nat_term_order cmp_term v w"
thus "le_of_nat_term_order cmp_term (t \<oplus> v) (t \<oplus> w)"
by (simp add: le_of_nat_term_order assms(6) nat_term_compare_splus)
next
fix v w
assume "le_of_nat_term_order cmp_term (?f (tp.pp_of_term v)) (?f (tp.pp_of_term w))"
hence 3: "nat_term_compare cmp_term (?f (tp.pp_of_term v)) (?f (tp.pp_of_term w)) \<noteq> Gt"
by (simp add: le_of_nat_term_order)
assume "tp.component_of_term v \<le> tp.component_of_term w"
hence 4: "snd (rep_nat_term v) \<le> snd (rep_nat_term w)"
by (simp add: tp.component_of_term_def assms(4))
note comparator_nat_term_compare nat_term_comp_nat_term_compare
moreover have "fst (rep_nat_term v) = fst (rep_nat_term (?f (tp.pp_of_term v)))"
by (simp add: assms(2) tp.pp_of_term_def tp.pair_term)
moreover have "fst (rep_nat_term w) = fst (rep_nat_term (?f (tp.pp_of_term w)))"
by (simp add: assms(2) tp.pp_of_term_def tp.pair_term)
moreover note 4
moreover have "snd (rep_nat_term (?f (tp.pp_of_term v))) = snd (rep_nat_term (?f (tp.pp_of_term w)))"
by (simp add: assms(3))
ultimately show "le_of_nat_term_order cmp_term v w" unfolding le_of_nat_term_order using 3
by (rule nat_term_compD4'')
qed
qed
qed
lemma gd_term_to_pair_unit:
"gd_term (to_pair_unit::'a::{nat_term_compare,nat_pp_term,graded_dickson_powerprod} \<Rightarrow> _) fst
(\<lambda>s t. le_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min)))
(\<lambda>s t. lt_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min)))
(le_of_nat_term_order cmp_term)
(lt_of_nat_term_order cmp_term)"
proof (intro gd_term.intro ordered_term.intro)
show "term_powerprod to_pair_unit fst" by unfold_locales
next
show "ordered_powerprod (\<lambda>s t. le_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min)))
(\<lambda>s t. lt_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min)))"
unfolding fst_conv using linorder_le_of_nat_term_order
proof (intro ordered_powerprod.intro)
from le_of_nat_term_order_zero_min show "ordered_powerprod_axioms (le_of_nat_term_order cmp_term)"
proof (unfold_locales)
fix s t u
assume "le_of_nat_term_order cmp_term s t"
hence "le_of_nat_term_order cmp_term (u + s) (u + t)" by (rule le_of_nat_term_order_plus_monotone)
thus "le_of_nat_term_order cmp_term (s + u) (t + u)" by (simp only: ac_simps)
qed
qed
next
show "class.linorder (le_of_nat_term_order cmp_term) (lt_of_nat_term_order cmp_term)"
by (fact linorder_le_of_nat_term_order)
next
show "ordered_term_axioms to_pair_unit fst (\<lambda>s t. le_of_nat_term_order cmp_term (fst (s, the_min)) (fst (t, the_min)))
(le_of_nat_term_order cmp_term)" by (unfold_locales, auto intro: le_of_nat_term_order_plus_monotone)
qed
corollary gd_nat_term_to_pair_unit:
"gd_nat_term (to_pair_unit::'a::{nat_term_compare,nat_pp_term,graded_dickson_powerprod} \<Rightarrow> _) fst cmp_term"
by (rule gd_nat_term.intro, fact gd_term_to_pair_unit, rule gd_nat_term_axioms.intro, simp add: splus_pp_term)
lemma gd_term_id:
"gd_term (\<lambda>x::('a::{nat_term_compare,nat_pp_compare,nat_pp_term,graded_dickson_powerprod} \<times> 'b::{nat,the_min}). x) (\<lambda>x. x)
(\<lambda>s t. le_of_nat_term_order cmp_term (s, the_min) (t, the_min))
(\<lambda>s t. lt_of_nat_term_order cmp_term (s, the_min) (t, the_min))
(le_of_nat_term_order cmp_term)
(lt_of_nat_term_order cmp_term)"
apply (rule term_powerprod_gd_term)
subgoal by unfold_locales
subgoal by (simp add: rep_nat_term_prod_def)
subgoal by (simp add: rep_nat_term_prod_def the_min_eq_zero)
subgoal by (simp add: rep_nat_term_prod_def ord_iff[symmetric])
subgoal by (simp add: splus_prod_def pprod.splus_def)
subgoal by (simp add: splus_prod_def)
done
corollary gd_nat_term_id: "gd_nat_term (\<lambda>x. x) (\<lambda>x. x) cmp_term"
for cmp_term :: "('a::{nat_term_compare,nat_pp_compare,nat_pp_term,graded_dickson_powerprod} \<times> 'c::{nat,the_min}) nat_term_order"
by (rule gd_nat_term.intro, fact gd_term_id, rule gd_nat_term_axioms.intro, simp add: splus_prod_def)
subsection \<open>Computations\<close>
type_synonym 'a mpoly_tc = "(nat, nat) pp \<Rightarrow>\<^sub>0 'a"
global_interpretation punit0: gd_nat_term "to_pair_unit::'a::{nat_term_compare,nat_pp_term,graded_dickson_powerprod} \<Rightarrow> _" fst cmp_term
rewrites "punit.adds_term = (adds)"
and "punit.pp_of_term = (\<lambda>x. x)"
and "punit.component_of_term = (\<lambda>_. ())"
for cmp_term
defines monom_mult_punit = punit.monom_mult
and mult_scalar_punit = punit.mult_scalar
and shift_map_keys_punit = punit0.shift_map_keys
and ord_pp_punit = punit0.ord_pp
and ord_pp_strict_punit = punit0.ord_pp_strict
and min_term_punit = punit0.min_term
and lt_punit = punit0.lt
and lc_punit = punit0.lc
and tail_punit = punit0.tail
and comp_opt_p_punit = punit0.comp_opt_p
and ord_p_punit = punit0.ord_p
and ord_strict_p_punit = punit0.ord_strict_p
and keys_to_list_punit = punit0.keys_to_list
subgoal by (fact gd_nat_term_to_pair_unit)
subgoal by (fact punit_adds_term)
subgoal by (fact punit_pp_of_term)
subgoal by (fact punit_component_of_term)
done
lemma shift_map_keys_punit_MP_oalist [code abstract]:
"list_of_oalist_ntm (shift_map_keys_punit t f xs) = map_raw (\<lambda>(k, v). (t + k, f v)) (list_of_oalist_ntm xs)"
by (simp add: punit0.list_of_oalist_shift_keys case_prod_beta')
lemmas [code] = punit0.mult_scalar_MP_oalist[unfolded mult_scalar_punit_def punit_mult_scalar]
punit0.punit_min_term
lemma ord_pp_punit_alt [code_unfold]: "ord_pp_punit = le_of_nat_term_order"
by (intro ext, simp add: punit0.ord_pp_def)
lemma ord_pp_strict_punit_alt [code_unfold]: "ord_pp_strict_punit = lt_of_nat_term_order"
by (intro ext, simp add: punit0.ord_pp_strict_def)
lemma gd_powerprod_ord_pp_punit: "gd_powerprod (ord_pp_punit cmp_term) (ord_pp_strict_punit cmp_term)"
unfolding punit0.ord_pp_def punit0.ord_pp_strict_def ..
locale trivariate\<^sub>0_rat
begin
abbreviation X::"rat mpoly_tc" where "X \<equiv> V\<^sub>0 (0::nat)"
abbreviation Y::"rat mpoly_tc" where "Y \<equiv> V\<^sub>0 (1::nat)"
abbreviation Z::"rat mpoly_tc" where "Z \<equiv> V\<^sub>0 (2::nat)"
end
experiment begin interpretation trivariate\<^sub>0_rat .
value [code] "X ^ 2"
value [code] "X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2"
value [code] "distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0::nat, 3::nat)], 1::rat)] = distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0, 3)], 1)]"
lemma
"ord_strict_p_punit DRLEX (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) (X\<^sup>2 * Z\<^sup>2 + 2 * Y ^ 3 * Z\<^sup>2)"
by eval
lemma
"tail_punit DLEX (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) = X\<^sup>2 * Z"
by eval
value [code] "min_term_punit::(nat, nat) pp"
value [code] "is_zero (distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0::nat, 3::nat)], 1::rat)])"
value [code] "lt_punit DRLEX (distr\<^sub>0 DRLEX [(sparse\<^sub>0 [(0::nat, 3::nat)], 1::rat)])"
lemma
"lt_punit DRLEX (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) = sparse\<^sub>0 [(1, 3), (2, 2)]"
by eval
lemma
"lt_punit DRLEX (X + Y + Z) = sparse\<^sub>0 [(2, 1)]"
by eval
lemma
"keys (X\<^sup>2 * Z ^ 3 + 2 * Y ^ 3 * Z\<^sup>2) =
{sparse\<^sub>0 [(0, 2), (2, 3)], sparse\<^sub>0 [(1, 3), (2, 2)]}"
by eval
lemma
"- 1 * X\<^sup>2 * Z ^ 7 + - 2 * Y ^ 3 * Z\<^sup>2 = - X\<^sup>2 * Z ^ 7 + - 2 * Y ^ 3 * Z\<^sup>2"
by eval
lemma
"X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 + X\<^sup>2 * Z ^ 4 + - 2 * Y ^ 3 * Z\<^sup>2 = X\<^sup>2 * Z ^ 7 + X\<^sup>2 * Z ^ 4"
by eval
lemma
"X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 - X\<^sup>2 * Z ^ 4 + - 2 * Y ^ 3 * Z\<^sup>2 =
X\<^sup>2 * Z ^ 7 - X\<^sup>2 * Z ^ 4"
by eval
lemma
"lookup (X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 + 2) (sparse\<^sub>0 [(0, 2), (2, 7)]) = 1"
by eval
lemma
"X\<^sup>2 * Z ^ 7 + 2 * Y ^ 3 * Z\<^sup>2 \<noteq>
X\<^sup>2 * Z ^ 4 + - 2 * Y ^ 3 * Z\<^sup>2"
by eval
lemma
"0 * X^2 * Z^7 + 0 * Y^3*Z\<^sup>2 = 0"
by eval
lemma
"monom_mult_punit 3 (sparse\<^sub>0 [(1, 2::nat)]) (X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) =
3 * Y\<^sup>2 * Z * X\<^sup>2 + 6 * Y ^ 5 * Z\<^sup>2"
by eval
lemma
"monomial (-4) (sparse\<^sub>0 [(0, 2::nat)]) = - 4 * X\<^sup>2"
by eval
lemma "monomial (0::rat) (sparse\<^sub>0 [(0::nat, 2::nat)]) = 0"
by eval
lemma
"(X\<^sup>2 * Z + 2 * Y ^ 3 * Z\<^sup>2) * (X\<^sup>2 * Z ^ 3 + - 2 * Y ^ 3 * Z\<^sup>2) =
X ^ 4 * Z ^ 4 + - 2 * X\<^sup>2 * Z ^ 3 * Y ^ 3 +
- 4 * Y ^ 6 * Z ^ 4 + 2 * Y ^ 3 * Z ^ 5 * X\<^sup>2"
by eval
end
subsection \<open>Code setup for type MPoly\<close>
text \<open>postprocessing from \<open>Var\<^sub>0, Const\<^sub>0\<close> to \<open>Var, Const\<close>.\<close>
lemmas [code_post] =
plus_mpoly.abs_eq[symmetric]
times_mpoly.abs_eq[symmetric]
one_mpoly_def[symmetric]
Var.abs_eq[symmetric]
Const.abs_eq[symmetric]
instantiation mpoly::("{equal, zero}")equal begin
lift_definition equal_mpoly:: "'a mpoly \<Rightarrow> 'a mpoly \<Rightarrow> bool" is HOL.equal .
instance proof standard qed (transfer, rule equal_eq)
end
end (* theory *)
diff --git a/thys/Polynomials/MPoly_Type_Class_Ordered.thy b/thys/Polynomials/MPoly_Type_Class_Ordered.thy
--- a/thys/Polynomials/MPoly_Type_Class_Ordered.thy
+++ b/thys/Polynomials/MPoly_Type_Class_Ordered.thy
@@ -1,3677 +1,3676 @@
(* Author: Fabian Immler, Alexander Maletzky *)
section \<open>Type-Class-Multivariate Polynomials in Ordered Terms\<close>
theory MPoly_Type_Class_Ordered
imports MPoly_Type_Class
begin
class the_min = linorder +
fixes the_min::'a
assumes the_min_min: "the_min \<le> x"
text \<open>Type class @{class the_min} guarantees that a least element exists. Instances of @{class the_min}
should provide @{emph \<open>computable\<close>} definitions of that element.\<close>
instantiation nat :: the_min
begin
definition "the_min_nat = (0::nat)"
instance by (standard, simp add: the_min_nat_def)
end
instantiation unit :: the_min
begin
definition "the_min_unit = ()"
instance by (standard, simp add: the_min_unit_def)
end
locale ordered_term =
term_powerprod pair_of_term term_of_pair +
ordered_powerprod ord ord_strict +
ord_term_lin: linorder ord_term ord_term_strict
for pair_of_term::"'t \<Rightarrow> ('a::comm_powerprod \<times> 'k::{the_min,wellorder})"
and term_of_pair::"('a \<times> 'k) \<Rightarrow> 't"
and ord::"'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<preceq>" 50)
and ord_strict (infixl "\<prec>" 50)
and ord_term::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<preceq>\<^sub>t" 50)
and ord_term_strict::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<prec>\<^sub>t" 50) +
assumes splus_mono: "v \<preceq>\<^sub>t w \<Longrightarrow> t \<oplus> v \<preceq>\<^sub>t t \<oplus> w"
assumes ord_termI: "pp_of_term v \<preceq> pp_of_term w \<Longrightarrow> component_of_term v \<le> component_of_term w \<Longrightarrow> v \<preceq>\<^sub>t w"
begin
abbreviation ord_term_conv (infixl "\<succeq>\<^sub>t" 50) where "ord_term_conv \<equiv> (\<preceq>\<^sub>t)\<inverse>\<inverse>"
abbreviation ord_term_strict_conv (infixl "\<succ>\<^sub>t" 50) where "ord_term_strict_conv \<equiv> (\<prec>\<^sub>t)\<inverse>\<inverse>"
text \<open>The definition of @{locale ordered_term} only covers TOP and POT orderings.
These two types of orderings are the only interesting ones.\<close>
definition "min_term \<equiv> term_of_pair (0, the_min)"
lemma min_term_min: "min_term \<preceq>\<^sub>t v"
proof (rule ord_termI)
show "pp_of_term min_term \<preceq> pp_of_term v" by (simp add: min_term_def zero_min term_simps)
next
show "component_of_term min_term \<le> component_of_term v" by (simp add: min_term_def the_min_min term_simps)
qed
lemma splus_mono_strict:
assumes "v \<prec>\<^sub>t w"
shows "t \<oplus> v \<prec>\<^sub>t t \<oplus> w"
proof -
from assms have "v \<preceq>\<^sub>t w" and "v \<noteq> w" by simp_all
from this(1) have "t \<oplus> v \<preceq>\<^sub>t t \<oplus> w" by (rule splus_mono)
moreover from \<open>v \<noteq> w\<close> have "t \<oplus> v \<noteq> t \<oplus> w" by (simp add: term_simps)
ultimately show ?thesis using ord_term_lin.antisym_conv1 by blast
qed
lemma splus_mono_left:
assumes "s \<preceq> t"
shows "s \<oplus> v \<preceq>\<^sub>t t \<oplus> v"
proof (rule ord_termI, simp_all add: term_simps)
from assms show "s + pp_of_term v \<preceq> t + pp_of_term v" by (rule plus_monotone)
qed
lemma splus_mono_strict_left:
assumes "s \<prec> t"
shows "s \<oplus> v \<prec>\<^sub>t t \<oplus> v"
proof -
from assms have "s \<preceq> t" and "s \<noteq> t" by simp_all
from this(1) have "s \<oplus> v \<preceq>\<^sub>t t \<oplus> v" by (rule splus_mono_left)
moreover from \<open>s \<noteq> t\<close> have "s \<oplus> v \<noteq> t \<oplus> v" by (simp add: term_simps)
ultimately show ?thesis using ord_term_lin.antisym_conv1 by blast
qed
lemma ord_term_canc:
assumes "t \<oplus> v \<preceq>\<^sub>t t \<oplus> w"
shows "v \<preceq>\<^sub>t w"
proof (rule ccontr)
assume "\<not> v \<preceq>\<^sub>t w"
hence "w \<prec>\<^sub>t v" by simp
hence "t \<oplus> w \<prec>\<^sub>t t \<oplus> v" by (rule splus_mono_strict)
with assms show False by simp
qed
lemma ord_term_strict_canc:
assumes "t \<oplus> v \<prec>\<^sub>t t \<oplus> w"
shows "v \<prec>\<^sub>t w"
proof (rule ccontr)
assume "\<not> v \<prec>\<^sub>t w"
hence "w \<preceq>\<^sub>t v" by simp
hence "t \<oplus> w \<preceq>\<^sub>t t \<oplus> v" by (rule splus_mono)
with assms show False by simp
qed
lemma ord_term_canc_left:
assumes "t \<oplus> v \<preceq>\<^sub>t s \<oplus> v"
shows "t \<preceq> s"
proof (rule ccontr)
assume "\<not> t \<preceq> s"
hence "s \<prec> t" by simp
hence "s \<oplus> v \<prec>\<^sub>t t \<oplus> v" by (rule splus_mono_strict_left)
with assms show False by simp
qed
lemma ord_term_strict_canc_left:
assumes "t \<oplus> v \<prec>\<^sub>t s \<oplus> v"
shows "t \<prec> s"
proof (rule ccontr)
assume "\<not> t \<prec> s"
hence "s \<preceq> t" by simp
hence "s \<oplus> v \<preceq>\<^sub>t t \<oplus> v" by (rule splus_mono_left)
with assms show False by simp
qed
lemma ord_adds_term:
assumes "u adds\<^sub>t v"
shows "u \<preceq>\<^sub>t v"
proof -
from assms have *: "component_of_term u \<le> component_of_term v" and "pp_of_term u adds pp_of_term v"
by (simp_all add: adds_term_def)
from this(2) have "pp_of_term u \<preceq> pp_of_term v" by (rule ord_adds)
from this * show ?thesis by (rule ord_termI)
qed
end
subsection \<open>Interpretations\<close>
context ordered_powerprod
begin
subsubsection \<open>Unit\<close>
sublocale punit: ordered_term to_pair_unit fst "(\<preceq>)" "(\<prec>)" "(\<preceq>)" "(\<prec>)"
apply standard
subgoal by (simp, fact plus_monotone_left)
subgoal by (simp only: punit_pp_of_term punit_component_of_term)
done
lemma punit_min_term [simp]: "punit.min_term = 0"
by (simp add: punit.min_term_def)
end
subsection \<open>Definitions\<close>
context ordered_term
begin
definition higher :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 't \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b::zero)"
where "higher p t = except p {s. s \<preceq>\<^sub>t t}"
definition lower :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 't \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b::zero)"
where "lower p t = except p {s. t \<preceq>\<^sub>t s}"
definition lt :: "('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 't"
where "lt p = (if p = 0 then min_term else ord_term_lin.Max (keys p))"
abbreviation "lp p \<equiv> pp_of_term (lt p)"
definition lc :: "('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 'b"
where "lc p = lookup p (lt p)"
definition tt :: "('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 't"
where "tt p = (if p = 0 then min_term else ord_term_lin.Min (keys p))"
abbreviation "tp p \<equiv> pp_of_term (tt p)"
definition tc :: "('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 'b"
where "tc p \<equiv> lookup p (tt p)"
definition tail :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b::zero)"
where "tail p \<equiv> lower p (lt p)"
subsection \<open>Leading Term and Leading Coefficient: @{const lt} and @{const lc}\<close>
lemma lt_zero [simp]: "lt 0 = min_term"
by (simp add: lt_def)
lemma lc_zero [simp]: "lc 0 = 0"
by (simp add: lc_def)
lemma lt_uminus [simp]: "lt (- p) = lt p"
by (simp add: lt_def keys_uminus)
lemma lc_uminus [simp]: "lc (- p) = - lc p"
by (simp add: lc_def)
lemma lt_alt:
assumes "p \<noteq> 0"
shows "lt p = ord_term_lin.Max (keys p)"
using assms unfolding lt_def by simp
lemma lt_max:
assumes "lookup p v \<noteq> 0"
shows "v \<preceq>\<^sub>t lt p"
proof -
from assms have t_in: "v \<in> keys p" by (simp add: in_keys_iff)
hence "keys p \<noteq> {}" by auto
hence "p \<noteq> 0" using keys_zero by blast
from lt_alt[OF this] ord_term_lin.Max_ge[OF finite_keys t_in] show ?thesis by simp
qed
lemma lt_eqI:
assumes "lookup p v \<noteq> 0" and "\<And>u. lookup p u \<noteq> 0 \<Longrightarrow> u \<preceq>\<^sub>t v"
shows "lt p = v"
proof -
from assms(1) have "v \<in> keys p" by (simp add: in_keys_iff)
hence "keys p \<noteq> {}" by auto
hence "p \<noteq> 0"
using keys_zero by blast
have "u \<preceq>\<^sub>t v" if "u \<in> keys p" for u
proof -
from that have "lookup p u \<noteq> 0" by (simp add: in_keys_iff)
thus "u \<preceq>\<^sub>t v" by (rule assms(2))
qed
from lt_alt[OF \<open>p \<noteq> 0\<close>] ord_term_lin.Max_eqI[OF finite_keys this \<open>v \<in> keys p\<close>] show ?thesis by simp
qed
lemma lt_less:
assumes "p \<noteq> 0" and "\<And>u. v \<preceq>\<^sub>t u \<Longrightarrow> lookup p u = 0"
shows "lt p \<prec>\<^sub>t v"
proof -
from \<open>p \<noteq> 0\<close> have "keys p \<noteq> {}"
by simp
have "\<forall>u\<in>keys p. u \<prec>\<^sub>t v"
proof
fix u::'t
assume "u \<in> keys p"
hence "lookup p u \<noteq> 0" by (simp add: in_keys_iff)
hence "\<not> v \<preceq>\<^sub>t u" using assms(2)[of u] by auto
thus "u \<prec>\<^sub>t v" by simp
qed
with lt_alt[OF assms(1)] ord_term_lin.Max_less_iff[OF finite_keys \<open>keys p \<noteq> {}\<close>] show ?thesis by simp
qed
lemma lt_le:
assumes "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = 0"
shows "lt p \<preceq>\<^sub>t v"
proof (cases "p = 0")
case True
show ?thesis by (simp add: True min_term_min)
next
case False
hence "keys p \<noteq> {}" by simp
have "\<forall>u\<in>keys p. u \<preceq>\<^sub>t v"
proof
fix u::'t
assume "u \<in> keys p"
hence "lookup p u \<noteq> 0" unfolding keys_def by simp
hence "\<not> v \<prec>\<^sub>t u" using assms[of u] by auto
thus "u \<preceq>\<^sub>t v" by simp
qed
with lt_alt[OF False] ord_term_lin.Max_le_iff[OF finite_keys[of p] \<open>keys p \<noteq> {}\<close>]
show ?thesis by simp
qed
lemma lt_gr:
assumes "lookup p s \<noteq> 0" and "t \<prec>\<^sub>t s"
shows "t \<prec>\<^sub>t lt p"
using assms lt_max ord_term_lin.order.strict_trans2 by blast
lemma lc_not_0:
assumes "p \<noteq> 0"
shows "lc p \<noteq> 0"
proof -
from keys_zero assms have "keys p \<noteq> {}" by auto
from lt_alt[OF assms] ord_term_lin.Max_in[OF finite_keys this] show ?thesis by (simp add: in_keys_iff lc_def)
qed
lemma lc_eq_zero_iff: "lc p = 0 \<longleftrightarrow> p = 0"
using lc_not_0 lc_zero by blast
lemma lt_in_keys:
assumes "p \<noteq> 0"
shows "lt p \<in> (keys p)"
by (metis assms in_keys_iff lc_def lc_not_0)
lemma lt_monomial:
- assumes "c \<noteq> 0"
- shows "lt (monomial c t) = t"
- by (metis assms lookup_single_eq lookup_single_not_eq lt_eqI ord_term_lin.eq_iff)
+ "lt (monomial c t) = t" if "c \<noteq> 0"
+ using that by (auto simp add: lt_def dest: monomial_0D)
lemma lc_monomial [simp]: "lc (monomial c t) = c"
proof (cases "c = 0")
case True
thus ?thesis by simp
next
case False
thus ?thesis by (simp add: lc_def lt_monomial)
qed
lemma lt_le_iff: "lt p \<preceq>\<^sub>t v \<longleftrightarrow> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = 0)" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
show ?R
proof (intro allI impI)
fix u
note \<open>lt p \<preceq>\<^sub>t v\<close>
also assume "v \<prec>\<^sub>t u"
finally have "lt p \<prec>\<^sub>t u" .
hence "\<not> u \<preceq>\<^sub>t lt p" by simp
with lt_max[of p u] show "lookup p u = 0" by blast
qed
next
assume ?R
thus ?L using lt_le by auto
qed
lemma lt_plus_eqI:
assumes "lt p \<prec>\<^sub>t lt q"
shows "lt (p + q) = lt q"
proof (cases "q = 0")
case True
with assms have "lt p \<prec>\<^sub>t min_term" by (simp add: lt_def)
with min_term_min[of "lt p"] show ?thesis by simp
next
case False
show ?thesis
proof (intro lt_eqI)
from lt_gr[of p "lt q" "lt p"] assms have "lookup p (lt q) = 0" by blast
with lookup_add[of p q "lt q"] lc_not_0[OF False] show "lookup (p + q) (lt q) \<noteq> 0"
unfolding lc_def by simp
next
fix u
assume "lookup (p + q) u \<noteq> 0"
show "u \<preceq>\<^sub>t lt q"
proof (rule ccontr)
assume "\<not> u \<preceq>\<^sub>t lt q"
hence qs: "lt q \<prec>\<^sub>t u" by simp
with assms have "lt p \<prec>\<^sub>t u" by simp
with lt_gr[of p u "lt p"] have "lookup p u = 0" by blast
moreover from qs lt_gr[of q u "lt q"] have "lookup q u = 0" by blast
ultimately show False using \<open>lookup (p + q) u \<noteq> 0\<close> lookup_add[of p q u] by auto
qed
qed
qed
lemma lt_plus_eqI_2:
assumes "lt q \<prec>\<^sub>t lt p"
shows "lt (p + q) = lt p"
proof (cases "p = 0")
case True
with assms have "lt q \<prec>\<^sub>t min_term" by (simp add: lt_def)
with min_term_min[of "lt q"] show ?thesis by simp
next
case False
show ?thesis
proof (intro lt_eqI)
from lt_gr[of q "lt p" "lt q"] assms have "lookup q (lt p) = 0" by blast
with lookup_add[of p q "lt p"] lc_not_0[OF False] show "lookup (p + q) (lt p) \<noteq> 0"
unfolding lc_def by simp
next
fix u
assume "lookup (p + q) u \<noteq> 0"
show "u \<preceq>\<^sub>t lt p"
proof (rule ccontr)
assume "\<not> u \<preceq>\<^sub>t lt p"
hence ps: "lt p \<prec>\<^sub>t u" by simp
with assms have "lt q \<prec>\<^sub>t u" by simp
with lt_gr[of q u "lt q"] have "lookup q u = 0" by blast
also from ps lt_gr[of p u "lt p"] have "lookup p u = 0" by blast
ultimately show False using \<open>lookup (p + q) u \<noteq> 0\<close> lookup_add[of p q u] by auto
qed
qed
qed
lemma lt_plus_eqI_3:
assumes "lt q = lt p" and "lc p + lc q \<noteq> 0"
shows "lt (p + q) = lt (p::'t \<Rightarrow>\<^sub>0 'b::monoid_add)"
proof (rule lt_eqI)
from assms(2) show "lookup (p + q) (lt p) \<noteq> 0" by (simp add: lookup_add lc_def assms(1))
next
fix u
assume "lookup (p + q) u \<noteq> 0"
hence "lookup p u + lookup q u \<noteq> 0" by (simp add: lookup_add)
hence "lookup p u \<noteq> 0 \<or> lookup q u \<noteq> 0" by auto
thus "u \<preceq>\<^sub>t lt p"
proof
assume "lookup p u \<noteq> 0"
thus ?thesis by (rule lt_max)
next
assume "lookup q u \<noteq> 0"
hence "u \<preceq>\<^sub>t lt q" by (rule lt_max)
thus ?thesis by (simp only: assms(1))
qed
qed
lemma lt_plus_lessE:
assumes "lt p \<prec>\<^sub>t lt (p + q)"
shows "lt p \<prec>\<^sub>t lt q"
proof (rule ccontr)
assume "\<not> lt p \<prec>\<^sub>t lt q"
hence "lt p = lt q \<or> lt q \<prec>\<^sub>t lt p" by auto
thus False
proof
assume lt_eq: "lt p = lt q"
have "lt (p + q) \<preceq>\<^sub>t lt p"
proof (rule lt_le)
fix u
assume "lt p \<prec>\<^sub>t u"
with lt_gr[of p u "lt p"] have "lookup p u = 0" by blast
from \<open>lt p \<prec>\<^sub>t u\<close> have "lt q \<prec>\<^sub>t u" using lt_eq by simp
with lt_gr[of q u "lt q"] have "lookup q u = 0" by blast
with \<open>lookup p u = 0\<close> show "lookup (p + q) u = 0" by (simp add: lookup_add)
qed
with assms show False by simp
next
assume "lt q \<prec>\<^sub>t lt p"
from lt_plus_eqI_2[OF this] assms show False by simp
qed
qed
lemma lt_plus_lessE_2:
assumes "lt q \<prec>\<^sub>t lt (p + q)"
shows "lt q \<prec>\<^sub>t lt p"
proof (rule ccontr)
assume "\<not> lt q \<prec>\<^sub>t lt p"
hence "lt q = lt p \<or> lt p \<prec>\<^sub>t lt q" by auto
thus False
proof
assume lt_eq: "lt q = lt p"
have "lt (p + q) \<preceq>\<^sub>t lt q"
proof (rule lt_le)
fix u
assume "lt q \<prec>\<^sub>t u"
with lt_gr[of q u "lt q"] have "lookup q u = 0" by blast
from \<open>lt q \<prec>\<^sub>t u\<close> have "lt p \<prec>\<^sub>t u" using lt_eq by simp
with lt_gr[of p u "lt p"] have "lookup p u = 0" by blast
with \<open>lookup q u = 0\<close> show "lookup (p + q) u = 0" by (simp add: lookup_add)
qed
with assms show False by simp
next
assume "lt p \<prec>\<^sub>t lt q"
from lt_plus_eqI[OF this] assms show False by simp
qed
qed
lemma lt_plus_lessI':
fixes p q :: "'t \<Rightarrow>\<^sub>0 'b::monoid_add"
assumes "p + q \<noteq> 0" and lt_eq: "lt q = lt p" and lc_eq: "lc p + lc q = 0"
shows "lt (p + q) \<prec>\<^sub>t lt p"
proof (rule ccontr)
assume "\<not> lt (p + q) \<prec>\<^sub>t lt p"
hence "lt (p + q) = lt p \<or> lt p \<prec>\<^sub>t lt (p + q)" by auto
thus False
proof
assume "lt (p + q) = lt p"
have "lookup (p + q) (lt p) = (lookup p (lt p)) + (lookup q (lt q))" unfolding lt_eq lookup_add ..
also have "... = lc p + lc q" unfolding lc_def ..
also have "... = 0" unfolding lc_eq by simp
finally have "lookup (p + q) (lt p) = 0" .
hence "lt (p + q) \<noteq> lt p" using lc_not_0[OF \<open>p + q \<noteq> 0\<close>] unfolding lc_def by auto
with \<open>lt (p + q) = lt p\<close> show False by simp
next
assume "lt p \<prec>\<^sub>t lt (p + q)"
have "lt p \<prec>\<^sub>t lt q" by (rule lt_plus_lessE, fact+)
hence "lt p \<noteq> lt q" by simp
with lt_eq show False by simp
qed
qed
corollary lt_plus_lessI:
fixes p q :: "'t \<Rightarrow>\<^sub>0 'b::group_add"
assumes "p + q \<noteq> 0" and "lt q = lt p" and "lc q = - lc p"
shows "lt (p + q) \<prec>\<^sub>t lt p"
using assms(1, 2)
proof (rule lt_plus_lessI')
from assms(3) show "lc p + lc q = 0" by simp
qed
lemma lt_plus_distinct_eq_max:
assumes "lt p \<noteq> lt q"
shows "lt (p + q) = ord_term_lin.max (lt p) (lt q)"
proof (rule ord_term_lin.linorder_cases)
assume a: "lt p \<prec>\<^sub>t lt q"
hence "lt (p + q) = lt q" by (rule lt_plus_eqI)
also from a have "... = ord_term_lin.max (lt p) (lt q)"
by (simp add: ord_term_lin.max.absorb2)
finally show ?thesis .
next
assume a: "lt q \<prec>\<^sub>t lt p"
hence "lt (p + q) = lt p" by (rule lt_plus_eqI_2)
also from a have "... = ord_term_lin.max (lt p) (lt q)"
by (simp add: ord_term_lin.max.absorb1)
finally show ?thesis .
next
assume "lt p = lt q"
with assms show ?thesis ..
qed
lemma lt_plus_le_max: "lt (p + q) \<preceq>\<^sub>t ord_term_lin.max (lt p) (lt q)"
proof (cases "lt p = lt q")
case True
show ?thesis
proof (rule lt_le)
fix u
assume "ord_term_lin.max (lt p) (lt q) \<prec>\<^sub>t u"
hence "lt p \<prec>\<^sub>t u" and "lt q \<prec>\<^sub>t u" by simp_all
hence "lookup p u = 0" and "lookup q u = 0" using lt_max ord_term_lin.leD by blast+
thus "lookup (p + q) u = 0" by (simp add: lookup_add)
qed
next
case False
hence "lt (p + q) = ord_term_lin.max (lt p) (lt q)" by (rule lt_plus_distinct_eq_max)
thus ?thesis by simp
qed
lemma lt_minus_eqI: "lt p \<prec>\<^sub>t lt q \<Longrightarrow> lt (p - q) = lt q" for p q :: "'t \<Rightarrow>\<^sub>0 'b::ab_group_add"
by (metis lt_plus_eqI_2 lt_uminus uminus_add_conv_diff)
lemma lt_minus_eqI_2: "lt q \<prec>\<^sub>t lt p \<Longrightarrow> lt (p - q) = lt p" for p q :: "'t \<Rightarrow>\<^sub>0 'b::ab_group_add"
by (metis lt_minus_eqI lt_uminus minus_diff_eq)
lemma lt_minus_eqI_3:
assumes "lt q = lt p" and "lc q \<noteq> lc p"
shows "lt (p - q) = lt (p::'t \<Rightarrow>\<^sub>0 'b::ab_group_add)"
proof (rule lt_eqI)
from assms(2) show "lookup (p - q) (lt p) \<noteq> 0" by (simp add: lookup_minus lc_def assms(1))
next
fix u
assume "lookup (p - q) u \<noteq> 0"
hence "lookup p u \<noteq> lookup q u" by (simp add: lookup_minus)
hence "lookup p u \<noteq> 0 \<or> lookup q u \<noteq> 0" by auto
thus "u \<preceq>\<^sub>t lt p"
proof
assume "lookup p u \<noteq> 0"
thus ?thesis by (rule lt_max)
next
assume "lookup q u \<noteq> 0"
hence "u \<preceq>\<^sub>t lt q" by (rule lt_max)
thus ?thesis by (simp only: assms(1))
qed
qed
lemma lt_minus_distinct_eq_max:
assumes "lt p \<noteq> lt (q::'t \<Rightarrow>\<^sub>0 'b::ab_group_add)"
shows "lt (p - q) = ord_term_lin.max (lt p) (lt q)"
proof (rule ord_term_lin.linorder_cases)
assume a: "lt p \<prec>\<^sub>t lt q"
hence "lt (p - q) = lt q" by (rule lt_minus_eqI)
also from a have "... = ord_term_lin.max (lt p) (lt q)"
by (simp add: ord_term_lin.max.absorb2)
finally show ?thesis .
next
assume a: "lt q \<prec>\<^sub>t lt p"
hence "lt (p - q) = lt p" by (rule lt_minus_eqI_2)
also from a have "... = ord_term_lin.max (lt p) (lt q)"
by (simp add: ord_term_lin.max.absorb1)
finally show ?thesis .
next
assume "lt p = lt q"
with assms show ?thesis ..
qed
lemma lt_minus_lessE: "lt p \<prec>\<^sub>t lt (p - q) \<Longrightarrow> lt p \<prec>\<^sub>t lt q" for p q :: "'t \<Rightarrow>\<^sub>0 'b::ab_group_add"
using lt_minus_eqI_2 by fastforce
lemma lt_minus_lessE_2: "lt q \<prec>\<^sub>t lt (p - q) \<Longrightarrow> lt q \<prec>\<^sub>t lt p" for p q :: "'t \<Rightarrow>\<^sub>0 'b::ab_group_add"
using lt_plus_eqI_2 by fastforce
lemma lt_minus_lessI: "p - q \<noteq> 0 \<Longrightarrow> lt q = lt p \<Longrightarrow> lc q = lc p \<Longrightarrow> lt (p - q) \<prec>\<^sub>t lt p"
for p q :: "'t \<Rightarrow>\<^sub>0 'b::ab_group_add"
by (metis (no_types, hide_lams) diff_diff_eq2 diff_self group_eq_aux lc_def lc_not_0 lookup_minus
lt_minus_eqI ord_term_lin.antisym_conv3)
lemma lt_max_keys:
assumes "v \<in> keys p"
shows "v \<preceq>\<^sub>t lt p"
proof (rule lt_max)
from assms show "lookup p v \<noteq> 0" by (simp add: in_keys_iff)
qed
lemma lt_eqI_keys:
assumes "v \<in> keys p" and a2: "\<And>u. u \<in> keys p \<Longrightarrow> u \<preceq>\<^sub>t v"
shows "lt p = v"
by (rule lt_eqI, simp_all only: in_keys_iff[symmetric], fact+)
lemma lt_gr_keys:
assumes "u \<in> keys p" and "v \<prec>\<^sub>t u"
shows "v \<prec>\<^sub>t lt p"
proof (rule lt_gr)
from assms(1) show "lookup p u \<noteq> 0" by (simp add: in_keys_iff)
qed fact
lemma lt_plus_eq_maxI:
assumes "lt p = lt q \<Longrightarrow> lc p + lc q \<noteq> 0"
shows "lt (p + q) = ord_term_lin.max (lt p) (lt q)"
proof (cases "lt p = lt q")
case True
show ?thesis
proof (rule lt_eqI_keys)
from True have "lc p + lc q \<noteq> 0" by (rule assms)
thus "ord_term_lin.max (lt p) (lt q) \<in> keys (p + q)"
by (simp add: in_keys_iff lc_def lookup_add True)
next
fix u
assume "u \<in> keys (p + q)"
hence "u \<preceq>\<^sub>t lt (p + q)" by (rule lt_max_keys)
also have "... \<preceq>\<^sub>t ord_term_lin.max (lt p) (lt q)" by (fact lt_plus_le_max)
finally show "u \<preceq>\<^sub>t ord_term_lin.max (lt p) (lt q)" .
qed
next
case False
thus ?thesis by (rule lt_plus_distinct_eq_max)
qed
lemma lt_monom_mult:
assumes "c \<noteq> (0::'b::semiring_no_zero_divisors)" and "p \<noteq> 0"
shows "lt (monom_mult c t p) = t \<oplus> lt p"
proof (intro lt_eqI)
from assms(1) show "lookup (monom_mult c t p) (t \<oplus> lt p) \<noteq> 0"
proof (simp add: lookup_monom_mult_plus)
show "lookup p (lt p) \<noteq> 0"
using assms(2) lt_in_keys by auto
qed
next
fix u::'t
assume "lookup (monom_mult c t p) u \<noteq> 0"
hence "u \<in> keys (monom_mult c t p)" by (simp add: in_keys_iff)
also have "... \<subseteq> (\<oplus>) t ` keys p" by (fact keys_monom_mult_subset)
finally obtain v where "v \<in> keys p" and "u = t \<oplus> v" ..
show "u \<preceq>\<^sub>t t \<oplus> lt p" unfolding \<open>u = t \<oplus> v\<close>
proof (rule splus_mono)
from \<open>v \<in> keys p\<close> show "v \<preceq>\<^sub>t lt p" by (rule lt_max_keys)
qed
qed
lemma lt_monom_mult_zero:
assumes "c \<noteq> (0::'b::semiring_no_zero_divisors)"
shows "lt (monom_mult c 0 p) = lt p"
proof (cases "p = 0")
case True
show ?thesis by (simp add: True)
next
case False
with assms show ?thesis by (simp add: lt_monom_mult term_simps)
qed
corollary lt_map_scale: "c \<noteq> (0::'b::semiring_no_zero_divisors) \<Longrightarrow> lt (c \<cdot> p) = lt p"
by (simp add: map_scale_eq_monom_mult lt_monom_mult_zero)
lemma lc_monom_mult [simp]: "lc (monom_mult c t p) = (c::'b::semiring_no_zero_divisors) * lc p"
proof (cases "c = 0")
case True
thus ?thesis by simp
next
case False
show ?thesis
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
with \<open>c \<noteq> 0\<close> show ?thesis by (simp add: lc_def lt_monom_mult lookup_monom_mult_plus)
qed
qed
corollary lc_map_scale [simp]: "lc (c \<cdot> p) = (c::'b::semiring_no_zero_divisors) * lc p"
by (simp add: map_scale_eq_monom_mult)
lemma (in ordered_term) lt_mult_scalar_monomial_right:
assumes "c \<noteq> (0::'b::semiring_no_zero_divisors)" and "p \<noteq> 0"
shows "lt (p \<odot> monomial c v) = punit.lt p \<oplus> v"
proof (intro lt_eqI)
from assms(1) show "lookup (p \<odot> monomial c v) (punit.lt p \<oplus> v) \<noteq> 0"
proof (simp add: lookup_mult_scalar_monomial_right_plus)
from assms(2) show "lookup p (punit.lt p) \<noteq> 0"
using in_keys_iff punit.lt_in_keys by fastforce
qed
next
fix u::'t
assume "lookup (p \<odot> monomial c v) u \<noteq> 0"
hence "u \<in> keys (p \<odot> monomial c v)" by (simp add: in_keys_iff)
also have "... \<subseteq> (\<lambda>t. t \<oplus> v) ` keys p" by (fact keys_mult_scalar_monomial_right_subset)
finally obtain t where "t \<in> keys p" and "u = t \<oplus> v" ..
show "u \<preceq>\<^sub>t punit.lt p \<oplus> v" unfolding \<open>u = t \<oplus> v\<close>
proof (rule splus_mono_left)
from \<open>t \<in> keys p\<close> show "t \<preceq> punit.lt p" by (rule punit.lt_max_keys)
qed
qed
lemma lc_mult_scalar_monomial_right:
"lc (p \<odot> monomial c v) = punit.lc p * (c::'b::semiring_no_zero_divisors)"
proof (cases "c = 0")
case True
thus ?thesis by simp
next
case False
show ?thesis
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
with \<open>c \<noteq> 0\<close> show ?thesis
by (simp add: punit.lc_def lc_def lt_mult_scalar_monomial_right lookup_mult_scalar_monomial_right_plus)
qed
qed
lemma lookup_monom_mult_eq_zero:
assumes "s \<oplus> lt p \<prec>\<^sub>t v"
shows "lookup (monom_mult (c::'b::semiring_no_zero_divisors) s p) v = 0"
by (metis assms aux lt_gr lt_monom_mult monom_mult_zero_left monom_mult_zero_right
ord_term_lin.order.strict_implies_not_eq)
lemma in_keys_monom_mult_le:
assumes "v \<in> keys (monom_mult c t p)"
shows "v \<preceq>\<^sub>t t \<oplus> lt p"
proof -
from keys_monom_mult_subset assms have "v \<in> (\<oplus>) t ` (keys p)" ..
then obtain u where "u \<in> keys p" and "v = t \<oplus> u" ..
from \<open>u \<in> keys p\<close> have "u \<preceq>\<^sub>t lt p" by (rule lt_max_keys)
thus "v \<preceq>\<^sub>t t \<oplus> lt p" unfolding \<open>v = t \<oplus> u\<close> by (rule splus_mono)
qed
lemma lt_monom_mult_le: "lt (monom_mult c t p) \<preceq>\<^sub>t t \<oplus> lt p"
by (metis aux in_keys_monom_mult_le lt_in_keys lt_le_iff)
lemma monom_mult_inj_2:
assumes "monom_mult c t1 p = monom_mult c t2 p"
and "c \<noteq> 0" and "(p::'t \<Rightarrow>\<^sub>0 'b::semiring_no_zero_divisors) \<noteq> 0"
shows "t1 = t2"
proof -
from assms(1) have "lt (monom_mult c t1 p) = lt (monom_mult c t2 p)" by simp
with \<open>c \<noteq> 0\<close> \<open>p \<noteq> 0\<close> have "t1 \<oplus> lt p = t2 \<oplus> lt p" by (simp add: lt_monom_mult)
thus ?thesis by (simp add: term_simps)
qed
subsection \<open>Trailing Term and Trailing Coefficient: @{const tt} and @{const tc}\<close>
lemma tt_zero [simp]: "tt 0 = min_term"
by (simp add: tt_def)
lemma tc_zero [simp]: "tc 0 = 0"
by (simp add: tc_def)
lemma tt_alt:
assumes "p \<noteq> 0"
shows "tt p = ord_term_lin.Min (keys p)"
using assms unfolding tt_def by simp
lemma tt_min_keys:
assumes "v \<in> keys p"
shows "tt p \<preceq>\<^sub>t v"
proof -
from assms have "keys p \<noteq> {}" by auto
hence "p \<noteq> 0" by simp
from tt_alt[OF this] ord_term_lin.Min_le[OF finite_keys assms] show ?thesis by simp
qed
lemma tt_min:
assumes "lookup p v \<noteq> 0"
shows "tt p \<preceq>\<^sub>t v"
proof -
from assms have "v \<in> keys p" unfolding keys_def by simp
thus ?thesis by (rule tt_min_keys)
qed
lemma tt_in_keys:
assumes "p \<noteq> 0"
shows "tt p \<in> keys p"
unfolding tt_alt[OF assms]
by (rule ord_term_lin.Min_in, fact finite_keys, simp add: assms)
lemma tt_eqI:
assumes "v \<in> keys p" and "\<And>u. u \<in> keys p \<Longrightarrow> v \<preceq>\<^sub>t u"
shows "tt p = v"
proof -
from assms(1) have "keys p \<noteq> {}" by auto
hence "p \<noteq> 0" by simp
from assms(1) have "tt p \<preceq>\<^sub>t v" by (rule tt_min_keys)
moreover have "v \<preceq>\<^sub>t tt p" by (rule assms(2), rule tt_in_keys, fact \<open>p \<noteq> 0\<close>)
ultimately show ?thesis by simp
qed
lemma tt_gr:
assumes "\<And>u. u \<in> keys p \<Longrightarrow> v \<prec>\<^sub>t u" and "p \<noteq> 0"
shows "v \<prec>\<^sub>t tt p"
proof -
from \<open>p \<noteq> 0\<close> have "keys p \<noteq> {}" by simp
show ?thesis by (rule assms(1), rule tt_in_keys, fact \<open>p \<noteq> 0\<close>)
qed
lemma tt_less:
assumes "u \<in> keys p" and "u \<prec>\<^sub>t v"
shows "tt p \<prec>\<^sub>t v"
proof -
from \<open>u \<in> keys p\<close> have "tt p \<preceq>\<^sub>t u" by (rule tt_min_keys)
also have "... \<prec>\<^sub>t v" by fact
finally show "tt p \<prec>\<^sub>t v" .
qed
lemma tt_ge:
assumes "\<And>u. u \<prec>\<^sub>t v \<Longrightarrow> lookup p u = 0" and "p \<noteq> 0"
shows "v \<preceq>\<^sub>t tt p"
proof -
from \<open>p \<noteq> 0\<close> have "keys p \<noteq> {}" by simp
have "\<forall>u\<in>keys p. v \<preceq>\<^sub>t u"
proof
fix u::'t
assume "u \<in> keys p"
hence "lookup p u \<noteq> 0" unfolding keys_def by simp
hence "\<not> u \<prec>\<^sub>t v" using assms(1)[of u] by auto
thus "v \<preceq>\<^sub>t u" by simp
qed
with tt_alt[OF \<open>p \<noteq> 0\<close>] ord_term_lin.Min_ge_iff[OF finite_keys[of p] \<open>keys p \<noteq> {}\<close>]
show ?thesis by simp
qed
lemma tt_ge_keys:
assumes "\<And>u. u \<in> keys p \<Longrightarrow> v \<preceq>\<^sub>t u" and "p \<noteq> 0"
shows "v \<preceq>\<^sub>t tt p"
by (rule assms(1), rule tt_in_keys, fact)
lemma tt_ge_iff: "v \<preceq>\<^sub>t tt p \<longleftrightarrow> ((p \<noteq> 0 \<or> v = min_term) \<and> (\<forall>u. u \<prec>\<^sub>t v \<longrightarrow> lookup p u = 0))"
(is "?L \<longleftrightarrow> (?A \<and> ?B)")
proof
assume ?L
show "?A \<and> ?B"
proof (intro conjI allI impI)
show "p \<noteq> 0 \<or> v = min_term"
proof (cases "p = 0")
case True
show ?thesis
proof (rule disjI2)
from \<open>?L\<close> True have "v \<preceq>\<^sub>t min_term" by (simp add: tt_def)
with min_term_min[of v] show "v = min_term" by simp
qed
next
case False
thus ?thesis ..
qed
next
fix u
assume "u \<prec>\<^sub>t v"
also note \<open>v \<preceq>\<^sub>t tt p\<close>
finally have "u \<prec>\<^sub>t tt p" .
hence "\<not> tt p \<preceq>\<^sub>t u" by simp
with tt_min[of p u] show "lookup p u = 0" by blast
qed
next
assume "?A \<and> ?B"
hence ?A and ?B by simp_all
show ?L
proof (cases "p = 0")
case True
with \<open>?A\<close> have "v = min_term" by simp
with True show ?thesis by (simp add: tt_def)
next
case False
from \<open>?B\<close> show ?thesis using tt_ge[OF _ False] by auto
qed
qed
lemma tc_not_0:
assumes "p \<noteq> 0"
shows "tc p \<noteq> 0"
unfolding tc_def in_keys_iff[symmetric] using assms by (rule tt_in_keys)
lemma tt_monomial:
assumes "c \<noteq> 0"
shows "tt (monomial c v) = v"
proof (rule tt_eqI)
from keys_of_monomial[OF assms, of v] show "v \<in> keys (monomial c v)" by simp
next
fix u
assume "u \<in> keys (monomial c v)"
with keys_of_monomial[OF assms, of v] have "u = v" by simp
thus "v \<preceq>\<^sub>t u" by simp
qed
lemma tc_monomial [simp]: "tc (monomial c t) = c"
proof (cases "c = 0")
case True
thus ?thesis by simp
next
case False
thus ?thesis by (simp add: tc_def tt_monomial)
qed
lemma tt_plus_eqI:
assumes "p \<noteq> 0" and "tt p \<prec>\<^sub>t tt q"
shows "tt (p + q) = tt p"
proof (intro tt_eqI)
from tt_less[of "tt p" q "tt q"] \<open>tt p \<prec>\<^sub>t tt q\<close> have "tt p \<notin> keys q" by blast
with lookup_add[of p q "tt p"] tc_not_0[OF \<open>p \<noteq> 0\<close>] show "tt p \<in> keys (p + q)"
unfolding in_keys_iff tc_def by simp
next
fix u
assume "u \<in> keys (p + q)"
show "tt p \<preceq>\<^sub>t u"
proof (rule ccontr)
assume "\<not> tt p \<preceq>\<^sub>t u"
hence sp: "u \<prec>\<^sub>t tt p" by simp
hence "u \<prec>\<^sub>t tt q" using \<open>tt p \<prec>\<^sub>t tt q\<close> by simp
with tt_less[of u q "tt q"] have "u \<notin> keys q" by blast
moreover from sp tt_less[of u p "tt p"] have "u \<notin> keys p" by blast
ultimately show False using \<open>u \<in> keys (p + q)\<close> Poly_Mapping.keys_add[of p q] by auto
qed
qed
lemma tt_plus_lessE:
fixes p q
assumes "p + q \<noteq> 0" and tt: "tt (p + q) \<prec>\<^sub>t tt p"
shows "tt q \<prec>\<^sub>t tt p"
proof (cases "p = 0")
case True
with tt show ?thesis by simp
next
case False
show ?thesis
proof (rule ccontr)
assume "\<not> tt q \<prec>\<^sub>t tt p"
hence "tt p = tt q \<or> tt p \<prec>\<^sub>t tt q" by auto
thus False
proof
assume tt_eq: "tt p = tt q"
have "tt p \<preceq>\<^sub>t tt (p + q)"
proof (rule tt_ge_keys)
fix u
assume "u \<in> keys (p + q)"
hence "u \<in> keys p \<union> keys q"
proof
show "keys (p + q) \<subseteq> keys p \<union> keys q" by (fact Poly_Mapping.keys_add)
qed
thus "tt p \<preceq>\<^sub>t u"
proof
assume "u \<in> keys p"
thus ?thesis by (rule tt_min_keys)
next
assume "u \<in> keys q"
thus ?thesis unfolding tt_eq by (rule tt_min_keys)
qed
qed (fact \<open>p + q \<noteq> 0\<close>)
with tt show False by simp
next
assume "tt p \<prec>\<^sub>t tt q"
from tt_plus_eqI[OF False this] tt show False by (simp add: ac_simps)
qed
qed
qed
lemma tt_plus_lessI:
fixes p q :: "_ \<Rightarrow>\<^sub>0 'b::ring"
assumes "p + q \<noteq> 0" and tt_eq: "tt q = tt p" and tc_eq: "tc q = - tc p"
shows "tt p \<prec>\<^sub>t tt (p + q)"
proof (rule ccontr)
assume "\<not> tt p \<prec>\<^sub>t tt (p + q)"
hence "tt p = tt (p + q) \<or> tt (p + q) \<prec>\<^sub>t tt p" by auto
thus False
proof
assume "tt p = tt (p + q)"
have "lookup (p + q) (tt p) = (lookup p (tt p)) + (lookup q (tt q))" unfolding tt_eq lookup_add ..
also have "... = tc p + tc q" unfolding tc_def ..
also have "... = 0" unfolding tc_eq by simp
finally have "lookup (p + q) (tt p) = 0" .
hence "tt (p + q) \<noteq> tt p" using tc_not_0[OF \<open>p + q \<noteq> 0\<close>] unfolding tc_def by auto
with \<open>tt p = tt (p + q)\<close> show False by simp
next
assume "tt (p + q) \<prec>\<^sub>t tt p"
have "tt q \<prec>\<^sub>t tt p" by (rule tt_plus_lessE, fact+)
hence "tt q \<noteq> tt p" by simp
with tt_eq show False by simp
qed
qed
lemma tt_uminus [simp]: "tt (- p) = tt p"
by (simp add: tt_def keys_uminus)
lemma tc_uminus [simp]: "tc (- p) = - tc p"
by (simp add: tc_def)
lemma tt_monom_mult:
assumes "c \<noteq> (0::'b::semiring_no_zero_divisors)" and "p \<noteq> 0"
shows "tt (monom_mult c t p) = t \<oplus> tt p"
proof (intro tt_eqI, rule keys_monom_multI, rule tt_in_keys, fact, fact)
fix u
assume "u \<in> keys (monom_mult c t p)"
then obtain v where "v \<in> keys p" and u: "u = t \<oplus> v" by (rule keys_monom_multE)
show "t \<oplus> tt p \<preceq>\<^sub>t u" unfolding u add.commute[of t] by (rule splus_mono, rule tt_min_keys, fact)
qed
lemma tt_map_scale: "c \<noteq> (0::'b::semiring_no_zero_divisors) \<Longrightarrow> tt (c \<cdot> p) = tt p"
by (cases "p = 0") (simp_all add: map_scale_eq_monom_mult tt_monom_mult term_simps)
lemma tc_monom_mult [simp]: "tc (monom_mult c t p) = (c::'b::semiring_no_zero_divisors) * tc p"
proof (cases "c = 0")
case True
thus ?thesis by simp
next
case False
show ?thesis
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
with \<open>c \<noteq> 0\<close> show ?thesis by (simp add: tc_def tt_monom_mult lookup_monom_mult_plus)
qed
qed
corollary tc_map_scale [simp]: "tc (c \<cdot> p) = (c::'b::semiring_no_zero_divisors) * tc p"
by (simp add: map_scale_eq_monom_mult)
lemma in_keys_monom_mult_ge:
assumes "v \<in> keys (monom_mult c t p)"
shows "t \<oplus> tt p \<preceq>\<^sub>t v"
proof -
from keys_monom_mult_subset assms have "v \<in> (\<oplus>) t ` (keys p)" ..
then obtain u where "u \<in> keys p" and "v = t \<oplus> u" ..
from \<open>u \<in> keys p\<close> have "tt p \<preceq>\<^sub>t u" by (rule tt_min_keys)
thus "t \<oplus> tt p \<preceq>\<^sub>t v" unfolding \<open>v = t \<oplus> u\<close> by (rule splus_mono)
qed
lemma lt_ge_tt: "tt p \<preceq>\<^sub>t lt p"
proof (cases "p = 0")
case True
show ?thesis unfolding True lt_def tt_def by simp
next
case False
show ?thesis by (rule lt_max_keys, rule tt_in_keys, fact False)
qed
lemma lt_eq_tt_monomial:
assumes "is_monomial p"
shows "lt p = tt p"
proof -
from assms obtain c v where "c \<noteq> 0" and p: "p = monomial c v" by (rule is_monomial_monomial)
from \<open>c \<noteq> 0\<close> have "lt p = v" and "tt p = v" unfolding p by (rule lt_monomial, rule tt_monomial)
thus ?thesis by simp
qed
subsection \<open>@{const higher} and @{const lower}\<close>
lemma lookup_higher: "lookup (higher p u) v = (if u \<prec>\<^sub>t v then lookup p v else 0)"
by (auto simp add: higher_def lookup_except)
lemma lookup_higher_when: "lookup (higher p u) v = (lookup p v when u \<prec>\<^sub>t v)"
by (auto simp add: lookup_higher when_def)
lemma higher_plus: "higher (p + q) v = higher p v + higher q v"
by (rule poly_mapping_eqI, simp add: lookup_add lookup_higher)
lemma higher_uminus [simp]: "higher (- p) v = -(higher p v)"
by (rule poly_mapping_eqI, simp add: lookup_higher)
lemma higher_minus: "higher (p - q) v = higher p v - higher q v"
by (auto intro!: poly_mapping_eqI simp: lookup_minus lookup_higher)
lemma higher_zero [simp]: "higher 0 t = 0"
by (rule poly_mapping_eqI, simp add: lookup_higher)
lemma higher_eq_iff: "higher p v = higher q v \<longleftrightarrow> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup q u)" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
show ?R
proof (intro allI impI)
fix u
assume "v \<prec>\<^sub>t u"
moreover from \<open>?L\<close> have "lookup (higher p v) u = lookup (higher q v) u" by simp
ultimately show "lookup p u = lookup q u" by (simp add: lookup_higher)
qed
next
assume ?R
show ?L
proof (rule poly_mapping_eqI, simp add: lookup_higher, rule)
fix u
assume "v \<prec>\<^sub>t u"
with \<open>?R\<close> show "lookup p u = lookup q u" by simp
qed
qed
lemma higher_eq_zero_iff: "higher p v = 0 \<longleftrightarrow> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = 0)"
proof -
have "higher p v = higher 0 v \<longleftrightarrow> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup 0 u)" by (rule higher_eq_iff)
thus ?thesis by simp
qed
lemma keys_higher: "keys (higher p v) = {u\<in>keys p. v \<prec>\<^sub>t u}"
by (rule set_eqI, simp only: in_keys_iff, simp add: lookup_higher)
lemma higher_higher: "higher (higher p u) v = higher p (ord_term_lin.max u v)"
by (rule poly_mapping_eqI, simp add: lookup_higher)
lemma lookup_lower: "lookup (lower p u) v = (if v \<prec>\<^sub>t u then lookup p v else 0)"
by (auto simp add: lower_def lookup_except)
lemma lookup_lower_when: "lookup (lower p u) v = (lookup p v when v \<prec>\<^sub>t u)"
by (auto simp add: lookup_lower when_def)
lemma lower_plus: "lower (p + q) v = lower p v + lower q v"
by (rule poly_mapping_eqI, simp add: lookup_add lookup_lower)
lemma lower_uminus [simp]: "lower (- p) v = - lower p v"
by (rule poly_mapping_eqI, simp add: lookup_lower)
lemma lower_minus: "lower (p - (q::_ \<Rightarrow>\<^sub>0 'b::ab_group_add)) v = lower p v - lower q v"
by (auto intro!: poly_mapping_eqI simp: lookup_minus lookup_lower)
lemma lower_zero [simp]: "lower 0 v = 0"
by (rule poly_mapping_eqI, simp add: lookup_lower)
lemma lower_eq_iff: "lower p v = lower q v \<longleftrightarrow> (\<forall>u. u \<prec>\<^sub>t v \<longrightarrow> lookup p u = lookup q u)" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
show ?R
proof (intro allI impI)
fix u
assume "u \<prec>\<^sub>t v"
moreover from \<open>?L\<close> have "lookup (lower p v) u = lookup (lower q v) u" by simp
ultimately show "lookup p u = lookup q u" by (simp add: lookup_lower)
qed
next
assume ?R
show ?L
proof (rule poly_mapping_eqI, simp add: lookup_lower, rule)
fix u
assume "u \<prec>\<^sub>t v"
with \<open>?R\<close> show "lookup p u = lookup q u" by simp
qed
qed
lemma lower_eq_zero_iff: "lower p v = 0 \<longleftrightarrow> (\<forall>u. u \<prec>\<^sub>t v \<longrightarrow> lookup p u = 0)"
proof -
have "lower p v = lower 0 v \<longleftrightarrow> (\<forall>u. u \<prec>\<^sub>t v \<longrightarrow> lookup p u = lookup 0 u)" by (rule lower_eq_iff)
thus ?thesis by simp
qed
lemma keys_lower: "keys (lower p v) = {u\<in>keys p. u \<prec>\<^sub>t v}"
by (rule set_eqI, simp only: in_keys_iff, simp add: lookup_lower)
lemma lower_lower: "lower (lower p u) v = lower p (ord_term_lin.min u v)"
by (rule poly_mapping_eqI, simp add: lookup_lower)
lemma lt_higher:
assumes "v \<prec>\<^sub>t lt p"
shows "lt (higher p v) = lt p"
proof (rule lt_eqI_keys, simp_all add: keys_higher, rule conjI, rule lt_in_keys, rule)
assume "p = 0"
hence "lt p = min_term" by (simp add: lt_def)
with min_term_min[of v] assms show False by simp
next
fix u
assume "u \<in> keys p \<and> v \<prec>\<^sub>t u"
hence "u \<in> keys p" ..
thus "u \<preceq>\<^sub>t lt p" by (rule lt_max_keys)
qed fact
lemma lc_higher:
assumes "v \<prec>\<^sub>t lt p"
shows "lc (higher p v) = lc p"
by (simp add: lc_def lt_higher assms lookup_higher)
lemma higher_eq_zero_iff': "higher p v = 0 \<longleftrightarrow> lt p \<preceq>\<^sub>t v"
by (simp add: higher_eq_zero_iff lt_le_iff)
lemma higher_id_iff: "higher p v = p \<longleftrightarrow> (p = 0 \<or> v \<prec>\<^sub>t tt p)" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
show ?R
proof (cases "p = 0")
case True
thus ?thesis ..
next
case False
show ?thesis
proof (rule disjI2, rule tt_gr)
fix u
assume "u \<in> keys p"
hence "lookup p u \<noteq> 0" by (simp add: in_keys_iff)
from \<open>?L\<close> have "lookup (higher p v) u = lookup p u" by simp
hence "lookup p u = (if v \<prec>\<^sub>t u then lookup p u else 0)" by (simp only: lookup_higher)
hence "\<not> v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = 0" by simp
with \<open>lookup p u \<noteq> 0\<close> show "v \<prec>\<^sub>t u" by auto
qed fact
qed
next
assume ?R
show ?L
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
with \<open>?R\<close> have "v \<prec>\<^sub>t tt p" by simp
show ?thesis
proof (rule poly_mapping_eqI, simp add: lookup_higher, intro impI)
fix u
assume "\<not> v \<prec>\<^sub>t u"
hence "u \<preceq>\<^sub>t v" by simp
from this \<open>v \<prec>\<^sub>t tt p\<close> have "u \<prec>\<^sub>t tt p" by simp
hence "\<not> tt p \<preceq>\<^sub>t u" by simp
with tt_min[of p u] show "lookup p u = 0" by blast
qed
qed
qed
lemma tt_lower:
assumes "tt p \<prec>\<^sub>t v"
shows "tt (lower p v) = tt p"
proof (cases "p = 0")
case True
thus ?thesis by simp
next
case False
show ?thesis
proof (rule tt_eqI, simp_all add: keys_lower, rule, rule tt_in_keys)
fix u
assume "u \<in> keys p \<and> u \<prec>\<^sub>t v"
hence "u \<in> keys p" ..
thus "tt p \<preceq>\<^sub>t u" by (rule tt_min_keys)
qed fact+
qed
lemma tc_lower:
assumes "tt p \<prec>\<^sub>t v"
shows "tc (lower p v) = tc p"
by (simp add: tc_def tt_lower assms lookup_lower)
lemma lt_lower: "lt (lower p v) \<preceq>\<^sub>t lt p"
proof (cases "lower p v = 0")
case True
thus ?thesis by (simp add: lt_def min_term_min)
next
case False
show ?thesis
proof (rule lt_le, simp add: lookup_lower, rule impI, rule ccontr)
fix u
assume "lookup p u \<noteq> 0"
hence "u \<preceq>\<^sub>t lt p" by (rule lt_max)
moreover assume "lt p \<prec>\<^sub>t u"
ultimately show False by simp
qed
qed
lemma lt_lower_less:
assumes "lower p v \<noteq> 0"
shows "lt (lower p v) \<prec>\<^sub>t v"
using assms
proof (rule lt_less)
fix u
assume "v \<preceq>\<^sub>t u"
thus "lookup (lower p v) u = 0" by (simp add: lookup_lower_when)
qed
lemma lt_lower_eq_iff: "lt (lower p v) = lt p \<longleftrightarrow> (lt p = min_term \<or> lt p \<prec>\<^sub>t v)" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
show ?R
proof (rule ccontr, simp, elim conjE)
assume "lt p \<noteq> min_term"
hence "min_term \<prec>\<^sub>t lt p" using min_term_min ord_term_lin.dual_order.not_eq_order_implies_strict
by blast
assume "\<not> lt p \<prec>\<^sub>t v"
hence "v \<preceq>\<^sub>t lt p" by simp
have "lt (lower p v) \<prec>\<^sub>t lt p"
proof (cases "lower p v = 0")
case True
thus ?thesis using \<open>min_term \<prec>\<^sub>t lt p\<close> by (simp add: lt_def)
next
case False
show ?thesis
proof (rule lt_less)
fix u
assume "lt p \<preceq>\<^sub>t u"
with \<open>v \<preceq>\<^sub>t lt p\<close> have "\<not> u \<prec>\<^sub>t v" by simp
thus "lookup (lower p v) u = 0" by (simp add: lookup_lower)
qed fact
qed
with \<open>?L\<close> show False by simp
qed
next
assume ?R
show ?L
proof (cases "lt p = min_term")
case True
hence "lt p \<preceq>\<^sub>t lt (lower p v)" by (simp add: min_term_min)
with lt_lower[of p v] show ?thesis by simp
next
case False
with \<open>?R\<close> have "lt p \<prec>\<^sub>t v" by simp
show ?thesis
proof (rule lt_eqI_keys, simp_all add: keys_lower, rule conjI, rule lt_in_keys, rule)
assume "p = 0"
hence "lt p = min_term" by (simp add: lt_def)
with False show False ..
next
fix u
assume "u \<in> keys p \<and> u \<prec>\<^sub>t v"
hence "u \<in> keys p" ..
thus "u \<preceq>\<^sub>t lt p" by (rule lt_max_keys)
qed fact
qed
qed
lemma tt_higher:
assumes "v \<prec>\<^sub>t lt p"
shows "tt p \<preceq>\<^sub>t tt (higher p v)"
proof (rule tt_ge_keys, simp add: keys_higher)
fix u
assume "u \<in> keys p \<and> v \<prec>\<^sub>t u"
hence "u \<in> keys p" ..
thus "tt p \<preceq>\<^sub>t u" by (rule tt_min_keys)
next
show "higher p v \<noteq> 0"
proof (simp add: higher_eq_zero_iff, intro exI conjI)
have "p \<noteq> 0"
proof
assume "p = 0"
hence "lt p \<preceq>\<^sub>t v" by (simp add: lt_def min_term_min)
with assms show False by simp
qed
thus "lookup p (lt p) \<noteq> 0"
using lt_in_keys by auto
qed fact
qed
lemma tt_higher_eq_iff:
"tt (higher p v) = tt p \<longleftrightarrow> ((lt p \<preceq>\<^sub>t v \<and> tt p = min_term) \<or> v \<prec>\<^sub>t tt p)" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
show ?R
proof (rule ccontr, simp, elim conjE)
assume a: "lt p \<preceq>\<^sub>t v \<longrightarrow> tt p \<noteq> min_term"
assume "\<not> v \<prec>\<^sub>t tt p"
hence "tt p \<preceq>\<^sub>t v" by simp
have "tt p \<prec>\<^sub>t tt (higher p v)"
proof (cases "higher p v = 0")
case True
with \<open>?L\<close> have "tt p = min_term" by (simp add: tt_def)
with a have "v \<prec>\<^sub>t lt p" by auto
have "lt p \<noteq> min_term"
proof
assume "lt p = min_term"
with \<open>v \<prec>\<^sub>t lt p\<close> show False using min_term_min[of v] by auto
qed
hence "p \<noteq> 0" by (auto simp add: lt_def)
from \<open>v \<prec>\<^sub>t lt p\<close> have "higher p v \<noteq> 0" by (simp add: higher_eq_zero_iff')
from this True show ?thesis ..
next
case False
show ?thesis
proof (rule tt_gr)
fix u
assume "u \<in> keys (higher p v)"
hence "v \<prec>\<^sub>t u" by (simp add: keys_higher)
with \<open>tt p \<preceq>\<^sub>t v\<close> show "tt p \<prec>\<^sub>t u" by simp
qed fact
qed
with \<open>?L\<close> show False by simp
qed
next
assume ?R
show ?L
proof (cases "lt p \<preceq>\<^sub>t v \<and> tt p = min_term")
case True
hence "lt p \<preceq>\<^sub>t v" and "tt p = min_term" by simp_all
from \<open>lt p \<preceq>\<^sub>t v\<close> have "higher p v = 0" by (simp add: higher_eq_zero_iff')
with \<open>tt p = min_term\<close> show ?thesis by (simp add: tt_def)
next
case False
with \<open>?R\<close> have "v \<prec>\<^sub>t tt p" by simp
show ?thesis
proof (rule tt_eqI, simp_all add: keys_higher, rule conjI, rule tt_in_keys, rule)
assume "p = 0"
hence "tt p = min_term" by (simp add: tt_def)
with \<open>v \<prec>\<^sub>t tt p\<close> min_term_min[of v] show False by simp
next
fix u
assume "u \<in> keys p \<and> v \<prec>\<^sub>t u"
hence "u \<in> keys p" ..
thus "tt p \<preceq>\<^sub>t u" by (rule tt_min_keys)
qed fact
qed
qed
lemma lower_eq_zero_iff': "lower p v = 0 \<longleftrightarrow> (p = 0 \<or> v \<preceq>\<^sub>t tt p)"
by (auto simp add: lower_eq_zero_iff tt_ge_iff)
lemma lower_id_iff: "lower p v = p \<longleftrightarrow> (p = 0 \<or> lt p \<prec>\<^sub>t v)" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
show ?R
proof (cases "p = 0")
case True
thus ?thesis ..
next
case False
show ?thesis
proof (rule disjI2, rule lt_less)
fix u
assume "v \<preceq>\<^sub>t u"
from \<open>?L\<close> have "lookup (lower p v) u = lookup p u" by simp
hence "lookup p u = (if u \<prec>\<^sub>t v then lookup p u else 0)" by (simp only: lookup_lower)
hence "v \<preceq>\<^sub>t u \<Longrightarrow> lookup p u = 0" by (meson ord_term_lin.leD)
with \<open>v \<preceq>\<^sub>t u\<close> show "lookup p u = 0" by simp
qed fact
qed
next
assume ?R
show ?L
proof (cases "p = 0", simp)
case False
with \<open>?R\<close> have "lt p \<prec>\<^sub>t v" by simp
show ?thesis
proof (rule poly_mapping_eqI, simp add: lookup_lower, intro impI)
fix u
assume "\<not> u \<prec>\<^sub>t v"
hence "v \<preceq>\<^sub>t u" by simp
with \<open>lt p \<prec>\<^sub>t v\<close> have "lt p \<prec>\<^sub>t u" by simp
hence "\<not> u \<preceq>\<^sub>t lt p" by simp
with lt_max[of p u] show "lookup p u = 0" by blast
qed
qed
qed
lemma lower_higher_commute: "higher (lower p s) t = lower (higher p t) s"
by (rule poly_mapping_eqI, simp add: lookup_higher lookup_lower)
lemma lt_lower_higher:
assumes "v \<prec>\<^sub>t lt (lower p u)"
shows "lt (lower (higher p v) u) = lt (lower p u)"
by (simp add: lower_higher_commute[symmetric] lt_higher[OF assms])
lemma lc_lower_higher:
assumes "v \<prec>\<^sub>t lt (lower p u)"
shows "lc (lower (higher p v) u) = lc (lower p u)"
using assms by (simp add: lc_def lt_lower_higher lookup_lower lookup_higher)
lemma trailing_monomial_higher:
assumes "p \<noteq> 0"
shows "p = (higher p (tt p)) + monomial (tc p) (tt p)"
proof (rule poly_mapping_eqI, simp only: lookup_add)
fix v
show "lookup p v = lookup (higher p (tt p)) v + lookup (monomial (tc p) (tt p)) v"
proof (cases "tt p \<preceq>\<^sub>t v")
case True
show ?thesis
proof (cases "v = tt p")
assume "v = tt p"
hence "\<not> tt p \<prec>\<^sub>t v" by simp
hence "lookup (higher p (tt p)) v = 0" by (simp add: lookup_higher)
moreover from \<open>v = tt p\<close> have "lookup (monomial (tc p) (tt p)) v = tc p" by (simp add: lookup_single)
moreover from \<open>v = tt p\<close> have "lookup p v = tc p" by (simp add: tc_def)
ultimately show ?thesis by simp
next
assume "v \<noteq> tt p"
from this True have "tt p \<prec>\<^sub>t v" by simp
hence "lookup (higher p (tt p)) v = lookup p v" by (simp add: lookup_higher)
moreover from \<open>v \<noteq> tt p\<close> have "lookup (monomial (tc p) (tt p)) v = 0" by (simp add: lookup_single)
ultimately show ?thesis by simp
qed
next
case False
hence "v \<prec>\<^sub>t tt p" by simp
hence "tt p \<noteq> v" by simp
from False have "\<not> tt p \<prec>\<^sub>t v" by simp
have "lookup p v = 0"
proof (rule ccontr)
assume "lookup p v \<noteq> 0"
from tt_min[OF this] False show False by simp
qed
moreover from \<open>tt p \<noteq> v\<close> have "lookup (monomial (tc p) (tt p)) v = 0" by (simp add: lookup_single)
moreover from \<open>\<not> tt p \<prec>\<^sub>t v\<close> have "lookup (higher p (tt p)) v = 0" by (simp add: lookup_higher)
ultimately show ?thesis by simp
qed
qed
lemma higher_lower_decomp: "higher p v + monomial (lookup p v) v + lower p v = p"
proof (rule poly_mapping_eqI)
fix u
show "lookup (higher p v + monomial (lookup p v) v + lower p v) u = lookup p u"
proof (rule ord_term_lin.linorder_cases)
assume "u \<prec>\<^sub>t v"
thus ?thesis by (simp add: lookup_add lookup_higher_when lookup_single lookup_lower_when)
next
assume "u = v"
thus ?thesis by (simp add: lookup_add lookup_higher_when lookup_single lookup_lower_when)
next
assume "v \<prec>\<^sub>t u"
thus ?thesis by (simp add: lookup_add lookup_higher_when lookup_single lookup_lower_when)
qed
qed
subsection \<open>@{const tail}\<close>
lemma lookup_tail: "lookup (tail p) v = (if v \<prec>\<^sub>t lt p then lookup p v else 0)"
by (simp add: lookup_lower tail_def)
lemma lookup_tail_when: "lookup (tail p) v = (lookup p v when v \<prec>\<^sub>t lt p)"
by (simp add: lookup_lower_when tail_def)
lemma lookup_tail_2: "lookup (tail p) v = (if v = lt p then 0 else lookup p v)"
proof (rule ord_term_lin.linorder_cases[of v "lt p"])
assume "v \<prec>\<^sub>t lt p"
hence "v \<noteq> lt p" by simp
from this \<open>v \<prec>\<^sub>t lt p\<close> lookup_tail[of p v] show ?thesis by simp
next
assume "v = lt p"
hence "\<not> v \<prec>\<^sub>t lt p" by simp
from \<open>v = lt p\<close> this lookup_tail[of p v] show ?thesis by simp
next
assume "lt p \<prec>\<^sub>t v"
hence "\<not> v \<preceq>\<^sub>t lt p" by simp
hence cp: "lookup p v = 0"
using lt_max by blast
from \<open>\<not> v \<preceq>\<^sub>t lt p\<close> have "\<not> v = lt p" and "\<not> v \<prec>\<^sub>t lt p" by simp_all
thus ?thesis using cp lookup_tail[of p v] by simp
qed
lemma leading_monomial_tail: "p = monomial (lc p) (lt p) + tail p" for p::"_ \<Rightarrow>\<^sub>0 'b::comm_monoid_add"
proof (rule poly_mapping_eqI)
fix v
have "lookup p v = lookup (monomial (lc p) (lt p)) v + lookup (tail p) v"
proof (cases "v \<preceq>\<^sub>t lt p")
case True
show ?thesis
proof (cases "v = lt p")
assume "v = lt p"
hence "\<not> v \<prec>\<^sub>t lt p" by simp
hence c3: "lookup (tail p) v = 0" unfolding lookup_tail[of p v] by simp
from \<open>v = lt p\<close> have c2: "lookup (monomial (lc p) (lt p)) v = lc p" by simp
from \<open>v = lt p\<close> have c1: "lookup p v = lc p" by (simp add: lc_def)
from c1 c2 c3 show ?thesis by simp
next
assume "v \<noteq> lt p"
from this True have "v \<prec>\<^sub>t lt p" by simp
hence c2: "lookup (tail p) v = lookup p v" unfolding lookup_tail[of p v] by simp
from \<open>v \<noteq> lt p\<close> have c1: "lookup (monomial (lc p) (lt p)) v = 0" by (simp add: lookup_single)
from c1 c2 show ?thesis by simp
qed
next
case False
hence "lt p \<prec>\<^sub>t v" by simp
hence "lt p \<noteq> v" by simp
from False have "\<not> v \<prec>\<^sub>t lt p" by simp
have c1: "lookup p v = 0"
proof (rule ccontr)
assume "lookup p v \<noteq> 0"
from lt_max[OF this] False show False by simp
qed
from \<open>lt p \<noteq> v\<close> have c2: "lookup (monomial (lc p) (lt p)) v = 0" by (simp add: lookup_single)
from \<open>\<not> v \<prec>\<^sub>t lt p\<close> lookup_tail[of p v] have c3: "lookup (tail p) v = 0" by simp
from c1 c2 c3 show ?thesis by simp
qed
thus "lookup p v = lookup (monomial (lc p) (lt p) + tail p) v" by (simp add: lookup_add)
qed
lemma tail_alt: "tail p = except p {lt p}"
by (rule poly_mapping_eqI, simp add: lookup_tail_2 lookup_except)
corollary tail_alt_2: "tail p = p - monomial (lc p) (lt p)"
proof -
have "p = monomial (lc p) (lt p) + tail p" by (fact leading_monomial_tail)
also have "... = tail p + monomial (lc p) (lt p)" by (simp only: add.commute)
finally have "p - monomial (lc p) (lt p) = (tail p + monomial (lc p) (lt p)) - monomial (lc p) (lt p)" by simp
thus ?thesis by simp
qed
lemma tail_zero [simp]: "tail 0 = 0"
by (simp only: tail_alt except_zero)
lemma lt_tail:
assumes "tail p \<noteq> 0"
shows "lt (tail p) \<prec>\<^sub>t lt p"
proof (intro lt_less)
fix u
assume "lt p \<preceq>\<^sub>t u"
hence "\<not> u \<prec>\<^sub>t lt p" by simp
thus "lookup (tail p) u = 0" unfolding lookup_tail[of p u] by simp
qed fact
lemma keys_tail: "keys (tail p) = keys p - {lt p}"
by (simp add: tail_alt keys_except)
lemma tail_monomial: "tail (monomial c v) = 0"
by (metis (no_types, lifting) lookup_tail_2 lookup_single_not_eq lt_less lt_monomial
ord_term_lin.dual_order.strict_implies_not_eq single_zero tail_zero)
lemma (in ordered_term) mult_scalar_tail_rec_left:
"p \<odot> q = monom_mult (punit.lc p) (punit.lt p) q + (punit.tail p) \<odot> q"
unfolding punit.lc_def punit.tail_alt by (fact mult_scalar_rec_left)
lemma mult_scalar_tail_rec_right: "p \<odot> q = p \<odot> monomial (lc q) (lt q) + p \<odot> tail q"
unfolding tail_alt lc_def by (rule mult_scalar_rec_right)
lemma lt_tail_max:
assumes "tail p \<noteq> 0" and "v \<in> keys p" and "v \<prec>\<^sub>t lt p"
shows "v \<preceq>\<^sub>t lt (tail p)"
proof (rule lt_max_keys, simp add: keys_tail assms(2))
from assms(3) show "v \<noteq> lt p" by auto
qed
lemma keys_tail_less_lt:
assumes "v \<in> keys (tail p)"
shows "v \<prec>\<^sub>t lt p"
using assms by (meson in_keys_iff lookup_tail)
lemma tt_tail:
assumes "tail p \<noteq> 0"
shows "tt (tail p) = tt p"
proof (rule tt_eqI, simp_all add: keys_tail)
from assms have "p \<noteq> 0" using tail_zero by auto
show "tt p \<in> keys p \<and> tt p \<noteq> lt p"
proof (rule conjI, rule tt_in_keys, fact)
have "tt p \<prec>\<^sub>t lt p"
by (metis assms lower_eq_zero_iff' tail_def ord_term_lin.le_less_linear)
thus "tt p \<noteq> lt p" by simp
qed
next
fix u
assume "u \<in> keys p \<and> u \<noteq> lt p"
hence "u \<in> keys p" ..
thus "tt p \<preceq>\<^sub>t u" by (rule tt_min_keys)
qed
lemma tc_tail:
assumes "tail p \<noteq> 0"
shows "tc (tail p) = tc p"
proof (simp add: tc_def tt_tail[OF assms] lookup_tail_2, rule)
assume "tt p = lt p"
moreover have "tt p \<prec>\<^sub>t lt p"
by (metis assms lower_eq_zero_iff' tail_def ord_term_lin.le_less_linear)
ultimately show "lookup p (lt p) = 0" by simp
qed
lemma tt_tail_min:
assumes "s \<in> keys p"
shows "tt (tail p) \<preceq>\<^sub>t s"
proof (cases "tail p = 0")
case True
hence "tt (tail p) = min_term" by (simp add: tt_def)
thus ?thesis by (simp add: min_term_min)
next
case False
from assms show ?thesis by (simp add: tt_tail[OF False], rule tt_min_keys)
qed
lemma tail_monom_mult:
"tail (monom_mult c t p) = monom_mult (c::'b::semiring_no_zero_divisors) t (tail p)"
proof (cases "p = 0")
case True
hence "tail p = 0" and "monom_mult c t p = 0" by simp_all
thus ?thesis by simp
next
case False
show ?thesis
proof (cases "c = 0")
case True
hence "monom_mult c t p = 0" and "monom_mult c t (tail p) = 0" by simp_all
thus ?thesis by simp
next
case False
let ?a = "monom_mult c t p"
let ?b = "monom_mult c t (tail p)"
from \<open>p \<noteq> 0\<close> False have "?a \<noteq> 0" by (simp add: monom_mult_eq_zero_iff)
from False \<open>p \<noteq> 0\<close> have lt_a: "lt ?a = t \<oplus> lt p" by (rule lt_monom_mult)
show ?thesis
proof (rule poly_mapping_eqI, simp add: lookup_tail lt_a, intro conjI impI)
fix u
assume "u \<prec>\<^sub>t t \<oplus> lt p"
show "lookup (monom_mult c t p) u = lookup (monom_mult c t (tail p)) u"
proof (cases "t adds\<^sub>p u")
case True
then obtain v where "u = t \<oplus> v" by (rule adds_ppE)
from \<open>u \<prec>\<^sub>t t \<oplus> lt p\<close> have "v \<prec>\<^sub>t lt p" unfolding \<open>u = t \<oplus> v\<close> by (rule ord_term_strict_canc)
hence "lookup p v = lookup (tail p) v" by (simp add: lookup_tail)
thus ?thesis by (simp add: \<open>u = t \<oplus> v\<close> lookup_monom_mult_plus)
next
case False
hence "lookup ?a u = 0" by (simp add: lookup_monom_mult)
moreover have "lookup ?b u = 0"
proof (rule ccontr, simp only: in_keys_iff[symmetric] keys_monom_mult[OF \<open>c \<noteq> 0\<close>])
assume "u \<in> (\<oplus>) t ` keys (tail p)"
then obtain v where "u = t \<oplus> v" by auto
hence "t adds\<^sub>p u" by (simp add: term_simps)
with False show False ..
qed
ultimately show ?thesis by simp
qed
next
fix u
assume "\<not> u \<prec>\<^sub>t t \<oplus> lt p"
hence "t \<oplus> lt p \<preceq>\<^sub>t u" by simp
show "lookup (monom_mult c t (tail p)) u = 0"
proof (rule ccontr, simp only: in_keys_iff[symmetric] keys_monom_mult[OF False])
assume "u \<in> (\<oplus>) t ` keys (tail p)"
then obtain v where "v \<in> keys (tail p)" and "u = t \<oplus> v" by auto
from \<open>t \<oplus> lt p \<preceq>\<^sub>t u\<close> have "lt p \<preceq>\<^sub>t v" unfolding \<open>u = t \<oplus> v\<close> by (rule ord_term_canc)
from \<open>v \<in> keys (tail p)\<close> have "v \<in> keys p" and "v \<noteq> lt p" by (simp_all add: keys_tail)
from \<open>v \<in> keys p\<close> have "v \<preceq>\<^sub>t lt p" by (rule lt_max_keys)
with \<open>lt p \<preceq>\<^sub>t v\<close> have "v = lt p " by simp
with \<open>v \<noteq> lt p\<close> show False ..
qed
qed
qed
qed
lemma keys_plus_eq_lt_tt_D:
assumes "keys (p + q) = {lt p, tt q}" and "lt q \<prec>\<^sub>t lt p" and "tt q \<prec>\<^sub>t tt (p::_ \<Rightarrow>\<^sub>0 'b::comm_monoid_add)"
shows "tail p + higher q (tt q) = 0"
proof -
note assms(3)
also have "... \<preceq>\<^sub>t lt p" by (rule lt_ge_tt)
finally have "tt q \<prec>\<^sub>t lt p" .
hence "lt p \<noteq> tt q" by simp
have "q \<noteq> 0"
proof
assume "q = 0"
hence "tt q = min_term" by (simp add: tt_def)
with \<open>q = 0\<close> assms(1) have "keys p = {lt p, min_term}" by simp
hence "min_term \<in> keys p" by simp
hence "tt p \<preceq>\<^sub>t tt q" unfolding \<open>tt q = min_term\<close> by (rule tt_min_keys)
with assms(3) show False by simp
qed
hence "tc q \<noteq> 0" by (rule tc_not_0)
have "p = monomial (lc p) (lt p) + tail p" by (rule leading_monomial_tail)
moreover from \<open>q \<noteq> 0\<close> have "q = higher q (tt q) + monomial (tc q) (tt q)" by (rule trailing_monomial_higher)
ultimately have pq: "p + q = (monomial (lc p) (lt p) + monomial (tc q) (tt q)) + (tail p + higher q (tt q))"
(is "_ = (?m1 + ?m2) + ?b") by (simp add: algebra_simps)
have keys_m1: "keys ?m1 = {lt p}"
proof (rule keys_of_monomial, rule lc_not_0, rule)
assume "p = 0"
with assms(2) have "lt q \<prec>\<^sub>t min_term" by (simp add: lt_def)
with min_term_min[of "lt q"] show False by simp
qed
moreover from \<open>tc q \<noteq> 0\<close> have keys_m2: "keys ?m2 = {tt q}" by (rule keys_of_monomial)
ultimately have keys_m1_m2: "keys (?m1 + ?m2) = {lt p, tt q}"
using \<open>lt p \<noteq> tt q\<close> keys_plus_eqI[of ?m1 ?m2] by auto
show ?thesis
proof (rule ccontr)
assume "?b \<noteq> 0"
hence "keys ?b \<noteq> {}" by simp
then obtain t where "t \<in> keys ?b" by blast
hence t_in: "t \<in> keys (tail p) \<union> keys (higher q (tt q))"
using Poly_Mapping.keys_add[of "tail p" "higher q (tt q)"] by blast
hence "t \<noteq> lt p"
proof (rule, simp add: keys_tail, simp add: keys_higher, elim conjE)
assume "t \<in> keys q"
hence "t \<preceq>\<^sub>t lt q" by (rule lt_max_keys)
from this assms(2) show ?thesis by simp
qed
moreover from t_in have "t \<noteq> tt q"
proof (rule, simp add: keys_tail, elim conjE)
assume "t \<in> keys p"
hence "tt p \<preceq>\<^sub>t t" by (rule tt_min_keys)
with assms(3) show ?thesis by simp
next
assume "t \<in> keys (higher q (tt q))"
thus ?thesis by (auto simp only: keys_higher)
qed
ultimately have "t \<notin> keys (?m1 + ?m2)" by (simp add: keys_m1_m2)
moreover from in_keys_plusI2[OF \<open>t \<in> keys ?b\<close> this] have "t \<in> keys (?m1 + ?m2)"
by (simp only: keys_m1_m2 pq[symmetric] assms(1))
ultimately show False ..
qed
qed
subsection \<open>Order Relation on Polynomials\<close>
definition ord_strict_p :: "('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool" (infixl "\<prec>\<^sub>p" 50) where
"p \<prec>\<^sub>p q \<longleftrightarrow> (\<exists>v. lookup p v = 0 \<and> lookup q v \<noteq> 0 \<and> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup q u))"
definition ord_p :: "('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool" (infixl "\<preceq>\<^sub>p" 50) where
"ord_p p q \<equiv> (p \<prec>\<^sub>p q \<or> p = q)"
lemma ord_strict_pI:
assumes "lookup p v = 0" and "lookup q v \<noteq> 0" and "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u"
shows "p \<prec>\<^sub>p q"
unfolding ord_strict_p_def using assms by blast
lemma ord_strict_pE:
assumes "p \<prec>\<^sub>p q"
obtains v where "lookup p v = 0" and "lookup q v \<noteq> 0" and "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u"
using assms unfolding ord_strict_p_def by blast
lemma not_ord_pI:
assumes "lookup p v \<noteq> lookup q v" and "lookup p v \<noteq> 0" and "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u"
shows "\<not> p \<preceq>\<^sub>p q"
proof
assume "p \<preceq>\<^sub>p q"
hence "p \<prec>\<^sub>p q \<or> p = q" by (simp only: ord_p_def)
thus False
proof
assume "p \<prec>\<^sub>p q"
then obtain v' where 1: "lookup p v' = 0" and 2: "lookup q v' \<noteq> 0"
and 3: "\<And>u. v' \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u" by (rule ord_strict_pE, blast)
from 1 2 have "lookup p v' \<noteq> lookup q v'" by simp
hence "\<not> v \<prec>\<^sub>t v'" using assms(3) by blast
hence "v' \<prec>\<^sub>t v \<or> v' = v" by auto
thus ?thesis
proof
assume "v' \<prec>\<^sub>t v"
hence "lookup p v = lookup q v" by (rule 3)
with assms(1) show ?thesis ..
next
assume "v' = v"
with assms(2) 1 show ?thesis by auto
qed
next
assume "p = q"
hence "lookup p v = lookup q v" by simp
with assms(1) show ?thesis ..
qed
qed
corollary not_ord_strict_pI:
assumes "lookup p v \<noteq> lookup q v" and "lookup p v \<noteq> 0" and "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u"
shows "\<not> p \<prec>\<^sub>p q"
proof -
from assms have "\<not> p \<preceq>\<^sub>p q" by (rule not_ord_pI)
thus ?thesis by (simp add: ord_p_def)
qed
lemma ord_strict_higher: "p \<prec>\<^sub>p q \<longleftrightarrow> (\<exists>v. lookup p v = 0 \<and> lookup q v \<noteq> 0 \<and> higher p v = higher q v)"
unfolding ord_strict_p_def higher_eq_iff ..
lemma ord_strict_p_asymmetric:
assumes "p \<prec>\<^sub>p q"
shows "\<not> q \<prec>\<^sub>p p"
using assms unfolding ord_strict_p_def
proof
fix v1::'t
assume "lookup p v1 = 0 \<and> lookup q v1 \<noteq> 0 \<and> (\<forall>u. v1 \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup q u)"
hence "lookup p v1 = 0" and "lookup q v1 \<noteq> 0" and v1: "\<forall>u. v1 \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup q u"
by auto
show "\<not> (\<exists>v. lookup q v = 0 \<and> lookup p v \<noteq> 0 \<and> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup q u = lookup p u))"
proof (intro notI, erule exE)
fix v2::'t
assume "lookup q v2 = 0 \<and> lookup p v2 \<noteq> 0 \<and> (\<forall>u. v2 \<prec>\<^sub>t u \<longrightarrow> lookup q u = lookup p u)"
hence "lookup q v2 = 0" and "lookup p v2 \<noteq> 0" and v2: "\<forall>u. v2 \<prec>\<^sub>t u \<longrightarrow> lookup q u = lookup p u"
by auto
show False
proof (rule ord_term_lin.linorder_cases)
assume "v1 \<prec>\<^sub>t v2"
from v1[rule_format, OF this] \<open>lookup q v2 = 0\<close> \<open>lookup p v2 \<noteq> 0\<close> show ?thesis by simp
next
assume "v1 = v2"
thus ?thesis using \<open>lookup p v1 = 0\<close> \<open>lookup p v2 \<noteq> 0\<close> by simp
next
assume "v2 \<prec>\<^sub>t v1"
from v2[rule_format, OF this] \<open>lookup p v1 = 0\<close> \<open>lookup q v1 \<noteq> 0\<close> show ?thesis by simp
qed
qed
qed
lemma ord_strict_p_irreflexive: "\<not> p \<prec>\<^sub>p p"
unfolding ord_strict_p_def
proof (intro notI, erule exE)
fix v::'t
assume "lookup p v = 0 \<and> lookup p v \<noteq> 0 \<and> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup p u)"
hence "lookup p v = 0" and "lookup p v \<noteq> 0" by auto
thus False by simp
qed
lemma ord_strict_p_transitive:
assumes "a \<prec>\<^sub>p b" and "b \<prec>\<^sub>p c"
shows "a \<prec>\<^sub>p c"
proof -
from \<open>a \<prec>\<^sub>p b\<close> obtain v1 where "lookup a v1 = 0"
and "lookup b v1 \<noteq> 0"
and v1[rule_format]: "(\<forall>u. v1 \<prec>\<^sub>t u \<longrightarrow> lookup a u = lookup b u)"
unfolding ord_strict_p_def by auto
from \<open>b \<prec>\<^sub>p c\<close> obtain v2 where "lookup b v2 = 0"
and "lookup c v2 \<noteq> 0"
and v2[rule_format]: "(\<forall>u. v2 \<prec>\<^sub>t u \<longrightarrow> lookup b u = lookup c u)"
unfolding ord_strict_p_def by auto
show "a \<prec>\<^sub>p c"
proof (rule ord_term_lin.linorder_cases)
assume "v1 \<prec>\<^sub>t v2"
show ?thesis unfolding ord_strict_p_def
proof
show "lookup a v2 = 0 \<and> lookup c v2 \<noteq> 0 \<and> (\<forall>u. v2 \<prec>\<^sub>t u \<longrightarrow> lookup a u = lookup c u)"
proof (intro conjI allI impI)
from \<open>lookup b v2 = 0\<close> v1[OF \<open>v1 \<prec>\<^sub>t v2\<close>] show "lookup a v2 = 0" by simp
next
from \<open>lookup c v2 \<noteq> 0\<close> show "lookup c v2 \<noteq> 0" .
next
fix u
assume "v2 \<prec>\<^sub>t u"
from ord_term_lin.less_trans[OF \<open>v1 \<prec>\<^sub>t v2\<close> this] have "v1 \<prec>\<^sub>t u" .
from v2[OF \<open>v2 \<prec>\<^sub>t u\<close>] v1[OF this] show "lookup a u = lookup c u" by simp
qed
qed
next
assume "v2 \<prec>\<^sub>t v1"
show ?thesis unfolding ord_strict_p_def
proof
show "lookup a v1 = 0 \<and> lookup c v1 \<noteq> 0 \<and> (\<forall>u. v1 \<prec>\<^sub>t u \<longrightarrow> lookup a u = lookup c u)"
proof (intro conjI allI impI)
from \<open>lookup a v1 = 0\<close> show "lookup a v1 = 0" .
next
from \<open>lookup b v1 \<noteq> 0\<close> v2[OF \<open>v2 \<prec>\<^sub>t v1\<close>] show "lookup c v1 \<noteq> 0" by simp
next
fix u
assume "v1 \<prec>\<^sub>t u"
from ord_term_lin.less_trans[OF \<open>v2 \<prec>\<^sub>t v1\<close> this] have "v2 \<prec>\<^sub>t u" .
from v1[OF \<open>v1 \<prec>\<^sub>t u\<close>] v2[OF this] show "lookup a u = lookup c u" by simp
qed
qed
next
assume "v1 = v2"
thus ?thesis using \<open>lookup b v1 \<noteq> 0\<close> \<open>lookup b v2 = 0\<close> by simp
qed
qed
sublocale order ord_p ord_strict_p
proof (intro order_strictI)
fix p q :: "'t \<Rightarrow>\<^sub>0 'b"
show "(p \<preceq>\<^sub>p q) = (p \<prec>\<^sub>p q \<or> p = q)" unfolding ord_p_def ..
next
fix p q :: "'t \<Rightarrow>\<^sub>0 'b"
assume "p \<prec>\<^sub>p q"
thus "\<not> q \<prec>\<^sub>p p" by (rule ord_strict_p_asymmetric)
next
fix p::"'t \<Rightarrow>\<^sub>0 'b"
show "\<not> p \<prec>\<^sub>p p" by (fact ord_strict_p_irreflexive)
next
fix a b c :: "'t \<Rightarrow>\<^sub>0 'b"
assume "a \<prec>\<^sub>p b" and "b \<prec>\<^sub>p c"
thus "a \<prec>\<^sub>p c" by (rule ord_strict_p_transitive)
qed
lemma ord_p_zero_min: "0 \<preceq>\<^sub>p p"
unfolding ord_p_def ord_strict_p_def
proof (cases "p = 0")
case True
thus "(\<exists>v. lookup 0 v = 0 \<and> lookup p v \<noteq> 0 \<and> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup 0 u = lookup p u)) \<or> 0 = p"
by auto
next
case False
show "(\<exists>v. lookup 0 v = 0 \<and> lookup p v \<noteq> 0 \<and> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup 0 u = lookup p u)) \<or> 0 = p"
proof
show "(\<exists>v. lookup 0 v = 0 \<and> lookup p v \<noteq> 0 \<and> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup 0 u = lookup p u))"
proof
show "lookup 0 (lt p) = 0 \<and> lookup p (lt p) \<noteq> 0 \<and> (\<forall>u. (lt p) \<prec>\<^sub>t u \<longrightarrow> lookup 0 u = lookup p u)"
proof (intro conjI allI impI)
show "lookup 0 (lt p) = 0" by (transfer, simp)
next
from lc_not_0[OF False] show "lookup p (lt p) \<noteq> 0" unfolding lc_def .
next
fix u
assume "lt p \<prec>\<^sub>t u"
hence "\<not> u \<preceq>\<^sub>t lt p" by simp
hence "lookup p u = 0" using lt_max[of p u] by metis
thus "lookup 0 u = lookup p u" by simp
qed
qed
qed
qed
lemma lt_ord_p:
assumes "lt p \<prec>\<^sub>t lt q"
shows "p \<prec>\<^sub>p q"
proof -
have "q \<noteq> 0"
proof
assume "q = 0"
with assms have "lt p \<prec>\<^sub>t min_term" by (simp add: lt_def)
with min_term_min[of "lt p"] show False by simp
qed
show ?thesis unfolding ord_strict_p_def
proof (intro exI conjI allI impI)
show "lookup p (lt q) = 0"
proof (rule ccontr)
assume "lookup p (lt q) \<noteq> 0"
from lt_max[OF this] \<open>lt p \<prec>\<^sub>t lt q\<close> show False by simp
qed
next
from lc_not_0[OF \<open>q \<noteq> 0\<close>] show "lookup q (lt q) \<noteq> 0" unfolding lc_def .
next
fix u
assume "lt q \<prec>\<^sub>t u"
hence "lt p \<prec>\<^sub>t u" using \<open>lt p \<prec>\<^sub>t lt q\<close> by simp
have c1: "lookup q u = 0"
proof (rule ccontr)
assume "lookup q u \<noteq> 0"
from lt_max[OF this] \<open>lt q \<prec>\<^sub>t u\<close> show False by simp
qed
have c2: "lookup p u = 0"
proof (rule ccontr)
assume "lookup p u \<noteq> 0"
from lt_max[OF this] \<open>lt p \<prec>\<^sub>t u\<close> show False by simp
qed
from c1 c2 show "lookup p u = lookup q u" by simp
qed
qed
lemma ord_p_lt:
assumes "p \<preceq>\<^sub>p q"
shows "lt p \<preceq>\<^sub>t lt q"
proof (rule ccontr)
assume "\<not> lt p \<preceq>\<^sub>t lt q"
hence "lt q \<prec>\<^sub>t lt p" by simp
from lt_ord_p[OF this] \<open>p \<preceq>\<^sub>p q\<close> show False by simp
qed
lemma ord_p_tail:
assumes "p \<noteq> 0" and "lt p = lt q" and "p \<prec>\<^sub>p q"
shows "tail p \<prec>\<^sub>p tail q"
using assms unfolding ord_strict_p_def
proof -
assume "p \<noteq> 0" and "lt p = lt q"
and "\<exists>v. lookup p v = 0 \<and> lookup q v \<noteq> 0 \<and> (\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup q u)"
then obtain v where "lookup p v = 0"
and "lookup q v \<noteq> 0"
and a: "\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup q u" by auto
from lt_max[OF \<open>lookup q v \<noteq> 0\<close>] \<open>lt p = lt q\<close> have "v \<prec>\<^sub>t lt p \<or> v = lt p" by auto
hence "v \<prec>\<^sub>t lt p"
proof
assume "v \<prec>\<^sub>t lt p"
thus ?thesis .
next
assume "v = lt p"
thus ?thesis using lc_not_0[OF \<open>p \<noteq> 0\<close>] \<open>lookup p v = 0\<close> unfolding lc_def by auto
qed
have pt: "lookup (tail p) v = lookup p v" using lookup_tail[of p v] \<open>v \<prec>\<^sub>t lt p\<close> by simp
have "q \<noteq> 0"
proof
assume "q = 0"
hence "p \<prec>\<^sub>p 0" using \<open>p \<prec>\<^sub>p q\<close> by simp
hence "\<not> 0 \<preceq>\<^sub>p p" by auto
thus False using ord_p_zero_min[of p] by simp
qed
have qt: "lookup (tail q) v = lookup q v"
using lookup_tail[of q v] \<open>v \<prec>\<^sub>t lt p\<close> \<open>lt p = lt q\<close> by simp
show "\<exists>w. lookup (tail p) w = 0 \<and> lookup (tail q) w \<noteq> 0 \<and>
(\<forall>u. w \<prec>\<^sub>t u \<longrightarrow> lookup (tail p) u = lookup (tail q) u)"
proof (intro exI conjI allI impI)
from pt \<open>lookup p v = 0\<close> show "lookup (tail p) v = 0" by simp
next
from qt \<open>lookup q v \<noteq> 0\<close> show "lookup (tail q) v \<noteq> 0" by simp
next
fix u
assume "v \<prec>\<^sub>t u"
from a[rule_format, OF \<open>v \<prec>\<^sub>t u\<close>] lookup_tail[of p u] lookup_tail[of q u]
\<open>lt p = lt q\<close> show "lookup (tail p) u = lookup (tail q) u" by simp
qed
qed
lemma tail_ord_p:
assumes "p \<noteq> 0"
shows "tail p \<prec>\<^sub>p p"
proof (cases "tail p = 0")
case True
with ord_p_zero_min[of p] \<open>p \<noteq> 0\<close> show ?thesis by simp
next
case False
from lt_tail[OF False] show ?thesis by (rule lt_ord_p)
qed
lemma higher_lookup_eq_zero:
assumes pt: "lookup p v = 0" and hp: "higher p v = 0" and le: "q \<preceq>\<^sub>p p"
shows "(lookup q v = 0) \<and> (higher q v) = 0"
using le unfolding ord_p_def
proof
assume "q \<prec>\<^sub>p p"
thus ?thesis unfolding ord_strict_p_def
proof
fix w
assume "lookup q w = 0 \<and> lookup p w \<noteq> 0 \<and> (\<forall>u. w \<prec>\<^sub>t u \<longrightarrow> lookup q u = lookup p u)"
hence qs: "lookup q w = 0" and ps: "lookup p w \<noteq> 0" and u: "\<forall>u. w \<prec>\<^sub>t u \<longrightarrow> lookup q u = lookup p u"
by auto
from hp have pu: "\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = 0" by (simp only: higher_eq_zero_iff)
from pu[rule_format, of w] ps have "\<not> v \<prec>\<^sub>t w" by auto
hence "w \<preceq>\<^sub>t v" by simp
hence "w \<prec>\<^sub>t v \<or> w = v" by auto
hence st: "w \<prec>\<^sub>t v"
proof (rule disjE, simp_all)
assume "w = v"
from this pt ps show False by simp
qed
show ?thesis
proof
from u[rule_format, OF st] pt show "lookup q v = 0" by simp
next
have "\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup q u = 0"
proof (intro allI, intro impI)
fix u
assume "v \<prec>\<^sub>t u"
from this st have "w \<prec>\<^sub>t u" by simp
from u[rule_format, OF this] pu[rule_format, OF \<open>v \<prec>\<^sub>t u\<close>] show "lookup q u = 0" by simp
qed
thus "higher q v = 0" by (simp only: higher_eq_zero_iff)
qed
qed
next
assume "q = p"
thus ?thesis using assms by simp
qed
lemma ord_strict_p_recI:
assumes "lt p = lt q" and "lc p = lc q" and tail: "tail p \<prec>\<^sub>p tail q"
shows "p \<prec>\<^sub>p q"
proof -
from tail obtain v where pt: "lookup (tail p) v = 0"
and qt: "lookup (tail q) v \<noteq> 0"
and a: "\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup (tail p) u = lookup (tail q) u"
unfolding ord_strict_p_def by auto
from qt lookup_zero[of v] have "tail q \<noteq> 0" by auto
from lt_max[OF qt] lt_tail[OF this] have "v \<prec>\<^sub>t lt q" by simp
hence "v \<prec>\<^sub>t lt p" using \<open>lt p = lt q\<close> by simp
show ?thesis unfolding ord_strict_p_def
proof (rule exI[of _ v], intro conjI allI impI)
from lookup_tail[of p v] \<open>v \<prec>\<^sub>t lt p\<close> pt show "lookup p v = 0" by simp
next
from lookup_tail[of q v] \<open>v \<prec>\<^sub>t lt q\<close> qt show "lookup q v \<noteq> 0" by simp
next
fix u
assume "v \<prec>\<^sub>t u"
from this a have s: "lookup (tail p) u = lookup (tail q) u" by simp
show "lookup p u = lookup q u"
proof (cases "u = lt p")
case True
from True \<open>lc p = lc q\<close> \<open>lt p = lt q\<close> show ?thesis unfolding lc_def by simp
next
case False
from False s lookup_tail_2[of p u] lookup_tail_2[of q u] \<open>lt p = lt q\<close> show ?thesis by simp
qed
qed
qed
lemma ord_strict_p_recE1:
assumes "p \<prec>\<^sub>p q"
shows "q \<noteq> 0"
proof
assume "q = 0"
from this assms ord_p_zero_min[of p] show False by simp
qed
lemma ord_strict_p_recE2:
assumes "p \<noteq> 0" and "p \<prec>\<^sub>p q" and "lt p = lt q"
shows "lc p = lc q"
proof -
from \<open>p \<prec>\<^sub>p q\<close> obtain v where pt: "lookup p v = 0"
and qt: "lookup q v \<noteq> 0"
and a: "\<forall>u. v \<prec>\<^sub>t u \<longrightarrow> lookup p u = lookup q u"
unfolding ord_strict_p_def by auto
show ?thesis
proof (cases "v \<prec>\<^sub>t lt p")
case True
from this a have "lookup p (lt p) = lookup q (lt p)" by simp
thus ?thesis using \<open>lt p = lt q\<close> unfolding lc_def by simp
next
case False
from this lt_max[OF qt] \<open>lt p = lt q\<close> have "v = lt p" by simp
from this lc_not_0[OF \<open>p \<noteq> 0\<close>] pt show ?thesis unfolding lc_def by auto
qed
qed
lemma ord_strict_p_rec [code]:
"p \<prec>\<^sub>p q =
(q \<noteq> 0 \<and>
(p = 0 \<or>
(let v1 = lt p; v2 = lt q in
(v1 \<prec>\<^sub>t v2 \<or> (v1 = v2 \<and> lookup p v1 = lookup q v2 \<and> lower p v1 \<prec>\<^sub>p lower q v2))
)
)
)"
(is "?L = ?R")
proof
assume ?L
show ?R
proof (intro conjI, rule ord_strict_p_recE1, fact)
have "((lt p = lt q \<and> lc p = lc q \<and> tail p \<prec>\<^sub>p tail q) \<or> lt p \<prec>\<^sub>t lt q) \<or> p = 0"
proof (intro disjCI)
assume "p \<noteq> 0" and nl: "\<not> lt p \<prec>\<^sub>t lt q"
from \<open>?L\<close> have "p \<preceq>\<^sub>p q" by simp
from ord_p_lt[OF this] nl have "lt p = lt q" by simp
show "lt p = lt q \<and> lc p = lc q \<and> tail p \<prec>\<^sub>p tail q"
by (intro conjI, fact, rule ord_strict_p_recE2, fact+, rule ord_p_tail, fact+)
qed
thus "p = 0 \<or>
(let v1 = lt p; v2 = lt q in
(v1 \<prec>\<^sub>t v2 \<or> v1 = v2 \<and> lookup p v1 = lookup q v2 \<and> lower p v1 \<prec>\<^sub>p lower q v2)
)"
unfolding lc_def tail_def by auto
qed
next
assume ?R
hence "q \<noteq> 0"
and dis: "p = 0 \<or>
(let v1 = lt p; v2 = lt q in
(v1 \<prec>\<^sub>t v2 \<or> v1 = v2 \<and> lookup p v1 = lookup q v2 \<and> lower p v1 \<prec>\<^sub>p lower q v2)
)"
by simp_all
show ?L
proof (cases "p = 0")
assume "p = 0"
hence "p \<preceq>\<^sub>p q" using ord_p_zero_min[of q] by simp
thus ?thesis using \<open>p = 0\<close> \<open>q \<noteq> 0\<close> by simp
next
assume "p \<noteq> 0"
hence "let v1 = lt p; v2 = lt q in
(v1 \<prec>\<^sub>t v2 \<or> v1 = v2 \<and> lookup p v1 = lookup q v2 \<and> lower p v1 \<prec>\<^sub>p lower q v2)"
using dis by simp
hence "lt p \<prec>\<^sub>t lt q \<or> (lt p = lt q \<and> lc p = lc q \<and> tail p \<prec>\<^sub>p tail q)"
unfolding lc_def tail_def by (simp add: Let_def)
thus ?thesis
proof
assume "lt p \<prec>\<^sub>t lt q"
from lt_ord_p[OF this] show ?thesis .
next
assume "lt p = lt q \<and> lc p = lc q \<and> tail p \<prec>\<^sub>p tail q"
hence "lt p = lt q" and "lc p = lc q" and "tail p \<prec>\<^sub>p tail q" by simp_all
thus ?thesis by (rule ord_strict_p_recI)
qed
qed
qed
lemma ord_strict_p_monomial_iff: "p \<prec>\<^sub>p monomial c v \<longleftrightarrow> (c \<noteq> 0 \<and> (p = 0 \<or> lt p \<prec>\<^sub>t v))"
proof -
from ord_p_zero_min[of "tail p"] have *: "\<not> tail p \<prec>\<^sub>p 0" by auto
show ?thesis
by (simp add: ord_strict_p_rec[of p] Let_def tail_def[symmetric] lc_def[symmetric]
monomial_0_iff tail_monomial *, simp add: lt_monomial cong: conj_cong)
qed
corollary ord_strict_p_monomial_plus:
assumes "p \<prec>\<^sub>p monomial c v" and "q \<prec>\<^sub>p monomial c v"
shows "p + q \<prec>\<^sub>p monomial c v"
proof -
from assms(1) have "c \<noteq> 0" and "p = 0 \<or> lt p \<prec>\<^sub>t v" by (simp_all add: ord_strict_p_monomial_iff)
from this(2) show ?thesis
proof
assume "p = 0"
with assms(2) show ?thesis by simp
next
assume "lt p \<prec>\<^sub>t v"
from assms(2) have "q = 0 \<or> lt q \<prec>\<^sub>t v" by (simp add: ord_strict_p_monomial_iff)
thus ?thesis
proof
assume "q = 0"
with assms(1) show ?thesis by simp
next
assume "lt q \<prec>\<^sub>t v"
with \<open>lt p \<prec>\<^sub>t v\<close> have "lt (p + q) \<prec>\<^sub>t v"
using lt_plus_le_max ord_term_lin.dual_order.strict_trans2 ord_term_lin.max_less_iff_conj
by blast
with \<open>c \<noteq> 0\<close> show ?thesis by (simp add: ord_strict_p_monomial_iff)
qed
qed
qed
lemma ord_strict_p_monom_mult:
assumes "p \<prec>\<^sub>p q" and "c \<noteq> (0::'b::semiring_no_zero_divisors)"
shows "monom_mult c t p \<prec>\<^sub>p monom_mult c t q"
proof -
from assms(1) obtain v where 1: "lookup p v = 0" and 2: "lookup q v \<noteq> 0"
and 3: "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u" unfolding ord_strict_p_def by auto
show ?thesis unfolding ord_strict_p_def
proof (intro exI conjI allI impI)
from 1 show "lookup (monom_mult c t p) (t \<oplus> v) = 0" by (simp add: lookup_monom_mult_plus)
next
from 2 assms(2) show "lookup (monom_mult c t q) (t \<oplus> v) \<noteq> 0" by (simp add: lookup_monom_mult_plus)
next
fix u
assume "t \<oplus> v \<prec>\<^sub>t u"
show "lookup (monom_mult c t p) u = lookup (monom_mult c t q) u"
proof (cases "t adds\<^sub>p u")
case True
then obtain w where u: "u = t \<oplus> w" ..
from \<open>t \<oplus> v \<prec>\<^sub>t u\<close> have "v \<prec>\<^sub>t w" unfolding u by (rule ord_term_strict_canc)
hence "lookup p w = lookup q w" by (rule 3)
thus ?thesis by (simp add: u lookup_monom_mult_plus)
next
case False
thus ?thesis by (simp add: lookup_monom_mult)
qed
qed
qed
lemma ord_strict_p_plus:
assumes "p \<prec>\<^sub>p q" and "keys r \<inter> keys q = {}"
shows "p + r \<prec>\<^sub>p q + r"
proof -
from assms(1) obtain v where 1: "lookup p v = 0" and 2: "lookup q v \<noteq> 0"
and 3: "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u" unfolding ord_strict_p_def by auto
have eq: "lookup r v = 0"
by (meson "2" assms(2) disjoint_iff_not_equal in_keys_iff)
show ?thesis unfolding ord_strict_p_def
proof (intro exI conjI allI impI, simp_all add: lookup_add)
from 1 show "lookup p v + lookup r v = 0" by (simp add: eq)
next
from 2 show "lookup q v + lookup r v \<noteq> 0" by (simp add: eq)
next
fix u
assume "v \<prec>\<^sub>t u"
hence "lookup p u = lookup q u" by (rule 3)
thus "lookup p u + lookup r u = lookup q u + lookup r u" by simp
qed
qed
lemma poly_mapping_tail_induct [case_names 0 tail]:
assumes "P 0" and "\<And>p. p \<noteq> 0 \<Longrightarrow> P (tail p) \<Longrightarrow> P p"
shows "P p"
proof (induct "card (keys p)" arbitrary: p)
case 0
with finite_keys[of p] have "keys p = {}" by simp
hence "p = 0" by simp
from \<open>P 0\<close> show ?case unfolding \<open>p = 0\<close> .
next
case ind: (Suc n)
from ind(2) have "keys p \<noteq> {}" by auto
hence "p \<noteq> 0" by simp
thus ?case
proof (rule assms(2))
show "P (tail p)"
proof (rule ind(1))
from \<open>p \<noteq> 0\<close> have "lt p \<in> keys p" by (rule lt_in_keys)
hence "card (keys (tail p)) = card (keys p) - 1" by (simp add: keys_tail)
also have "... = n" unfolding ind(2)[symmetric] by simp
finally show "n = card (keys (tail p))" by simp
qed
qed
qed
lemma poly_mapping_neqE:
assumes "p \<noteq> q"
obtains v where "v \<in> keys p \<union> keys q" and "lookup p v \<noteq> lookup q v"
and "\<And>u. v \<prec>\<^sub>t u \<Longrightarrow> lookup p u = lookup q u"
proof -
let ?A = "{v. lookup p v \<noteq> lookup q v}"
define v where "v = ord_term_lin.Max ?A"
have "?A \<subseteq> keys p \<union> keys q"
using UnI2 in_keys_iff by fastforce
also have "finite ..." by (rule finite_UnI) (fact finite_keys)+
finally(finite_subset) have fin: "finite ?A" .
moreover have "?A \<noteq> {}"
proof
assume "?A = {}"
hence "p = q"
using poly_mapping_eqI by fastforce
with assms show False ..
qed
ultimately have "v \<in> ?A" unfolding v_def by (rule ord_term_lin.Max_in)
show ?thesis
proof
from \<open>?A \<subseteq> keys p \<union> keys q\<close> \<open>v \<in> ?A\<close> show "v \<in> keys p \<union> keys q" ..
next
from \<open>v \<in> ?A\<close> show "lookup p v \<noteq> lookup q v" by simp
next
fix u
assume "v \<prec>\<^sub>t u"
show "lookup p u = lookup q u"
proof (rule ccontr)
assume "lookup p u \<noteq> lookup q u"
hence "u \<in> ?A" by simp
with fin have "u \<preceq>\<^sub>t v" unfolding v_def by (rule ord_term_lin.Max_ge)
with \<open>v \<prec>\<^sub>t u\<close> show False by simp
qed
qed
qed
subsection \<open>Monomials\<close>
lemma keys_monomial:
assumes "is_monomial p"
shows "keys p = {lt p}"
using assms by (metis is_monomial_monomial lt_monomial keys_of_monomial)
lemma monomial_eq_itself:
assumes "is_monomial p"
shows "monomial (lc p) (lt p) = p"
proof -
from assms have "p \<noteq> 0" by (rule monomial_not_0)
hence "lc p \<noteq> 0" by (rule lc_not_0)
hence keys1: "keys (monomial (lc p) (lt p)) = {lt p}" by (rule keys_of_monomial)
show ?thesis
by (rule poly_mapping_keys_eqI, simp only: keys_monomial[OF assms] keys1,
simp only: keys1 lookup_single Poly_Mapping.when_def, auto simp add: lc_def)
qed
lemma lt_eq_min_term_monomial:
assumes "lt p = min_term"
shows "monomial (lc p) min_term = p"
proof (rule poly_mapping_eqI)
fix v
from min_term_min[of v] have "v = min_term \<or> min_term \<prec>\<^sub>t v" by auto
thus "lookup (monomial (lc p) min_term) v = lookup p v"
proof
assume "v = min_term"
thus ?thesis by (simp add: lookup_single lc_def assms)
next
assume "min_term \<prec>\<^sub>t v"
moreover have "v \<notin> keys p"
proof
assume "v \<in> keys p"
hence "v \<preceq>\<^sub>t lt p" by (rule lt_max_keys)
with \<open>min_term \<prec>\<^sub>t v\<close> show False by (simp add: assms)
qed
ultimately show ?thesis by (simp add: lookup_single in_keys_iff)
qed
qed
lemma is_monomial_monomial_ordered:
assumes "is_monomial p"
obtains c v where "c \<noteq> 0" and "lc p = c" and "lt p = v" and "p = monomial c v"
proof -
from assms obtain c v where "c \<noteq> 0" and p_eq: "p = monomial c v" by (rule is_monomial_monomial)
note this(1)
moreover have "lc p = c" unfolding p_eq by (rule lc_monomial)
moreover from \<open>c \<noteq> 0\<close> have "lt p = v" unfolding p_eq by (rule lt_monomial)
ultimately show ?thesis using p_eq ..
qed
lemma monomial_plus_not_0:
assumes "c \<noteq> 0" and "lt p \<prec>\<^sub>t v"
shows "monomial c v + p \<noteq> 0"
proof
assume "monomial c v + p = 0"
hence "0 = lookup (monomial c v + p) v" by simp
also have "... = c + lookup p v" by (simp add: lookup_add)
also have "... = c"
proof -
from assms(2) have "\<not> v \<preceq>\<^sub>t lt p" by simp
with lt_max[of p v] have "lookup p v = 0" by blast
thus ?thesis by simp
qed
finally show False using \<open>c \<noteq> 0\<close> by simp
qed
lemma lt_monomial_plus:
assumes "c \<noteq> (0::'b::comm_monoid_add)" and "lt p \<prec>\<^sub>t v"
shows "lt (monomial c v + p) = v"
proof -
have eq: "lt (monomial c v) = v" by (simp only: lt_monomial[OF \<open>c \<noteq> 0\<close>])
moreover have "lt (p + monomial c v) = lt (monomial c v)" by (rule lt_plus_eqI, simp only: eq, fact)
ultimately show ?thesis by (simp add: add.commute)
qed
lemma lc_monomial_plus:
assumes "c \<noteq> (0::'b::comm_monoid_add)" and "lt p \<prec>\<^sub>t v"
shows "lc (monomial c v + p) = c"
proof -
from assms(2) have "\<not> v \<preceq>\<^sub>t lt p" by simp
with lt_max[of p v] have "lookup p v = 0" by blast
thus ?thesis by (simp add: lc_def lt_monomial_plus[OF assms] lookup_add)
qed
lemma tt_monomial_plus:
assumes "p \<noteq> (0::_ \<Rightarrow>\<^sub>0 'b::comm_monoid_add)" and "lt p \<prec>\<^sub>t v"
shows "tt (monomial c v + p) = tt p"
proof (cases "c = 0")
case True
thus ?thesis by (simp add: monomial_0I)
next
case False
have eq: "tt (monomial c v) = v" by (simp only: tt_monomial[OF \<open>c \<noteq> 0\<close>])
moreover have "tt (p + monomial c v) = tt p"
proof (rule tt_plus_eqI, fact, simp only: eq)
from lt_ge_tt[of p] assms(2) show "tt p \<prec>\<^sub>t v" by simp
qed
ultimately show ?thesis by (simp add: ac_simps)
qed
lemma tc_monomial_plus:
assumes "p \<noteq> (0::_ \<Rightarrow>\<^sub>0 'b::comm_monoid_add)" and "lt p \<prec>\<^sub>t v"
shows "tc (monomial c v + p) = tc p"
proof (simp add: tc_def tt_monomial_plus[OF assms] lookup_add lookup_single Poly_Mapping.when_def,
rule impI)
assume "v = tt p"
with assms(2) have "lt p \<prec>\<^sub>t tt p" by simp
with lt_ge_tt[of p] show "c + lookup p (tt p) = lookup p (tt p)" by simp
qed
lemma tail_monomial_plus:
assumes "c \<noteq> (0::'b::comm_monoid_add)" and "lt p \<prec>\<^sub>t v"
shows "tail (monomial c v + p) = p" (is "tail ?q = _")
proof -
from assms have "lt ?q = v" by (rule lt_monomial_plus)
moreover have "lower (monomial c v) v = 0"
by (simp add: lower_eq_zero_iff', rule disjI2, simp add: tt_monomial[OF \<open>c \<noteq> 0\<close>])
ultimately show ?thesis by (simp add: tail_def lower_plus lower_id_iff, intro disjI2 assms(2))
qed
subsection \<open>Lists of Keys\<close>
text \<open>In algorithms one very often needs to compute the sorted list of all terms appearing
in a list of polynomials.\<close>
definition pps_to_list :: "'t set \<Rightarrow> 't list" where
"pps_to_list S = rev (ord_term_lin.sorted_list_of_set S)"
definition keys_to_list :: "('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 't list"
where "keys_to_list p = pps_to_list (keys p)"
definition Keys_to_list :: "('t \<Rightarrow>\<^sub>0 'b::zero) list \<Rightarrow> 't list"
where "Keys_to_list ps = fold (\<lambda>p ts. merge_wrt (\<succ>\<^sub>t) (keys_to_list p) ts) ps []"
text \<open>Function @{const pps_to_list} turns finite sets of terms into sorted lists, where the
lists are sorted descending (i.\,e. greater elements come before smaller ones).\<close>
lemma distinct_pps_to_list: "distinct (pps_to_list S)"
unfolding pps_to_list_def distinct_rev by (rule ord_term_lin.distinct_sorted_list_of_set)
lemma set_pps_to_list:
assumes "finite S"
shows "set (pps_to_list S) = S"
unfolding pps_to_list_def set_rev using assms by simp
lemma length_pps_to_list: "length (pps_to_list S) = card S"
proof (cases "finite S")
case True
from distinct_card[OF distinct_pps_to_list] have "length (pps_to_list S) = card (set (pps_to_list S))"
by simp
also from True have "... = card S" by (simp only: set_pps_to_list)
finally show ?thesis .
next
case False
thus ?thesis by (simp add: pps_to_list_def)
qed
lemma pps_to_list_sorted_wrt: "sorted_wrt (\<succ>\<^sub>t) (pps_to_list S)"
proof -
have "sorted_wrt (\<succeq>\<^sub>t) (pps_to_list S)"
proof -
have tr: "transp (\<preceq>\<^sub>t)" using transp_def by fastforce
have *: "(\<lambda>x y. y \<succeq>\<^sub>t x) = (\<preceq>\<^sub>t)" by simp
show ?thesis
by (simp only: * pps_to_list_def sorted_wrt_rev ord_term_lin.sorted_sorted_wrt[symmetric],
rule ord_term_lin.sorted_sorted_list_of_set)
qed
with distinct_pps_to_list have "sorted_wrt (\<lambda>x y. x \<succeq>\<^sub>t y \<and> x \<noteq> y) (pps_to_list S)"
by (rule distinct_sorted_wrt_imp_sorted_wrt_strict)
moreover have "(\<succ>\<^sub>t) = (\<lambda>x y. x \<succeq>\<^sub>t y \<and> x \<noteq> y)"
using ord_term_lin.dual_order.order_iff_strict by auto
ultimately show ?thesis by simp
qed
lemma pps_to_list_nth_leI:
assumes "j \<le> i" and "i < card S"
shows "(pps_to_list S) ! i \<preceq>\<^sub>t (pps_to_list S) ! j"
proof (cases "j = i")
case True
show ?thesis by (simp add: True)
next
case False
with assms(1) have "j < i" by simp
let ?ts = "pps_to_list S"
from pps_to_list_sorted_wrt \<open>j < i\<close> have "(\<prec>\<^sub>t)\<inverse>\<inverse> (?ts ! j) (?ts ! i)"
proof (rule sorted_wrt_nth_less)
from assms(2) show "i < length ?ts" by (simp only: length_pps_to_list)
qed
thus ?thesis by simp
qed
lemma pps_to_list_nth_lessI:
assumes "j < i" and "i < card S"
shows "(pps_to_list S) ! i \<prec>\<^sub>t (pps_to_list S) ! j"
proof -
let ?ts = "pps_to_list S"
from assms(1) have "j \<le> i" and "i \<noteq> j" by simp_all
with assms(2) have "i < length ?ts" and "j < length ?ts" by (simp_all only: length_pps_to_list)
show ?thesis
proof (rule ord_term_lin.neq_le_trans)
from \<open>i \<noteq> j\<close> show "?ts ! i \<noteq> ?ts ! j"
by (simp add: nth_eq_iff_index_eq[OF distinct_pps_to_list \<open>i < length ?ts\<close> \<open>j < length ?ts\<close>])
next
from \<open>j \<le> i\<close> assms(2) show "?ts ! i \<preceq>\<^sub>t ?ts ! j" by (rule pps_to_list_nth_leI)
qed
qed
lemma pps_to_list_nth_leD:
assumes "(pps_to_list S) ! i \<preceq>\<^sub>t (pps_to_list S) ! j" and "j < card S"
shows "j \<le> i"
proof (rule ccontr)
assume "\<not> j \<le> i"
hence "i < j" by simp
from this \<open>j < card S\<close> have "(pps_to_list S) ! j \<prec>\<^sub>t (pps_to_list S) ! i" by (rule pps_to_list_nth_lessI)
with assms(1) show False by simp
qed
lemma pps_to_list_nth_lessD:
assumes "(pps_to_list S) ! i \<prec>\<^sub>t (pps_to_list S) ! j" and "j < card S"
shows "j < i"
proof (rule ccontr)
assume "\<not> j < i"
hence "i \<le> j" by simp
from this \<open>j < card S\<close> have "(pps_to_list S) ! j \<preceq>\<^sub>t (pps_to_list S) ! i" by (rule pps_to_list_nth_leI)
with assms(1) show False by simp
qed
lemma set_keys_to_list: "set (keys_to_list p) = keys p"
by (simp add: keys_to_list_def set_pps_to_list)
lemma length_keys_to_list: "length (keys_to_list p) = card (keys p)"
by (simp only: keys_to_list_def length_pps_to_list)
lemma keys_to_list_zero [simp]: "keys_to_list 0 = []"
by (simp add: keys_to_list_def pps_to_list_def)
lemma Keys_to_list_Nil [simp]: "Keys_to_list [] = []"
by (simp add: Keys_to_list_def)
lemma set_Keys_to_list: "set (Keys_to_list ps) = Keys (set ps)"
proof -
have "set (Keys_to_list ps) = (\<Union>p\<in>set ps. set (keys_to_list p)) \<union> set []"
unfolding Keys_to_list_def by (rule set_fold, simp only: set_merge_wrt)
also have "... = Keys (set ps)" by (simp add: Keys_def set_keys_to_list)
finally show ?thesis .
qed
lemma Keys_to_list_sorted_wrt_aux:
assumes "sorted_wrt (\<succ>\<^sub>t) ts"
shows "sorted_wrt (\<succ>\<^sub>t) (fold (\<lambda>p ts. merge_wrt (\<succ>\<^sub>t) (keys_to_list p) ts) ps ts)"
using assms
proof (induct ps arbitrary: ts)
case Nil
thus ?case by simp
next
case (Cons p ps)
show ?case
proof (simp only: fold.simps o_def, rule Cons(1), rule sorted_merge_wrt)
show "transp (\<succ>\<^sub>t)" unfolding transp_def by fastforce
next
fix x y :: 't
assume "x \<noteq> y"
thus "x \<succ>\<^sub>t y \<or> y \<succ>\<^sub>t x" by auto
next
show "sorted_wrt (\<succ>\<^sub>t) (keys_to_list p)" unfolding keys_to_list_def
by (fact pps_to_list_sorted_wrt)
qed fact
qed
corollary Keys_to_list_sorted_wrt: "sorted_wrt (\<succ>\<^sub>t) (Keys_to_list ps)"
unfolding Keys_to_list_def
proof (rule Keys_to_list_sorted_wrt_aux)
show "sorted_wrt (\<succ>\<^sub>t) []" by simp
qed
corollary distinct_Keys_to_list: "distinct (Keys_to_list ps)"
proof (rule distinct_sorted_wrt_irrefl)
show "irreflp (\<succ>\<^sub>t)" by (simp add: irreflp_def)
next
show "transp (\<succ>\<^sub>t)" unfolding transp_def by fastforce
next
show "sorted_wrt (\<succ>\<^sub>t) (Keys_to_list ps)" by (fact Keys_to_list_sorted_wrt)
qed
lemma length_Keys_to_list: "length (Keys_to_list ps) = card (Keys (set ps))"
proof -
from distinct_Keys_to_list have "card (set (Keys_to_list ps)) = length (Keys_to_list ps)"
by (rule distinct_card)
thus ?thesis by (simp only: set_Keys_to_list)
qed
lemma Keys_to_list_eq_pps_to_list: "Keys_to_list ps = pps_to_list (Keys (set ps))"
using _ Keys_to_list_sorted_wrt distinct_Keys_to_list pps_to_list_sorted_wrt distinct_pps_to_list
proof (rule sorted_wrt_distinct_set_unique)
show "antisymp (\<succ>\<^sub>t)" unfolding antisymp_def by fastforce
next
from finite_set have fin: "finite (Keys (set ps))" by (rule finite_Keys)
show "set (Keys_to_list ps) = set (pps_to_list (Keys (set ps)))"
by (simp add: set_Keys_to_list set_pps_to_list[OF fin])
qed
subsection \<open>Multiplication\<close>
lemma in_keys_mult_scalar_le:
assumes "v \<in> keys (p \<odot> q)"
shows "v \<preceq>\<^sub>t punit.lt p \<oplus> lt q"
proof -
from assms obtain t u where "t \<in> keys p" and "u \<in> keys q" and "v = t \<oplus> u"
by (rule in_keys_mult_scalarE)
from \<open>t \<in> keys p\<close> have "t \<preceq> punit.lt p" by (rule punit.lt_max_keys)
from \<open>u \<in> keys q\<close> have "u \<preceq>\<^sub>t lt q" by (rule lt_max_keys)
hence "v \<preceq>\<^sub>t t \<oplus> lt q" unfolding \<open>v = t \<oplus> u\<close> by (rule splus_mono)
also from \<open>t \<preceq> punit.lt p\<close> have "... \<preceq>\<^sub>t punit.lt p \<oplus> lt q" by (rule splus_mono_left)
finally show ?thesis .
qed
lemma in_keys_mult_scalar_ge:
assumes "v \<in> keys (p \<odot> q)"
shows "punit.tt p \<oplus> tt q \<preceq>\<^sub>t v"
proof -
from assms obtain t u where "t \<in> keys p" and "u \<in> keys q" and "v = t \<oplus> u"
by (rule in_keys_mult_scalarE)
from \<open>t \<in> keys p\<close> have "punit.tt p \<preceq> t" by (rule punit.tt_min_keys)
from \<open>u \<in> keys q\<close> have "tt q \<preceq>\<^sub>t u" by (rule tt_min_keys)
hence "punit.tt p \<oplus> tt q \<preceq>\<^sub>t punit.tt p \<oplus> u" by (rule splus_mono)
also from \<open>punit.tt p \<preceq> t\<close> have "... \<preceq>\<^sub>t v" unfolding \<open>v = t \<oplus> u\<close> by (rule splus_mono_left)
finally show ?thesis .
qed
lemma (in ordered_term) lookup_mult_scalar_lt_lt:
"lookup (p \<odot> q) (punit.lt p \<oplus> lt q) = punit.lc p * lc q"
proof (induct p rule: punit.poly_mapping_tail_induct)
case 0
show ?case by simp
next
case step: (tail p)
from step(1) have "punit.lc p \<noteq> 0" by (rule punit.lc_not_0)
let ?t = "punit.lt p \<oplus> lt q"
show ?case
proof (cases "is_monomial p")
case True
then obtain c t where "c \<noteq> 0" and "punit.lt p = t" and "punit.lc p = c" and p_eq: "p = monomial c t"
by (rule punit.is_monomial_monomial_ordered)
hence "p \<odot> q = monom_mult (punit.lc p) (punit.lt p) q" by (simp add: mult_scalar_monomial)
thus ?thesis by (simp add: lookup_monom_mult_plus lc_def)
next
case False
have "punit.lt (punit.tail p) \<noteq> punit.lt p"
proof (simp add: punit.tail_def punit.lt_lower_eq_iff, rule)
assume "punit.lt p = 0"
have "keys p \<subseteq> {punit.lt p}"
proof (rule, simp)
fix s
assume "s \<in> keys p"
hence "s \<preceq> punit.lt p" by (rule punit.lt_max_keys)
moreover have "punit.lt p \<preceq> s" unfolding \<open>punit.lt p = 0\<close> by (rule zero_min)
ultimately show "s = punit.lt p" by simp
qed
hence "card (keys p) = 0 \<or> card (keys p) = 1" using subset_singletonD by fastforce
thus False
proof
assume "card (keys p) = 0"
hence "p = 0" by (meson card_0_eq keys_eq_empty finite_keys)
with step(1) show False ..
next
assume "card (keys p) = 1"
with False show False unfolding is_monomial_def ..
qed
qed
with punit.lt_lower[of p "punit.lt p"] have "punit.lt (punit.tail p) \<prec> punit.lt p"
by (simp add: punit.tail_def)
have eq: "lookup ((punit.tail p) \<odot> q) ?t = 0"
proof (rule ccontr)
assume "lookup ((punit.tail p) \<odot> q) ?t \<noteq> 0"
hence "?t \<preceq>\<^sub>t punit.lt (punit.tail p) \<oplus> lt q"
by (meson in_keys_mult_scalar_le lookup_not_eq_zero_eq_in_keys)
hence "punit.lt p \<preceq> punit.lt (punit.tail p)" by (rule ord_term_canc_left)
also have "... \<prec> punit.lt p" by fact
finally show False ..
qed
from step(2) have "lookup (monom_mult (punit.lc p) (punit.lt p) q) ?t = punit.lc p * lc q"
by (simp only: lookup_monom_mult_plus lc_def)
thus ?thesis by (simp add: mult_scalar_tail_rec_left[of p q] lookup_add eq)
qed
qed
lemma lookup_mult_scalar_tt_tt: "lookup (p \<odot> q) (punit.tt p \<oplus> tt q) = punit.tc p * tc q"
proof (induct p rule: punit.poly_mapping_tail_induct)
case 0
show ?case by simp
next
case step: (tail p)
from step(1) have "punit.lc p \<noteq> 0" by (rule punit.lc_not_0)
let ?t = "punit.tt p \<oplus> tt q"
show ?case
proof (cases "is_monomial p")
case True
then obtain c t where "c \<noteq> 0" and "punit.lt p = t" and "punit.lc p = c" and p_eq: "p = monomial c t"
by (rule punit.is_monomial_monomial_ordered)
from \<open>c \<noteq> 0\<close> have "punit.tt p = t" and "punit.tc p = c" by (simp_all add: p_eq punit.tt_monomial)
with p_eq have "p \<odot> q = monom_mult (punit.tc p) (punit.tt p) q" by (simp add: mult_scalar_monomial)
thus ?thesis by (simp add: lookup_monom_mult_plus tc_def)
next
case False
from step(1) have "keys p \<noteq> {}" by simp
with finite_keys have "card (keys p) \<noteq> 0" by auto
with False have "2 \<le> card (keys p)" unfolding is_monomial_def by linarith
then obtain s t where "s \<in> keys p" and "t \<in> keys p" and "s \<prec> t"
by (metis (mono_tags, lifting) card.empty card.infinite card_insert_disjoint card_mono empty_iff
finite.emptyI insertCI lessI not_less numeral_2_eq_2 ordered_powerprod_lin.infinite_growing
ordered_powerprod_lin.neqE preorder_class.less_le_trans subsetI)
from this(1) this(3) have "punit.tt p \<prec> t" by (rule punit.tt_less)
also from \<open>t \<in> keys p\<close> have "t \<preceq> punit.lt p" by (rule punit.lt_max_keys)
finally have "punit.tt p \<prec> punit.lt p" .
hence tt_tail: "punit.tt (punit.tail p) = punit.tt p" and tc_tail: "punit.tc (punit.tail p) = punit.tc p"
unfolding punit.tail_def by (rule punit.tt_lower, rule punit.tc_lower)
have eq: "lookup (monom_mult (punit.lc p) (punit.lt p) q) ?t = 0"
proof (rule ccontr)
assume "lookup (monom_mult (punit.lc p) (punit.lt p) q) ?t \<noteq> 0"
hence "punit.lt p \<oplus> tt q \<preceq>\<^sub>t ?t"
by (meson in_keys_iff in_keys_monom_mult_ge)
hence "punit.lt p \<preceq> punit.tt p" by (rule ord_term_canc_left)
also have "... \<prec> punit.lt p" by fact
finally show False ..
qed
from step(2) have "lookup (punit.tail p \<odot> q) ?t = punit.tc p * tc q" by (simp only: tt_tail tc_tail)
thus ?thesis by (simp add: mult_scalar_tail_rec_left[of p q] lookup_add eq)
qed
qed
lemma lt_mult_scalar:
assumes "p \<noteq> 0" and "q \<noteq> (0::'t \<Rightarrow>\<^sub>0 'b::semiring_no_zero_divisors)"
shows "lt (p \<odot> q) = punit.lt p \<oplus> lt q"
proof (rule lt_eqI_keys, simp only: in_keys_iff lookup_mult_scalar_lt_lt)
from assms(1) have "punit.lc p \<noteq> 0" by (rule punit.lc_not_0)
moreover from assms(2) have "lc q \<noteq> 0" by (rule lc_not_0)
ultimately show "punit.lc p * lc q \<noteq> 0" by simp
qed (rule in_keys_mult_scalar_le)
lemma tt_mult_scalar:
assumes "p \<noteq> 0" and "q \<noteq> (0::'t \<Rightarrow>\<^sub>0 'b::semiring_no_zero_divisors)"
shows "tt (p \<odot> q) = punit.tt p \<oplus> tt q"
proof (rule tt_eqI, simp only: in_keys_iff lookup_mult_scalar_tt_tt)
from assms(1) have "punit.tc p \<noteq> 0" by (rule punit.tc_not_0)
moreover from assms(2) have "tc q \<noteq> 0" by (rule tc_not_0)
ultimately show "punit.tc p * tc q \<noteq> 0" by simp
qed (rule in_keys_mult_scalar_ge)
lemma lc_mult_scalar: "lc (p \<odot> q) = punit.lc p * lc (q::'t \<Rightarrow>\<^sub>0 'b::semiring_no_zero_divisors)"
proof (cases "p = 0")
case True
thus ?thesis by (simp add: lc_def)
next
case False
show ?thesis
proof (cases "q = 0")
case True
thus ?thesis by (simp add: lc_def)
next
case False
with \<open>p \<noteq> 0\<close> show ?thesis by (simp add: lc_def lt_mult_scalar lookup_mult_scalar_lt_lt)
qed
qed
lemma tc_mult_scalar: "tc (p \<odot> q) = punit.tc p * tc (q::'t \<Rightarrow>\<^sub>0 'b::semiring_no_zero_divisors)"
proof (cases "p = 0")
case True
thus ?thesis by (simp add: tc_def)
next
case False
show ?thesis
proof (cases "q = 0")
case True
thus ?thesis by (simp add: tc_def)
next
case False
with \<open>p \<noteq> 0\<close> show ?thesis by (simp add: tc_def tt_mult_scalar lookup_mult_scalar_tt_tt)
qed
qed
lemma mult_scalar_not_zero:
assumes "p \<noteq> 0" and "q \<noteq> (0::'t \<Rightarrow>\<^sub>0 'b::semiring_no_zero_divisors)"
shows "p \<odot> q \<noteq> 0"
proof
assume "p \<odot> q = 0"
hence "0 = lc (p \<odot> q)" by (simp add: lc_def)
also have "... = punit.lc p * lc q" by (rule lc_mult_scalar)
finally have "punit.lc p * lc q = 0" by simp
moreover from assms(1) have "punit.lc p \<noteq> 0" by (rule punit.lc_not_0)
moreover from assms(2) have "lc q \<noteq> 0" by (rule lc_not_0)
ultimately show False by simp
qed
end (* ordered_term_powerprod *)
context ordered_powerprod
begin
lemmas in_keys_times_le = punit.in_keys_mult_scalar_le[simplified]
lemmas in_keys_times_ge = punit.in_keys_mult_scalar_ge[simplified]
lemmas lookup_times_lp_lp = punit.lookup_mult_scalar_lt_lt[simplified]
lemmas lookup_times_tp_tp = punit.lookup_mult_scalar_tt_tt[simplified]
lemmas lookup_times_monomial_right_plus = punit.lookup_mult_scalar_monomial_right_plus[simplified]
lemmas lookup_times_monomial_right = punit.lookup_mult_scalar_monomial_right[simplified]
lemmas lp_times = punit.lt_mult_scalar[simplified]
lemmas tp_times = punit.tt_mult_scalar[simplified]
lemmas lc_times = punit.lc_mult_scalar[simplified]
lemmas tc_times = punit.tc_mult_scalar[simplified]
lemmas times_not_zero = punit.mult_scalar_not_zero[simplified]
lemmas times_tail_rec_left = punit.mult_scalar_tail_rec_left[simplified]
lemmas times_tail_rec_right = punit.mult_scalar_tail_rec_right[simplified]
lemmas punit_in_keys_monom_mult_le = punit.in_keys_monom_mult_le[simplified]
lemmas punit_in_keys_monom_mult_ge = punit.in_keys_monom_mult_ge[simplified]
lemmas lp_monom_mult = punit.lt_monom_mult[simplified]
lemmas tp_monom_mult = punit.tt_monom_mult[simplified]
end
subsection \<open>@{term dgrad_p_set} and @{term dgrad_p_set_le}\<close>
locale gd_term =
ordered_term pair_of_term term_of_pair ord ord_strict ord_term ord_term_strict
for pair_of_term::"'t \<Rightarrow> ('a::graded_dickson_powerprod \<times> 'k::{the_min,wellorder})"
and term_of_pair::"('a \<times> 'k) \<Rightarrow> 't"
and ord::"'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<preceq>" 50)
and ord_strict (infixl "\<prec>" 50)
and ord_term::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<preceq>\<^sub>t" 50)
and ord_term_strict::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<prec>\<^sub>t" 50)
begin
sublocale gd_powerprod ..
lemma adds_term_antisym:
assumes "u adds\<^sub>t v" and "v adds\<^sub>t u"
shows "u = v"
using assms unfolding adds_term_def using adds_antisym by (metis term_of_pair_pair)
definition dgrad_p_set :: "('a \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b::zero) set"
where "dgrad_p_set d m = {p. pp_of_term ` keys p \<subseteq> dgrad_set d m}"
definition dgrad_p_set_le :: "('a \<Rightarrow> nat) \<Rightarrow> (('t \<Rightarrow>\<^sub>0 'b) set) \<Rightarrow> (('t \<Rightarrow>\<^sub>0 'b::zero) set) \<Rightarrow> bool"
where "dgrad_p_set_le d F G \<longleftrightarrow> (dgrad_set_le d (pp_of_term ` Keys F) (pp_of_term ` Keys G))"
lemma in_dgrad_p_set_iff: "p \<in> dgrad_p_set d m \<longleftrightarrow> (\<forall>v\<in>keys p. d (pp_of_term v) \<le> m)"
by (auto simp add: dgrad_p_set_def dgrad_set_def)
lemma dgrad_p_setI [intro]:
assumes "\<And>v. v \<in> keys p \<Longrightarrow> d (pp_of_term v) \<le> m"
shows "p \<in> dgrad_p_set d m"
using assms by (auto simp: in_dgrad_p_set_iff)
lemma dgrad_p_setD:
assumes "p \<in> dgrad_p_set d m" and "v \<in> keys p"
shows "d (pp_of_term v) \<le> m"
using assms by (simp only: in_dgrad_p_set_iff)
lemma zero_in_dgrad_p_set: "0 \<in> dgrad_p_set d m"
by (rule, simp)
lemma dgrad_p_set_zero [simp]: "dgrad_p_set (\<lambda>_. 0) m = UNIV"
by auto
lemma subset_dgrad_p_set_zero: "F \<subseteq> dgrad_p_set (\<lambda>_. 0) m"
by simp
lemma dgrad_p_set_subset:
assumes "m \<le> n"
shows "dgrad_p_set d m \<subseteq> dgrad_p_set d n"
using assms by (auto simp: dgrad_p_set_def dgrad_set_def)
lemma dgrad_p_setD_lp:
assumes "p \<in> dgrad_p_set d m" and "p \<noteq> 0"
shows "d (lp p) \<le> m"
by (rule dgrad_p_setD, fact, rule lt_in_keys, fact)
lemma dgrad_p_set_exhaust_expl:
assumes "finite F"
shows "F \<subseteq> dgrad_p_set d (Max (d ` pp_of_term ` Keys F))"
proof
fix f
assume "f \<in> F"
from assms have "finite (Keys F)" by (rule finite_Keys)
have fin: "finite (d ` pp_of_term ` Keys F)" by (intro finite_imageI, fact)
show "f \<in> dgrad_p_set d (Max (d ` pp_of_term ` Keys F))"
proof (rule dgrad_p_setI)
fix v
assume "v \<in> keys f"
from this \<open>f \<in> F\<close> have "v \<in> Keys F" by (rule in_KeysI)
hence "d (pp_of_term v) \<in> d ` pp_of_term ` Keys F" by simp
with fin show "d (pp_of_term v) \<le> Max (d ` pp_of_term ` Keys F)" by (rule Max_ge)
qed
qed
lemma dgrad_p_set_exhaust:
assumes "finite F"
obtains m where "F \<subseteq> dgrad_p_set d m"
proof
from assms show "F \<subseteq> dgrad_p_set d (Max (d ` pp_of_term ` Keys F))" by (rule dgrad_p_set_exhaust_expl)
qed
lemma dgrad_p_set_insert:
assumes "F \<subseteq> dgrad_p_set d m"
obtains n where "m \<le> n" and "f \<in> dgrad_p_set d n" and "F \<subseteq> dgrad_p_set d n"
proof -
have "finite {f}" by simp
then obtain m1 where "{f} \<subseteq> dgrad_p_set d m1" by (rule dgrad_p_set_exhaust)
hence "f \<in> dgrad_p_set d m1" by simp
define n where "n = ord_class.max m m1"
have "m \<le> n" and "m1 \<le> n" by (simp_all add: n_def)
from this(1) show ?thesis
proof
from \<open>m1 \<le> n\<close> have "dgrad_p_set d m1 \<subseteq> dgrad_p_set d n" by (rule dgrad_p_set_subset)
with \<open>f \<in> dgrad_p_set d m1\<close> show "f \<in> dgrad_p_set d n" ..
next
from \<open>m \<le> n\<close> have "dgrad_p_set d m \<subseteq> dgrad_p_set d n" by (rule dgrad_p_set_subset)
with assms show "F \<subseteq> dgrad_p_set d n" by (rule subset_trans)
qed
qed
lemma dgrad_p_set_leI:
assumes "\<And>f. f \<in> F \<Longrightarrow> dgrad_p_set_le d {f} G"
shows "dgrad_p_set_le d F G"
unfolding dgrad_p_set_le_def dgrad_set_le_def
proof
fix s
assume "s \<in> pp_of_term ` Keys F"
then obtain v where "v \<in> Keys F" and "s = pp_of_term v" ..
from this(1) obtain f where "f \<in> F" and "v \<in> keys f" by (rule in_KeysE)
from this(2) have "s \<in> pp_of_term ` Keys {f}" by (simp add: \<open>s = pp_of_term v\<close> Keys_insert)
from \<open>f \<in> F\<close> have "dgrad_p_set_le d {f} G" by (rule assms)
from this \<open>s \<in> pp_of_term ` Keys {f}\<close> show "\<exists>t\<in>pp_of_term ` Keys G. d s \<le> d t"
unfolding dgrad_p_set_le_def dgrad_set_le_def ..
qed
lemma dgrad_p_set_le_trans [trans]:
assumes "dgrad_p_set_le d F G" and "dgrad_p_set_le d G H"
shows "dgrad_p_set_le d F H"
using assms unfolding dgrad_p_set_le_def by (rule dgrad_set_le_trans)
lemma dgrad_p_set_le_subset:
assumes "F \<subseteq> G"
shows "dgrad_p_set_le d F G"
unfolding dgrad_p_set_le_def by (rule dgrad_set_le_subset, rule image_mono, rule Keys_mono, fact)
lemma dgrad_p_set_leI_insert_keys:
assumes "dgrad_p_set_le d F G" and "dgrad_set_le d (pp_of_term ` keys f) (pp_of_term ` Keys G)"
shows "dgrad_p_set_le d (insert f F) G"
using assms by (simp add: dgrad_p_set_le_def Keys_insert dgrad_set_le_Un image_Un)
lemma dgrad_p_set_leI_insert:
assumes "dgrad_p_set_le d F G" and "dgrad_p_set_le d {f} G"
shows "dgrad_p_set_le d (insert f F) G"
using assms by (simp add: dgrad_p_set_le_def Keys_insert dgrad_set_le_Un image_Un)
lemma dgrad_p_set_leI_Un:
assumes "dgrad_p_set_le d F1 G" and "dgrad_p_set_le d F2 G"
shows "dgrad_p_set_le d (F1 \<union> F2) G"
using assms by (auto simp: dgrad_p_set_le_def dgrad_set_le_def Keys_Un)
lemma dgrad_p_set_le_dgrad_p_set:
assumes "dgrad_p_set_le d F G" and "G \<subseteq> dgrad_p_set d m"
shows "F \<subseteq> dgrad_p_set d m"
proof
fix f
assume "f \<in> F"
show "f \<in> dgrad_p_set d m"
proof (rule dgrad_p_setI)
fix v
assume "v \<in> keys f"
from this \<open>f \<in> F\<close> have "v \<in> Keys F" by (rule in_KeysI)
hence "pp_of_term v \<in> pp_of_term ` Keys F" by simp
with assms(1) obtain s where "s \<in> pp_of_term ` Keys G" and "d (pp_of_term v) \<le> d s"
unfolding dgrad_p_set_le_def by (rule dgrad_set_leE)
from this(1) obtain u where "u \<in> Keys G" and s: "s = pp_of_term u" ..
from this(1) obtain g where "g \<in> G" and "u \<in> keys g" by (rule in_KeysE)
from this(1) assms(2) have "g \<in> dgrad_p_set d m" ..
from this \<open>u \<in> keys g\<close> have "d s \<le> m" unfolding s by (rule dgrad_p_setD)
with \<open>d (pp_of_term v) \<le> d s\<close> show "d (pp_of_term v) \<le> m" by (rule le_trans)
qed
qed
lemma dgrad_p_set_le_except: "dgrad_p_set_le d {except p S} {p}"
by (auto simp add: dgrad_p_set_le_def Keys_insert keys_except intro: dgrad_set_le_subset)
lemma dgrad_p_set_le_tail: "dgrad_p_set_le d {tail p} {p}"
by (simp only: tail_def lower_def, fact dgrad_p_set_le_except)
lemma dgrad_p_set_le_plus: "dgrad_p_set_le d {p + q} {p, q}"
by (simp add: dgrad_p_set_le_def Keys_insert, rule dgrad_set_le_subset, rule image_mono, fact Poly_Mapping.keys_add)
lemma dgrad_p_set_le_uminus: "dgrad_p_set_le d {-p} {p}"
by (simp add: dgrad_p_set_le_def Keys_insert keys_uminus, fact dgrad_set_le_refl)
lemma dgrad_p_set_le_minus: "dgrad_p_set_le d {p - q} {p, q}"
by (simp add: dgrad_p_set_le_def Keys_insert, rule dgrad_set_le_subset, rule image_mono, fact keys_minus)
lemma dgrad_set_le_monom_mult:
assumes "dickson_grading d"
shows "dgrad_set_le d (pp_of_term ` keys (monom_mult c t p)) (insert t (pp_of_term ` keys p))"
proof (rule dgrad_set_leI)
fix s
assume "s \<in> pp_of_term ` keys (monom_mult c t p)"
with keys_monom_mult_subset have "s \<in> pp_of_term ` ((\<oplus>) t ` keys p)" by fastforce
then obtain v where "v \<in> keys p" and s: "s = pp_of_term (t \<oplus> v)" by fastforce
have "d s = ord_class.max (d t) (d (pp_of_term v))"
by (simp only: s pp_of_term_splus dickson_gradingD1[OF assms(1)])
hence "d s = d t \<or> d s = d (pp_of_term v)" by auto
thus "\<exists>t\<in>insert t (pp_of_term ` keys p). d s \<le> d t"
proof
assume "d s = d t"
thus ?thesis by simp
next
assume "d s = d (pp_of_term v)"
show ?thesis
proof
from \<open>d s = d (pp_of_term v)\<close> show "d s \<le> d (pp_of_term v)" by simp
next
from \<open>v \<in> keys p\<close> show "pp_of_term v \<in> insert t (pp_of_term ` keys p)" by simp
qed
qed
qed
lemma dgrad_p_set_closed_plus:
assumes "p \<in> dgrad_p_set d m" and "q \<in> dgrad_p_set d m"
shows "p + q \<in> dgrad_p_set d m"
proof -
from dgrad_p_set_le_plus have "{p + q} \<subseteq> dgrad_p_set d m"
proof (rule dgrad_p_set_le_dgrad_p_set)
from assms show "{p, q} \<subseteq> dgrad_p_set d m" by simp
qed
thus ?thesis by simp
qed
lemma dgrad_p_set_closed_uminus:
assumes "p \<in> dgrad_p_set d m"
shows "-p \<in> dgrad_p_set d m"
proof -
from dgrad_p_set_le_uminus have "{-p} \<subseteq> dgrad_p_set d m"
proof (rule dgrad_p_set_le_dgrad_p_set)
from assms show "{p} \<subseteq> dgrad_p_set d m" by simp
qed
thus ?thesis by simp
qed
lemma dgrad_p_set_closed_minus:
assumes "p \<in> dgrad_p_set d m" and "q \<in> dgrad_p_set d m"
shows "p - q \<in> dgrad_p_set d m"
proof -
from dgrad_p_set_le_minus have "{p - q} \<subseteq> dgrad_p_set d m"
proof (rule dgrad_p_set_le_dgrad_p_set)
from assms show "{p, q} \<subseteq> dgrad_p_set d m" by simp
qed
thus ?thesis by simp
qed
lemma dgrad_p_set_closed_monom_mult:
assumes "dickson_grading d" and "d t \<le> m" and "p \<in> dgrad_p_set d m"
shows "monom_mult c t p \<in> dgrad_p_set d m"
proof (rule dgrad_p_setI)
fix v
assume "v \<in> keys (monom_mult c t p)"
hence "pp_of_term v \<in> pp_of_term ` keys (monom_mult c t p)" by simp
with dgrad_set_le_monom_mult[OF assms(1)] obtain s where "s \<in> insert t (pp_of_term ` keys p)"
and "d (pp_of_term v) \<le> d s" by (rule dgrad_set_leE)
from this(1) have "s = t \<or> s \<in> pp_of_term ` keys p" by simp
thus "d (pp_of_term v) \<le> m"
proof
assume "s = t"
with \<open>d (pp_of_term v) \<le> d s\<close> assms(2) show ?thesis by simp
next
assume "s \<in> pp_of_term ` keys p"
then obtain u where "u \<in> keys p" and "s = pp_of_term u" ..
from assms(3) this(1) have "d s \<le> m" unfolding \<open>s = pp_of_term u\<close> by (rule dgrad_p_setD)
with \<open>d (pp_of_term v) \<le> d s\<close> show ?thesis by (rule le_trans)
qed
qed
lemma dgrad_p_set_closed_monom_mult_zero:
assumes "p \<in> dgrad_p_set d m"
shows "monom_mult c 0 p \<in> dgrad_p_set d m"
proof (rule dgrad_p_setI)
fix v
assume "v \<in> keys (monom_mult c 0 p)"
hence "pp_of_term v \<in> pp_of_term ` keys (monom_mult c 0 p)" by simp
then obtain u where "u \<in> keys (monom_mult c 0 p)" and eq: "pp_of_term v = pp_of_term u" ..
from this(1) have "u \<in> keys p" by (metis keys_monom_multE splus_zero)
with assms have "d (pp_of_term u) \<le> m" by (rule dgrad_p_setD)
thus "d (pp_of_term v) \<le> m" by (simp only: eq)
qed
lemma dgrad_p_set_closed_except:
assumes "p \<in> dgrad_p_set d m"
shows "except p S \<in> dgrad_p_set d m"
by (rule dgrad_p_setI, rule dgrad_p_setD, rule assms, simp add: keys_except)
lemma dgrad_p_set_closed_tail:
assumes "p \<in> dgrad_p_set d m"
shows "tail p \<in> dgrad_p_set d m"
unfolding tail_def lower_def using assms by (rule dgrad_p_set_closed_except)
subsection \<open>Dickson's Lemma for Sequences of Terms\<close>
lemma Dickson_term:
assumes "dickson_grading d" and "finite K"
shows "almost_full_on (adds\<^sub>t) {t. pp_of_term t \<in> dgrad_set d m \<and> component_of_term t \<in> K}"
(is "almost_full_on _ ?A")
proof (rule almost_full_onI)
fix seq :: "nat \<Rightarrow> 't"
assume *: "\<forall>i. seq i \<in> ?A"
define seq' where "seq' = (\<lambda>i. (pp_of_term (seq i), component_of_term (seq i)))"
have "pp_of_term ` ?A \<subseteq> {x. d x \<le> m}" by (auto dest: dgrad_setD)
moreover from assms(1) have "almost_full_on (adds) {x. d x \<le> m}" by (rule dickson_gradingD2)
ultimately have "almost_full_on (adds) (pp_of_term ` ?A)" by (rule almost_full_on_subset)
moreover have "almost_full_on (=) (component_of_term ` ?A)"
proof (rule eq_almost_full_on_finite_set)
have "component_of_term ` ?A \<subseteq> K" by blast
thus "finite (component_of_term ` ?A)" using assms(2) by (rule finite_subset)
qed
ultimately have "almost_full_on (prod_le (adds) (=)) (pp_of_term ` ?A \<times> component_of_term ` ?A)"
by (rule almost_full_on_Sigma)
moreover from * have "\<And>i. seq' i \<in> pp_of_term ` ?A \<times> component_of_term ` ?A" by (simp add: seq'_def)
ultimately obtain i j where "i < j" and "prod_le (adds) (=) (seq' i) (seq' j)"
by (rule almost_full_onD)
from this(2) have "seq i adds\<^sub>t seq j" by (simp add: seq'_def prod_le_def adds_term_def)
with \<open>i < j\<close> show "good (adds\<^sub>t) seq" by (rule goodI)
qed
corollary Dickson_termE:
assumes "dickson_grading d" and "finite (component_of_term ` range (f::nat \<Rightarrow> 't))"
and "pp_of_term ` range f \<subseteq> dgrad_set d m"
obtains i j where "i < j" and "f i adds\<^sub>t f j"
proof -
let ?A = "{t. pp_of_term t \<in> dgrad_set d m \<and> component_of_term t \<in> component_of_term ` range f}"
from assms(1, 2) have "almost_full_on (adds\<^sub>t) ?A" by (rule Dickson_term)
moreover from assms(3) have "\<And>i. f i \<in> ?A" by blast
ultimately obtain i j where "i < j" and "f i adds\<^sub>t f j" by (rule almost_full_onD)
thus ?thesis ..
qed
lemma ex_finite_adds_term:
assumes "dickson_grading d" and "finite (component_of_term ` S)" and "pp_of_term ` S \<subseteq> dgrad_set d m"
obtains T where "finite T" and "T \<subseteq> S" and "\<And>s. s \<in> S \<Longrightarrow> (\<exists>t\<in>T. t adds\<^sub>t s)"
proof -
let ?A = "{t. pp_of_term t \<in> dgrad_set d m \<and> component_of_term t \<in> component_of_term ` S}"
have "reflp ((adds\<^sub>t)::'t \<Rightarrow> _)" by (simp add: reflp_def adds_term_refl)
moreover have "almost_full_on (adds\<^sub>t) S"
proof (rule almost_full_on_subset)
from assms(3) show "S \<subseteq> ?A" by blast
next
from assms(1, 2) show "almost_full_on (adds\<^sub>t) ?A" by (rule Dickson_term)
qed
ultimately obtain T where "finite T" and "T \<subseteq> S" and "\<And>s. s \<in> S \<Longrightarrow> (\<exists>t\<in>T. t adds\<^sub>t s)"
by (rule almost_full_on_finite_subsetE, blast)
thus ?thesis ..
qed
subsection \<open>Well-foundedness\<close>
definition dickson_less_v :: "('a \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> 't \<Rightarrow> 't \<Rightarrow> bool"
where "dickson_less_v d m v u \<longleftrightarrow> (d (pp_of_term v) \<le> m \<and> d (pp_of_term u) \<le> m \<and> v \<prec>\<^sub>t u)"
definition dickson_less_p :: "('a \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> bool"
where "dickson_less_p d m p q \<longleftrightarrow> ({p, q} \<subseteq> dgrad_p_set d m \<and> p \<prec>\<^sub>p q)"
lemma dickson_less_vI:
assumes "d (pp_of_term v) \<le> m" and "d (pp_of_term u) \<le> m" and "v \<prec>\<^sub>t u"
shows "dickson_less_v d m v u"
using assms by (simp add: dickson_less_v_def)
lemma dickson_less_vD1:
assumes "dickson_less_v d m v u"
shows "d (pp_of_term v) \<le> m"
using assms by (simp add: dickson_less_v_def)
lemma dickson_less_vD2:
assumes "dickson_less_v d m v u"
shows "d (pp_of_term u) \<le> m"
using assms by (simp add: dickson_less_v_def)
lemma dickson_less_vD3:
assumes "dickson_less_v d m v u"
shows "v \<prec>\<^sub>t u"
using assms by (simp add: dickson_less_v_def)
lemma dickson_less_v_irrefl: "\<not> dickson_less_v d m v v"
by (simp add: dickson_less_v_def)
lemma dickson_less_v_trans:
assumes "dickson_less_v d m v u" and "dickson_less_v d m u w"
shows "dickson_less_v d m v w"
using assms by (auto simp add: dickson_less_v_def)
lemma wf_dickson_less_v_aux1:
assumes "dickson_grading d" and "\<And>i::nat. dickson_less_v d m (seq (Suc i)) (seq i)"
obtains i where "\<And>j. j > i \<Longrightarrow> component_of_term (seq j) < component_of_term (seq i)"
proof -
let ?Q = "pp_of_term ` range seq"
have "pp_of_term (seq 0) \<in> ?Q" by simp
with wf_dickson_less[OF assms(1)] obtain t where "t \<in> ?Q" and *: "\<And>s. dickson_less d m s t \<Longrightarrow> s \<notin> ?Q"
by (rule wfE_min[to_pred], blast)
from this(1) obtain i where t: "t = pp_of_term (seq i)" by fastforce
show ?thesis
proof
fix j
assume "i < j"
with _ assms(2) have dlv: "dickson_less_v d m (seq j) (seq i)"
proof (rule transp_sequence)
from dickson_less_v_trans show "transp (dickson_less_v d m)" by (rule transpI)
qed
hence "seq j \<prec>\<^sub>t seq i" by (rule dickson_less_vD3)
define s where "s = pp_of_term (seq j)"
have "pp_of_term (seq j) \<in> ?Q" by simp
hence "\<not> dickson_less d m s t" unfolding s_def using * by blast
moreover from dlv have "d s \<le> m" and "d t \<le> m" unfolding s_def t
by (rule dickson_less_vD1, rule dickson_less_vD2)
ultimately have "t \<preceq> s" by (simp add: dickson_less_def)
show "component_of_term (seq j) < component_of_term (seq i)"
proof (rule ccontr, simp only: not_less)
assume "component_of_term (seq i) \<le> component_of_term (seq j)"
with \<open>t \<preceq> s\<close> have "seq i \<preceq>\<^sub>t seq j" unfolding s_def t by (rule ord_termI)
moreover from dlv have "seq j \<prec>\<^sub>t seq i" by (rule dickson_less_vD3)
ultimately show False by simp
qed
qed
qed
lemma wf_dickson_less_v_aux2:
assumes "dickson_grading d" and "\<And>i::nat. dickson_less_v d m (seq (Suc i)) (seq i)"
and "\<And>i::nat. component_of_term (seq i) < k"
shows thesis
using assms(2, 3)
proof (induct k arbitrary: seq thesis rule: less_induct)
case (less k)
from assms(1) less(2) obtain i where *: "\<And>j. j > i \<Longrightarrow> component_of_term (seq j) < component_of_term (seq i)"
by (rule wf_dickson_less_v_aux1, blast)
define seq1 where "seq1 = (\<lambda>j. seq (Suc (i + j)))"
from less(3) show ?case
proof (rule less(1))
fix j
show "dickson_less_v d m (seq1 (Suc j)) (seq1 j)" by (simp add: seq1_def, fact less(2))
next
fix j
show "component_of_term (seq1 j) < component_of_term (seq i)" by (simp add: seq1_def *)
qed
qed
lemma wf_dickson_less_v:
assumes "dickson_grading d"
shows "wfP (dickson_less_v d m)"
proof (rule wfP_chain, rule, elim exE)
fix seq::"nat \<Rightarrow> 't"
assume "\<forall>i. dickson_less_v d m (seq (Suc i)) (seq i)"
hence *: "\<And>i. dickson_less_v d m (seq (Suc i)) (seq i)" ..
with assms obtain i where **: "\<And>j. j > i \<Longrightarrow> component_of_term (seq j) < component_of_term (seq i)"
by (rule wf_dickson_less_v_aux1, blast)
define seq1 where "seq1 = (\<lambda>j. seq (Suc (i + j)))"
from assms show False
proof (rule wf_dickson_less_v_aux2)
fix j
show "dickson_less_v d m (seq1 (Suc j)) (seq1 j)" by (simp add: seq1_def, fact *)
next
fix j
show "component_of_term (seq1 j) < component_of_term (seq i)" by (simp add: seq1_def **)
qed
qed
lemma dickson_less_v_zero: "dickson_less_v (\<lambda>_. 0) m = (\<prec>\<^sub>t)"
by (rule, rule, simp add: dickson_less_v_def)
lemma dickson_less_pI:
assumes "p \<in> dgrad_p_set d m" and "q \<in> dgrad_p_set d m" and "p \<prec>\<^sub>p q"
shows "dickson_less_p d m p q"
using assms by (simp add: dickson_less_p_def)
lemma dickson_less_pD1:
assumes "dickson_less_p d m p q"
shows "p \<in> dgrad_p_set d m"
using assms by (simp add: dickson_less_p_def)
lemma dickson_less_pD2:
assumes "dickson_less_p d m p q"
shows "q \<in> dgrad_p_set d m"
using assms by (simp add: dickson_less_p_def)
lemma dickson_less_pD3:
assumes "dickson_less_p d m p q"
shows "p \<prec>\<^sub>p q"
using assms by (simp add: dickson_less_p_def)
lemma dickson_less_p_irrefl: "\<not> dickson_less_p d m p p"
by (simp add: dickson_less_p_def)
lemma dickson_less_p_trans:
assumes "dickson_less_p d m p q" and "dickson_less_p d m q r"
shows "dickson_less_p d m p r"
using assms by (auto simp add: dickson_less_p_def)
lemma dickson_less_p_mono:
assumes "dickson_less_p d m p q" and "m \<le> n"
shows "dickson_less_p d n p q"
proof -
from assms(2) have "dgrad_p_set d m \<subseteq> dgrad_p_set d n" by (rule dgrad_p_set_subset)
moreover from assms(1) have "p \<in> dgrad_p_set d m" and "q \<in> dgrad_p_set d m" and "p \<prec>\<^sub>p q"
by (rule dickson_less_pD1, rule dickson_less_pD2, rule dickson_less_pD3)
ultimately have "p \<in> dgrad_p_set d n" and "q \<in> dgrad_p_set d n" by auto
from this \<open>p \<prec>\<^sub>p q\<close> show ?thesis by (rule dickson_less_pI)
qed
lemma dickson_less_p_zero: "dickson_less_p (\<lambda>_. 0) m = (\<prec>\<^sub>p)"
by (rule, rule, simp add: dickson_less_p_def)
lemma wf_dickson_less_p_aux:
assumes "dickson_grading d"
assumes "x \<in> Q" and "\<forall>y\<in>Q. y \<noteq> 0 \<longrightarrow> (y \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt y) u)"
shows "\<exists>p\<in>Q. (\<forall>q\<in>Q. \<not> dickson_less_p d m q p)"
using assms(2) assms(3)
proof (induct u arbitrary: x Q rule: wfP_induct[OF wf_dickson_less_v, OF assms(1)])
fix u::'t and x::"'t \<Rightarrow>\<^sub>0 'b" and Q::"('t \<Rightarrow>\<^sub>0 'b) set"
assume hyp: "\<forall>u0. dickson_less_v d m u0 u \<longrightarrow> (\<forall>x0 Q0::('t \<Rightarrow>\<^sub>0 'b) set. x0 \<in> Q0 \<longrightarrow>
(\<forall>y\<in>Q0. y \<noteq> 0 \<longrightarrow> (y \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt y) u0)) \<longrightarrow>
(\<exists>p\<in>Q0. \<forall>q\<in>Q0. \<not> dickson_less_p d m q p))"
assume "x \<in> Q"
assume "\<forall>y\<in>Q. y \<noteq> 0 \<longrightarrow> (y \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt y) u)"
hence bounded: "\<And>y. y \<in> Q \<Longrightarrow> y \<noteq> 0 \<Longrightarrow> (y \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt y) u)" by auto
show "\<exists>p\<in>Q. \<forall>q\<in>Q. \<not> dickson_less_p d m q p"
proof (cases "0 \<in> Q")
case True
show ?thesis
proof (rule, rule, rule)
fix q::"'t \<Rightarrow>\<^sub>0 'b"
assume "dickson_less_p d m q 0"
hence "q \<prec>\<^sub>p 0" by (rule dickson_less_pD3)
thus False using ord_p_zero_min[of q] by simp
next
from True show "0 \<in> Q" .
qed
next
case False
define Q1 where "Q1 = {lt p | p. p \<in> Q}"
from \<open>x \<in> Q\<close> have "lt x \<in> Q1" unfolding Q1_def by auto
with wf_dickson_less_v[OF assms(1)] obtain v
where "v \<in> Q1" and v_min_1: "\<And>q. dickson_less_v d m q v \<Longrightarrow> q \<notin> Q1"
by (rule wfE_min[to_pred], auto)
have v_min: "\<And>q. q \<in> Q \<Longrightarrow> \<not> dickson_less_v d m (lt q) v"
proof -
fix q
assume "q \<in> Q"
hence "lt q \<in> Q1" unfolding Q1_def by auto
thus "\<not> dickson_less_v d m (lt q) v" using v_min_1 by auto
qed
from \<open>v \<in> Q1\<close> obtain p where "lt p = v" and "p \<in> Q" unfolding Q1_def by auto
hence "p \<noteq> 0" using False by auto
with \<open>p \<in> Q\<close> have "p \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt p) u" by (rule bounded)
hence "p \<in> dgrad_p_set d m" and "dickson_less_v d m (lt p) u" by simp_all
moreover from this(1) \<open>p \<noteq> 0\<close> have "d (pp_of_term (lt p)) \<le> m" by (rule dgrad_p_setD_lp)
ultimately have "d (pp_of_term v) \<le> m" by (simp only: \<open>lt p = v\<close>)
define Q2 where "Q2 = {tail p | p. p \<in> Q \<and> lt p = v}"
from \<open>p \<in> Q\<close> \<open>lt p = v\<close> have "tail p \<in> Q2" unfolding Q2_def by auto
have "\<forall>q\<in>Q2. q \<noteq> 0 \<longrightarrow> (q \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt q) (lt p))"
proof (intro ballI impI)
fix q
assume "q \<in> Q2"
then obtain q0 where q: "q = tail q0" and "lt q0 = lt p" and "q0 \<in> Q"
using \<open>lt p = v\<close> by (auto simp add: Q2_def)
assume "q \<noteq> 0"
hence "tail q0 \<noteq> 0" using \<open>q = tail q0\<close> by simp
hence "q0 \<noteq> 0" by auto
with \<open>q0 \<in> Q\<close> have "q0 \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt q0) u" by (rule bounded)
hence "q0 \<in> dgrad_p_set d m" and "dickson_less_v d m (lt q0) u" by simp_all
from this(1) have "q \<in> dgrad_p_set d m" unfolding q by (rule dgrad_p_set_closed_tail)
show "q \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt q) (lt p)"
proof
show "dickson_less_v d m (lt q) (lt p)"
proof (rule dickson_less_vI)
from \<open>q \<in> dgrad_p_set d m\<close> \<open>q \<noteq> 0\<close> show "d (pp_of_term (lt q)) \<le> m" by (rule dgrad_p_setD_lp)
next
from \<open>dickson_less_v d m (lt p) u\<close> show "d (pp_of_term (lt p)) \<le> m" by (rule dickson_less_vD1)
next
from lt_tail[OF \<open>tail q0 \<noteq> 0\<close>] \<open>q = tail q0\<close> \<open>lt q0 = lt p\<close> show "lt q \<prec>\<^sub>t lt p" by simp
qed
qed fact
qed
with hyp \<open>dickson_less_v d m (lt p) u\<close> \<open>tail p \<in> Q2\<close> have "\<exists>p\<in>Q2. \<forall>q\<in>Q2. \<not> dickson_less_p d m q p"
by blast
then obtain q where "q \<in> Q2" and q_min: "\<forall>r\<in>Q2. \<not> dickson_less_p d m r q" ..
from \<open>q \<in> Q2\<close> obtain q0 where "q = tail q0" and "q0 \<in> Q" and "lt q0 = v" unfolding Q2_def by auto
from q_min \<open>q = tail q0\<close> have q0_tail_min: "\<And>r. r \<in> Q2 \<Longrightarrow> \<not> dickson_less_p d m r (tail q0)" by simp
from \<open>q0 \<in> Q\<close> show ?thesis
proof
show "\<forall>r\<in>Q. \<not> dickson_less_p d m r q0"
proof (intro ballI notI)
fix r
assume "dickson_less_p d m r q0"
hence "r \<in> dgrad_p_set d m" and "q0 \<in> dgrad_p_set d m" and "r \<prec>\<^sub>p q0"
by (rule dickson_less_pD1, rule dickson_less_pD2, rule dickson_less_pD3)
from this(3) have "lt r \<preceq>\<^sub>t lt q0" by (simp add: ord_p_lt)
with \<open>lt q0 = v\<close> have "lt r \<preceq>\<^sub>t v" by simp
assume "r \<in> Q"
hence "\<not> dickson_less_v d m (lt r) v" by (rule v_min)
from False \<open>r \<in> Q\<close> have "r \<noteq> 0" using False by blast
with \<open>r \<in> dgrad_p_set d m\<close> have "d (pp_of_term (lt r)) \<le> m" by (rule dgrad_p_setD_lp)
have "\<not> lt r \<prec>\<^sub>t v"
proof
assume "lt r \<prec>\<^sub>t v"
with \<open>d (pp_of_term (lt r)) \<le> m\<close> \<open>d (pp_of_term v) \<le> m\<close> have "dickson_less_v d m (lt r) v"
by (rule dickson_less_vI)
with \<open>\<not> dickson_less_v d m (lt r) v\<close> show False ..
qed
with \<open>lt r \<preceq>\<^sub>t v\<close> have "lt r = v" by simp
with \<open>r \<in> Q\<close> have "tail r \<in> Q2" by (auto simp add: Q2_def)
have "dickson_less_p d m (tail r) (tail q0)"
proof (rule dickson_less_pI)
show "tail r \<in> dgrad_p_set d m" by (rule dgrad_p_set_closed_tail, fact)
next
show "tail q0 \<in> dgrad_p_set d m" by (rule dgrad_p_set_closed_tail, fact)
next
have "lt r = lt q0" by (simp only: \<open>lt r = v\<close> \<open>lt q0 = v\<close>)
from \<open>r \<noteq> 0\<close> this \<open>r \<prec>\<^sub>p q0\<close> show "tail r \<prec>\<^sub>p tail q0" by (rule ord_p_tail)
qed
with q0_tail_min[OF \<open>tail r \<in> Q2\<close>] show False ..
qed
qed
qed
qed
theorem wf_dickson_less_p:
assumes "dickson_grading d"
shows "wfP (dickson_less_p d m)"
proof (rule wfI_min[to_pred])
fix Q::"('t \<Rightarrow>\<^sub>0 'b) set" and x
assume "x \<in> Q"
show "\<exists>z\<in>Q. \<forall>y. dickson_less_p d m y z \<longrightarrow> y \<notin> Q"
proof (cases "0 \<in> Q")
case True
show ?thesis
proof (rule, rule, rule)
from True show "0 \<in> Q" .
next
fix q::"'t \<Rightarrow>\<^sub>0 'b"
assume "dickson_less_p d m q 0"
hence "q \<prec>\<^sub>p 0" by (rule dickson_less_pD3)
thus "q \<notin> Q" using ord_p_zero_min[of q] by simp
qed
next
case False
show ?thesis
proof (cases "Q \<subseteq> dgrad_p_set d m")
case True
let ?L = "lt ` Q"
from \<open>x \<in> Q\<close> have "lt x \<in> ?L" by simp
with wf_dickson_less_v[OF assms] obtain v where "v \<in> ?L"
and v_min: "\<And>u. dickson_less_v d m u v \<Longrightarrow> u \<notin> ?L" by (rule wfE_min[to_pred], blast)
from this(1) obtain x1 where "x1 \<in> Q" and "v = lt x1" ..
from this(1) True False have "x1 \<in> dgrad_p_set d m" and "x1 \<noteq> 0" by auto
hence "d (pp_of_term v) \<le> m" unfolding \<open>v = lt x1\<close> by (rule dgrad_p_setD_lp)
define Q1 where "Q1 = {tail p | p. p \<in> Q \<and> lt p = v}"
from \<open>x1 \<in> Q\<close> have "tail x1 \<in> Q1" by (auto simp add: Q1_def \<open>v = lt x1\<close>)
with assms have "\<exists>p\<in>Q1. \<forall>q\<in>Q1. \<not> dickson_less_p d m q p"
proof (rule wf_dickson_less_p_aux)
show "\<forall>y\<in>Q1. y \<noteq> 0 \<longrightarrow> y \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt y) v"
proof (intro ballI impI)
fix y
assume "y \<in> Q1" and "y \<noteq> 0"
from this(1) obtain y1 where "y1 \<in> Q" and "v = lt y1" and "y = tail y1" unfolding Q1_def
by blast
from this(1) True have "y1 \<in> dgrad_p_set d m" ..
hence "y \<in> dgrad_p_set d m" unfolding \<open>y = tail y1\<close> by (rule dgrad_p_set_closed_tail)
thus "y \<in> dgrad_p_set d m \<and> dickson_less_v d m (lt y) v"
proof
show "dickson_less_v d m (lt y) v"
proof (rule dickson_less_vI)
from \<open>y \<in> dgrad_p_set d m\<close> \<open>y \<noteq> 0\<close> show "d (pp_of_term (lt y)) \<le> m"
by (rule dgrad_p_setD_lp)
next
from \<open>y \<noteq> 0\<close> show "lt y \<prec>\<^sub>t v" unfolding \<open>v = lt y1\<close> \<open>y = tail y1\<close> by (rule lt_tail)
qed fact
qed
qed
qed
then obtain p0 where "p0 \<in> Q1" and p0_min: "\<And>q. q \<in> Q1 \<Longrightarrow> \<not> dickson_less_p d m q p0" by blast
from this(1) obtain p where "p \<in> Q" and "v = lt p" and "p0 = tail p" unfolding Q1_def
by blast
from this(1) False have "p \<noteq> 0" by blast
show ?thesis
proof (intro bexI allI impI notI)
fix y
assume "y \<in> Q"
hence "y \<noteq> 0" using False by blast
assume "dickson_less_p d m y p"
hence "y \<in> dgrad_p_set d m" and "p \<in> dgrad_p_set d m" and "y \<prec>\<^sub>p p"
by (rule dickson_less_pD1, rule dickson_less_pD2, rule dickson_less_pD3)
from this(3) have "y \<preceq>\<^sub>p p" by simp
hence "lt y \<preceq>\<^sub>t lt p" by (rule ord_p_lt)
moreover have "\<not> lt y \<prec>\<^sub>t lt p"
proof
assume "lt y \<prec>\<^sub>t lt p"
have "dickson_less_v d m (lt y) v" unfolding \<open>v = lt p\<close>
by (rule dickson_less_vI, rule dgrad_p_setD_lp, fact+, rule dgrad_p_setD_lp, fact+)
hence "lt y \<notin> ?L" by (rule v_min)
hence "y \<notin> Q" by fastforce
from this \<open>y \<in> Q\<close> show False ..
qed
ultimately have "lt y = lt p" by simp
from \<open>y \<noteq> 0\<close> this \<open>y \<prec>\<^sub>p p\<close> have "tail y \<prec>\<^sub>p tail p" by (rule ord_p_tail)
from \<open>y \<in> Q\<close> have "tail y \<in> Q1" by (auto simp add: Q1_def \<open>v = lt p\<close> \<open>lt y = lt p\<close>[symmetric])
hence "\<not> dickson_less_p d m (tail y) p0" by (rule p0_min)
moreover have "dickson_less_p d m (tail y) p0" unfolding \<open>p0 = tail p\<close>
by (rule dickson_less_pI, rule dgrad_p_set_closed_tail, fact, rule dgrad_p_set_closed_tail, fact+)
ultimately show False ..
qed fact
next
case False
then obtain q where "q \<in> Q" and "q \<notin> dgrad_p_set d m" by blast
from this(1) show ?thesis
proof
show "\<forall>y. dickson_less_p d m y q \<longrightarrow> y \<notin> Q"
proof (intro allI impI)
fix y
assume "dickson_less_p d m y q"
hence "q \<in> dgrad_p_set d m" by (rule dickson_less_pD2)
with \<open>q \<notin> dgrad_p_set d m\<close> show "y \<notin> Q" ..
qed
qed
qed
qed
qed
corollary ord_p_minimum_dgrad_p_set:
assumes "dickson_grading d" and "x \<in> Q" and "Q \<subseteq> dgrad_p_set d m"
obtains q where "q \<in> Q" and "\<And>y. y \<prec>\<^sub>p q \<Longrightarrow> y \<notin> Q"
proof -
from assms(1) have "wfP (dickson_less_p d m)" by (rule wf_dickson_less_p)
from this assms(2) obtain q where "q \<in> Q" and *: "\<And>y. dickson_less_p d m y q \<Longrightarrow> y \<notin> Q"
by (rule wfE_min[to_pred], auto)
from assms(3) \<open>q \<in> Q\<close> have "q \<in> dgrad_p_set d m" ..
from \<open>q \<in> Q\<close> show ?thesis
proof
fix y
assume "y \<prec>\<^sub>p q"
show "y \<notin> Q"
proof
assume "y \<in> Q"
with assms(3) have "y \<in> dgrad_p_set d m" ..
from this \<open>q \<in> dgrad_p_set d m\<close> \<open>y \<prec>\<^sub>p q\<close> have "dickson_less_p d m y q"
by (rule dickson_less_pI)
hence "y \<notin> Q" by (rule *)
from this \<open>y \<in> Q\<close> show False ..
qed
qed
qed
lemma ord_term_minimum_dgrad_set:
assumes "dickson_grading d" and "v \<in> V" and "pp_of_term ` V \<subseteq> dgrad_set d m"
obtains u where "u \<in> V" and "\<And>w. w \<prec>\<^sub>t u \<Longrightarrow> w \<notin> V"
proof -
from assms(1) have "wfP (dickson_less_v d m)" by (rule wf_dickson_less_v)
then obtain u where "u \<in> V" and *: "\<And>w. dickson_less_v d m w u \<Longrightarrow> w \<notin> V" using assms(2)
by (rule wfE_min[to_pred]) blast
from this(1) have "pp_of_term u \<in> pp_of_term ` V" by (rule imageI)
with assms(3) have "pp_of_term u \<in> dgrad_set d m" ..
hence "d (pp_of_term u) \<le> m" by (rule dgrad_setD)
from \<open>u \<in> V\<close> show ?thesis
proof
fix w
assume "w \<prec>\<^sub>t u"
show "w \<notin> V"
proof
assume "w \<in> V"
hence "pp_of_term w \<in> pp_of_term ` V" by (rule imageI)
with assms(3) have "pp_of_term w \<in> dgrad_set d m" ..
hence "d (pp_of_term w) \<le> m" by (rule dgrad_setD)
from this \<open>d (pp_of_term u) \<le> m\<close> \<open>w \<prec>\<^sub>t u\<close> have "dickson_less_v d m w u"
by (rule dickson_less_vI)
hence "w \<notin> V" by (rule *)
from this \<open>w \<in> V\<close> show False ..
qed
qed
qed
end (* gd_term *)
subsection \<open>More Interpretations\<close>
context gd_powerprod
begin
sublocale punit: gd_term to_pair_unit fst "(\<preceq>)" "(\<prec>)" "(\<preceq>)" "(\<prec>)" ..
end
locale od_term =
ordered_term pair_of_term term_of_pair ord ord_strict ord_term ord_term_strict
for pair_of_term::"'t \<Rightarrow> ('a::dickson_powerprod \<times> 'k::{the_min,wellorder})"
and term_of_pair::"('a \<times> 'k) \<Rightarrow> 't"
and ord::"'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<preceq>" 50)
and ord_strict (infixl "\<prec>" 50)
and ord_term::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<preceq>\<^sub>t" 50)
and ord_term_strict::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<prec>\<^sub>t" 50)
begin
sublocale gd_term ..
lemma ord_p_wf: "wfP (\<prec>\<^sub>p)"
proof -
from dickson_grading_zero have "wfP (dickson_less_p (\<lambda>_. 0) 0)" by (rule wf_dickson_less_p)
thus ?thesis by (simp only: dickson_less_p_zero)
qed
end (* od_term *)
end (* theory *)
diff --git a/thys/Polynomials/OAlist.thy b/thys/Polynomials/OAlist.thy
--- a/thys/Polynomials/OAlist.thy
+++ b/thys/Polynomials/OAlist.thy
@@ -1,4279 +1,4282 @@
(* Author: Florian Haftmann, TU Muenchen *)
(* Author: Andreas Lochbihler, ETH Zurich *)
(* Author: Alexander Maletzky, RISC Linz *)
section \<open>Associative Lists with Sorted Keys\<close>
theory OAlist
imports Deriving.Comparator
begin
text \<open>We define the type of @{emph \<open>ordered associative lists\<close>} (oalist). An oalist is an associative
list (i.\,e. a list of pairs) such that the keys are distinct and sorted wrt. some linear
order relation, and no key is mapped to @{term 0}. The latter invariant allows to implement various
functions operating on oalists more efficiently.
The ordering of the keys in an oalist \<open>xs\<close> is encoded as an additional parameter of \<open>xs\<close>.
This means that oalists may be ordered wrt. different orderings, even if they are of the same type.
Operations operating on more than one oalists, like \<open>map2_val\<close>, typically ensure that the orderings
of their arguments are identical by re-ordering one argument wrt. the order relation of the other.
This, however, implies that equality of order relations must be effectively decidable if executable
code is to be generated.\<close>
subsection \<open>Preliminaries\<close>
fun min_list_param :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a" where
"min_list_param rel (x # xs) = (case xs of [] \<Rightarrow> x | _ \<Rightarrow> (let m = min_list_param rel xs in if rel x m then x else m))"
lemma min_list_param_in:
assumes "xs \<noteq> []"
shows "min_list_param rel xs \<in> set xs"
using assms
proof (induct xs)
case Nil
thus ?case by simp
next
case (Cons x xs)
show ?case
proof (simp add: min_list_param.simps[of rel x xs] Let_def del: min_list_param.simps set_simps(2) split: list.split,
intro conjI impI allI, simp, simp)
fix y ys
assume xs: "xs = y # ys"
have "min_list_param rel (y # ys) = min_list_param rel xs" by (simp only: xs)
also have "... \<in> set xs" by (rule Cons(1), simp add: xs)
also have "... \<subseteq> set (x # y # ys)" by (auto simp: xs)
finally show "min_list_param rel (y # ys) \<in> set (x # y # ys)" .
qed
qed
lemma min_list_param_minimal:
assumes "transp rel" and "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> rel x y \<or> rel y x"
and "z \<in> set xs"
shows "rel (min_list_param rel xs) z"
using assms(2, 3)
proof (induct xs)
case Nil
from Nil(2) show ?case by simp
next
case (Cons x xs)
from Cons(3) have disj1: "z = x \<or> z \<in> set xs" by simp
have "x \<in> set (x # xs)" by simp
hence disj2: "rel x z \<or> rel z x" using Cons(3) by (rule Cons(2))
have *: "rel (min_list_param rel xs) z" if "z \<in> set xs" using _ that
proof (rule Cons(1))
fix a b
assume "a \<in> set xs" and "b \<in> set xs"
hence "a \<in> set (x # xs)" and "b \<in> set (x # xs)" by simp_all
thus "rel a b \<or> rel b a" by (rule Cons(2))
qed
show ?case
proof (simp add: min_list_param.simps[of rel x xs] Let_def del: min_list_param.simps set_simps(2) split: list.split,
intro conjI impI allI)
assume "xs = []"
with disj1 disj2 show "rel x z" by simp
next
fix y ys
assume "xs = y # ys" and "rel x (min_list_param rel (y # ys))"
hence "rel x (min_list_param rel xs)" by simp
from disj1 show "rel x z"
proof
assume "z = x"
thus ?thesis using disj2 by simp
next
assume "z \<in> set xs"
hence "rel (min_list_param rel xs) z" by (rule *)
with assms(1) \<open>rel x (min_list_param rel xs)\<close> show ?thesis by (rule transpD)
qed
next
fix y ys
assume xs: "xs = y # ys" and "\<not> rel x (min_list_param rel (y # ys))"
from disj1 show "rel (min_list_param rel (y # ys)) z"
proof
assume "z = x"
have "min_list_param rel (y # ys) \<in> set (y # ys)" by (rule min_list_param_in, simp)
hence "min_list_param rel (y # ys) \<in> set (x # xs)" by (simp add: xs)
with \<open>x \<in> set (x # xs)\<close> have "rel x (min_list_param rel (y # ys)) \<or> rel (min_list_param rel (y # ys)) x"
by (rule Cons(2))
with \<open>\<not> rel x (min_list_param rel (y # ys))\<close> have "rel (min_list_param rel (y # ys)) x" by simp
thus ?thesis by (simp only: \<open>z = x\<close>)
next
assume "z \<in> set xs"
hence "rel (min_list_param rel xs) z" by (rule *)
thus ?thesis by (simp only: xs)
qed
qed
qed
definition comp_of_ord :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a comparator" where
"comp_of_ord le x y = (if le x y then if x = y then Eq else Lt else Gt)"
lemma comp_of_ord_eq_comp_of_ords:
assumes "antisymp le"
shows "comp_of_ord le = comp_of_ords le (\<lambda>x y. le x y \<and> \<not> le y x)"
by (intro ext, auto simp: comp_of_ord_def comp_of_ords_def intro: assms antisympD)
lemma comparator_converse:
assumes "comparator cmp"
shows "comparator (\<lambda>x y. cmp y x)"
proof -
- from assms interpret comp: comparator cmp .
- show ?thesis by (unfold_locales, auto simp: comp.eq comp.sym intro: comp.trans)
+ from assms interpret comp?: comparator cmp .
+ show ?thesis by (unfold_locales, auto simp: comp.eq comp.sym intro: comp_trans)
qed
lemma comparator_composition:
assumes "comparator cmp" and "inj f"
shows "comparator (\<lambda>x y. cmp (f x) (f y))"
proof -
- from assms(1) interpret comp: comparator cmp .
+ from assms(1) interpret comp?: comparator cmp .
from assms(2) have *: "x = y" if "f x = f y" for x y using that by (rule injD)
- show ?thesis by (unfold_locales, auto simp: comp.sym comp.eq * intro: comp.trans)
+ show ?thesis by (unfold_locales, auto simp: comp.sym comp.eq * intro: comp_trans)
qed
(*
subsection \<open>Syntactic Type Class for Default Elements\<close>
text \<open>We do not want to use the existing type-class @{class default}, but instead introduce a new one:\<close>
class oalist_dflt = fixes dflt::'a
simproc_setup reorient_dflt ("dflt = x") = Reorient_Proc.proc
*)
subsection \<open>Type \<open>key_order\<close>\<close>
typedef 'a key_order = "{compare :: 'a comparator. comparator compare}"
morphisms key_compare Abs_key_order
proof -
from well_order_on obtain r where "well_order_on (UNIV::'a set) r" ..
hence "linear_order r" by (simp only: well_order_on_def)
hence lin: "(x, y) \<in> r \<or> (y, x) \<in> r" for x y
by (metis Diff_iff Linear_order_in_diff_Id UNIV_I \<open>well_order r\<close> well_order_on_Field)
have antisym: "(x, y) \<in> r \<Longrightarrow> (y, x) \<in> r \<Longrightarrow> x = y" for x y
by (meson \<open>linear_order r\<close> antisymD linear_order_on_def partial_order_on_def)
have trans: "(x, y) \<in> r \<Longrightarrow> (y, z) \<in> r \<Longrightarrow> (x, z) \<in> r" for x y z
by (meson \<open>linear_order r\<close> linear_order_on_def order_on_defs(1) partial_order_on_def trans_def)
define comp where "comp = (\<lambda>x y. if (x, y) \<in> r then if (y, x) \<in> r then Eq else Lt else Gt)"
show ?thesis
proof (rule, simp)
show "comparator comp"
proof (standard, simp_all add: comp_def split: if_splits, intro impI)
fix x y
assume "(x, y) \<in> r" and "(y, x) \<in> r"
thus "x = y" by (rule antisym)
next
fix x y
assume "(x, y) \<notin> r"
with lin show "(y, x) \<in> r" by blast
next
fix x y z
assume "(y, x) \<notin> r" and "(z, y) \<notin> r"
assume "(x, y) \<in> r" and "(y, z) \<in> r"
hence "(x, z) \<in> r" by (rule trans)
moreover have "(z, x) \<notin> r"
proof
assume "(z, x) \<in> r"
with \<open>(x, z) \<in> r\<close> have "x = z" by (rule antisym)
from \<open>(z, y) \<notin> r\<close> \<open>(x, y) \<in> r\<close> show False unfolding \<open>x = z\<close> ..
qed
ultimately show "(z, x) \<notin> r \<and> ((z, x) \<notin> r \<longrightarrow> (x, z) \<in> r)" by simp
qed
qed
qed
lemma comparator_key_compare [simp, intro!]: "comparator (key_compare ko)"
using key_compare[of ko] by simp
instantiation key_order :: (type) equal
begin
definition equal_key_order :: "'a key_order \<Rightarrow> 'a key_order \<Rightarrow> bool" where "equal_key_order = (=)"
instance by (standard, simp add: equal_key_order_def)
end
setup_lifting type_definition_key_order
instantiation key_order :: (type) uminus
begin
lift_definition uminus_key_order :: "'a key_order \<Rightarrow> 'a key_order" is "\<lambda>c x y. c y x"
by (fact comparator_converse)
instance ..
end
lift_definition le_of_key_order :: "'a key_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" is "\<lambda>cmp. le_of_comp cmp" .
lift_definition lt_of_key_order :: "'a key_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" is "\<lambda>cmp. lt_of_comp cmp" .
definition key_order_of_ord :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a key_order"
where "key_order_of_ord ord = Abs_key_order (comp_of_ord ord)"
lift_definition key_order_of_le :: "'a::linorder key_order" is comparator_of
by (fact comparator_of)
interpretation key_order_lin: linorder "le_of_key_order ko" "lt_of_key_order ko"
proof transfer
fix comp::"'a comparator"
assume "comparator comp"
then interpret comp: comparator comp .
show "class.linorder comp.le comp.lt" by (fact comp.linorder)
qed
lemma le_of_key_order_alt: "le_of_key_order ko x y = (key_compare ko x y \<noteq> Gt)"
by (transfer, simp add: comparator.nGt_le_conv)
lemma lt_of_key_order_alt: "lt_of_key_order ko x y = (key_compare ko x y = Lt)"
by (transfer, meson comparator.Lt_lt_conv)
lemma key_compare_Gt: "key_compare ko x y = Gt \<longleftrightarrow> key_compare ko y x = Lt"
by (transfer, meson comparator.nGt_le_conv comparator.nLt_le_conv)
lemma key_compare_Eq: "key_compare ko x y = Eq \<longleftrightarrow> x = y"
by (transfer, simp add: comparator.eq)
lemma key_compare_same [simp]: "key_compare ko x x = Eq"
by (simp add: key_compare_Eq)
lemma uminus_key_compare [simp]: "invert_order (key_compare ko x y) = key_compare ko y x"
by (transfer, simp add: comparator.sym)
lemma key_compare_uminus [simp]: "key_compare (- ko) x y = key_compare ko y x"
by (transfer, rule refl)
lemma uminus_key_order_sameD:
assumes "- ko = (ko::'a key_order)"
shows "x = (y::'a)"
proof (rule ccontr)
assume "x \<noteq> y"
hence "key_compare ko x y \<noteq> Eq" by (simp add: key_compare_Eq)
hence "key_compare ko x y \<noteq> invert_order (key_compare ko x y)"
by (metis invert_order.elims order.distinct(5))
also have "invert_order (key_compare ko x y) = key_compare (- ko) x y" by simp
finally have "- ko \<noteq> ko" by (auto simp del: key_compare_uminus)
thus False using assms ..
qed
lemma key_compare_key_order_of_ord:
assumes "antisymp ord" and "transp ord" and "\<And>x y. ord x y \<or> ord y x"
shows "key_compare (key_order_of_ord ord) = (\<lambda>x y. if ord x y then if x = y then Eq else Lt else Gt)"
proof -
have eq: "key_compare (key_order_of_ord ord) = comp_of_ord ord"
unfolding key_order_of_ord_def comp_of_ord_eq_comp_of_ords[OF assms(1)]
proof (rule Abs_key_order_inverse, simp, rule comp_of_ords, unfold_locales)
fix x
from assms(3) show "ord x x" by blast
next
fix x y z
assume "ord x y" and "ord y z"
with assms(2) show "ord x z" by (rule transpD)
next
fix x y
assume "ord x y" and "ord y x"
with assms(1) show "x = y" by (rule antisympD)
qed (rule refl, rule assms(3))
have *: "x = y" if "ord x y" and "ord y x" for x y using assms(1) that by (rule antisympD)
show ?thesis by (rule, rule, auto simp: eq comp_of_ord_def intro: *)
qed
lemma key_compare_key_order_of_le:
"key_compare key_order_of_le = (\<lambda>x y. if x < y then Lt else if x = y then Eq else Gt)"
by (transfer, intro ext, fact comparator_of_def)
subsection \<open>Invariant in Context @{locale comparator}\<close>
context comparator
begin
definition oalist_inv_raw :: "('a \<times> 'b::zero) list \<Rightarrow> bool"
where "oalist_inv_raw xs \<longleftrightarrow> (0 \<notin> snd ` set xs \<and> sorted_wrt lt (map fst xs))"
lemma oalist_inv_rawI:
assumes "0 \<notin> snd ` set xs" and "sorted_wrt lt (map fst xs)"
shows "oalist_inv_raw xs"
unfolding oalist_inv_raw_def using assms unfolding fst_conv snd_conv by blast
lemma oalist_inv_rawD1:
assumes "oalist_inv_raw xs"
shows "0 \<notin> snd ` set xs"
using assms unfolding oalist_inv_raw_def fst_conv by blast
lemma oalist_inv_rawD2:
assumes "oalist_inv_raw xs"
shows "sorted_wrt lt (map fst xs)"
using assms unfolding oalist_inv_raw_def fst_conv snd_conv by blast
lemma oalist_inv_raw_Nil: "oalist_inv_raw []"
by (simp add: oalist_inv_raw_def)
lemma oalist_inv_raw_singleton: "oalist_inv_raw [(k, v)] \<longleftrightarrow> (v \<noteq> 0)"
by (auto simp: oalist_inv_raw_def)
lemma oalist_inv_raw_ConsI:
assumes "oalist_inv_raw xs" and "v \<noteq> 0" and "xs \<noteq> [] \<Longrightarrow> lt k (fst (hd xs))"
shows "oalist_inv_raw ((k, v) # xs)"
proof (rule oalist_inv_rawI)
from assms(1) have "0 \<notin> snd ` set xs" by (rule oalist_inv_rawD1)
with assms(2) show "0 \<notin> snd ` set ((k, v) # xs)" by simp
next
show "sorted_wrt lt (map fst ((k, v) # xs))"
proof (cases "xs = []")
case True
thus ?thesis by simp
next
case False
then obtain k' v' xs' where xs: "xs = (k', v') # xs'" by (metis list.exhaust prod.exhaust)
from assms(3)[OF False] have "lt k k'" by (simp add: xs)
moreover from assms(1) have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2)
ultimately show "sorted_wrt lt (map fst ((k, v) # xs))"
by (simp add: xs sorted_wrt2[OF transp_less] del: sorted_wrt.simps)
qed
qed
lemma oalist_inv_raw_ConsD1:
assumes "oalist_inv_raw (x # xs)"
shows "oalist_inv_raw xs"
proof (rule oalist_inv_rawI)
from assms have "0 \<notin> snd ` set (x # xs)" by (rule oalist_inv_rawD1)
thus "0 \<notin> snd ` set xs" by simp
next
from assms have "sorted_wrt lt (map fst (x # xs))" by (rule oalist_inv_rawD2)
thus "sorted_wrt lt (map fst xs)" by simp
qed
lemma oalist_inv_raw_ConsD2:
assumes "oalist_inv_raw ((k, v) # xs)"
shows "v \<noteq> 0"
proof -
from assms have "0 \<notin> snd ` set ((k, v) # xs)" by (rule oalist_inv_rawD1)
thus ?thesis by auto
qed
lemma oalist_inv_raw_ConsD3:
assumes "oalist_inv_raw ((k, v) # xs)" and "k' \<in> fst ` set xs"
shows "lt k k'"
proof -
from assms(2) obtain x where "x \<in> set xs" and "k' = fst x" by fastforce
from assms(1) have "sorted_wrt lt (map fst ((k, v) # xs))" by (rule oalist_inv_rawD2)
hence "\<forall>x\<in>set xs. lt k (fst x)" by simp
hence "lt k (fst x)" using \<open>x \<in> set xs\<close> ..
thus ?thesis by (simp only: \<open>k' = fst x\<close>)
qed
lemma oalist_inv_raw_tl:
assumes "oalist_inv_raw xs"
shows "oalist_inv_raw (tl xs)"
proof (rule oalist_inv_rawI)
from assms have "0 \<notin> snd ` set xs" by (rule oalist_inv_rawD1)
thus "0 \<notin> snd ` set (tl xs)" by (metis (no_types, lifting) image_iff list.set_sel(2) tl_Nil)
next
show "sorted_wrt lt (map fst (tl xs))"
by (metis hd_Cons_tl oalist_inv_rawD2 oalist_inv_raw_ConsD1 assms tl_Nil)
qed
lemma oalist_inv_raw_filter:
assumes "oalist_inv_raw xs"
shows "oalist_inv_raw (filter P xs)"
proof (rule oalist_inv_rawI)
from assms have "0 \<notin> snd ` set xs" by (rule oalist_inv_rawD1)
thus "0 \<notin> snd ` set (filter P xs)" by auto
next
from assms have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2)
thus "sorted_wrt lt (map fst (filter P xs))" by (induct xs, simp, simp)
qed
lemma oalist_inv_raw_map:
assumes "oalist_inv_raw xs"
and "\<And>a. snd (f a) = 0 \<Longrightarrow> snd a = 0"
and "\<And>a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)"
shows "oalist_inv_raw (map f xs)"
proof (rule oalist_inv_rawI)
show "0 \<notin> snd ` set (map f xs)"
proof (simp, rule)
assume "0 \<in> snd ` f ` set xs"
then obtain a where "a \<in> set xs" and "snd (f a) = 0" by fastforce
from this(2) have "snd a = 0" by (rule assms(2))
from assms(1) have "0 \<notin> snd ` set xs" by (rule oalist_inv_rawD1)
moreover from \<open>a \<in> set xs\<close> have "0 \<in> snd ` set xs" by (simp add: \<open>snd a = 0\<close>[symmetric])
ultimately show False ..
qed
next
from assms(1) have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2)
hence "sorted_wrt (\<lambda>x y. comp (fst x) (fst y) = Lt) xs"
by (simp only: sorted_wrt_map Lt_lt_conv)
thus "sorted_wrt lt (map fst (map f xs))"
by (simp add: sorted_wrt_map Lt_lt_conv[symmetric] assms(3))
qed
lemma oalist_inv_raw_induct [consumes 1, case_names Nil Cons]:
assumes "oalist_inv_raw xs"
assumes "P []"
assumes "\<And>k v xs. oalist_inv_raw ((k, v) # xs) \<Longrightarrow> oalist_inv_raw xs \<Longrightarrow> v \<noteq> 0 \<Longrightarrow>
(\<And>k'. k' \<in> fst ` set xs \<Longrightarrow> lt k k') \<Longrightarrow> P xs \<Longrightarrow> P ((k, v) # xs)"
shows "P xs"
using assms(1)
proof (induct xs)
case Nil
from assms(2) show ?case .
next
case (Cons x xs)
obtain k v where x: "x = (k, v)" by fastforce
from Cons(2) have "oalist_inv_raw ((k, v) # xs)" and "oalist_inv_raw xs" and "v \<noteq> 0" unfolding x
by (this, rule oalist_inv_raw_ConsD1, rule oalist_inv_raw_ConsD2)
moreover from Cons(2) have "lt k k'" if "k' \<in> fst ` set xs" for k' using that
unfolding x by (rule oalist_inv_raw_ConsD3)
moreover from \<open>oalist_inv_raw xs\<close> have "P xs" by (rule Cons(1))
ultimately show ?case unfolding x by (rule assms(3))
qed
subsection \<open>Operations on Lists of Pairs in Context @{locale comparator}\<close>
type_synonym (in -) ('a, 'b) comp_opt = "'a \<Rightarrow> 'b \<Rightarrow> (order option)"
definition (in -) lookup_dflt :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b::zero"
where "lookup_dflt xs k = (case map_of xs k of Some v \<Rightarrow> v | None \<Rightarrow> 0)"
text \<open>@{const lookup_dflt} is only an auxiliary function needed for proving some lemmas.\<close>
fun lookup_pair :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b::zero"
where
"lookup_pair [] x = 0"|
"lookup_pair ((k, v) # xs) x =
(case comp x k of
Lt \<Rightarrow> 0
| Eq \<Rightarrow> v
| Gt \<Rightarrow> lookup_pair xs x)"
fun update_by_pair :: "('a \<times> 'b) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b::zero) list"
where
"update_by_pair (k, v) [] = (if v = 0 then [] else [(k, v)])"
| "update_by_pair (k, v) ((k', v') # xs) =
(case comp k k' of Lt \<Rightarrow> (if v = 0 then (k', v') # xs else (k, v) # (k', v') # xs)
| Eq \<Rightarrow> (if v = 0 then xs else (k, v) # xs)
| Gt \<Rightarrow> (k', v') # update_by_pair (k, v) xs)"
(* TODO: Add update_by_gr_pair. *)
definition sort_oalist :: "('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b::zero) list"
where "sort_oalist xs = foldr update_by_pair xs []"
fun update_by_fun_pair :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b::zero) list"
where
"update_by_fun_pair k f [] = (let v = f 0 in if v = 0 then [] else [(k, v)])"
| "update_by_fun_pair k f ((k', v') # xs) =
(case comp k k' of Lt \<Rightarrow> (let v = f 0 in if v = 0 then (k', v') # xs else (k, v) # (k', v') # xs)
| Eq \<Rightarrow> (let v = f v' in if v = 0 then xs else (k, v) # xs)
| Gt \<Rightarrow> (k', v') # update_by_fun_pair k f xs)"
definition update_by_fun_gr_pair :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'b::zero) list"
where "update_by_fun_gr_pair k f xs =
(if xs = [] then
(let v = f 0 in if v = 0 then [] else [(k, v)])
else if comp k (fst (last xs)) = Gt then
(let v = f 0 in if v = 0 then xs else xs @ [(k, v)])
else
update_by_fun_pair k f xs
)"
fun (in -) map_pair :: "(('a \<times> 'b) \<Rightarrow> ('a \<times> 'c)) \<Rightarrow> ('a \<times> 'b::zero) list \<Rightarrow> ('a \<times> 'c::zero) list"
where
"map_pair f [] = []"
| "map_pair f (kv # xs) =
(let (k, v) = f kv; aux = map_pair f xs in if v = 0 then aux else (k, v) # aux)"
text \<open>The difference between @{const List.map} and @{const map_pair} is that the latter removes
@{term 0} values, whereas the former does not.\<close>
abbreviation (in -) map_val_pair :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b::zero) list \<Rightarrow> ('a \<times> 'c::zero) list"
where "map_val_pair f \<equiv> map_pair (\<lambda>(k, v). (k, f k v))"
fun map2_val_pair :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> (('a \<times> 'b) list \<Rightarrow> ('a \<times> 'd) list) \<Rightarrow>
(('a \<times> 'c) list \<Rightarrow> ('a \<times> 'd) list) \<Rightarrow>
('a \<times> 'b::zero) list \<Rightarrow> ('a \<times> 'c::zero) list \<Rightarrow> ('a \<times> 'd::zero) list"
where
"map2_val_pair f g h xs [] = g xs"
| "map2_val_pair f g h [] ys = h ys"
| "map2_val_pair f g h ((kx, vx) # xs) ((ky, vy) # ys) =
(case comp kx ky of
Lt \<Rightarrow> (let v = f kx vx 0; aux = map2_val_pair f g h xs ((ky, vy) # ys) in if v = 0 then aux else (kx, v) # aux)
| Eq \<Rightarrow> (let v = f kx vx vy; aux = map2_val_pair f g h xs ys in if v = 0 then aux else (kx, v) # aux)
| Gt \<Rightarrow> (let v = f ky 0 vy; aux = map2_val_pair f g h ((kx, vx) # xs) ys in if v = 0 then aux else (ky, v) # aux))"
fun lex_ord_pair :: "('a \<Rightarrow> (('b, 'c) comp_opt)) \<Rightarrow> (('a \<times> 'b::zero) list, ('a \<times> 'c::zero) list) comp_opt"
where
"lex_ord_pair f [] [] = Some Eq"|
"lex_ord_pair f [] ((ky, vy) # ys) =
(let aux = f ky 0 vy in if aux = Some Eq then lex_ord_pair f [] ys else aux)"|
"lex_ord_pair f ((kx, vx) # xs) [] =
(let aux = f kx vx 0 in if aux = Some Eq then lex_ord_pair f xs [] else aux)"|
"lex_ord_pair f ((kx, vx) # xs) ((ky, vy) # ys) =
(case comp kx ky of
Lt \<Rightarrow> (let aux = f kx vx 0 in if aux = Some Eq then lex_ord_pair f xs ((ky, vy) # ys) else aux)
| Eq \<Rightarrow> (let aux = f kx vx vy in if aux = Some Eq then lex_ord_pair f xs ys else aux)
| Gt \<Rightarrow> (let aux = f ky 0 vy in if aux = Some Eq then lex_ord_pair f ((kx, vx) # xs) ys else aux))"
fun prod_ord_pair :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'b::zero) list \<Rightarrow> ('a \<times> 'c::zero) list \<Rightarrow> bool"
where
"prod_ord_pair f [] [] = True"|
"prod_ord_pair f [] ((ky, vy) # ys) = (f ky 0 vy \<and> prod_ord_pair f [] ys)"|
"prod_ord_pair f ((kx, vx) # xs) [] = (f kx vx 0 \<and> prod_ord_pair f xs [])"|
"prod_ord_pair f ((kx, vx) # xs) ((ky, vy) # ys) =
(case comp kx ky of
Lt \<Rightarrow> (f kx vx 0 \<and> prod_ord_pair f xs ((ky, vy) # ys))
| Eq \<Rightarrow> (f kx vx vy \<and> prod_ord_pair f xs ys)
| Gt \<Rightarrow> (f ky 0 vy \<and> prod_ord_pair f ((kx, vx) # xs) ys))"
text \<open>@{const prod_ord_pair} is actually just a special case of @{const lex_ord_pair}, as proved below
in lemma \<open>prod_ord_pair_eq_lex_ord_pair\<close>.\<close>
subsubsection \<open>@{const lookup_pair}\<close>
lemma lookup_pair_eq_0:
assumes "oalist_inv_raw xs"
shows "lookup_pair xs k = 0 \<longleftrightarrow> (k \<notin> fst ` set xs)"
using assms
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k' v' xs)
show ?case
proof (simp add: Cons(3) eq split: order.splits, rule, simp_all only: atomize_imp[symmetric])
assume "comp k k' = Lt"
hence "k \<noteq> k'" by auto
moreover have "k \<notin> fst ` set xs"
proof
assume "k \<in> fst ` set xs"
hence "lt k' k" by (rule Cons(4))
with \<open>comp k k' = Lt\<close> show False by (simp add: Lt_lt_conv)
qed
ultimately show "k \<noteq> k' \<and> k \<notin> fst ` set xs" ..
next
assume "comp k k' = Gt"
hence "k \<noteq> k'" by auto
thus "(lookup_pair xs k = 0) = (k \<noteq> k' \<and> k \<notin> fst ` set xs)" by (simp add: Cons(5))
qed
qed
lemma lookup_pair_eq_value:
assumes "oalist_inv_raw xs" and "v \<noteq> 0"
shows "lookup_pair xs k = v \<longleftrightarrow> ((k, v) \<in> set xs)"
using assms(1)
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
from assms(2) show ?case by simp
next
case (Cons k' v' xs)
have *: "(k', u) \<notin> set xs" for u
proof
assume "(k', u) \<in> set xs"
hence "fst (k', u) \<in> fst ` set xs" by fastforce
hence "k' \<in> fst ` set xs" by simp
hence "lt k' k'" by (rule Cons(4))
thus False by (simp add: lt_of_key_order_alt[symmetric])
qed
show ?case
proof (simp add: assms(2) Cons(5) eq split: order.split, intro conjI impI)
assume "comp k k' = Lt"
show "(k, v) \<notin> set xs"
proof
assume "(k, v) \<in> set xs"
hence "fst (k, v) \<in> fst ` set xs" by fastforce
hence "k \<in> fst ` set xs" by simp
hence "lt k' k" by (rule Cons(4))
with \<open>comp k k' = Lt\<close> show False by (simp add: Lt_lt_conv)
qed
qed (auto simp: *)
qed
lemma lookup_pair_eq_valueI:
assumes "oalist_inv_raw xs" and "(k, v) \<in> set xs"
shows "lookup_pair xs k = v"
proof -
from assms(2) have "v \<in> snd ` set xs" by force
moreover from assms(1) have "0 \<notin> snd ` set xs" by (rule oalist_inv_rawD1)
ultimately have "v \<noteq> 0" by blast
with assms show ?thesis by (simp add: lookup_pair_eq_value)
qed
lemma lookup_dflt_eq_lookup_pair:
assumes "oalist_inv_raw xs"
shows "lookup_dflt xs = lookup_pair xs"
proof (rule, simp add: lookup_dflt_def split: option.split, intro conjI impI allI)
fix k
assume "map_of xs k = None"
with assms show "lookup_pair xs k = 0" by (simp add: lookup_pair_eq_0 map_of_eq_None_iff)
next
fix k v
assume "map_of xs k = Some v"
hence "(k, v) \<in> set xs" by (rule map_of_SomeD)
with assms have "lookup_pair xs k = v" by (rule lookup_pair_eq_valueI)
thus "v = lookup_pair xs k" by (rule HOL.sym)
qed
lemma lookup_pair_inj:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "lookup_pair xs = lookup_pair ys"
shows "xs = ys"
using assms
proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct)
case Nil
thus ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k' v' ys)
have "v' = lookup_pair ((k', v') # ys) k'" by simp
also have "... = lookup_pair [] k'" by (simp only: Cons(6))
also have "... = 0" by simp
finally have "v' = 0" .
with Cons(3) show ?case ..
qed
next
case *: (Cons k v xs)
from *(6, 7) show ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
have "v = lookup_pair ((k, v) # xs) k" by simp
also have "... = lookup_pair [] k" by (simp only: Nil)
also have "... = 0" by simp
finally have "v = 0" .
with *(3) show ?case ..
next
case (Cons k' v' ys)
show ?case
proof (cases "comp k k'")
case Lt
hence "\<not> lt k' k" by (simp add: Lt_lt_conv)
with Cons(4) have "k \<notin> fst ` set ys" by blast
moreover from Lt have "k \<noteq> k'" by auto
ultimately have "k \<notin> fst ` set ((k', v') # ys)" by simp
hence "0 = lookup_pair ((k', v') # ys) k"
by (simp add: lookup_pair_eq_0[OF Cons(1)] del: lookup_pair.simps)
also have "... = lookup_pair ((k, v) # xs) k" by (simp only: Cons(6))
also have "... = v" by simp
finally have "v = 0" by simp
with *(3) show ?thesis ..
next
case Eq
hence "k' = k" by (simp only: eq)
have "v' = lookup_pair ((k', v') # ys) k'" by simp
also have "... = lookup_pair ((k, v) # xs) k" by (simp only: Cons(6) \<open>k' = k\<close>)
also have "... = v" by simp
finally have "v' = v" .
moreover note \<open>k' = k\<close>
moreover from Cons(2) have "xs = ys"
proof (rule *(5))
show "lookup_pair xs = lookup_pair ys"
proof
fix k0
show "lookup_pair xs k0 = lookup_pair ys k0"
proof (cases "lt k k0")
case True
hence eq: "comp k0 k = Gt"
by (simp add: Gt_lt_conv)
have "lookup_pair xs k0 = lookup_pair ((k, v) # xs) k0" by (simp add: eq)
also have "... = lookup_pair ((k, v') # ys) k0" by (simp only: Cons(6) \<open>k' = k\<close>)
also have "... = lookup_pair ys k0" by (simp add: eq)
finally show ?thesis .
next
case False
with *(4) have "k0 \<notin> fst ` set xs" by blast
with *(2) have eq: "lookup_pair xs k0 = 0" by (simp add: lookup_pair_eq_0)
from False Cons(4) have "k0 \<notin> fst ` set ys" unfolding \<open>k' = k\<close> by blast
with Cons(2) have "lookup_pair ys k0 = 0" by (simp add: lookup_pair_eq_0)
with eq show ?thesis by simp
qed
qed
qed
ultimately show ?thesis by simp
next
case Gt
hence "\<not> lt k k'" by (simp add: Gt_lt_conv)
with *(4) have "k' \<notin> fst ` set xs" by blast
moreover from Gt have "k' \<noteq> k" by auto
ultimately have "k' \<notin> fst ` set ((k, v) # xs)" by simp
hence "0 = lookup_pair ((k, v) # xs) k'"
by (simp add: lookup_pair_eq_0[OF *(1)] del: lookup_pair.simps)
also have "... = lookup_pair ((k', v') # ys) k'" by (simp only: Cons(6))
also have "... = v'" by simp
finally have "v' = 0" by simp
with Cons(3) show ?thesis ..
qed
qed
qed
lemma lookup_pair_tl:
assumes "oalist_inv_raw xs"
shows "lookup_pair (tl xs) k = (if (\<forall>k'\<in>fst ` set xs. le k k') then 0 else lookup_pair xs k)"
proof -
from assms have 1: "oalist_inv_raw (tl xs)" by (rule oalist_inv_raw_tl)
show ?thesis
proof (split if_split, intro conjI impI)
assume *: "\<forall>x\<in>fst ` set xs. le k x"
show "lookup_pair (tl xs) k = 0"
proof (simp add: lookup_pair_eq_0[OF 1], rule)
assume k_in: "k \<in> fst ` set (tl xs)"
hence "xs \<noteq> []" by auto
then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis
have "k' \<in> fst ` set xs" unfolding xs by fastforce
with * have "le k k'" ..
from assms have "oalist_inv_raw ((k', v') # ys)" by (simp only: xs)
moreover from k_in have "k \<in> fst ` set ys" by (simp add: xs)
ultimately have "lt k' k" by (rule oalist_inv_raw_ConsD3)
with \<open>le k k'\<close> show False by simp
qed
next
assume "\<not> (\<forall>k'\<in>fst ` set xs. le k k')"
hence "\<exists>x\<in>fst ` set xs. \<not> le k x" by simp
then obtain k'' where k''_in: "k'' \<in> fst ` set xs" and "\<not> le k k''" ..
from this(2) have "lt k'' k" by simp
from k''_in have "xs \<noteq> []" by auto
then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis
from k''_in have "k'' = k' \<or> k'' \<in> fst ` set ys" by (simp add: xs)
hence "lt k' k"
proof
assume "k'' = k'"
with \<open>lt k'' k\<close> show ?thesis by simp
next
from assms have "oalist_inv_raw ((k', v') # ys)" by (simp only: xs)
moreover assume "k'' \<in> fst ` set ys"
ultimately have "lt k' k''" by (rule oalist_inv_raw_ConsD3)
thus ?thesis using \<open>lt k'' k\<close> by (rule less_trans)
qed
hence "comp k k' = Gt" by (simp add: Gt_lt_conv)
thus "lookup_pair (tl xs) k = lookup_pair xs k" by (simp add: xs lt_of_key_order_alt)
qed
qed
lemma lookup_pair_tl':
assumes "oalist_inv_raw xs"
shows "lookup_pair (tl xs) k = (if k = fst (hd xs) then 0 else lookup_pair xs k)"
proof -
from assms have 1: "oalist_inv_raw (tl xs)" by (rule oalist_inv_raw_tl)
show ?thesis
proof (split if_split, intro conjI impI)
assume k: "k = fst (hd xs)"
show "lookup_pair (tl xs) k = 0"
proof (simp add: lookup_pair_eq_0[OF 1], rule)
assume k_in: "k \<in> fst ` set (tl xs)"
hence "xs \<noteq> []" by auto
then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis
from assms have "oalist_inv_raw ((k', v') # ys)" by (simp only: xs)
moreover from k_in have "k' \<in> fst ` set ys" by (simp add: k xs)
ultimately have "lt k' k'" by (rule oalist_inv_raw_ConsD3)
thus False by simp
qed
next
assume "k \<noteq> fst (hd xs)"
show "lookup_pair (tl xs) k = lookup_pair xs k"
proof (cases "xs = []")
case True
show ?thesis by (simp add: True)
next
case False
then obtain k' v' ys where xs: "xs = (k', v') # ys" using prod.exhaust list.exhaust by metis
show ?thesis
proof (simp add: xs eq Lt_lt_conv split: order.split, intro conjI impI)
from \<open>k \<noteq> fst (hd xs)\<close> have "k \<noteq> k'" by (simp add: xs)
moreover assume "k = k'"
ultimately show "lookup_pair ys k' = v'" ..
next
assume "lt k k'"
from assms have "oalist_inv_raw ys" unfolding xs by (rule oalist_inv_raw_ConsD1)
moreover have "k \<notin> fst ` set ys"
proof
assume "k \<in> fst ` set ys"
with assms have "lt k' k" unfolding xs by (rule oalist_inv_raw_ConsD3)
with \<open>lt k k'\<close> show False by simp
qed
ultimately show "lookup_pair ys k = 0" by (simp add: lookup_pair_eq_0)
qed
qed
qed
qed
lemma lookup_pair_filter:
assumes "oalist_inv_raw xs"
shows "lookup_pair (filter P xs) k = (let v = lookup_pair xs k in if P (k, v) then v else 0)"
using assms
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k' v' xs)
show ?case
proof (simp add: Cons(5) Let_def eq split: order.split, intro conjI impI)
show "lookup_pair xs k' = 0"
proof (simp add: lookup_pair_eq_0 Cons(2), rule)
assume "k' \<in> fst ` set xs"
hence "lt k' k'" by (rule Cons(4))
thus False by simp
qed
next
assume "comp k k' = Lt"
hence "lt k k'" by (simp only: Lt_lt_conv)
show "lookup_pair xs k = 0"
proof (simp add: lookup_pair_eq_0 Cons(2), rule)
assume "k \<in> fst ` set xs"
hence "lt k' k" by (rule Cons(4))
with \<open>lt k k'\<close> show False by simp
qed
qed
qed
lemma lookup_pair_map:
assumes "oalist_inv_raw xs"
and "\<And>k'. snd (f (k', 0)) = 0"
and "\<And>a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)"
shows "lookup_pair (map f xs) (fst (f (k, v))) = snd (f (k, lookup_pair xs k))"
using assms(1)
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by (simp add: assms(2))
next
case (Cons k' v' xs)
obtain k'' v'' where f: "f (k', v') = (k'', v'')" by fastforce
have "comp k k' = comp (fst (f (k, v))) (fst (f (k', v')))"
by (simp add: assms(3))
also have "... = comp (fst (f (k, v))) k''" by (simp add: f)
finally have eq0: "comp k k' = comp (fst (f (k, v))) k''" .
show ?case
proof (simp add: assms(2) split: order.split, intro conjI impI, simp add: eq)
assume "k = k'"
hence "lookup_pair (f (k', v') # map f xs) (fst (f (k', v))) =
lookup_pair (f (k', v') # map f xs) (fst (f (k, v)))" by simp
also have "... = snd (f (k', v'))" by (simp add: f eq0[symmetric], simp add: \<open>k = k'\<close>)
finally show "lookup_pair (f (k', v') # map f xs) (fst (f (k', v))) = snd (f (k', v'))" .
qed (simp_all add: f eq0 Cons(5))
qed
lemma lookup_pair_Cons:
assumes "oalist_inv_raw ((k, v) # xs)"
shows "lookup_pair ((k, v) # xs) k0 = (if k = k0 then v else lookup_pair xs k0)"
proof (simp add: eq split: order.split, intro impI)
assume "comp k0 k = Lt"
from assms have inv: "oalist_inv_raw xs" by (rule oalist_inv_raw_ConsD1)
show "lookup_pair xs k0 = 0"
proof (simp only: lookup_pair_eq_0[OF inv], rule)
assume "k0 \<in> fst ` set xs"
with assms have "lt k k0" by (rule oalist_inv_raw_ConsD3)
with \<open>comp k0 k = Lt\<close> show False by (simp add: Lt_lt_conv)
qed
qed
lemma lookup_pair_single: "lookup_pair [(k, v)] k0 = (if k = k0 then v else 0)"
by (simp add: eq split: order.split)
subsubsection \<open>@{const update_by_pair}\<close>
lemma set_update_by_pair_subset: "set (update_by_pair kv xs) \<subseteq> insert kv (set xs)"
proof (induct xs arbitrary: kv)
case Nil
obtain k v where kv: "kv = (k, v)" by fastforce
thus ?case by simp
next
case (Cons x xs)
obtain k' v' where x: "x = (k', v')" by fastforce
obtain k v where kv: "kv = (k, v)" by fastforce
have 1: "set xs \<subseteq> insert a (insert b (set xs))" for a b by auto
have 2: "set (update_by_pair kv xs) \<subseteq> insert kv (insert (k', v') (set xs))" for kv
using Cons by blast
show ?case by (simp add: x kv 1 2 split: order.split)
qed
lemma update_by_pair_sorted:
assumes "sorted_wrt lt (map fst xs)"
shows "sorted_wrt lt (map fst (update_by_pair kv xs))"
using assms
proof (induct xs arbitrary: kv)
case Nil
obtain k v where kv: "kv = (k, v)" by fastforce
thus ?case by simp
next
case (Cons x xs)
obtain k' v' where x: "x = (k', v')" by fastforce
obtain k v where kv: "kv = (k, v)" by fastforce
from Cons(2) have 1: "sorted_wrt lt (k' # (map fst xs))" by (simp add: x)
hence 2: "sorted_wrt lt (map fst xs)" using sorted_wrt.elims(3) by fastforce
hence 3: "sorted_wrt lt (map fst (update_by_pair (k, u) xs))" for u by (rule Cons(1))
have 4: "sorted_wrt lt (k' # map fst (update_by_pair (k, u) xs))"
if *: "comp k k' = Gt" for u
proof (simp, intro conjI ballI)
fix y
assume "y \<in> set (update_by_pair (k, u) xs)"
also from set_update_by_pair_subset have "... \<subseteq> insert (k, u) (set xs)" .
finally have "y = (k, u) \<or> y \<in> set xs" by simp
thus "lt k' (fst y)"
proof
assume "y = (k, u)"
hence "fst y = k" by simp
with * show ?thesis by (simp only: Gt_lt_conv)
next
from 1 have 5: "\<forall>y \<in> fst ` set xs. lt k' y" by simp
assume "y \<in> set xs"
hence "fst y \<in> fst ` set xs" by simp
with 5 show ?thesis ..
qed
qed (fact 3)
show ?case
by (simp add: kv x 1 2 4 sorted_wrt2 split: order.split del: sorted_wrt.simps,
intro conjI impI, simp add: 1 eq del: sorted_wrt.simps, simp add: Lt_lt_conv)
qed
lemma update_by_pair_not_0:
assumes "0 \<notin> snd ` set xs"
shows "0 \<notin> snd ` set (update_by_pair kv xs)"
using assms
proof (induct xs arbitrary: kv)
case Nil
obtain k v where kv: "kv = (k, v)" by fastforce
thus ?case by simp
next
case (Cons x xs)
obtain k' v' where x: "x = (k', v')" by fastforce
obtain k v where kv: "kv = (k, v)" by fastforce
from Cons(2) have 1: "v' \<noteq> 0" and 2: "0 \<notin> snd ` set xs" by (auto simp: x)
from 2 have 3: "0 \<notin> snd ` set (update_by_pair (k, u) xs)" for u by (rule Cons(1))
show ?case by (auto simp: kv x 1 2 3 split: order.split)
qed
corollary oalist_inv_raw_update_by_pair:
assumes "oalist_inv_raw xs"
shows "oalist_inv_raw (update_by_pair kv xs)"
proof (rule oalist_inv_rawI)
from assms have "0 \<notin> snd ` set xs" by (rule oalist_inv_rawD1)
thus "0 \<notin> snd ` set (update_by_pair kv xs)" by (rule update_by_pair_not_0)
next
from assms have "sorted_wrt lt (map fst xs)" by (rule oalist_inv_rawD2)
thus "sorted_wrt lt (map fst (update_by_pair kv xs))" by (rule update_by_pair_sorted)
qed
lemma update_by_pair_less:
assumes "v \<noteq> 0" and "xs = [] \<or> comp k (fst (hd xs)) = Lt"
shows "update_by_pair (k, v) xs = (k, v) # xs"
using assms(2)
proof (induct xs)
case Nil
from assms(1) show ?case by simp
next
case (Cons x xs)
obtain k' v' where x: "x = (k', v')" by fastforce
from Cons(2) have "comp k k' = Lt" by (simp add: x)
with assms(1) show ?case by (simp add: x)
qed
lemma lookup_pair_update_by_pair:
assumes "oalist_inv_raw xs"
shows "lookup_pair (update_by_pair (k1, v) xs) k2 = (if k1 = k2 then v else lookup_pair xs k2)"
using assms
proof (induct xs arbitrary: v rule: oalist_inv_raw_induct)
case Nil
show ?case by (simp split: order.split, simp add: eq)
next
case (Cons k' v' xs)
show ?case
proof (split if_split, intro conjI impI)
assume "k1 = k2"
with Cons(5) have eq0: "lookup_pair (update_by_pair (k2, u) xs) k2 = u" for u
by (simp del: update_by_pair.simps)
show "lookup_pair (update_by_pair (k1, v) ((k', v') # xs)) k2 = v"
proof (simp add: \<open>k1 = k2\<close> eq0 split: order.split, intro conjI impI)
assume "comp k2 k' = Eq"
hence "\<not> lt k' k2" by (simp add: eq)
with Cons(4) have "k2 \<notin> fst ` set xs" by auto
thus "lookup_pair xs k2 = 0" using Cons(2) by (simp add: lookup_pair_eq_0)
qed
next
assume "k1 \<noteq> k2"
with Cons(5) have eq0: "lookup_pair (update_by_pair (k1, u) xs) k2 = lookup_pair xs k2" for u
by (simp del: update_by_pair.simps)
have *: "lookup_pair xs k2 = 0" if "lt k2 k'"
proof -
from \<open>lt k2 k'\<close> have "\<not> lt k' k2" by auto
with Cons(4) have "k2 \<notin> fst ` set xs" by auto
thus "lookup_pair xs k2 = 0" using Cons(2) by (simp add: lookup_pair_eq_0)
qed
show "lookup_pair (update_by_pair (k1, v) ((k', v') # xs)) k2 = lookup_pair ((k', v') # xs) k2"
by (simp add: \<open>k1 \<noteq> k2\<close> eq0 split: order.split,
auto intro: * simp: \<open>k1 \<noteq> k2\<close>[symmetric] eq Gt_lt_conv Lt_lt_conv)
qed
qed
corollary update_by_pair_id:
assumes "oalist_inv_raw xs" and "lookup_pair xs k = v"
shows "update_by_pair (k, v) xs = xs"
proof (rule lookup_pair_inj, rule oalist_inv_raw_update_by_pair)
show "lookup_pair (update_by_pair (k, v) xs) = lookup_pair xs"
proof
fix k0
from assms(2) show "lookup_pair (update_by_pair (k, v) xs) k0 = lookup_pair xs k0"
by (auto simp: lookup_pair_update_by_pair[OF assms(1)])
qed
qed fact+
lemma set_update_by_pair:
assumes "oalist_inv_raw xs" and "v \<noteq> 0"
shows "set (update_by_pair (k, v) xs) = insert (k, v) (set xs - range (Pair k))" (is "?A = ?B")
proof (rule set_eqI)
fix x::"'a \<times> 'b"
obtain k' v' where x: "x = (k', v')" by fastforce
from assms(1) have inv: "oalist_inv_raw (update_by_pair (k, v) xs)"
by (rule oalist_inv_raw_update_by_pair)
show "(x \<in> ?A) \<longleftrightarrow> (x \<in> ?B)"
proof (cases "v' = 0")
case True
have "0 \<notin> snd ` set (update_by_pair (k, v) xs)" and "0 \<notin> snd ` set xs"
by (rule oalist_inv_rawD1, fact)+
hence "(k', 0) \<notin> set (update_by_pair (k, v) xs)" and "(k', 0) \<notin> set xs"
using image_iff by fastforce+
thus ?thesis by (simp add: x True assms(2))
next
case False
show ?thesis
by (auto simp: x lookup_pair_eq_value[OF inv False, symmetric] lookup_pair_eq_value[OF assms(1) False]
lookup_pair_update_by_pair[OF assms(1)])
qed
qed
lemma set_update_by_pair_zero:
assumes "oalist_inv_raw xs"
shows "set (update_by_pair (k, 0) xs) = set xs - range (Pair k)" (is "?A = ?B")
proof (rule set_eqI)
fix x::"'a \<times> 'b"
obtain k' v' where x: "x = (k', v')" by fastforce
from assms(1) have inv: "oalist_inv_raw (update_by_pair (k, 0) xs)"
by (rule oalist_inv_raw_update_by_pair)
show "(x \<in> ?A) \<longleftrightarrow> (x \<in> ?B)"
proof (cases "v' = 0")
case True
have "0 \<notin> snd ` set (update_by_pair (k, 0) xs)" and "0 \<notin> snd ` set xs"
by (rule oalist_inv_rawD1, fact)+
hence "(k', 0) \<notin> set (update_by_pair (k, 0) xs)" and "(k', 0) \<notin> set xs"
using image_iff by fastforce+
thus ?thesis by (simp add: x True)
next
case False
show ?thesis
by (auto simp: x lookup_pair_eq_value[OF inv False, symmetric] lookup_pair_eq_value[OF assms False]
lookup_pair_update_by_pair[OF assms] False)
qed
qed
subsubsection \<open>@{const update_by_fun_pair} and @{const update_by_fun_gr_pair}\<close>
lemma update_by_fun_pair_eq_update_by_pair:
assumes "oalist_inv_raw xs"
shows "update_by_fun_pair k f xs = update_by_pair (k, f (lookup_pair xs k)) xs"
using assms by (induct xs rule: oalist_inv_raw_induct, simp, simp split: order.split)
corollary oalist_inv_raw_update_by_fun_pair:
assumes "oalist_inv_raw xs"
shows "oalist_inv_raw (update_by_fun_pair k f xs)"
unfolding update_by_fun_pair_eq_update_by_pair[OF assms] using assms by (rule oalist_inv_raw_update_by_pair)
corollary lookup_pair_update_by_fun_pair:
assumes "oalist_inv_raw xs"
shows "lookup_pair (update_by_fun_pair k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_pair xs k2)"
by (simp add: update_by_fun_pair_eq_update_by_pair[OF assms] lookup_pair_update_by_pair[OF assms])
lemma update_by_fun_pair_gr:
assumes "oalist_inv_raw xs" and "xs = [] \<or> comp k (fst (last xs)) = Gt"
shows "update_by_fun_pair k f xs = xs @ (if f 0 = 0 then [] else [(k, f 0)])"
using assms
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k' v' xs)
from Cons(6) have 1: "comp k (fst (last ((k', v') # xs))) = Gt" by simp
have eq1: "comp k k' = Gt"
proof (cases "xs = []")
case True
with 1 show ?thesis by simp
next
case False
have "lt k' (fst (last xs))" by (rule Cons(4), simp add: False)
from False 1 have "comp k (fst (last xs)) = Gt" by simp
moreover from \<open>lt k' (fst (last xs))\<close> have "comp (fst (last xs)) k' = Gt"
by (simp add: Gt_lt_conv)
ultimately show ?thesis
by (meson Gt_lt_conv less_trans Lt_lt_conv[symmetric])
qed
have eq2: "update_by_fun_pair k f xs = xs @ (if f 0 = 0 then [] else [(k, f 0)])"
proof (rule Cons(5), simp only: disj_commute[of "xs = []"], rule disjCI)
assume "xs \<noteq> []"
with 1 show "comp k (fst (last xs)) = Gt" by simp
qed
show ?case by (simp split: order.split add: Let_def eq1 eq2)
qed
corollary update_by_fun_gr_pair_eq_update_by_fun_pair:
assumes "oalist_inv_raw xs"
shows "update_by_fun_gr_pair k f xs = update_by_fun_pair k f xs"
by (simp add: update_by_fun_gr_pair_def Let_def update_by_fun_pair_gr[OF assms] split: order.split)
corollary oalist_inv_raw_update_by_fun_gr_pair:
assumes "oalist_inv_raw xs"
shows "oalist_inv_raw (update_by_fun_gr_pair k f xs)"
unfolding update_by_fun_pair_eq_update_by_pair[OF assms] update_by_fun_gr_pair_eq_update_by_fun_pair[OF assms]
using assms by (rule oalist_inv_raw_update_by_pair)
corollary lookup_pair_update_by_fun_gr_pair:
assumes "oalist_inv_raw xs"
shows "lookup_pair (update_by_fun_gr_pair k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_pair xs k2)"
by (simp add: update_by_fun_pair_eq_update_by_pair[OF assms]
update_by_fun_gr_pair_eq_update_by_fun_pair[OF assms] lookup_pair_update_by_pair[OF assms])
subsubsection \<open>@{const map_pair}\<close>
lemma map_pair_cong:
assumes "\<And>kv. kv \<in> set xs \<Longrightarrow> f kv = g kv"
shows "map_pair f xs = map_pair g xs"
using assms
proof (induct xs)
case Nil
show ?case by simp
next
case (Cons x xs)
have "f x = g x" by (rule Cons(2), simp)
moreover have "map_pair f xs = map_pair g xs" by (rule Cons(1), rule Cons(2), simp)
ultimately show ?case by simp
qed
lemma map_pair_subset: "set (map_pair f xs) \<subseteq> f ` set xs"
proof (induct xs rule: map_pair.induct)
case (1 f)
show ?case by simp
next
case (2 f kv xs)
obtain k v where f: "f kv = (k, v)" by fastforce
- from f[symmetric] refl have *: "set (map_pair f xs) \<subseteq> f ` set xs" by (rule 2)
+ from f[symmetric] HOL.refl have *: "set (map_pair f xs) \<subseteq> f ` set xs"
+ by (rule 2)
show ?case by (simp add: f Let_def, intro conjI impI subset_insertI2 *)
qed
lemma oalist_inv_raw_map_pair:
assumes "oalist_inv_raw xs"
and "\<And>a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)"
shows "oalist_inv_raw (map_pair f xs)"
using assms(1)
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
from oalist_inv_raw_Nil show ?case by simp
next
case (Cons k v xs)
obtain k' v' where f: "f (k, v) = (k', v')" by fastforce
show ?case
proof (simp add: f Let_def Cons(5), rule)
assume "v' \<noteq> 0"
with Cons(5) show "oalist_inv_raw ((k', v') # map_pair f xs)"
proof (rule oalist_inv_raw_ConsI)
assume "map_pair f xs \<noteq> []"
hence "hd (map_pair f xs) \<in> set (map_pair f xs)" by simp
also have "... \<subseteq> f ` set xs" by (fact map_pair_subset)
finally obtain x where "x \<in> set xs" and eq: "hd (map_pair f xs) = f x" ..
from this(1) have "fst x \<in> fst ` set xs" by fastforce
hence "lt k (fst x)" by (rule Cons(4))
hence "lt (fst (f (k, v))) (fst (f x))"
by (simp add: Lt_lt_conv[symmetric] assms(2))
thus "lt k' (fst (hd (map_pair f xs)))" by (simp add: f eq)
qed
qed
qed
lemma lookup_pair_map_pair:
assumes "oalist_inv_raw xs" and "snd (f (k, 0)) = 0"
and "\<And>a b. comp (fst (f a)) (fst (f b)) = comp (fst a) (fst b)"
shows "lookup_pair (map_pair f xs) (fst (f (k, v))) = snd (f (k, lookup_pair xs k))"
using assms(1)
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by (simp add: assms(2))
next
case (Cons k' v' xs)
obtain k'' v'' where f: "f (k', v') = (k'', v'')" by fastforce
have "comp (fst (f (k, v))) k'' = comp (fst (f (k, v))) (fst (f (k', v')))"
by (simp add: f)
also have "... = comp k k'"
by (simp add: assms(3))
finally have eq0: "comp (fst (f (k, v))) k'' = comp k k'" .
have *: "lookup_pair xs k = 0" if "comp k k' \<noteq> Gt"
proof (simp add: lookup_pair_eq_0[OF Cons(2)], rule)
assume "k \<in> fst ` set xs"
hence "lt k' k" by (rule Cons(4))
hence "comp k k' = Gt" by (simp add: Gt_lt_conv)
with \<open>comp k k' \<noteq> Gt\<close> show False ..
qed
show ?case
proof (simp add: assms(2) f Let_def eq0 Cons(5) split: order.split, intro conjI impI)
assume "comp k k' = Lt"
hence "comp k k' \<noteq> Gt" by simp
hence "lookup_pair xs k = 0" by (rule *)
thus "snd (f (k, lookup_pair xs k)) = 0" by (simp add: assms(2))
next
assume "v'' = 0"
assume "comp k k' = Eq"
hence "k = k'" and "comp k k' \<noteq> Gt" by (simp only: eq, simp)
from this(2) have "lookup_pair xs k = 0" by (rule *)
hence "snd (f (k, lookup_pair xs k)) = 0" by (simp add: assms(2))
also have "... = snd (f (k, v'))" by (simp add: \<open>k = k'\<close> f \<open>v'' = 0\<close>)
finally show "snd (f (k, lookup_pair xs k)) = snd (f (k, v'))" .
qed (simp add: f eq)
qed
lemma lookup_dflt_map_pair:
assumes "distinct (map fst xs)" and "snd (f (k, 0)) = 0"
and "\<And>a b. (fst (f a) = fst (f b)) \<longleftrightarrow> (fst a = fst b)"
shows "lookup_dflt (map_pair f xs) (fst (f (k, v))) = snd (f (k, lookup_dflt xs k))"
using assms(1)
proof (induct xs)
case Nil
show ?case by (simp add: lookup_dflt_def assms(2))
next
case (Cons x xs)
obtain k' v' where x: "x = (k', v')" by fastforce
obtain k'' v'' where f: "f (k', v') = (k'', v'')" by fastforce
from Cons(2) have "distinct (map fst xs)" and "k' \<notin> fst ` set xs" by (simp_all add: x)
from this(1) have eq1: "lookup_dflt (map_pair f xs) (fst (f (k, v))) = snd (f (k, lookup_dflt xs k))"
by (rule Cons(1))
have eq2: "lookup_dflt ((a, b) # ys) c = (if c = a then b else lookup_dflt ys c)"
for a b c and ys::"('b \<times> 'e::zero) list" by (simp add: lookup_dflt_def map_of_Cons_code)
from \<open>k' \<notin> fst ` set xs\<close> have "map_of xs k' = None" by (simp add: map_of_eq_None_iff)
hence eq3: "lookup_dflt xs k' = 0" by (simp add: lookup_dflt_def)
show ?case
proof (simp add: f Let_def x eq1 eq2 eq3, intro conjI impI)
assume "k = k'"
hence "snd (f (k', 0)) = snd (f (k, 0))" by simp
also have "... = 0" by (fact assms(2))
finally show "snd (f (k', 0)) = 0" .
next
assume "fst (f (k', v)) \<noteq> k''"
hence "fst (f (k', v)) \<noteq> fst (f (k', v'))" by (simp add: f)
thus "snd (f (k', 0)) = v''" by (simp add: assms(3))
next
assume "k \<noteq> k'"
assume "fst (f (k, v)) = k''"
also have "... = fst (f (k', v'))" by (simp add: f)
finally have "k = k'" by (simp add: assms(3))
with \<open>k \<noteq> k'\<close> show "v'' = snd (f (k, lookup_dflt xs k))" ..
qed
qed
lemma distinct_map_pair:
assumes "distinct (map fst xs)" and "\<And>a b. fst (f a) = fst (f b) \<Longrightarrow> fst a = fst b"
shows "distinct (map fst (map_pair f xs))"
using assms(1)
proof (induct xs)
case Nil
show ?case by simp
next
case (Cons x xs)
obtain k v where x: "x = (k, v)" by fastforce
obtain k' v' where f: "f (k, v) = (k', v')" by fastforce
from Cons(2) have "distinct (map fst xs)" and "k \<notin> fst ` set xs" by (simp_all add: x)
from this(1) have 1: "distinct (map fst (map_pair f xs))" by (rule Cons(1))
show ?case
proof (simp add: x f Let_def 1, intro impI notI)
assume "v' \<noteq> 0"
assume "k' \<in> fst ` set (map_pair f xs)"
then obtain y where "y \<in> set (map_pair f xs)" and "k' = fst y" ..
from this(1) map_pair_subset have "y \<in> f ` set xs" ..
then obtain z where "z \<in> set xs" and "y = f z" ..
from this(2) have "fst (f z) = k'" by (simp add: \<open>k' = fst y\<close>)
also have "... = fst (f (k, v))" by (simp add: f)
finally have "fst z = fst (k, v)" by (rule assms(2))
also have "... = k" by simp
finally have "k \<in> fst ` set xs" using \<open>z \<in> set xs\<close> by blast
with \<open>k \<notin> fst ` set xs\<close> show False ..
qed
qed
lemma map_val_pair_cong:
assumes "\<And>k v. (k, v) \<in> set xs \<Longrightarrow> f k v = g k v"
shows "map_val_pair f xs = map_val_pair g xs"
proof (rule map_pair_cong)
fix kv
assume "kv \<in> set xs"
moreover obtain k v where "kv = (k, v)" by fastforce
ultimately show "(case kv of (k, v) \<Rightarrow> (k, f k v)) = (case kv of (k, v) \<Rightarrow> (k, g k v))"
by (simp add: assms)
qed
lemma oalist_inv_raw_map_val_pair:
assumes "oalist_inv_raw xs"
shows "oalist_inv_raw (map_val_pair f xs)"
by (rule oalist_inv_raw_map_pair, fact assms, auto)
lemma lookup_pair_map_val_pair:
assumes "oalist_inv_raw xs" and "f k 0 = 0"
shows "lookup_pair (map_val_pair f xs) k = f k (lookup_pair xs k)"
proof -
let ?f = "\<lambda>(k', v'). (k', f k' v')"
have "lookup_pair (map_val_pair f xs) k = lookup_pair (map_val_pair f xs) (fst (?f (k, 0)))"
by simp
also have "... = snd (?f (k, local.lookup_pair xs k))"
by (rule lookup_pair_map_pair, fact assms(1), auto simp: assms(2))
also have "... = f k (lookup_pair xs k)" by simp
finally show ?thesis .
qed
lemma map_pair_id:
assumes "oalist_inv_raw xs"
shows "map_pair id xs = xs"
using assms
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k v xs')
show ?case by (simp add: Let_def Cons(3, 5) id_def[symmetric])
qed
subsubsection \<open>@{const map2_val_pair}\<close>
definition map2_val_compat :: "(('a \<times> 'b::zero) list \<Rightarrow> ('a \<times> 'c::zero) list) \<Rightarrow> bool"
where "map2_val_compat f \<longleftrightarrow> (\<forall>zs. (oalist_inv_raw zs \<longrightarrow>
(oalist_inv_raw (f zs) \<and> fst ` set (f zs) \<subseteq> fst ` set zs)))"
lemma map2_val_compatI:
assumes "\<And>zs. oalist_inv_raw zs \<Longrightarrow> oalist_inv_raw (f zs)"
and "\<And>zs. oalist_inv_raw zs \<Longrightarrow> fst ` set (f zs) \<subseteq> fst ` set zs"
shows "map2_val_compat f"
unfolding map2_val_compat_def using assms by blast
lemma map2_val_compatD1:
assumes "map2_val_compat f" and "oalist_inv_raw zs"
shows "oalist_inv_raw (f zs)"
using assms unfolding map2_val_compat_def by blast
lemma map2_val_compatD2:
assumes "map2_val_compat f" and "oalist_inv_raw zs"
shows "fst ` set (f zs) \<subseteq> fst ` set zs"
using assms unfolding map2_val_compat_def by blast
lemma map2_val_compat_Nil:
assumes "map2_val_compat (f::('a \<times> 'b::zero) list \<Rightarrow> ('a \<times> 'c::zero) list)"
shows "f [] = []"
proof -
from assms oalist_inv_raw_Nil have "fst ` set (f []) \<subseteq> fst ` set ([]::('a \<times> 'b) list)"
by (rule map2_val_compatD2)
thus ?thesis by simp
qed
lemma map2_val_compat_id: "map2_val_compat id"
by (rule map2_val_compatI, auto)
lemma map2_val_compat_map_val_pair: "map2_val_compat (map_val_pair f)"
proof (rule map2_val_compatI, erule oalist_inv_raw_map_val_pair)
fix zs
from map_pair_subset image_iff show "fst ` set (map_val_pair f zs) \<subseteq> fst ` set zs" by fastforce
qed
lemma fst_map2_val_pair_subset:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys"
assumes "map2_val_compat g" and "map2_val_compat h"
shows "fst ` set (map2_val_pair f g h xs ys) \<subseteq> fst ` set xs \<union> fst ` set ys"
using assms
proof (induct f g h xs ys rule: map2_val_pair.induct)
case (1 f g h xs)
show ?case by (simp, rule map2_val_compatD2, fact+)
next
case (2 f g h v va)
show ?case by (simp del: set_simps(2), rule map2_val_compatD2, fact+)
next
case (3 f g h kx vx xs ky vy ys)
from 3(4) have "oalist_inv_raw xs" by (rule oalist_inv_raw_ConsD1)
from 3(5) have "oalist_inv_raw ys" by (rule oalist_inv_raw_ConsD1)
show ?case
proof (simp split: order.split, intro conjI impI)
assume "comp kx ky = Lt"
hence "fst ` set (map2_val_pair f g h xs ((ky, vy) # ys)) \<subseteq> fst ` set xs \<union> fst ` set ((ky, vy) # ys)"
- using refl \<open>oalist_inv_raw xs\<close> 3(5, 6, 7) by (rule 3(2))
+ using HOL.refl \<open>oalist_inv_raw xs\<close> 3(5, 6, 7) by (rule 3(2))
thus "fst ` set (let v = f kx vx 0; aux = map2_val_pair f g h xs ((ky, vy) # ys)
in if v = 0 then aux else (kx, v) # aux)
\<subseteq> insert ky (insert kx (fst ` set xs \<union> fst ` set ys))" by (auto simp: Let_def)
next
assume "comp kx ky = Eq"
hence "fst ` set (map2_val_pair f g h xs ys) \<subseteq> fst ` set xs \<union> fst ` set ys"
- using refl \<open>oalist_inv_raw xs\<close> \<open>oalist_inv_raw ys\<close> 3(6, 7) by (rule 3(1))
+ using HOL.refl \<open>oalist_inv_raw xs\<close> \<open>oalist_inv_raw ys\<close> 3(6, 7) by (rule 3(1))
thus "fst ` set (let v = f kx vx vy; aux = map2_val_pair f g h xs ys in if v = 0 then aux else (kx, v) # aux)
\<subseteq> insert ky (insert kx (fst ` set xs \<union> fst ` set ys))" by (auto simp: Let_def)
next
assume "comp kx ky = Gt"
hence "fst ` set (map2_val_pair f g h ((kx, vx) # xs) ys) \<subseteq> fst ` set ((kx, vx) # xs) \<union> fst ` set ys"
- using refl 3(4) \<open>oalist_inv_raw ys\<close> 3(6, 7) by (rule 3(3))
+ using HOL.refl 3(4) \<open>oalist_inv_raw ys\<close> 3(6, 7) by (rule 3(3))
thus "fst ` set (let v = f ky 0 vy; aux = map2_val_pair f g h ((kx, vx) # xs) ys
in if v = 0 then aux else (ky, v) # aux)
\<subseteq> insert ky (insert kx (fst ` set xs \<union> fst ` set ys))" by (auto simp: Let_def)
qed
qed
lemma oalist_inv_raw_map2_val_pair:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys"
assumes "map2_val_compat g" and "map2_val_compat h"
shows "oalist_inv_raw (map2_val_pair f g h xs ys)"
using assms(1, 2)
proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct)
case Nil
show ?case
proof (cases ys)
case Nil
show ?thesis by (simp add: Nil, rule map2_val_compatD1, fact assms(3), fact oalist_inv_raw_Nil)
next
case (Cons y ys')
show ?thesis by (simp add: Cons, rule map2_val_compatD1, fact assms(4), simp only: Cons[symmetric], fact Nil)
qed
next
case *: (Cons k v xs)
from *(6) show ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
show ?case by (simp, rule map2_val_compatD1, fact assms(3), fact *(1))
next
case (Cons k' v' ys)
show ?case
proof (simp split: order.split, intro conjI impI)
assume "comp k k' = Lt"
hence 0: "lt k k'" by (simp only: Lt_lt_conv)
from Cons(1) have 1: "oalist_inv_raw (map2_val_pair f g h xs ((k', v') # ys))" by (rule *(5))
show "oalist_inv_raw (let v = f k v 0; aux = map2_val_pair f g h xs ((k', v') # ys)
in if v = 0 then aux else (k, v) # aux)"
proof (simp add: Let_def, intro conjI impI)
assume "f k v 0 \<noteq> 0"
with 1 show "oalist_inv_raw ((k, f k v 0) # map2_val_pair f g h xs ((k', v') # ys))"
proof (rule oalist_inv_raw_ConsI)
define k0 where "k0 = fst (hd (local.map2_val_pair f g h xs ((k', v') # ys)))"
assume "map2_val_pair f g h xs ((k', v') # ys) \<noteq> []"
hence "k0 \<in> fst ` set (map2_val_pair f g h xs ((k', v') # ys))" by (simp add: k0_def)
also from *(2) Cons(1) assms(3, 4) have "... \<subseteq> fst ` set xs \<union> fst ` set ((k', v') # ys)"
by (rule fst_map2_val_pair_subset)
finally have "k0 \<in> fst ` set xs \<or> k0 = k' \<or> k0 \<in> fst ` set ys" by auto
thus "lt k k0"
proof (elim disjE)
assume "k0 = k'"
with 0 show ?thesis by simp
next
assume "k0 \<in> fst ` set ys"
hence "lt k' k0" by (rule Cons(4))
with 0 show ?thesis by (rule less_trans)
qed (rule *(4))
qed
qed (rule 1)
next
assume "comp k k' = Eq"
hence "k = k'" by (simp only: eq)
from Cons(2) have 1: "oalist_inv_raw (map2_val_pair f g h xs ys)" by (rule *(5))
show "oalist_inv_raw (let v = f k v v'; aux = map2_val_pair f g h xs ys in if v = 0 then aux else (k, v) # aux)"
proof (simp add: Let_def, intro conjI impI)
assume "f k v v' \<noteq> 0"
with 1 show "oalist_inv_raw ((k, f k v v') # map2_val_pair f g h xs ys)"
proof (rule oalist_inv_raw_ConsI)
define k0 where "k0 = fst (hd (map2_val_pair f g h xs ys))"
assume "map2_val_pair f g h xs ys \<noteq> []"
hence "k0 \<in> fst ` set (map2_val_pair f g h xs ys)" by (simp add: k0_def)
also from *(2) Cons(2) assms(3, 4) have "... \<subseteq> fst ` set xs \<union> fst ` set ys"
by (rule fst_map2_val_pair_subset)
finally show "lt k k0"
proof
assume "k0 \<in> fst ` set ys"
hence "lt k' k0" by (rule Cons(4))
thus ?thesis by (simp only: \<open>k = k'\<close>)
qed (rule *(4))
qed
qed (rule 1)
next
assume "comp k k' = Gt"
hence 0: "lt k' k" by (simp only: Gt_lt_conv)
show "oalist_inv_raw (let va = f k' 0 v'; aux = map2_val_pair f g h ((k, v) # xs) ys
in if va = 0 then aux else (k', va) # aux)"
proof (simp add: Let_def, intro conjI impI)
assume "f k' 0 v' \<noteq> 0"
with Cons(5) show "oalist_inv_raw ((k', f k' 0 v') # map2_val_pair f g h ((k, v) # xs) ys)"
proof (rule oalist_inv_raw_ConsI)
define k0 where "k0 = fst (hd (map2_val_pair f g h ((k, v) # xs) ys))"
assume "map2_val_pair f g h ((k, v) # xs) ys \<noteq> []"
hence "k0 \<in> fst ` set (map2_val_pair f g h ((k, v) # xs) ys)" by (simp add: k0_def)
also from *(1) Cons(2) assms(3, 4) have "... \<subseteq> fst ` set ((k, v) # xs) \<union> fst ` set ys"
by (rule fst_map2_val_pair_subset)
finally have "k0 = k \<or> k0 \<in> fst ` set xs \<or> k0 \<in> fst ` set ys" by auto
thus "lt k' k0"
proof (elim disjE)
assume "k0 = k"
with 0 show ?thesis by simp
next
assume "k0 \<in> fst ` set xs"
hence "lt k k0" by (rule *(4))
with 0 show ?thesis by (rule less_trans)
qed (rule Cons(4))
qed
qed (rule Cons(5))
qed
qed
qed
lemma lookup_pair_map2_val_pair:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys"
assumes "map2_val_compat g" and "map2_val_compat h"
assumes "\<And>zs. oalist_inv_raw zs \<Longrightarrow> g zs = map_val_pair (\<lambda>k v. f k v 0) zs"
and "\<And>zs. oalist_inv_raw zs \<Longrightarrow> h zs = map_val_pair (\<lambda>k. f k 0) zs"
and "\<And>k. f k 0 0 = 0"
shows "lookup_pair (map2_val_pair f g h xs ys) k0 = f k0 (lookup_pair xs k0) (lookup_pair ys k0)"
using assms(1, 2)
proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct)
case Nil
show ?case
proof (cases ys)
case Nil
show ?thesis by (simp add: Nil map2_val_compat_Nil[OF assms(3)] assms(7))
next
case (Cons y ys')
then obtain k v ys' where ys: "ys = (k, v) # ys'" by fastforce
from Nil have "lookup_pair (h ys) k0 = lookup_pair (map_val_pair (\<lambda>k. f k 0) ys) k0"
by (simp only: assms(6))
also have "... = f k0 0 (lookup_pair ys k0)" by (rule lookup_pair_map_val_pair, fact Nil, fact assms(7))
finally have "lookup_pair (h ((k, v) # ys')) k0 = f k0 0 (lookup_pair ((k, v) # ys') k0)"
by (simp only: ys)
thus ?thesis by (simp add: ys)
qed
next
case *: (Cons k v xs)
from *(6) show ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
from *(1) have "lookup_pair (g ((k, v) # xs)) k0 = lookup_pair (map_val_pair (\<lambda>k v. f k v 0) ((k, v) # xs)) k0"
by (simp only: assms(5))
also have "... = f k0 (lookup_pair ((k, v) # xs) k0) 0"
by (rule lookup_pair_map_val_pair, fact *(1), fact assms(7))
finally show ?case by simp
next
case (Cons k' v' ys)
show ?case
proof (cases "comp k0 k = Lt \<and> comp k0 k' = Lt")
case True
hence 1: "comp k0 k = Lt" and 2: "comp k0 k' = Lt" by simp_all
hence eq: "f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0) = 0"
by (simp add: assms(7))
from *(1) Cons(1) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h ((k, v) # xs) ((k', v') # ys))"
by (rule oalist_inv_raw_map2_val_pair)
show ?thesis
proof (simp only: eq lookup_pair_eq_0[OF inv], rule)
assume "k0 \<in> fst ` set (local.map2_val_pair f g h ((k, v) # xs) ((k', v') # ys))"
also from *(1) Cons(1) assms(3, 4) have "... \<subseteq> fst ` set ((k, v) # xs) \<union> fst ` set ((k', v') # ys)"
by (rule fst_map2_val_pair_subset)
finally have "k0 \<in> fst ` set xs \<or> k0 \<in> fst ` set ys" using 1 2 by auto
thus False
proof
assume "k0 \<in> fst ` set xs"
hence "lt k k0" by (rule *(4))
with 1 show ?thesis by (simp add: Lt_lt_conv)
next
assume "k0 \<in> fst ` set ys"
hence "lt k' k0" by (rule Cons(4))
with 2 show ?thesis by (simp add: Lt_lt_conv)
qed
qed
next
case False
show ?thesis
proof (simp split: order.split del: lookup_pair.simps, intro conjI impI)
assume "comp k k' = Lt"
with False have "comp k0 k \<noteq> Lt" by (auto simp: Lt_lt_conv)
show "lookup_pair (let v = f k v 0; aux = map2_val_pair f g h xs ((k', v') # ys)
in if v = 0 then aux else (k, v) # aux) k0 =
f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)"
proof (cases "comp k0 k")
case Lt
with \<open>comp k0 k \<noteq> Lt\<close> show ?thesis ..
next
case Eq
hence "k0 = k" by (simp only: eq)
with \<open>comp k k' = Lt\<close> have "comp k0 k' = Lt" by simp
hence eq1: "lookup_pair ((k', v') # ys) k = 0" by (simp add: \<open>k0 = k\<close>)
have eq2: "lookup_pair ((k, v) # xs) k = v" by simp
show ?thesis
proof (simp add: Let_def eq1 eq2 \<open>k0 = k\<close> del: lookup_pair.simps, intro conjI impI)
from *(2) Cons(1) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h xs ((k', v') # ys))"
by (rule oalist_inv_raw_map2_val_pair)
show "lookup_pair (map2_val_pair f g h xs ((k', v') # ys)) k = 0"
proof (simp only: lookup_pair_eq_0[OF inv], rule)
assume "k \<in> fst ` set (local.map2_val_pair f g h xs ((k', v') # ys))"
also from *(2) Cons(1) assms(3, 4) have "... \<subseteq> fst ` set xs \<union> fst ` set ((k', v') # ys)"
by (rule fst_map2_val_pair_subset)
finally have "k \<in> fst ` set xs \<or> k \<in> fst ` set ys" using \<open>comp k k' = Lt\<close>
by auto
thus False
proof
assume "k \<in> fst ` set xs"
hence "lt k k" by (rule *(4))
thus ?thesis by simp
next
assume "k \<in> fst ` set ys"
hence "lt k' k" by (rule Cons(4))
with \<open>comp k k' = Lt\<close> show ?thesis by (simp add: Lt_lt_conv)
qed
qed
qed simp
next
case Gt
hence eq1: "lookup_pair ((k, v) # xs) k0 = lookup_pair xs k0"
and eq2: "lookup_pair ((k, f k v 0) # map2_val_pair f g h xs ((k', v') # ys)) k0 =
lookup_pair (map2_val_pair f g h xs ((k', v') # ys)) k0" by simp_all
show ?thesis
by (simp add: Let_def eq1 eq2 del: lookup_pair.simps, rule *(5), fact Cons(1))
qed
next
assume "comp k k' = Eq"
hence "k = k'" by (simp only: eq)
with False have "comp k0 k' \<noteq> Lt" by (auto simp: Lt_lt_conv)
show "lookup_pair (let v = f k v v'; aux = map2_val_pair f g h xs ys in
if v = 0 then aux else (k, v) # aux) k0 =
f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)"
proof (cases "comp k0 k'")
case Lt
with \<open>comp k0 k' \<noteq> Lt\<close> show ?thesis ..
next
case Eq
hence "k0 = k'" by (simp only: eq)
show ?thesis
proof (simp add: Let_def \<open>k = k'\<close> \<open>k0 = k'\<close>, intro impI)
from *(2) Cons(2) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h xs ys)"
by (rule oalist_inv_raw_map2_val_pair)
show "lookup_pair (map2_val_pair f g h xs ys) k' = 0"
proof (simp only: lookup_pair_eq_0[OF inv], rule)
assume "k' \<in> fst ` set (map2_val_pair f g h xs ys)"
also from *(2) Cons(2) assms(3, 4) have "... \<subseteq> fst ` set xs \<union> fst ` set ys"
by (rule fst_map2_val_pair_subset)
finally show False
proof
assume "k' \<in> fst ` set ys"
hence "lt k' k'" by (rule Cons(4))
thus ?thesis by simp
next
assume "k' \<in> fst ` set xs"
hence "lt k k'" by (rule *(4))
thus ?thesis by (simp add: \<open>k = k'\<close>)
qed
qed
qed
next
case Gt
hence eq1: "lookup_pair ((k, v) # xs) k0 = lookup_pair xs k0"
and eq2: "lookup_pair ((k', v') # ys) k0 = lookup_pair ys k0"
and eq3: "lookup_pair ((k, f k v v') # map2_val_pair f g h xs ys) k0 =
lookup_pair (map2_val_pair f g h xs ys) k0" by (simp_all add: \<open>k = k'\<close>)
show ?thesis by (simp add: Let_def eq1 eq2 eq3 del: lookup_pair.simps, rule *(5), fact Cons(2))
qed
next
assume "comp k k' = Gt"
hence "comp k' k = Lt" by (simp only: Gt_lt_conv Lt_lt_conv)
with False have "comp k0 k' \<noteq> Lt" by (auto simp: Lt_lt_conv)
show "lookup_pair (let va = f k' 0 v'; aux = map2_val_pair f g h ((k, v) # xs) ys
in if va = 0 then aux else (k', va) # aux) k0 =
f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)"
proof (cases "comp k0 k'")
case Lt
with \<open>comp k0 k' \<noteq> Lt\<close> show ?thesis ..
next
case Eq
hence "k0 = k'" by (simp only: eq)
with \<open>comp k' k = Lt\<close> have "comp k0 k = Lt" by simp
hence eq1: "lookup_pair ((k, v) # xs) k' = 0" by (simp add: \<open>k0 = k'\<close>)
have eq2: "lookup_pair ((k', v') # ys) k' = v'" by simp
show ?thesis
proof (simp add: Let_def eq1 eq2 \<open>k0 = k'\<close> del: lookup_pair.simps, intro conjI impI)
from *(1) Cons(2) assms(3, 4) have inv: "oalist_inv_raw (map2_val_pair f g h ((k, v) # xs) ys)"
by (rule oalist_inv_raw_map2_val_pair)
show "lookup_pair (map2_val_pair f g h ((k, v) # xs) ys) k' = 0"
proof (simp only: lookup_pair_eq_0[OF inv], rule)
assume "k' \<in> fst ` set (map2_val_pair f g h ((k, v) # xs) ys)"
also from *(1) Cons(2) assms(3, 4) have "... \<subseteq> fst ` set ((k, v) # xs) \<union> fst ` set ys"
by (rule fst_map2_val_pair_subset)
finally have "k' \<in> fst ` set xs \<or> k' \<in> fst ` set ys" using \<open>comp k' k = Lt\<close>
by auto
thus False
proof
assume "k' \<in> fst ` set ys"
hence "lt k' k'" by (rule Cons(4))
thus ?thesis by simp
next
assume "k' \<in> fst ` set xs"
hence "lt k k'" by (rule *(4))
with \<open>comp k' k = Lt\<close> show ?thesis by (simp add: Lt_lt_conv)
qed
qed
qed simp
next
case Gt
hence eq1: "lookup_pair ((k', v') # ys) k0 = lookup_pair ys k0"
and eq2: "lookup_pair ((k', f k' 0 v') # map2_val_pair f g h ((k, v) # xs) ys) k0 =
lookup_pair (map2_val_pair f g h ((k, v) # xs) ys) k0" by simp_all
show ?thesis by (simp add: Let_def eq1 eq2 del: lookup_pair.simps, rule Cons(5))
qed
qed
qed
qed
qed
lemma map2_val_pair_singleton_eq_update_by_fun_pair:
assumes "oalist_inv_raw xs"
assumes "\<And>k x. f k x 0 = x" and "\<And>zs. oalist_inv_raw zs \<Longrightarrow> g zs = zs"
and "h [(k, v)] = map_val_pair (\<lambda>k. f k 0) [(k, v)]"
shows "map2_val_pair f g h xs [(k, v)] = update_by_fun_pair k (\<lambda>x. f k x v) xs"
using assms(1)
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by (simp add: Let_def assms(4))
next
case (Cons k' v' xs)
show ?case
proof (cases "comp k' k")
case Lt
hence gr: "comp k k' = Gt" by (simp only: Gt_lt_conv Lt_lt_conv)
show ?thesis by (simp add: Lt gr Let_def assms(2) Cons(3, 5))
next
case Eq
hence eq1: "comp k k' = Eq" and eq2: "k = k'" by (simp_all only: eq)
show ?thesis by (simp add: Eq eq1 eq2 Let_def assms(3)[OF Cons(2)])
next
case Gt
hence less: "comp k k' = Lt" by (simp only: Gt_lt_conv Lt_lt_conv)
show ?thesis by (simp add: Gt less Let_def assms(3)[OF Cons(1)])
qed
qed
subsubsection \<open>@{const lex_ord_pair}\<close>
lemma lex_ord_pair_EqI:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys"
and "\<And>k. k \<in> fst ` set xs \<union> fst ` set ys \<Longrightarrow> f k (lookup_pair xs k) (lookup_pair ys k) = Some Eq"
shows "lex_ord_pair f xs ys = Some Eq"
using assms
proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct)
case Nil
thus ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k v ys)
show ?case
proof (simp add: Let_def, intro conjI impI, rule Cons(5))
fix k0
assume "k0 \<in> fst ` set [] \<union> fst ` set ys"
hence "k0 \<in> fst ` set ys" by simp
hence "lt k k0" by (rule Cons(4))
hence "f k0 (lookup_pair [] k0) (lookup_pair ys k0) = f k0 (lookup_pair [] k0) (lookup_pair ((k, v) # ys) k0)"
by (auto simp add: lookup_pair_Cons[OF Cons(1)] simp del: lookup_pair.simps)
also have "... = Some Eq" by (rule Cons(6), simp add: \<open>k0 \<in> fst ` set ys\<close>)
finally show "f k0 (lookup_pair [] k0) (lookup_pair ys k0) = Some Eq" .
next
have "f k 0 v = f k (lookup_pair [] k) (lookup_pair ((k, v) # ys) k)" by simp
also have "... = Some Eq" by (rule Cons(6), simp)
finally show "f k 0 v = Some Eq" .
qed
qed
next
case *: (Cons k v xs)
from *(6, 7) show ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
show ?case
proof (simp add: Let_def, intro conjI impI, rule *(5), rule oalist_inv_raw_Nil)
fix k0
assume "k0 \<in> fst ` set xs \<union> fst ` set []"
hence "k0 \<in> fst ` set xs" by simp
hence "lt k k0" by (rule *(4))
hence "f k0 (lookup_pair xs k0) (lookup_pair [] k0) = f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair [] k0)"
by (auto simp add: lookup_pair_Cons[OF *(1)] simp del: lookup_pair.simps)
also have "... = Some Eq" by (rule Nil, simp add: \<open>k0 \<in> fst ` set xs\<close>)
finally show "f k0 (lookup_pair xs k0) (lookup_pair [] k0) = Some Eq" .
next
have "f k v 0 = f k (lookup_pair ((k, v) # xs) k) (lookup_pair [] k)" by simp
also have "... = Some Eq" by (rule Nil, simp)
finally show "f k v 0 = Some Eq" .
qed
next
case (Cons k' v' ys)
show ?case
proof (simp split: order.split, intro conjI impI)
assume "comp k k' = Lt"
show "(let aux = f k v 0 in if aux = Some Eq then lex_ord_pair f xs ((k', v') # ys) else aux) = Some Eq"
proof (simp add: Let_def, intro conjI impI, rule *(5), rule Cons(1))
fix k0
assume k0_in: "k0 \<in> fst ` set xs \<union> fst ` set ((k', v') # ys)"
hence "k0 \<in> fst ` set xs \<or> k0 = k' \<or> k0 \<in> fst ` set ys" by auto
hence "k0 \<noteq> k"
proof (elim disjE)
assume "k0 \<in> fst ` set xs"
hence "lt k k0" by (rule *(4))
thus ?thesis by simp
next
assume "k0 = k'"
with \<open>comp k k' = Lt\<close> show ?thesis by auto
next
assume "k0 \<in> fst ` set ys"
hence "lt k' k0" by (rule Cons(4))
with \<open>comp k k' = Lt\<close> show ?thesis by (simp add: Lt_lt_conv)
qed
hence "f k0 (lookup_pair xs k0) (lookup_pair ((k', v') # ys) k0) =
f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)"
by (auto simp add: lookup_pair_Cons[OF *(1)] simp del: lookup_pair.simps)
also have "... = Some Eq" by (rule Cons(6), rule rev_subsetD, fact k0_in, auto)
finally show "f k0 (lookup_pair xs k0) (lookup_pair ((k', v') # ys) k0) = Some Eq" .
next
have "f k v 0 = f k (lookup_pair ((k, v) # xs) k) (lookup_pair ((k', v') # ys) k)"
by (simp add: \<open>comp k k' = Lt\<close>)
also have "... = Some Eq" by (rule Cons(6), simp)
finally show "f k v 0 = Some Eq" .
qed
next
assume "comp k k' = Eq"
hence "k = k'" by (simp only: eq)
show "(let aux = f k v v' in if aux = Some Eq then lex_ord_pair f xs ys else aux) = Some Eq"
proof (simp add: Let_def, intro conjI impI, rule *(5), rule Cons(2))
fix k0
assume k0_in: "k0 \<in> fst ` set xs \<union> fst ` set ys"
hence "k0 \<noteq> k'"
proof
assume "k0 \<in> fst ` set xs"
hence "lt k k0" by (rule *(4))
thus ?thesis by (simp add: \<open>k = k'\<close>)
next
assume "k0 \<in> fst ` set ys"
hence "lt k' k0" by (rule Cons(4))
thus ?thesis by simp
qed
hence "f k0 (lookup_pair xs k0) (lookup_pair ys k0) =
f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)"
by (simp add: lookup_pair_Cons[OF *(1)] lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps,
auto simp: \<open>k = k'\<close>)
also have "... = Some Eq" by (rule Cons(6), rule rev_subsetD, fact k0_in, auto)
finally show "f k0 (lookup_pair xs k0) (lookup_pair ys k0) = Some Eq" .
next
have "f k v v' = f k (lookup_pair ((k, v) # xs) k) (lookup_pair ((k', v') # ys) k)"
by (simp add: \<open>k = k'\<close>)
also have "... = Some Eq" by (rule Cons(6), simp)
finally show "f k v v' = Some Eq" .
qed
next
assume "comp k k' = Gt"
hence "comp k' k = Lt" by (simp only: Gt_lt_conv Lt_lt_conv)
show "(let aux = f k' 0 v' in if aux = Some Eq then lex_ord_pair f ((k, v) # xs) ys else aux) = Some Eq"
proof (simp add: Let_def, intro conjI impI, rule Cons(5))
fix k0
assume k0_in: "k0 \<in> fst ` set ((k, v) # xs) \<union> fst ` set ys"
hence "k0 \<in> fst ` set xs \<or> k0 = k \<or> k0 \<in> fst ` set ys" by auto
hence "k0 \<noteq> k'"
proof (elim disjE)
assume "k0 \<in> fst ` set xs"
hence "lt k k0" by (rule *(4))
with \<open>comp k' k = Lt\<close> show ?thesis by (simp add: Lt_lt_conv)
next
assume "k0 = k"
with \<open>comp k' k = Lt\<close> show ?thesis by auto
next
assume "k0 \<in> fst ` set ys"
hence "lt k' k0" by (rule Cons(4))
thus ?thesis by simp
qed
hence "f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ys k0) =
f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ((k', v') # ys) k0)"
by (auto simp add: lookup_pair_Cons[OF Cons(1)] simp del: lookup_pair.simps)
also have "... = Some Eq" by (rule Cons(6), rule rev_subsetD, fact k0_in, auto)
finally show "f k0 (lookup_pair ((k, v) # xs) k0) (lookup_pair ys k0) = Some Eq" .
next
have "f k' 0 v' = f k' (lookup_pair ((k, v) # xs) k') (lookup_pair ((k', v') # ys) k')"
by (simp add: \<open>comp k' k = Lt\<close>)
also have "... = Some Eq" by (rule Cons(6), simp)
finally show "f k' 0 v' = Some Eq" .
qed
qed
qed
qed
lemma lex_ord_pair_valI:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "aux \<noteq> Some Eq"
assumes "k \<in> fst ` set xs \<union> fst ` set ys" and "aux = f k (lookup_pair xs k) (lookup_pair ys k)"
and "\<And>k'. k' \<in> fst ` set xs \<union> fst ` set ys \<Longrightarrow> lt k' k \<Longrightarrow>
f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq"
shows "lex_ord_pair f xs ys = aux"
using assms(1, 2, 4, 5, 6)
proof (induct xs arbitrary: ys rule: oalist_inv_raw_induct)
case Nil
thus ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
from Nil(1) show ?case by simp
next
case (Cons k' v' ys)
from Cons(6) have "k = k' \<or> k \<in> fst ` set ys" by simp
thus ?case
proof
assume "k = k'"
with Cons(7) have "f k' 0 v' = aux" by simp
thus ?thesis by (simp add: Let_def \<open>k = k'\<close> assms(3))
next
assume "k \<in> fst `set ys"
hence "lt k' k" by (rule Cons(4))
hence "comp k k' = Gt" by (simp add: Gt_lt_conv)
hence eq1: "lookup_pair ((k', v') # ys) k = lookup_pair ys k" by simp
have "f k' (lookup_pair [] k') (lookup_pair ((k', v') # ys) k') = Some Eq"
by (rule Cons(8), simp, fact)
hence eq2: "f k' 0 v' = Some Eq" by simp
show ?thesis
proof (simp add: Let_def eq2, rule Cons(5))
from \<open>k \<in> fst `set ys\<close> show "k \<in> fst ` set [] \<union> fst ` set ys" by simp
next
show "aux = f k (lookup_pair [] k) (lookup_pair ys k)" by (simp only: Cons(7) eq1)
next
fix k0
assume "lt k0 k"
assume "k0 \<in> fst ` set [] \<union> fst ` set ys"
hence k0_in: "k0 \<in> fst ` set ys" by simp
hence "lt k' k0" by (rule Cons(4))
hence "comp k0 k' = Gt" by (simp add: Gt_lt_conv)
hence "f k0 (lookup_pair [] k0) (lookup_pair ys k0) =
f k0 (lookup_pair [] k0) (lookup_pair ((k', v') # ys) k0)" by simp
also have "... = Some Eq" by (rule Cons(8), simp add: k0_in, fact)
finally show "f k0 (lookup_pair [] k0) (lookup_pair ys k0) = Some Eq" .
qed
qed
qed
next
case *: (Cons k' v' xs)
from *(6, 7, 8, 9) show ?case
proof (induct ys rule: oalist_inv_raw_induct)
case Nil
from Nil(1) have "k = k' \<or> k \<in> fst ` set xs" by simp
thus ?case
proof
assume "k = k'"
with Nil(2) have "f k' v' 0 = aux" by simp
thus ?thesis by (simp add: Let_def \<open>k = k'\<close> assms(3))
next
assume "k \<in> fst ` set xs"
hence "lt k' k" by (rule *(4))
hence "comp k k' = Gt" by (simp add: Gt_lt_conv)
hence eq1: "lookup_pair ((k', v') # xs) k = lookup_pair xs k" by simp
have "f k' (lookup_pair ((k', v') # xs) k') (lookup_pair [] k') = Some Eq"
by (rule Nil(3), simp, fact)
hence eq2: "f k' v' 0 = Some Eq" by simp
show ?thesis
proof (simp add: Let_def eq2, rule *(5), fact oalist_inv_raw_Nil)
from \<open>k \<in> fst `set xs\<close> show "k \<in> fst ` set xs \<union> fst ` set []" by simp
next
show "aux = f k (lookup_pair xs k) (lookup_pair [] k)" by (simp only: Nil(2) eq1)
next
fix k0
assume "lt k0 k"
assume "k0 \<in> fst ` set xs \<union> fst ` set []"
hence k0_in: "k0 \<in> fst ` set xs" by simp
hence "lt k' k0" by (rule *(4))
hence "comp k0 k' = Gt" by (simp add: Gt_lt_conv)
hence "f k0 (lookup_pair xs k0) (lookup_pair [] k0) =
f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair [] k0)" by simp
also have "... = Some Eq" by (rule Nil(3), simp add: k0_in, fact)
finally show "f k0 (lookup_pair xs k0) (lookup_pair [] k0) = Some Eq" .
qed
qed
next
case (Cons k'' v'' ys)
have 0: thesis if 1: "lt k k'" and 2: "lt k k''" for thesis
proof -
from 1 have "k \<noteq> k'" by simp
moreover from 2 have "k \<noteq> k''" by simp
ultimately have "k \<in> fst ` set xs \<or> k \<in> fst ` set ys" using Cons(6) by simp
thus ?thesis
proof
assume "k \<in> fst ` set xs"
hence "lt k' k" by (rule *(4))
with 1 show ?thesis by simp
next
assume "k \<in> fst ` set ys"
hence "lt k'' k" by (rule Cons(4))
with 2 show ?thesis by simp
qed
qed
show ?case
proof (simp split: order.split, intro conjI impI)
assume Lt: "comp k' k'' = Lt"
show "(let aux = f k' v' 0 in if aux = Some Eq then lex_ord_pair f xs ((k'', v'') # ys) else aux) = aux"
proof (simp add: Let_def split: order.split, intro conjI impI)
assume "f k' v' 0 = Some Eq"
have "k \<noteq> k'"
proof
assume "k = k'"
have "aux = f k v' 0" by (simp add: Cons(7) \<open>k = k'\<close> Lt)
with \<open>f k' v' 0 = Some Eq\<close> assms(3) show False by (simp add: \<open>k = k'\<close>)
qed
from Cons(1) show "lex_ord_pair f xs ((k'', v'') # ys) = aux"
proof (rule *(5))
from Cons(6) \<open>k \<noteq> k'\<close> show "k \<in> fst ` set xs \<union> fst ` set ((k'', v'') # ys)" by simp
next
show "aux = f k (lookup_pair xs k) (lookup_pair ((k'', v'') # ys) k)"
by (simp add: Cons(7) lookup_pair_Cons[OF *(1)] \<open>k \<noteq> k'\<close>[symmetric] del: lookup_pair.simps)
next
fix k0
assume "lt k0 k"
assume k0_in: "k0 \<in> fst ` set xs \<union> fst ` set ((k'', v'') # ys)"
also have "... \<subseteq> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" by fastforce
finally have k0_in': "k0 \<in> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" .
have "k' \<noteq> k0"
proof
assume "k' = k0"
with k0_in have "k' \<in> fst ` set xs \<union> fst ` set ((k'', v'') # ys)" by simp
with Lt have "k' \<in> fst ` set xs \<or> k' \<in> fst ` set ys" by auto
thus False
proof
assume "k' \<in> fst ` set xs"
hence "lt k' k'" by (rule *(4))
thus ?thesis by simp
next
assume "k' \<in> fst ` set ys"
hence "lt k'' k'" by (rule Cons(4))
with Lt show ?thesis by (simp add: Lt_lt_conv)
qed
qed
hence "f k0 (lookup_pair xs k0) (lookup_pair ((k'', v'') # ys) k0) =
f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ((k'', v'') # ys) k0)"
by (simp add: lookup_pair_Cons[OF *(1)] del: lookup_pair.simps)
also from k0_in' \<open>lt k0 k\<close> have "... = Some Eq" by (rule Cons(8))
finally show "f k0 (lookup_pair xs k0) (lookup_pair ((k'', v'') # ys) k0) = Some Eq" .
qed
next
assume "f k' v' 0 \<noteq> Some Eq"
have "\<not> lt k' k"
proof
have "k' \<in> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" by simp
moreover assume "lt k' k"
ultimately have "f k' (lookup_pair ((k', v') # xs) k') (lookup_pair ((k'', v'') # ys) k') = Some Eq"
by (rule Cons(8))
hence "f k' v' 0 = Some Eq" by (simp add: Lt)
with \<open>f k' v' 0 \<noteq> Some Eq\<close> show False ..
qed
moreover have "\<not> lt k k'"
proof
assume "lt k k'"
moreover from this Lt have "lt k k''" by (simp add: Lt_lt_conv)
ultimately show False by (rule 0)
qed
ultimately have "k = k'" by simp
show "f k' v' 0 = aux" by (simp add: Cons(7) \<open>k = k'\<close> Lt)
qed
next
assume "comp k' k'' = Eq"
hence "k' = k''" by (simp only: eq)
show "(let aux = f k' v' v'' in if aux = Some Eq then lex_ord_pair f xs ys else aux) = aux"
proof (simp add: Let_def \<open>k' = k''\<close> split: order.split, intro conjI impI)
assume "f k'' v' v'' = Some Eq"
have "k \<noteq> k''"
proof
assume "k = k''"
have "aux = f k v' v''" by (simp add: Cons(7) \<open>k = k''\<close> \<open>k' = k''\<close>)
with \<open>f k'' v' v'' = Some Eq\<close> assms(3) show False by (simp add: \<open>k = k''\<close>)
qed
from Cons(2) show "lex_ord_pair f xs ys = aux"
proof (rule *(5))
from Cons(6) \<open>k \<noteq> k''\<close> show "k \<in> fst ` set xs \<union> fst ` set ys" by (simp add: \<open>k' = k''\<close>)
next
show "aux = f k (lookup_pair xs k) (lookup_pair ys k)"
by (simp add: Cons(7) lookup_pair_Cons[OF *(1)] lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps,
simp add: \<open>k' = k''\<close> \<open>k \<noteq> k''\<close>[symmetric])
next
fix k0
assume "lt k0 k"
assume k0_in: "k0 \<in> fst ` set xs \<union> fst ` set ys"
also have "... \<subseteq> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" by fastforce
finally have k0_in': "k0 \<in> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" .
have "k'' \<noteq> k0"
proof
assume "k'' = k0"
with k0_in have "k'' \<in> fst ` set xs \<union> fst ` set ys" by simp
thus False
proof
assume "k'' \<in> fst ` set xs"
hence "lt k' k''" by (rule *(4))
thus ?thesis by (simp add: \<open>k' = k''\<close>)
next
assume "k'' \<in> fst ` set ys"
hence "lt k'' k''" by (rule Cons(4))
thus ?thesis by simp
qed
qed
hence "f k0 (lookup_pair xs k0) (lookup_pair ys k0) =
f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ((k'', v'') # ys) k0)"
by (simp add: lookup_pair_Cons[OF *(1)] lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps,
simp add: \<open>k' = k''\<close>)
also from k0_in' \<open>lt k0 k\<close> have "... = Some Eq" by (rule Cons(8))
finally show "f k0 (lookup_pair xs k0) (lookup_pair ys k0) = Some Eq" .
qed
next
assume "f k'' v' v'' \<noteq> Some Eq"
have "\<not> lt k'' k"
proof
have "k'' \<in> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" by simp
moreover assume "lt k'' k"
ultimately have "f k'' (lookup_pair ((k', v') # xs) k'') (lookup_pair ((k'', v'') # ys) k'') = Some Eq"
by (rule Cons(8))
hence "f k'' v' v'' = Some Eq" by (simp add: \<open>k' = k''\<close>)
with \<open>f k'' v' v'' \<noteq> Some Eq\<close> show False ..
qed
moreover have "\<not> lt k k''"
proof
assume "lt k k''"
hence "lt k k'" by (simp only: \<open>k' = k''\<close>)
thus False using \<open>lt k k''\<close> by (rule 0)
qed
ultimately have "k = k''" by simp
show "f k'' v' v'' = aux" by (simp add: Cons(7) \<open>k = k''\<close> \<open>k' = k''\<close>)
qed
next
assume Gt: "comp k' k'' = Gt"
hence Lt: "comp k'' k' = Lt" by (simp only: Gt_lt_conv Lt_lt_conv)
show "(let aux = f k'' 0 v'' in if aux = Some Eq then lex_ord_pair f ((k', v') # xs) ys else aux) = aux"
proof (simp add: Let_def split: order.split, intro conjI impI)
assume "f k'' 0 v'' = Some Eq"
have "k \<noteq> k''"
proof
assume "k = k''"
have "aux = f k 0 v''" by (simp add: Cons(7) \<open>k = k''\<close> Lt)
with \<open>f k'' 0 v'' = Some Eq\<close> assms(3) show False by (simp add: \<open>k = k''\<close>)
qed
show "lex_ord_pair f ((k', v') # xs) ys = aux"
proof (rule Cons(5))
from Cons(6) \<open>k \<noteq> k''\<close> show "k \<in> fst ` set ((k', v') # xs) \<union> fst ` set ys" by simp
next
show "aux = f k (lookup_pair ((k', v') # xs) k) (lookup_pair ys k)"
by (simp add: Cons(7) lookup_pair_Cons[OF Cons(1)] \<open>k \<noteq> k''\<close>[symmetric] del: lookup_pair.simps)
next
fix k0
assume "lt k0 k"
assume k0_in: "k0 \<in> fst ` set ((k', v') # xs) \<union> fst ` set ys"
also have "... \<subseteq> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" by fastforce
finally have k0_in': "k0 \<in> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" .
have "k'' \<noteq> k0"
proof
assume "k'' = k0"
with k0_in have "k'' \<in> fst ` set ((k', v') # xs) \<union> fst ` set ys" by simp
with Lt have "k'' \<in> fst ` set xs \<or> k'' \<in> fst ` set ys" by auto
thus False
proof
assume "k'' \<in> fst ` set xs"
hence "lt k' k''" by (rule *(4))
with Lt show ?thesis by (simp add: Lt_lt_conv)
next
assume "k'' \<in> fst ` set ys"
hence "lt k'' k''" by (rule Cons(4))
thus ?thesis by simp
qed
qed
hence "f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ys k0) =
f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ((k'', v'') # ys) k0)"
by (simp add: lookup_pair_Cons[OF Cons(1)] del: lookup_pair.simps)
also from k0_in' \<open>lt k0 k\<close> have "... = Some Eq" by (rule Cons(8))
finally show "f k0 (lookup_pair ((k', v') # xs) k0) (lookup_pair ys k0) = Some Eq" .
qed
next
assume "f k'' 0 v'' \<noteq> Some Eq"
have "\<not> lt k'' k"
proof
have "k'' \<in> fst ` set ((k', v') # xs) \<union> fst ` set ((k'', v'') # ys)" by simp
moreover assume "lt k'' k"
ultimately have "f k'' (lookup_pair ((k', v') # xs) k'') (lookup_pair ((k'', v'') # ys) k'') = Some Eq"
by (rule Cons(8))
hence "f k'' 0 v'' = Some Eq" by (simp add: Lt)
with \<open>f k'' 0 v'' \<noteq> Some Eq\<close> show False ..
qed
moreover have "\<not> lt k k''"
proof
assume "lt k k''"
with Lt have "lt k k'" by (simp add: Lt_lt_conv)
thus False using \<open>lt k k''\<close> by (rule 0)
qed
ultimately have "k = k''" by simp
show "f k'' 0 v'' = aux" by (simp add: Cons(7) \<open>k = k''\<close> Lt)
qed
qed
qed
qed
lemma lex_ord_pair_EqD:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "lex_ord_pair f xs ys = Some Eq"
and "k \<in> fst ` set xs \<union> fst ` set ys"
shows "f k (lookup_pair xs k) (lookup_pair ys k) = Some Eq"
proof (rule ccontr)
let ?A = "(fst ` set xs \<union> fst ` set ys) \<inter> {k. f k (lookup_pair xs k) (lookup_pair ys k) \<noteq> Some Eq}"
define k0 where "k0 = Min ?A"
have "finite ?A" by auto
assume "f k (lookup_pair xs k) (lookup_pair ys k) \<noteq> Some Eq"
with assms(4) have "k \<in> ?A" by simp
hence "?A \<noteq> {}" by blast
with \<open>finite ?A\<close> have "k0 \<in> ?A" unfolding k0_def by (rule Min_in)
hence k0_in: "k0 \<in> fst ` set xs \<union> fst ` set ys"
and neq: "f k0 (lookup_pair xs k0) (lookup_pair ys k0) \<noteq> Some Eq" by simp_all
have "le k0 k'" if "k' \<in> ?A" for k' unfolding k0_def using \<open>finite ?A\<close> that
by (rule Min_le)
hence "f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq"
if "k' \<in> fst ` set xs \<union> fst ` set ys" and "lt k' k0" for k' using that by fastforce
- with assms(1, 2) neq k0_in refl have "lex_ord_pair f xs ys = f k0 (lookup_pair xs k0) (lookup_pair ys k0)"
+ with assms(1, 2) neq k0_in HOL.refl have "lex_ord_pair f xs ys = f k0 (lookup_pair xs k0) (lookup_pair ys k0)"
by (rule lex_ord_pair_valI)
with assms(3) neq show False by simp
qed
lemma lex_ord_pair_valE:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "lex_ord_pair f xs ys = aux"
and "aux \<noteq> Some Eq"
obtains k where "k \<in> fst ` set xs \<union> fst ` set ys" and "aux = f k (lookup_pair xs k) (lookup_pair ys k)"
and "\<And>k'. k' \<in> fst ` set xs \<union> fst ` set ys \<Longrightarrow> lt k' k \<Longrightarrow>
f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq"
proof -
let ?A = "(fst ` set xs \<union> fst ` set ys) \<inter> {k. f k (lookup_pair xs k) (lookup_pair ys k) \<noteq> Some Eq}"
define k where "k = Min ?A"
have "finite ?A" by auto
have "\<exists>k \<in> fst ` set xs \<union> fst ` set ys. f k (lookup_pair xs k) (lookup_pair ys k) \<noteq> Some Eq" (is ?prop)
proof (rule ccontr)
assume "\<not> ?prop"
hence "f k (lookup_pair xs k) (lookup_pair ys k) = Some Eq"
if "k \<in> fst ` set xs \<union> fst ` set ys" for k using that by auto
with assms(1, 2) have "lex_ord_pair f xs ys = Some Eq" by (rule lex_ord_pair_EqI)
with assms(3, 4) show False by simp
qed
then obtain k0 where "k0 \<in> fst ` set xs \<union> fst ` set ys"
and "f k0 (lookup_pair xs k0) (lookup_pair ys k0) \<noteq> Some Eq" ..
hence "k0 \<in> ?A" by simp
hence "?A \<noteq> {}" by blast
with \<open>finite ?A\<close> have "k \<in> ?A" unfolding k_def by (rule Min_in)
hence k_in: "k \<in> fst ` set xs \<union> fst ` set ys"
and neq: "f k (lookup_pair xs k) (lookup_pair ys k) \<noteq> Some Eq" by simp_all
have "le k k'" if "k' \<in> ?A" for k' unfolding k_def using \<open>finite ?A\<close> that
by (rule Min_le)
hence *: "\<And>k'. k' \<in> fst ` set xs \<union> fst ` set ys \<Longrightarrow> lt k' k \<Longrightarrow>
f k' (lookup_pair xs k') (lookup_pair ys k') = Some Eq" by fastforce
- with assms(1, 2) neq k_in refl have "lex_ord_pair f xs ys = f k (lookup_pair xs k) (lookup_pair ys k)"
+ with assms(1, 2) neq k_in HOL.refl have "lex_ord_pair f xs ys = f k (lookup_pair xs k) (lookup_pair ys k)"
by (rule lex_ord_pair_valI)
hence "aux = f k (lookup_pair xs k) (lookup_pair ys k)" by (simp only: assms(3))
with k_in show ?thesis using * ..
qed
subsubsection \<open>@{const prod_ord_pair}\<close>
lemma prod_ord_pair_eq_lex_ord_pair:
"prod_ord_pair P xs ys = (lex_ord_pair (\<lambda>k x y. if P k x y then Some Eq else None) xs ys = Some Eq)"
proof (induct P xs ys rule: prod_ord_pair.induct)
case (1 P)
show ?case by simp
next
case (2 P ky vy ys)
thus ?case by simp
next
case (3 P kx vx xs)
thus ?case by simp
next
case (4 P kx vx xs ky vy ys)
show ?case
proof (cases "comp kx ky")
case Lt
thus ?thesis by (simp add: 4(2)[OF Lt])
next
case Eq
thus ?thesis by (simp add: 4(1)[OF Eq])
next
case Gt
thus ?thesis by (simp add: 4(3)[OF Gt])
qed
qed
lemma prod_ord_pairI:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys"
and "\<And>k. k \<in> fst ` set xs \<union> fst ` set ys \<Longrightarrow> P k (lookup_pair xs k) (lookup_pair ys k)"
shows "prod_ord_pair P xs ys"
unfolding prod_ord_pair_eq_lex_ord_pair by (rule lex_ord_pair_EqI, fact, fact, simp add: assms(3))
lemma prod_ord_pairD:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys" and "prod_ord_pair P xs ys"
and "k \<in> fst ` set xs \<union> fst ` set ys"
shows "P k (lookup_pair xs k) (lookup_pair ys k)"
proof -
from assms have "(if P k (lookup_pair xs k) (lookup_pair ys k) then Some Eq else None) = Some Eq"
unfolding prod_ord_pair_eq_lex_ord_pair by (rule lex_ord_pair_EqD)
thus ?thesis by (simp split: if_splits)
qed
corollary prod_ord_pair_alt:
assumes "oalist_inv_raw xs" and "oalist_inv_raw ys"
shows "(prod_ord_pair P xs ys) \<longleftrightarrow> (\<forall>k\<in>fst ` set xs \<union> fst ` set ys. P k (lookup_pair xs k) (lookup_pair ys k))"
using prod_ord_pairI[OF assms] prod_ord_pairD[OF assms] by meson
subsubsection \<open>@{const sort_oalist}\<close>
lemma oalist_inv_raw_foldr_update_by_pair:
assumes "oalist_inv_raw ys"
shows "oalist_inv_raw (foldr update_by_pair xs ys)"
proof (induct xs)
case Nil
from assms show ?case by simp
next
case (Cons x xs)
hence "oalist_inv_raw (update_by_pair x (foldr update_by_pair xs ys))"
by (rule oalist_inv_raw_update_by_pair)
thus ?case by simp
qed
corollary oalist_inv_raw_sort_oalist: "oalist_inv_raw (sort_oalist xs)"
proof -
from oalist_inv_raw_Nil have "oalist_inv_raw (foldr local.update_by_pair xs [])"
by (rule oalist_inv_raw_foldr_update_by_pair)
thus "oalist_inv_raw (sort_oalist xs)" by (simp only: sort_oalist_def)
qed
lemma sort_oalist_id:
assumes "oalist_inv_raw xs"
shows "sort_oalist xs = xs"
proof -
have "foldr update_by_pair xs ys = xs @ ys" if "oalist_inv_raw (xs @ ys)" for ys using assms that
proof (induct xs rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k v xs)
from Cons(6) have *: "oalist_inv_raw ((k, v) # (xs @ ys))" by simp
hence 1: "oalist_inv_raw (xs @ ys)" by (rule oalist_inv_raw_ConsD1)
hence 2: "foldr update_by_pair xs ys = xs @ ys" by (rule Cons(5))
show ?case
proof (simp add: 2, rule update_by_pair_less)
from * show "v \<noteq> 0" by (auto simp: oalist_inv_raw_def)
next
have "comp k (fst (hd (xs @ ys))) = Lt \<or> xs @ ys = []"
proof (rule disjCI)
assume "xs @ ys \<noteq> []"
then obtain k'' v'' zs where eq0: "xs @ ys = (k'', v'') # zs"
using list.exhaust prod.exhaust by metis
from * have "lt k k''" by (simp add: eq0 oalist_inv_raw_def)
thus "comp k (fst (hd (xs @ ys))) = Lt" by (simp add: eq0 Lt_lt_conv)
qed
thus "xs @ ys = [] \<or> comp k (fst (hd (xs @ ys))) = Lt" by auto
qed
qed
with assms show ?thesis by (simp add: sort_oalist_def)
qed
lemma set_sort_oalist:
assumes "distinct (map fst xs)"
shows "set (sort_oalist xs) = {kv. kv \<in> set xs \<and> snd kv \<noteq> 0}"
using assms
proof (induct xs)
case Nil
show ?case by (simp add: sort_oalist_def)
next
case (Cons x xs)
obtain k v where x: "x = (k, v)" by fastforce
from Cons(2) have "distinct (map fst xs)" and "k \<notin> fst ` set xs" by (simp_all add: x)
from this(1) have "set (sort_oalist xs) = {kv \<in> set xs. snd kv \<noteq> 0}" by (rule Cons(1))
with \<open>k \<notin> fst ` set xs\<close> have eq: "set (sort_oalist xs) - range (Pair k) = {kv \<in> set xs. snd kv \<noteq> 0}"
by (auto simp: image_iff)
have "set (sort_oalist (x # xs)) = set (update_by_pair (k, v) (sort_oalist xs))"
by (simp add: sort_oalist_def x)
also have "... = {kv \<in> set (x # xs). snd kv \<noteq> 0}"
proof (cases "v = 0")
case True
have "set (update_by_pair (k, v) (sort_oalist xs)) = set (sort_oalist xs) - range (Pair k)"
unfolding True using oalist_inv_raw_sort_oalist by (rule set_update_by_pair_zero)
also have "... = {kv \<in> set (x # xs). snd kv \<noteq> 0}" by (auto simp: eq x True)
finally show ?thesis .
next
case False
with oalist_inv_raw_sort_oalist
have "set (update_by_pair (k, v) (sort_oalist xs)) = insert (k, v) (set (sort_oalist xs) - range (Pair k))"
by (rule set_update_by_pair)
also have "... = {kv \<in> set (x # xs). snd kv \<noteq> 0}" by (auto simp: eq x False)
finally show ?thesis .
qed
finally show ?case .
qed
lemma lookup_pair_sort_oalist':
assumes "distinct (map fst xs)"
shows "lookup_pair (sort_oalist xs) = lookup_dflt xs"
using assms
proof (induct xs)
case Nil
show ?case by (simp add: sort_oalist_def lookup_dflt_def)
next
case (Cons x xs)
obtain k v where x: "x = (k, v)" by fastforce
from Cons(2) have "distinct (map fst xs)" and "k \<notin> fst ` set xs" by (simp_all add: x)
from this(1) have eq1: "lookup_pair (sort_oalist xs) = lookup_dflt xs" by (rule Cons(1))
have eq2: "sort_oalist (x # xs) = update_by_pair (k, v) (sort_oalist xs)" by (simp add: x sort_oalist_def)
show ?case
proof
fix k'
have "lookup_pair (sort_oalist (x # xs)) k' = (if k = k' then v else lookup_dflt xs k')"
by (simp add: eq1 eq2 lookup_pair_update_by_pair[OF oalist_inv_raw_sort_oalist])
also have "... = lookup_dflt (x # xs) k'" by (simp add: x lookup_dflt_def)
finally show "lookup_pair (sort_oalist (x # xs)) k' = lookup_dflt (x # xs) k'" .
qed
qed
end
locale comparator2 = comparator comp1 + cmp2: comparator comp2 for comp1 comp2 :: "'a comparator"
begin
lemma set_sort_oalist:
assumes "cmp2.oalist_inv_raw xs"
shows "set (sort_oalist xs) = set xs"
proof -
have rl: "set (foldr update_by_pair xs ys) = set xs \<union> set ys"
if "oalist_inv_raw ys" and "fst ` set xs \<inter> fst ` set ys = {}" for ys
using assms that(2)
proof (induct xs rule: cmp2.oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k v xs)
from Cons(6) have "k \<notin> fst ` set ys" and "fst ` set xs \<inter> fst ` set ys = {}" by simp_all
from this(2) have eq1: "set (foldr update_by_pair xs ys) = set xs \<union> set ys" by (rule Cons(5))
have "\<not> cmp2.lt k k" by auto
with Cons(4) have "k \<notin> fst ` set xs" by blast
with \<open>k \<notin> fst ` set ys\<close> have "k \<notin> fst ` (set xs \<union> set ys)" by (simp add: image_Un)
hence "(set xs \<union> set ys) \<inter> range (Pair k) = {}" by (smt Int_emptyI fstI image_iff)
hence eq2: "(set xs \<union> set ys) - range (Pair k) = set xs \<union> set ys" by (rule Diff_triv)
from \<open>oalist_inv_raw ys\<close> have "oalist_inv_raw (foldr update_by_pair xs ys)"
by (rule oalist_inv_raw_foldr_update_by_pair)
hence "set (update_by_pair (k, v) (foldr update_by_pair xs ys)) =
insert (k, v) (set (foldr update_by_pair xs ys) - range (Pair k))"
using Cons(3) by (rule set_update_by_pair)
also have "... = insert (k, v) (set xs \<union> set ys)" by (simp only: eq1 eq2)
finally show ?case by simp
qed
have "set (foldr update_by_pair xs []) = set xs \<union> set []"
by (rule rl, fact oalist_inv_raw_Nil, simp)
thus ?thesis by (simp add: sort_oalist_def)
qed
lemma lookup_pair_eqI:
assumes "oalist_inv_raw xs" and "cmp2.oalist_inv_raw ys" and "set xs = set ys"
shows "lookup_pair xs = cmp2.lookup_pair ys"
proof
fix k
show "lookup_pair xs k = cmp2.lookup_pair ys k"
proof (cases "cmp2.lookup_pair ys k = 0")
case True
with assms(2) have "k \<notin> fst ` set ys" by (simp add: cmp2.lookup_pair_eq_0)
with assms(1) show ?thesis by (simp add: True assms(3)[symmetric] lookup_pair_eq_0)
next
case False
define v where "v = cmp2.lookup_pair ys k"
from False have "v \<noteq> 0" by (simp add: v_def)
with assms(2) v_def[symmetric] have "(k, v) \<in> set ys" by (simp add: cmp2.lookup_pair_eq_value)
with assms(1) \<open>v \<noteq> 0\<close> have "lookup_pair xs k = v"
by (simp add: assms(3)[symmetric] lookup_pair_eq_value)
thus ?thesis by (simp only: v_def)
qed
qed
corollary lookup_pair_sort_oalist:
assumes "cmp2.oalist_inv_raw xs"
shows "lookup_pair (sort_oalist xs) = cmp2.lookup_pair xs"
by (rule lookup_pair_eqI, rule oalist_inv_raw_sort_oalist, fact, rule set_sort_oalist, fact)
end (* comparator2 *)
subsection \<open>Invariant on Pairs\<close>
type_synonym ('a, 'b, 'c) oalist_raw = "('a \<times> 'b) list \<times> 'c"
locale oalist_raw = fixes rep_key_order::"'o \<Rightarrow> 'a key_order"
begin
sublocale comparator "key_compare (rep_key_order x)"
by (fact comparator_key_compare)
definition oalist_inv :: "('a, 'b::zero, 'o) oalist_raw \<Rightarrow> bool"
where "oalist_inv xs \<longleftrightarrow> oalist_inv_raw (snd xs) (fst xs)"
lemma oalist_inv_alt: "oalist_inv (xs, ko) \<longleftrightarrow> oalist_inv_raw ko xs"
by (simp add: oalist_inv_def)
subsection \<open>Operations on Raw Ordered Associative Lists\<close>
fun sort_oalist_aux :: "'o \<Rightarrow> ('a, 'b, 'o) oalist_raw \<Rightarrow> ('a \<times> 'b::zero) list"
where "sort_oalist_aux ko (xs, ox) = (if ko = ox then xs else sort_oalist ko xs)"
fun lookup_raw :: "('a, 'b, 'o) oalist_raw \<Rightarrow> 'a \<Rightarrow> 'b::zero"
where "lookup_raw (xs, ko) = lookup_pair ko xs"
definition sorted_domain_raw :: "'o \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw \<Rightarrow> 'a list"
where "sorted_domain_raw ko xs = map fst (sort_oalist_aux ko xs)"
fun tl_raw :: "('a, 'b, 'o) oalist_raw \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw"
where "tl_raw (xs, ko) = (List.tl xs, ko)"
fun min_key_val_raw :: "'o \<Rightarrow> ('a, 'b, 'o) oalist_raw \<Rightarrow> ('a \<times> 'b::zero)"
where "min_key_val_raw ko (xs, ox) =
(if ko = ox then List.hd else min_list_param (\<lambda>x y. le ko (fst x) (fst y))) xs"
fun update_by_raw :: "('a \<times> 'b) \<Rightarrow> ('a, 'b, 'o) oalist_raw \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw"
where "update_by_raw kv (xs, ko) = (update_by_pair ko kv xs, ko)"
fun update_by_fun_raw :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b, 'o) oalist_raw \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw"
where "update_by_fun_raw k f (xs, ko) = (update_by_fun_pair ko k f xs, ko)"
fun update_by_fun_gr_raw :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b, 'o) oalist_raw \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw"
where "update_by_fun_gr_raw k f (xs, ko) = (update_by_fun_gr_pair ko k f xs, ko)"
fun (in -) filter_raw :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'b) \<Rightarrow> ('a list \<times> 'b)"
where "filter_raw P (xs, ko) = (filter P xs, ko)"
fun (in -) map_raw :: "(('a \<times> 'b) \<Rightarrow> ('a \<times> 'c)) \<Rightarrow> (('a \<times> 'b::zero) list \<times> 'd) \<Rightarrow> ('a \<times> 'c::zero) list \<times> 'd"
where "map_raw f (xs, ko) = (map_pair f xs, ko)"
abbreviation (in -) "map_val_raw f \<equiv> map_raw (\<lambda>(k, v). (k, f k v))"
fun map2_val_raw :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> (('a, 'b, 'o) oalist_raw \<Rightarrow> ('a, 'd, 'o) oalist_raw) \<Rightarrow>
(('a, 'c, 'o) oalist_raw \<Rightarrow> ('a, 'd, 'o) oalist_raw) \<Rightarrow>
('a, 'b::zero, 'o) oalist_raw \<Rightarrow> ('a, 'c::zero, 'o) oalist_raw \<Rightarrow>
('a, 'd::zero, 'o) oalist_raw"
where "map2_val_raw f g h (xs, ox) ys =
(map2_val_pair ox f (\<lambda>zs. fst (g (zs, ox))) (\<lambda>zs. fst (h (zs, ox)))
xs (sort_oalist_aux ox ys), ox)"
definition lex_ord_raw :: "'o \<Rightarrow> ('a \<Rightarrow> (('b, 'c) comp_opt)) \<Rightarrow>
(('a, 'b::zero, 'o) oalist_raw, ('a, 'c::zero, 'o) oalist_raw) comp_opt"
where "lex_ord_raw ko f xs ys = lex_ord_pair ko f (sort_oalist_aux ko xs) (sort_oalist_aux ko ys)"
fun prod_ord_raw :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw \<Rightarrow>
('a, 'c::zero, 'o) oalist_raw \<Rightarrow> bool"
where "prod_ord_raw f (xs, ox) ys = prod_ord_pair ox f xs (sort_oalist_aux ox ys)"
fun oalist_eq_raw :: "('a, 'b, 'o) oalist_raw \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw \<Rightarrow> bool"
where "oalist_eq_raw (xs, ox) ys = (xs = (sort_oalist_aux ox ys))"
fun sort_oalist_raw :: "('a, 'b, 'o) oalist_raw \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw"
where "sort_oalist_raw (xs, ko) = (sort_oalist ko xs, ko)"
subsubsection \<open>@{const sort_oalist_aux}\<close>
lemma set_sort_oalist_aux:
assumes "oalist_inv xs"
shows "set (sort_oalist_aux ko xs) = set (fst xs)"
proof -
obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce
interpret ko2: comparator2 "key_compare (rep_key_order ko)" "key_compare (rep_key_order ko')" ..
from assms show ?thesis by (simp add: xs oalist_inv_alt ko2.set_sort_oalist)
qed
lemma oalist_inv_raw_sort_oalist_aux:
assumes "oalist_inv xs"
shows "oalist_inv_raw ko (sort_oalist_aux ko xs)"
proof -
obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce
from assms show ?thesis by (simp add: xs oalist_inv_alt oalist_inv_raw_sort_oalist)
qed
lemma oalist_inv_sort_oalist_aux:
assumes "oalist_inv xs"
shows "oalist_inv (sort_oalist_aux ko xs, ko)"
unfolding oalist_inv_alt using assms by (rule oalist_inv_raw_sort_oalist_aux)
lemma lookup_pair_sort_oalist_aux:
assumes "oalist_inv xs"
shows "lookup_pair ko (sort_oalist_aux ko xs) = lookup_raw xs"
proof -
obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce
interpret ko2: comparator2 "key_compare (rep_key_order ko)" "key_compare (rep_key_order ko')" ..
from assms show ?thesis by (simp add: xs oalist_inv_alt ko2.lookup_pair_sort_oalist)
qed
subsubsection \<open>@{const lookup_raw}\<close>
lemma lookup_raw_eq_value:
assumes "oalist_inv xs" and "v \<noteq> 0"
shows "lookup_raw xs k = v \<longleftrightarrow> ((k, v) \<in> set (fst xs))"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_def)
show ?thesis by (simp add: xs, rule lookup_pair_eq_value, fact+)
qed
lemma lookup_raw_eq_valueI:
assumes "oalist_inv xs" and "(k, v) \<in> set (fst xs)"
shows "lookup_raw xs k = v"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_def)
from assms(2) have "(k, v) \<in> set xs'" by (simp add: xs)
show ?thesis by (simp add: xs, rule lookup_pair_eq_valueI, fact+)
qed
lemma lookup_raw_inj:
assumes "oalist_inv (xs, ko)" and "oalist_inv (ys, ko)" and "lookup_raw (xs, ko) = lookup_raw (ys, ko)"
shows "xs = ys"
using assms unfolding lookup_raw.simps oalist_inv_alt by (rule lookup_pair_inj)
subsubsection \<open>@{const sorted_domain_raw}\<close>
lemma set_sorted_domain_raw:
assumes "oalist_inv xs"
shows "set (sorted_domain_raw ko xs) = fst ` set (fst xs)"
using assms by (simp add: sorted_domain_raw_def set_sort_oalist_aux)
corollary in_sorted_domain_raw_iff_lookup_raw:
assumes "oalist_inv xs"
shows "k \<in> set (sorted_domain_raw ko xs) \<longleftrightarrow> (lookup_raw xs k \<noteq> 0)"
unfolding set_sorted_domain_raw[OF assms]
proof -
obtain xs' ko' where xs: "xs = (xs', ko')" by fastforce
from assms show "k \<in> fst ` set (fst xs) \<longleftrightarrow> (lookup_raw xs k \<noteq> 0)"
by (simp add: xs oalist_inv_alt lookup_pair_eq_0)
qed
lemma sorted_sorted_domain_raw:
assumes "oalist_inv xs"
shows "sorted_wrt (lt_of_key_order (rep_key_order ko)) (sorted_domain_raw ko xs)"
unfolding sorted_domain_raw_def oalist_inv_alt lt_of_key_order.rep_eq
by (rule oalist_inv_rawD2, rule oalist_inv_raw_sort_oalist_aux, fact)
subsubsection \<open>@{const tl_raw}\<close>
lemma oalist_inv_tl_raw:
assumes "oalist_inv xs"
shows "oalist_inv (tl_raw xs)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis unfolding xs tl_raw.simps oalist_inv_alt by (rule oalist_inv_raw_tl)
qed
lemma lookup_raw_tl_raw:
assumes "oalist_inv xs"
shows "lookup_raw (tl_raw xs) k =
(if (\<forall>k'\<in>fst ` set (fst xs). le (snd xs) k k') then 0 else lookup_raw xs k)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis by (simp add: xs lookup_pair_tl oalist_inv_alt split del: if_split cong: if_cong)
qed
lemma lookup_raw_tl_raw':
assumes "oalist_inv xs"
shows "lookup_raw (tl_raw xs) k = (if k = fst (List.hd (fst xs)) then 0 else lookup_raw xs k)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis by (simp add: xs lookup_pair_tl' oalist_inv_alt)
qed
subsubsection \<open>@{const min_key_val_raw}\<close>
lemma min_key_val_raw_alt:
assumes "oalist_inv xs" and "fst xs \<noteq> []"
shows "min_key_val_raw ko xs = List.hd (sort_oalist_aux ko xs)"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(2) have "xs' \<noteq> []" by (simp add: xs)
interpret ko2: comparator2 "key_compare (rep_key_order ko)" "key_compare (rep_key_order ox)" ..
from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt)
hence set_sort: "set (sort_oalist ko xs') = set xs'" by (rule ko2.set_sort_oalist)
also from \<open>xs' \<noteq> []\<close> have "... \<noteq> {}" by simp
finally have "sort_oalist ko xs' \<noteq> []" by simp
then obtain k v xs'' where eq: "sort_oalist ko xs' = (k, v) # xs''"
using prod.exhaust list.exhaust by metis
hence "(k, v) \<in> set xs'" by (simp add: set_sort[symmetric])
have *: "le ko k k'" if "k' \<in> fst ` set xs'" for k'
proof -
from that have "k' = k \<or> k' \<in> fst ` set xs''" by (simp add: set_sort[symmetric] eq)
thus ?thesis
proof
assume "k' = k"
thus ?thesis by simp
next
have "oalist_inv_raw ko ((k, v) # xs'')" unfolding eq[symmetric] by (fact oalist_inv_raw_sort_oalist)
moreover assume "k' \<in> fst ` set xs''"
ultimately have "lt ko k k'" by (rule oalist_inv_raw_ConsD3)
thus ?thesis by simp
qed
qed
from \<open>xs' \<noteq> []\<close> have "min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs' \<in> set xs'" by (rule min_list_param_in)
with \<open>oalist_inv_raw ox xs'\<close> have "lookup_pair ox xs' (fst (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs')) =
snd (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs')" by (auto intro: lookup_pair_eq_valueI)
moreover have 1: "fst (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs') = k"
proof (rule antisym)
- from order.trans
+ from order_trans
have "transp (\<lambda>x y. le ko (fst x) (fst y))" by (rule transpI)
hence "le ko (fst (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs')) (fst (k, v))"
using linear \<open>(k, v) \<in> set xs'\<close> by (rule min_list_param_minimal)
thus "le ko (fst (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs')) k" by simp
next
show "le ko k (fst (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs'))" by (rule *, rule imageI, fact)
qed
ultimately have "snd (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs') = lookup_pair ox xs' k"
by simp
also from \<open>oalist_inv_raw ox xs'\<close> \<open>(k, v) \<in> set xs'\<close> have "... = v" by (rule lookup_pair_eq_valueI)
finally have "snd (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs') = v" .
with 1 have "min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs' = (k, v)" by auto
thus ?thesis by (simp add: xs eq)
qed
lemma min_key_val_raw_in:
assumes "fst xs \<noteq> []"
shows "min_key_val_raw ko xs \<in> set (fst xs)"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms have "xs' \<noteq> []" by (simp add: xs)
show ?thesis unfolding xs
proof (simp, intro conjI impI)
from \<open>xs' \<noteq> []\<close> show "hd xs' \<in> set xs'" by simp
next
from \<open>xs' \<noteq> []\<close> show "min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs' \<in> set xs'"
by (rule min_list_param_in)
qed
qed
lemma snd_min_key_val_raw:
assumes "oalist_inv xs" and "fst xs \<noteq> []"
shows "snd (min_key_val_raw ko xs) = lookup_raw xs (fst (min_key_val_raw ko xs))"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt)
from assms(2) have "min_key_val_raw ko xs \<in> set (fst xs)" by (rule min_key_val_raw_in)
hence *: "min_key_val_raw ko (xs', ox) \<in> set xs'" by (simp add: xs)
show ?thesis unfolding xs lookup_raw.simps
by (rule HOL.sym, rule lookup_pair_eq_valueI, fact, simp add: * del: min_key_val_raw.simps)
qed
lemma min_key_val_raw_minimal:
assumes "oalist_inv xs" and "z \<in> set (fst xs)"
shows "le ko (fst (min_key_val_raw ko xs)) (fst z)"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms have "oalist_inv (xs', ox)" and "z \<in> set xs'" by (simp_all add: xs)
show ?thesis unfolding xs
proof (simp, intro conjI impI)
from \<open>z \<in> set xs'\<close> have "xs' \<noteq> []" by auto
then obtain k v ys where xs': "xs' = (k, v) # ys" using prod.exhaust list.exhaust by metis
from \<open>z \<in> set xs'\<close> have "z = (k, v) \<or> z \<in> set ys" by (simp add: xs')
thus "le ox (fst (hd xs')) (fst z)"
proof
assume "z = (k, v)"
show ?thesis by (simp add: xs' \<open>z = (k, v)\<close>)
next
assume "z \<in> set ys"
hence "fst z \<in> fst ` set ys" by fastforce
with \<open>oalist_inv (xs', ox)\<close> have "lt ox k (fst z)"
unfolding xs' oalist_inv_alt lt_of_key_order.rep_eq by (rule oalist_inv_raw_ConsD3)
thus ?thesis by (simp add: xs')
qed
next
show "le ko (fst (min_list_param (\<lambda>x y. le ko (fst x) (fst y)) xs')) (fst z)"
proof (rule min_list_param_minimal[of "\<lambda>x y. le ko (fst x) (fst y)"])
- show "transp (\<lambda>x y. le ko (fst x) (fst y))" by (metis (no_types, lifting) order.trans transpI)
+ thm trans local.trans order.trans local.order_trans
+ print_context
+ show "transp (\<lambda>x y. le ko (fst x) (fst y))" by (metis (no_types, lifting) order_trans transpI)
qed (auto intro: \<open>z \<in> set xs'\<close>)
qed
qed
subsubsection \<open>@{const filter_raw}\<close>
lemma oalist_inv_filter_raw:
assumes "oalist_inv xs"
shows "oalist_inv (filter_raw P xs)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis unfolding xs filter_raw.simps oalist_inv_alt
by (rule oalist_inv_raw_filter)
qed
lemma lookup_raw_filter_raw:
assumes "oalist_inv xs"
shows "lookup_raw (filter_raw P xs) k = (let v = lookup_raw xs k in if P (k, v) then v else 0)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis unfolding xs lookup_raw.simps filter_raw.simps oalist_inv_alt
by (rule lookup_pair_filter)
qed
subsubsection \<open>@{const update_by_raw}\<close>
lemma oalist_inv_update_by_raw:
assumes "oalist_inv xs"
shows "oalist_inv (update_by_raw kv xs)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis unfolding xs update_by_raw.simps oalist_inv_alt
by (rule oalist_inv_raw_update_by_pair)
qed
lemma lookup_raw_update_by_raw:
assumes "oalist_inv xs"
shows "lookup_raw (update_by_raw (k1, v) xs) k2 = (if k1 = k2 then v else lookup_raw xs k2)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis unfolding xs lookup_raw.simps update_by_raw.simps oalist_inv_alt
by (rule lookup_pair_update_by_pair)
qed
subsubsection \<open>@{const update_by_fun_raw} and @{const update_by_fun_gr_raw}\<close>
lemma update_by_fun_raw_eq_update_by_raw:
assumes "oalist_inv xs"
shows "update_by_fun_raw k f xs = update_by_raw (k, f (lookup_raw xs k)) xs"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt)
show ?thesis unfolding xs update_by_fun_raw.simps lookup_raw.simps update_by_raw.simps
- by (rule, rule conjI, rule update_by_fun_pair_eq_update_by_pair, fact, fact refl)
+ by (rule, rule conjI, rule update_by_fun_pair_eq_update_by_pair, fact, fact HOL.refl)
qed
corollary oalist_inv_update_by_fun_raw:
assumes "oalist_inv xs"
shows "oalist_inv (update_by_fun_raw k f xs)"
unfolding update_by_fun_raw_eq_update_by_raw[OF assms] using assms by (rule oalist_inv_update_by_raw)
corollary lookup_raw_update_by_fun_raw:
assumes "oalist_inv xs"
shows "lookup_raw (update_by_fun_raw k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_raw xs k2)"
using assms by (simp add: update_by_fun_raw_eq_update_by_raw lookup_raw_update_by_raw)
lemma update_by_fun_gr_raw_eq_update_by_fun_raw:
assumes "oalist_inv xs"
shows "update_by_fun_gr_raw k f xs = update_by_fun_raw k f xs"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt)
show ?thesis unfolding xs update_by_fun_raw.simps lookup_raw.simps update_by_fun_gr_raw.simps
- by (rule, rule conjI, rule update_by_fun_gr_pair_eq_update_by_fun_pair, fact, fact refl)
+ by (rule, rule conjI, rule update_by_fun_gr_pair_eq_update_by_fun_pair, fact, fact HOL.refl)
qed
corollary oalist_inv_update_by_fun_gr_raw:
assumes "oalist_inv xs"
shows "oalist_inv (update_by_fun_gr_raw k f xs)"
unfolding update_by_fun_gr_raw_eq_update_by_fun_raw[OF assms] using assms by (rule oalist_inv_update_by_fun_raw)
corollary lookup_raw_update_by_fun_gr_raw:
assumes "oalist_inv xs"
shows "lookup_raw (update_by_fun_gr_raw k1 f xs) k2 = (if k1 = k2 then f else id) (lookup_raw xs k2)"
using assms by (simp add: update_by_fun_gr_raw_eq_update_by_fun_raw lookup_raw_update_by_fun_raw)
subsubsection \<open>@{const map_raw} and @{const map_val_raw}\<close>
lemma map_raw_cong:
assumes "\<And>kv. kv \<in> set (fst xs) \<Longrightarrow> f kv = g kv"
shows "map_raw f xs = map_raw g xs"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms have "f kv = g kv" if "kv \<in> set xs'" for kv using that by (simp add: xs)
thus ?thesis by (simp add: xs, rule map_pair_cong)
qed
lemma map_raw_subset: "set (fst (map_raw f xs)) \<subseteq> f ` set (fst xs)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
show ?thesis by (simp add: xs map_pair_subset)
qed
lemma oalist_inv_map_raw:
assumes "oalist_inv xs"
and "\<And>a b. key_compare (rep_key_order (snd xs)) (fst (f a)) (fst (f b)) = key_compare (rep_key_order (snd xs)) (fst a) (fst b)"
shows "oalist_inv (map_raw f xs)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms(1) have "oalist_inv (xs', ko)" by (simp only: xs)
moreover from assms(2)
have "\<And>a b. key_compare (rep_key_order ko) (fst (f a)) (fst (f b)) = key_compare (rep_key_order ko) (fst a) (fst b)"
by (simp add: xs)
ultimately show ?thesis unfolding xs map_raw.simps oalist_inv_alt by (rule oalist_inv_raw_map_pair)
qed
lemma lookup_raw_map_raw:
assumes "oalist_inv xs" and "snd (f (k, 0)) = 0"
and "\<And>a b. key_compare (rep_key_order (snd xs)) (fst (f a)) (fst (f b)) = key_compare (rep_key_order (snd xs)) (fst a) (fst b)"
shows "lookup_raw (map_raw f xs) (fst (f (k, v))) = snd (f (k, lookup_raw xs k))"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms(1) have "oalist_inv (xs', ko)" by (simp only: xs)
moreover note assms(2)
moreover from assms(3)
have "\<And>a b. key_compare (rep_key_order ko) (fst (f a)) (fst (f b)) = key_compare (rep_key_order ko) (fst a) (fst b)"
by (simp add: xs)
ultimately show ?thesis unfolding xs lookup_raw.simps map_raw.simps oalist_inv_alt
by (rule lookup_pair_map_pair)
qed
lemma map_raw_id:
assumes "oalist_inv xs"
shows "map_raw id xs = xs"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt)
hence "map_pair id xs' = xs'"
proof (induct xs' rule: oalist_inv_raw_induct)
case Nil
show ?case by simp
next
case (Cons k v xs')
show ?case by (simp add: Let_def Cons(3, 5) id_def[symmetric])
qed
thus ?thesis by (simp add: xs)
qed
lemma map_val_raw_cong:
assumes "\<And>k v. (k, v) \<in> set (fst xs) \<Longrightarrow> f k v = g k v"
shows "map_val_raw f xs = map_val_raw g xs"
proof (rule map_raw_cong)
fix kv
assume "kv \<in> set (fst xs)"
moreover obtain k v where "kv = (k, v)" by fastforce
ultimately show "(case kv of (k, v) \<Rightarrow> (k, f k v)) = (case kv of (k, v) \<Rightarrow> (k, g k v))"
by (simp add: assms)
qed
lemma oalist_inv_map_val_raw:
assumes "oalist_inv xs"
shows "oalist_inv (map_val_raw f xs)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis unfolding xs map_raw.simps oalist_inv_alt by (rule oalist_inv_raw_map_val_pair)
qed
lemma lookup_raw_map_val_raw:
assumes "oalist_inv xs" and "f k 0 = 0"
shows "lookup_raw (map_val_raw f xs) k = f k (lookup_raw xs k)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms show ?thesis unfolding xs map_raw.simps lookup_raw.simps oalist_inv_alt
by (rule lookup_pair_map_val_pair)
qed
subsubsection \<open>@{const map2_val_raw}\<close>
definition map2_val_compat' :: "(('a, 'b::zero, 'o) oalist_raw \<Rightarrow> ('a, 'c::zero, 'o) oalist_raw) \<Rightarrow> bool"
where "map2_val_compat' f \<longleftrightarrow>
(\<forall>zs. (oalist_inv zs \<longrightarrow> (oalist_inv (f zs) \<and> snd (f zs) = snd zs \<and> fst ` set (fst (f zs)) \<subseteq> fst ` set (fst zs))))"
lemma map2_val_compat'I:
assumes "\<And>zs. oalist_inv zs \<Longrightarrow> oalist_inv (f zs)"
and "\<And>zs. oalist_inv zs \<Longrightarrow> snd (f zs) = snd zs"
and "\<And>zs. oalist_inv zs \<Longrightarrow> fst ` set (fst (f zs)) \<subseteq> fst ` set (fst zs)"
shows "map2_val_compat' f"
unfolding map2_val_compat'_def using assms by blast
lemma map2_val_compat'D1:
assumes "map2_val_compat' f" and "oalist_inv zs"
shows "oalist_inv (f zs)"
using assms unfolding map2_val_compat'_def by blast
lemma map2_val_compat'D2:
assumes "map2_val_compat' f" and "oalist_inv zs"
shows "snd (f zs) = snd zs"
using assms unfolding map2_val_compat'_def by blast
lemma map2_val_compat'D3:
assumes "map2_val_compat' f" and "oalist_inv zs"
shows "fst ` set (fst (f zs)) \<subseteq> fst ` set (fst zs)"
using assms unfolding map2_val_compat'_def by blast
lemma map2_val_compat'_map_val_raw: "map2_val_compat' (map_val_raw f)"
proof (rule map2_val_compat'I, erule oalist_inv_map_val_raw)
fix zs::"('a, 'b, 'o) oalist_raw"
obtain zs' ko where "zs = (zs', ko)" by fastforce
thus "snd (map_val_raw f zs) = snd zs" by simp
next
fix zs::"('a, 'b, 'o) oalist_raw"
obtain zs' ko where zs: "zs = (zs', ko)" by fastforce
show "fst ` set (fst (map_val_raw f zs)) \<subseteq> fst ` set (fst zs)"
proof (simp add: zs)
from map_pair_subset have "fst ` set (map_val_pair f zs') \<subseteq> fst ` (\<lambda>(k, v). (k, f k v)) ` set zs'"
by (rule image_mono)
also have "... = fst ` set zs'" by force
finally show "fst ` set (map_val_pair f zs') \<subseteq> fst ` set zs'" .
qed
qed
lemma map2_val_compat'_id: "map2_val_compat' id"
by (rule map2_val_compat'I, auto)
lemma map2_val_compat'_imp_map2_val_compat:
assumes "map2_val_compat' g"
shows "map2_val_compat ko (\<lambda>zs. fst (g (zs, ko)))"
proof (rule map2_val_compatI)
fix zs::"('a \<times> 'b) list"
assume a: "oalist_inv_raw ko zs"
hence "oalist_inv (zs, ko)" by (simp only: oalist_inv_alt)
with assms have "oalist_inv (g (zs, ko))" by (rule map2_val_compat'D1)
hence "oalist_inv (fst (g (zs, ko)), snd (g (zs, ko)))" by simp
thus "oalist_inv_raw ko (fst (g (zs, ko)))" using assms a by (simp add: oalist_inv_alt map2_val_compat'D2)
next
fix zs::"('a \<times> 'b) list"
assume a: "oalist_inv_raw ko zs"
hence "oalist_inv (zs, ko)" by (simp only: oalist_inv_alt)
with assms have "fst ` set (fst (g (zs, ko))) \<subseteq> fst ` set (fst (zs, ko))" by (rule map2_val_compat'D3)
thus "fst ` set (fst (g (zs, ko))) \<subseteq> fst ` set zs" by simp
qed
lemma oalist_inv_map2_val_raw:
assumes "oalist_inv xs" and "oalist_inv ys"
assumes "map2_val_compat' g" and "map2_val_compat' h"
shows "oalist_inv (map2_val_raw f g h xs ys)"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
let ?ys = "sort_oalist_aux ox ys"
from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_alt)
moreover from assms(2) have "oalist_inv_raw ox (sort_oalist_aux ox ys)"
by (rule oalist_inv_raw_sort_oalist_aux)
moreover from assms(3) have "map2_val_compat ko (\<lambda>zs. fst (g (zs, ko)))" for ko
by (rule map2_val_compat'_imp_map2_val_compat)
moreover from assms(4) have "map2_val_compat ko (\<lambda>zs. fst (h (zs, ko)))" for ko
by (rule map2_val_compat'_imp_map2_val_compat)
ultimately have "oalist_inv_raw ox (map2_val_pair ox f (\<lambda>zs. fst (g (zs, ox))) (\<lambda>zs. fst (h (zs, ox))) xs' ?ys)"
by (rule oalist_inv_raw_map2_val_pair)
thus ?thesis by (simp add: xs oalist_inv_alt)
qed
lemma lookup_raw_map2_val_raw:
assumes "oalist_inv xs" and "oalist_inv ys"
assumes "map2_val_compat' g" and "map2_val_compat' h"
assumes "\<And>zs. oalist_inv zs \<Longrightarrow> g zs = map_val_raw (\<lambda>k v. f k v 0) zs"
and "\<And>zs. oalist_inv zs \<Longrightarrow> h zs = map_val_raw (\<lambda>k. f k 0) zs"
and "\<And>k. f k 0 0 = 0"
shows "lookup_raw (map2_val_raw f g h xs ys) k0 = f k0 (lookup_raw xs k0) (lookup_raw ys k0)"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
let ?ys = "sort_oalist_aux ox ys"
from assms(1) have "oalist_inv_raw ox xs'" by (simp add: xs oalist_inv_alt)
moreover from assms(2) have "oalist_inv_raw ox (sort_oalist_aux ox ys)" by (rule oalist_inv_raw_sort_oalist_aux)
moreover from assms(3) have "map2_val_compat ko (\<lambda>zs. fst (g (zs, ko)))" for ko
by (rule map2_val_compat'_imp_map2_val_compat)
moreover from assms(4) have "map2_val_compat ko (\<lambda>zs. fst (h (zs, ko)))" for ko
by (rule map2_val_compat'_imp_map2_val_compat)
ultimately have "lookup_pair ox (map2_val_pair ox f (\<lambda>zs. fst (g (zs, ox))) (\<lambda>zs. fst (h (zs, ox))) xs' ?ys) k0 =
f k0 (lookup_pair ox xs' k0) (lookup_pair ox ?ys k0)" using _ _ assms(7)
proof (rule lookup_pair_map2_val_pair)
fix zs::"('a \<times> 'b) list"
assume "oalist_inv_raw ox zs"
hence "oalist_inv (zs, ox)" by (simp only: oalist_inv_alt)
hence "g (zs, ox) = map_val_raw (\<lambda>k v. f k v 0) (zs, ox)" by (rule assms(5))
thus "fst (g (zs, ox)) = map_val_pair (\<lambda>k v. f k v 0) zs" by simp
next
fix zs::"('a \<times> 'c) list"
assume "oalist_inv_raw ox zs"
hence "oalist_inv (zs, ox)" by (simp only: oalist_inv_alt)
hence "h (zs, ox) = map_val_raw (\<lambda>k. f k 0) (zs, ox)" by (rule assms(6))
thus "fst (h (zs, ox)) = map_val_pair (\<lambda>k. f k 0) zs" by simp
qed
also from assms(2) have "... = f k0 (lookup_pair ox xs' k0) (lookup_raw ys k0)"
by (simp only: lookup_pair_sort_oalist_aux)
finally have *: "lookup_pair ox (map2_val_pair ox f (\<lambda>zs. fst (g (zs, ox))) (\<lambda>zs. fst (h (zs, ox))) xs' ?ys) k0 =
f k0 (lookup_pair ox xs' k0) (lookup_raw ys k0)" .
thus ?thesis by (simp add: xs)
qed
lemma map2_val_raw_singleton_eq_update_by_fun_raw:
assumes "oalist_inv xs"
assumes "\<And>k x. f k x 0 = x" and "\<And>zs. oalist_inv zs \<Longrightarrow> g zs = zs"
and "\<And>ko. h ([(k, v)], ko) = map_val_raw (\<lambda>k. f k 0) ([(k, v)], ko)"
shows "map2_val_raw f g h xs ([(k, v)], ko) = update_by_fun_raw k (\<lambda>x. f k x v) xs"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
let ?ys = "sort_oalist ox [(k, v)]"
from assms(1) have inv: "oalist_inv (xs', ox)" by (simp only: xs)
hence inv_raw: "oalist_inv_raw ox xs'" by (simp only: oalist_inv_alt)
show ?thesis
proof (simp add: xs, intro conjI impI)
assume "ox = ko"
from inv_raw have "oalist_inv_raw ko xs'" by (simp only: \<open>ox = ko\<close>)
thus "map2_val_pair ko f (\<lambda>zs. fst (g (zs, ko))) (\<lambda>zs. fst (h (zs, ko))) xs' [(k, v)] =
update_by_fun_pair ko k (\<lambda>x. f k x v) xs'"
using assms(2)
proof (rule map2_val_pair_singleton_eq_update_by_fun_pair)
fix zs::"('a \<times> 'b) list"
assume "oalist_inv_raw ko zs"
hence "oalist_inv (zs, ko)" by (simp only: oalist_inv_alt)
hence "g (zs, ko) = (zs, ko)" by (rule assms(3))
thus "fst (g (zs, ko)) = zs" by simp
next
show "fst (h ([(k, v)], ko)) = map_val_pair (\<lambda>k. f k 0) [(k, v)]" by (simp add: assms(4))
qed
next
show "map2_val_pair ox f (\<lambda>zs. fst (g (zs, ox))) (\<lambda>zs. fst (h (zs, ox))) xs' (sort_oalist ox [(k, v)]) =
update_by_fun_pair ox k (\<lambda>x. f k x v) xs'"
proof (cases "v = 0")
case True
have eq1: "sort_oalist ox [(k, 0)] = []" by (simp add: sort_oalist_def)
from inv have eq2: "g (xs', ox) = (xs', ox)" by (rule assms(3))
show ?thesis
by (simp add: True eq1 eq2 assms(2) update_by_fun_pair_eq_update_by_pair[OF inv_raw],
- rule HOL.sym, rule update_by_pair_id, fact inv_raw, fact refl)
+ rule HOL.sym, rule update_by_pair_id, fact inv_raw, fact HOL.refl)
next
case False
hence "oalist_inv_raw ox [(k, v)]" by (simp add: oalist_inv_raw_singleton)
hence eq: "sort_oalist ox [(k, v)] = [(k, v)]" by (rule sort_oalist_id)
show ?thesis unfolding eq using inv_raw assms(2)
proof (rule map2_val_pair_singleton_eq_update_by_fun_pair)
fix zs::"('a \<times> 'b) list"
assume "oalist_inv_raw ox zs"
hence "oalist_inv (zs, ox)" by (simp only: oalist_inv_alt)
hence "g (zs, ox) = (zs, ox)" by (rule assms(3))
thus "fst (g (zs, ox)) = zs" by simp
next
show "fst (h ([(k, v)], ox)) = map_val_pair (\<lambda>k. f k 0) [(k, v)]" by (simp add: assms(4))
qed
qed
qed
qed
subsubsection \<open>@{const lex_ord_raw}\<close>
lemma lex_ord_raw_EqI:
assumes "oalist_inv xs" and "oalist_inv ys"
and "\<And>k. k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys) \<Longrightarrow> f k (lookup_raw xs k) (lookup_raw ys k) = Some Eq"
shows "lex_ord_raw ko f xs ys = Some Eq"
unfolding lex_ord_raw_def
by (rule lex_ord_pair_EqI, simp_all add: assms oalist_inv_raw_sort_oalist_aux lookup_pair_sort_oalist_aux set_sort_oalist_aux)
lemma lex_ord_raw_valI:
assumes "oalist_inv xs" and "oalist_inv ys" and "aux \<noteq> Some Eq"
assumes "k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)" and "aux = f k (lookup_raw xs k) (lookup_raw ys k)"
and "\<And>k'. k' \<in> fst ` set (fst xs) \<union> fst ` set (fst ys) \<Longrightarrow> lt ko k' k \<Longrightarrow>
f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq"
shows "lex_ord_raw ko f xs ys = aux"
unfolding lex_ord_raw_def
using oalist_inv_sort_oalist_aux[OF assms(1)] oalist_inv_raw_sort_oalist_aux[OF assms(2)] assms(3)
unfolding oalist_inv_alt
proof (rule lex_ord_pair_valI)
from assms(1, 2, 4) show "k \<in> fst ` set (sort_oalist_aux ko xs) \<union> fst ` set (sort_oalist_aux ko ys)"
by (simp add: set_sort_oalist_aux)
next
from assms(1, 2, 5) show "aux = f k (lookup_pair ko (sort_oalist_aux ko xs) k) (lookup_pair ko (sort_oalist_aux ko ys) k)"
by (simp add: lookup_pair_sort_oalist_aux)
next
fix k'
assume "k' \<in> fst ` set (sort_oalist_aux ko xs) \<union> fst ` set (sort_oalist_aux ko ys)"
with assms(1, 2) have "k' \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)" by (simp add: set_sort_oalist_aux)
moreover assume "lt ko k' k"
ultimately have "f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq" by (rule assms(6))
with assms(1, 2) show "f k' (lookup_pair ko (sort_oalist_aux ko xs) k') (lookup_pair ko (sort_oalist_aux ko ys) k') = Some Eq"
by (simp add: lookup_pair_sort_oalist_aux)
qed
lemma lex_ord_raw_EqD:
assumes "oalist_inv xs" and "oalist_inv ys" and "lex_ord_raw ko f xs ys = Some Eq"
and "k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)"
shows "f k (lookup_raw xs k) (lookup_raw ys k) = Some Eq"
proof -
have "f k (lookup_pair ko (sort_oalist_aux ko xs) k) (lookup_pair ko (sort_oalist_aux ko ys) k) = Some Eq"
by (rule lex_ord_pair_EqD[where f=f],
simp_all add: oalist_inv_raw_sort_oalist_aux assms lex_ord_raw_def[symmetric] set_sort_oalist_aux del: Un_iff)
with assms(1, 2) show ?thesis by (simp add: lookup_pair_sort_oalist_aux)
qed
lemma lex_ord_raw_valE:
assumes "oalist_inv xs" and "oalist_inv ys" and "lex_ord_raw ko f xs ys = aux"
and "aux \<noteq> Some Eq"
obtains k where "k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)"
and "aux = f k (lookup_raw xs k) (lookup_raw ys k)"
and "\<And>k'. k' \<in> fst ` set (fst xs) \<union> fst ` set (fst ys) \<Longrightarrow> lt ko k' k \<Longrightarrow>
f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq"
proof -
let ?xs = "sort_oalist_aux ko xs"
let ?ys = "sort_oalist_aux ko ys"
from assms(3) have "lex_ord_pair ko f ?xs ?ys = aux" by (simp only: lex_ord_raw_def)
with oalist_inv_sort_oalist_aux[OF assms(1)] oalist_inv_sort_oalist_aux[OF assms(2)]
obtain k where a: "k \<in> fst ` set ?xs \<union> fst ` set ?ys"
and b: "aux = f k (lookup_pair ko ?xs k) (lookup_pair ko ?ys k)"
and c: "\<And>k'. k' \<in> fst ` set ?xs \<union> fst ` set ?ys \<Longrightarrow> lt ko k' k \<Longrightarrow>
f k' (lookup_pair ko ?xs k') (lookup_pair ko ?ys k') = Some Eq"
using assms(4) unfolding oalist_inv_alt by (rule lex_ord_pair_valE, blast)
from a have "k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)"
by (simp add: set_sort_oalist_aux assms(1, 2))
moreover from b have "aux = f k (lookup_raw xs k) (lookup_raw ys k)"
by (simp add: lookup_pair_sort_oalist_aux assms(1, 2))
moreover have "f k' (lookup_raw xs k') (lookup_raw ys k') = Some Eq"
if k'_in: "k' \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)" and k'_less: "lt ko k' k" for k'
proof -
have "f k' (lookup_raw xs k') (lookup_raw ys k') = f k' (lookup_pair ko ?xs k') (lookup_pair ko ?ys k')"
by (simp add: lookup_pair_sort_oalist_aux assms(1, 2))
also have "... = Some Eq"
proof (rule c)
from k'_in show "k' \<in> fst ` set ?xs \<union> fst ` set ?ys"
by (simp add: set_sort_oalist_aux assms(1, 2))
next
from k'_less show "lt ko k' k" by (simp only: lt_of_key_order.rep_eq)
qed
finally show ?thesis .
qed
ultimately show ?thesis ..
qed
subsubsection \<open>@{const prod_ord_raw}\<close>
lemma prod_ord_rawI:
assumes "oalist_inv xs" and "oalist_inv ys"
and "\<And>k. k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys) \<Longrightarrow> P k (lookup_raw xs k) (lookup_raw ys k)"
shows "prod_ord_raw P xs ys"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt)
hence *: "prod_ord_pair ox P xs' (sort_oalist_aux ox ys)" using oalist_inv_raw_sort_oalist_aux
proof (rule prod_ord_pairI)
fix k
assume "k \<in> fst ` set xs' \<union> fst ` set (sort_oalist_aux ox ys)"
hence "k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)" by (simp add: xs set_sort_oalist_aux assms(2))
hence "P k (lookup_raw xs k) (lookup_raw ys k)" by (rule assms(3))
thus "P k (lookup_pair ox xs' k) (lookup_pair ox (sort_oalist_aux ox ys) k)"
by (simp add: xs lookup_pair_sort_oalist_aux assms(2))
qed fact
thus ?thesis by (simp add: xs)
qed
lemma prod_ord_rawD:
assumes "oalist_inv xs" and "oalist_inv ys" and "prod_ord_raw P xs ys"
and "k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)"
shows "P k (lookup_raw xs k) (lookup_raw ys k)"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt)
moreover note oalist_inv_raw_sort_oalist_aux[OF assms(2)]
moreover from assms(3) have "prod_ord_pair ox P xs' (sort_oalist_aux ox ys)" by (simp add: xs)
moreover from assms(4) have "k \<in> fst ` set xs' \<union> fst ` set (sort_oalist_aux ox ys)"
by (simp add: xs set_sort_oalist_aux assms(2))
ultimately have *: "P k (lookup_pair ox xs' k) (lookup_pair ox (sort_oalist_aux ox ys) k)"
by (rule prod_ord_pairD)
thus ?thesis by (simp add: xs lookup_pair_sort_oalist_aux assms(2))
qed
corollary prod_ord_raw_alt:
assumes "oalist_inv xs" and "oalist_inv ys"
shows "prod_ord_raw P xs ys \<longleftrightarrow>
(\<forall>k\<in>fst ` set (fst xs) \<union> fst ` set (fst ys). P k (lookup_raw xs k) (lookup_raw ys k))"
using prod_ord_rawI[OF assms] prod_ord_rawD[OF assms] by meson
subsubsection \<open>@{const oalist_eq_raw}\<close>
lemma oalist_eq_rawI:
assumes "oalist_inv xs" and "oalist_inv ys"
and "\<And>k. k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys) \<Longrightarrow> lookup_raw xs k = lookup_raw ys k"
shows "oalist_eq_raw xs ys"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(1) have "oalist_inv_raw ox xs'" by (simp only: xs oalist_inv_alt)
hence *: "xs' = sort_oalist_aux ox ys" using oalist_inv_raw_sort_oalist_aux[OF assms(2)]
proof (rule lookup_pair_inj)
show "lookup_pair ox xs' = lookup_pair ox (sort_oalist_aux ox ys)"
proof
fix k
show "lookup_pair ox xs' k = lookup_pair ox (sort_oalist_aux ox ys) k"
proof (cases "k \<in> fst ` set xs' \<union> fst ` set (sort_oalist_aux ox ys)")
case True
hence "k \<in> fst ` set (fst xs) \<union> fst ` set (fst ys)" by (simp add: xs set_sort_oalist_aux assms(2))
hence "lookup_raw xs k = lookup_raw ys k" by (rule assms(3))
thus ?thesis by (simp add: xs lookup_pair_sort_oalist_aux assms(2))
next
case False
hence "k \<notin> fst ` set xs'" and "k \<notin> fst ` set (sort_oalist_aux ox ys)" by simp_all
with \<open>oalist_inv_raw ox xs'\<close> oalist_inv_raw_sort_oalist_aux[OF assms(2)]
have "lookup_pair ox xs' k = 0" and "lookup_pair ox (sort_oalist_aux ox ys) k = 0"
by (simp_all add: lookup_pair_eq_0)
thus ?thesis by simp
qed
qed
qed
thus ?thesis by (simp add: xs)
qed
lemma oalist_eq_rawD:
assumes "oalist_inv ys" and "oalist_eq_raw xs ys"
shows "lookup_raw xs = lookup_raw ys"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms(2) have "xs' = sort_oalist_aux ox ys" by (simp add: xs)
hence "lookup_pair ox xs' = lookup_pair ox (sort_oalist_aux ox ys)" by simp
thus ?thesis by (simp add: xs lookup_pair_sort_oalist_aux assms(1))
qed
lemma oalist_eq_raw_alt:
assumes "oalist_inv xs" and "oalist_inv ys"
shows "oalist_eq_raw xs ys \<longleftrightarrow> (lookup_raw xs = lookup_raw ys)"
using oalist_eq_rawI[OF assms] oalist_eq_rawD[OF assms(2)] by metis
subsubsection \<open>@{const sort_oalist_raw}\<close>
lemma oalist_inv_sort_oalist_raw: "oalist_inv (sort_oalist_raw xs)"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
show ?thesis by (simp add: xs oalist_inv_raw_sort_oalist oalist_inv_alt)
qed
lemma sort_oalist_raw_id:
assumes "oalist_inv xs"
shows "sort_oalist_raw xs = xs"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms have "oalist_inv_raw ko xs'" by (simp only: xs oalist_inv_alt)
hence "sort_oalist ko xs' = xs'" by (rule sort_oalist_id)
thus ?thesis by (simp add: xs)
qed
lemma set_sort_oalist_raw:
assumes "distinct (map fst (fst xs))"
shows "set (fst (sort_oalist_raw xs)) = {kv. kv \<in> set (fst xs) \<and> snd kv \<noteq> 0}"
proof -
obtain xs' ko where xs: "xs = (xs', ko)" by fastforce
from assms have "distinct (map fst xs')" by (simp add: xs)
hence "set (sort_oalist ko xs') = {kv \<in> set xs'. snd kv \<noteq> 0}" by (rule set_sort_oalist)
thus ?thesis by (simp add: xs)
qed
end (* oalist_raw *)
subsection \<open>Fundamental Operations on One List\<close>
locale oalist_abstract = oalist_raw rep_key_order for rep_key_order::"'o \<Rightarrow> 'a key_order" +
fixes list_of_oalist :: "'x \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw"
fixes oalist_of_list :: "('a, 'b, 'o) oalist_raw \<Rightarrow> 'x"
assumes oalist_inv_list_of_oalist: "oalist_inv (list_of_oalist x)"
and list_of_oalist_of_list: "list_of_oalist (oalist_of_list xs) = sort_oalist_raw xs"
and oalist_of_list_of_oalist: "oalist_of_list (list_of_oalist x) = x"
begin
lemma list_of_oalist_of_list_id:
assumes "oalist_inv xs"
shows "list_of_oalist (oalist_of_list xs) = xs"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
from assms show ?thesis by (simp add: xs list_of_oalist_of_list sort_oalist_id oalist_inv_alt)
qed
definition lookup :: "'x \<Rightarrow> 'a \<Rightarrow> 'b"
where "lookup xs = lookup_raw (list_of_oalist xs)"
definition sorted_domain :: "'o \<Rightarrow> 'x \<Rightarrow> 'a list"
where "sorted_domain ko xs = sorted_domain_raw ko (list_of_oalist xs)"
definition empty :: "'o \<Rightarrow> 'x"
where "empty ko = oalist_of_list ([], ko)"
definition reorder :: "'o \<Rightarrow> 'x \<Rightarrow> 'x"
where "reorder ko xs = oalist_of_list (sort_oalist_aux ko (list_of_oalist xs), ko)"
definition tl :: "'x \<Rightarrow> 'x"
where "tl xs = oalist_of_list (tl_raw (list_of_oalist xs))"
definition hd :: "'x \<Rightarrow> ('a \<times> 'b)"
where "hd xs = List.hd (fst (list_of_oalist xs))"
definition except_min :: "'o \<Rightarrow> 'x \<Rightarrow> 'x"
where "except_min ko xs = tl (reorder ko xs)"
definition min_key_val :: "'o \<Rightarrow> 'x \<Rightarrow> ('a \<times> 'b)"
where "min_key_val ko xs = min_key_val_raw ko (list_of_oalist xs)"
definition insert :: "('a \<times> 'b) \<Rightarrow> 'x \<Rightarrow> 'x"
where "insert x xs = oalist_of_list (update_by_raw x (list_of_oalist xs))"
definition update_by_fun :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> 'x \<Rightarrow> 'x"
where "update_by_fun k f xs = oalist_of_list (update_by_fun_raw k f (list_of_oalist xs))"
definition update_by_fun_gr :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> 'x \<Rightarrow> 'x"
where "update_by_fun_gr k f xs = oalist_of_list (update_by_fun_gr_raw k f (list_of_oalist xs))"
definition filter :: "(('a \<times> 'b) \<Rightarrow> bool) \<Rightarrow> 'x \<Rightarrow> 'x"
where "filter P xs = oalist_of_list (filter_raw P (list_of_oalist xs))"
definition map2_val_neutr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'x \<Rightarrow> 'x \<Rightarrow> 'x"
where "map2_val_neutr f xs ys = oalist_of_list (map2_val_raw f id id (list_of_oalist xs) (list_of_oalist ys))"
definition oalist_eq :: "'x \<Rightarrow> 'x \<Rightarrow> bool"
where "oalist_eq xs ys = oalist_eq_raw (list_of_oalist xs) (list_of_oalist ys)"
subsubsection \<open>Invariant\<close>
lemma zero_notin_list_of_oalist: "0 \<notin> snd ` set (fst (list_of_oalist xs))"
proof -
from oalist_inv_list_of_oalist have "oalist_inv_raw (snd (list_of_oalist xs)) (fst (list_of_oalist xs))"
by (simp only: oalist_inv_def)
thus ?thesis by (rule oalist_inv_rawD1)
qed
lemma list_of_oalist_sorted: "sorted_wrt (lt (snd (list_of_oalist xs))) (map fst (fst (list_of_oalist xs)))"
proof -
from oalist_inv_list_of_oalist have "oalist_inv_raw (snd (list_of_oalist xs)) (fst (list_of_oalist xs))"
by (simp only: oalist_inv_def)
thus ?thesis by (rule oalist_inv_rawD2)
qed
subsubsection \<open>@{const lookup}\<close>
lemma lookup_eq_value: "v \<noteq> 0 \<Longrightarrow> lookup xs k = v \<longleftrightarrow> ((k, v) \<in> set (fst (list_of_oalist xs)))"
unfolding lookup_def using oalist_inv_list_of_oalist by (rule lookup_raw_eq_value)
lemma lookup_eq_valueI: "(k, v) \<in> set (fst (list_of_oalist xs)) \<Longrightarrow> lookup xs k = v"
unfolding lookup_def using oalist_inv_list_of_oalist by (rule lookup_raw_eq_valueI)
lemma lookup_oalist_of_list:
"distinct (map fst xs) \<Longrightarrow> lookup (oalist_of_list (xs, ko)) = lookup_dflt xs"
by (simp add: list_of_oalist_of_list lookup_def lookup_pair_sort_oalist')
subsubsection \<open>@{const sorted_domain}\<close>
lemma set_sorted_domain: "set (sorted_domain ko xs) = fst ` set (fst (list_of_oalist xs))"
unfolding sorted_domain_def using oalist_inv_list_of_oalist by (rule set_sorted_domain_raw)
lemma in_sorted_domain_iff_lookup: "k \<in> set (sorted_domain ko xs) \<longleftrightarrow> (lookup xs k \<noteq> 0)"
unfolding sorted_domain_def lookup_def using oalist_inv_list_of_oalist
by (rule in_sorted_domain_raw_iff_lookup_raw)
lemma sorted_sorted_domain: "sorted_wrt (lt ko) (sorted_domain ko xs)"
unfolding sorted_domain_def lt_of_key_order.rep_eq[symmetric]
using oalist_inv_list_of_oalist by (rule sorted_sorted_domain_raw)
subsubsection \<open>@{const empty} and Singletons\<close>
lemma list_of_oalist_empty [simp, code abstract]: "list_of_oalist (empty ko) = ([], ko)"
by (simp add: empty_def sort_oalist_def list_of_oalist_of_list)
lemma lookup_empty: "lookup (empty ko) k = 0"
by (simp add: lookup_def)
lemma lookup_oalist_of_list_single:
"lookup (oalist_of_list ([(k, v)], ko)) k' = (if k = k' then v else 0)"
by (simp add: lookup_def list_of_oalist_of_list sort_oalist_def key_compare_Eq split: order.split)
subsubsection \<open>@{const reorder}\<close>
lemma list_of_oalist_reorder [simp, code abstract]:
"list_of_oalist (reorder ko xs) = (sort_oalist_aux ko (list_of_oalist xs), ko)"
unfolding reorder_def
by (rule list_of_oalist_of_list_id, simp add: oalist_inv_def, rule oalist_inv_raw_sort_oalist_aux, fact oalist_inv_list_of_oalist)
lemma lookup_reorder: "lookup (reorder ko xs) k = lookup xs k"
by (simp add: lookup_def lookup_pair_sort_oalist_aux oalist_inv_list_of_oalist)
subsubsection \<open>@{const hd} and @{const tl}\<close>
lemma list_of_oalist_tl [simp, code abstract]: "list_of_oalist (tl xs) = tl_raw (list_of_oalist xs)"
unfolding tl_def
by (rule list_of_oalist_of_list_id, rule oalist_inv_tl_raw, fact oalist_inv_list_of_oalist)
lemma lookup_tl:
"lookup (tl xs) k =
(if (\<forall>k'\<in>fst ` set (fst (list_of_oalist xs)). le (snd (list_of_oalist xs)) k k') then 0 else lookup xs k)"
by (simp add: lookup_def lookup_raw_tl_raw oalist_inv_list_of_oalist)
lemma hd_in:
assumes "fst (list_of_oalist xs) \<noteq> []"
shows "hd xs \<in> set (fst (list_of_oalist xs))"
unfolding hd_def using assms by (rule hd_in_set)
lemma snd_hd:
assumes "fst (list_of_oalist xs) \<noteq> []"
shows "snd (hd xs) = lookup xs (fst (hd xs))"
proof -
from assms have *: "hd xs \<in> set (fst (list_of_oalist xs))" by (rule hd_in)
show ?thesis by (rule HOL.sym, rule lookup_eq_valueI, simp add: *)
qed
lemma lookup_tl': "lookup (tl xs) k = (if k = fst (hd xs) then 0 else lookup xs k)"
by (simp add: lookup_def lookup_raw_tl_raw' oalist_inv_list_of_oalist hd_def)
lemma hd_tl:
assumes "fst (list_of_oalist xs) \<noteq> []"
shows "list_of_oalist xs = ((hd xs) # (fst (list_of_oalist (tl xs))), snd (list_of_oalist (tl xs)))"
proof -
obtain xs' ko where xs: "list_of_oalist xs = (xs', ko)" by fastforce
from assms obtain x xs'' where xs': "xs' = x # xs''" unfolding xs fst_conv using list.exhaust by blast
show ?thesis by (simp add: xs xs' hd_def)
qed
subsubsection \<open>@{const min_key_val}\<close>
lemma min_key_val_alt:
assumes "fst (list_of_oalist xs) \<noteq> []"
shows "min_key_val ko xs = hd (reorder ko xs)"
using assms oalist_inv_list_of_oalist by (simp add: min_key_val_def hd_def min_key_val_raw_alt)
lemma min_key_val_in:
assumes "fst (list_of_oalist xs) \<noteq> []"
shows "min_key_val ko xs \<in> set (fst (list_of_oalist xs))"
unfolding min_key_val_def using assms by (rule min_key_val_raw_in)
lemma snd_min_key_val:
assumes "fst (list_of_oalist xs) \<noteq> []"
shows "snd (min_key_val ko xs) = lookup xs (fst (min_key_val ko xs))"
unfolding lookup_def min_key_val_def using oalist_inv_list_of_oalist assms by (rule snd_min_key_val_raw)
lemma min_key_val_minimal:
assumes "z \<in> set (fst (list_of_oalist xs))"
shows "le ko (fst (min_key_val ko xs)) (fst z)"
unfolding min_key_val_def
by (rule min_key_val_raw_minimal, fact oalist_inv_list_of_oalist, fact)
subsubsection \<open>@{const except_min}\<close>
lemma list_of_oalist_except_min [simp, code abstract]:
"list_of_oalist (except_min ko xs) = (List.tl (sort_oalist_aux ko (list_of_oalist xs)), ko)"
by (simp add: except_min_def)
lemma except_min_Nil:
assumes "fst (list_of_oalist xs) = []"
shows "fst (list_of_oalist (except_min ko xs)) = []"
proof -
obtain xs' ox where eq: "list_of_oalist xs = (xs', ox)" by fastforce
from assms have "xs' = []" by (simp add: eq)
show ?thesis by (simp add: eq \<open>xs' = []\<close> sort_oalist_def)
qed
lemma lookup_except_min:
"lookup (except_min ko xs) k =
(if (\<forall>k'\<in>fst ` set (fst (list_of_oalist xs)). le ko k k') then 0 else lookup xs k)"
by (simp add: except_min_def lookup_tl set_sort_oalist_aux oalist_inv_list_of_oalist lookup_reorder)
lemma lookup_except_min':
"lookup (except_min ko xs) k = (if k = fst (min_key_val ko xs) then 0 else lookup xs k)"
proof (cases "fst (list_of_oalist xs) = []")
case True
hence "lookup xs k = 0" by (metis empty_def lookup_empty oalist_of_list_of_oalist prod.collapse)
thus ?thesis by (simp add: lookup_except_min True)
next
case False
thus ?thesis by (simp add: except_min_def lookup_tl' min_key_val_alt lookup_reorder)
qed
subsubsection \<open>@{const insert}\<close>
lemma list_of_oalist_insert [simp, code abstract]:
"list_of_oalist (insert x xs) = update_by_raw x (list_of_oalist xs)"
unfolding insert_def
by (rule list_of_oalist_of_list_id, rule oalist_inv_update_by_raw, fact oalist_inv_list_of_oalist)
lemma lookup_insert: "lookup (insert (k, v) xs) k' = (if k = k' then v else lookup xs k')"
by (simp add: lookup_def lookup_raw_update_by_raw oalist_inv_list_of_oalist split del: if_split cong: if_cong)
subsubsection \<open>@{const update_by_fun} and @{const update_by_fun_gr}\<close>
lemma list_of_oalist_update_by_fun [simp, code abstract]:
"list_of_oalist (update_by_fun k f xs) = update_by_fun_raw k f (list_of_oalist xs)"
unfolding update_by_fun_def
by (rule list_of_oalist_of_list_id, rule oalist_inv_update_by_fun_raw, fact oalist_inv_list_of_oalist)
lemma lookup_update_by_fun:
"lookup (update_by_fun k f xs) k' = (if k = k' then f else id) (lookup xs k')"
by (simp add: lookup_def lookup_raw_update_by_fun_raw oalist_inv_list_of_oalist split del: if_split cong: if_cong)
lemma list_of_oalist_update_by_fun_gr [simp, code abstract]:
"list_of_oalist (update_by_fun_gr k f xs) = update_by_fun_gr_raw k f (list_of_oalist xs)"
unfolding update_by_fun_gr_def
by (rule list_of_oalist_of_list_id, rule oalist_inv_update_by_fun_gr_raw, fact oalist_inv_list_of_oalist)
lemma update_by_fun_gr_eq_update_by_fun: "update_by_fun_gr = update_by_fun"
by (rule, rule, rule,
simp add: update_by_fun_gr_def update_by_fun_def update_by_fun_gr_raw_eq_update_by_fun_raw oalist_inv_list_of_oalist)
subsubsection \<open>@{const filter}\<close>
lemma list_of_oalist_filter [simp, code abstract]:
"list_of_oalist (filter P xs) = filter_raw P (list_of_oalist xs)"
unfolding filter_def
by (rule list_of_oalist_of_list_id, rule oalist_inv_filter_raw, fact oalist_inv_list_of_oalist)
lemma lookup_filter: "lookup (filter P xs) k = (let v = lookup xs k in if P (k, v) then v else 0)"
by (simp add: lookup_def lookup_raw_filter_raw oalist_inv_list_of_oalist)
subsubsection \<open>@{const map2_val_neutr}\<close>
lemma list_of_oalist_map2_val_neutr [simp, code abstract]:
"list_of_oalist (map2_val_neutr f xs ys) = map2_val_raw f id id (list_of_oalist xs) (list_of_oalist ys)"
unfolding map2_val_neutr_def
by (rule list_of_oalist_of_list_id, rule oalist_inv_map2_val_raw,
fact oalist_inv_list_of_oalist, fact oalist_inv_list_of_oalist,
fact map2_val_compat'_id, fact map2_val_compat'_id)
lemma lookup_map2_val_neutr:
assumes "\<And>k x. f k x 0 = x" and "\<And>k x. f k 0 x = x"
shows "lookup (map2_val_neutr f xs ys) k = f k (lookup xs k) (lookup ys k)"
proof (simp add: lookup_def, rule lookup_raw_map2_val_raw)
fix zs::"('a, 'b, 'o) oalist_raw"
assume "oalist_inv zs"
thus "id zs = map_val_raw (\<lambda>k v. f k v 0) zs" by (simp add: assms(1) map_raw_id)
next
fix zs::"('a, 'b, 'o) oalist_raw"
assume "oalist_inv zs"
thus "id zs = map_val_raw (\<lambda>k. f k 0) zs" by (simp add: assms(2) map_raw_id)
qed (fact oalist_inv_list_of_oalist, fact oalist_inv_list_of_oalist,
fact map2_val_compat'_id, fact map2_val_compat'_id, simp only: assms(1))
subsubsection \<open>@{const oalist_eq}\<close>
lemma oalist_eq_alt: "oalist_eq xs ys \<longleftrightarrow> (lookup xs = lookup ys)"
by (simp add: oalist_eq_def lookup_def oalist_eq_raw_alt oalist_inv_list_of_oalist)
end (* oalist_abstract *)
subsection \<open>Fundamental Operations on Three Lists\<close>
locale oalist_abstract3 =
oalist_abstract rep_key_order list_of_oalistx oalist_of_listx +
oay: oalist_abstract rep_key_order list_of_oalisty oalist_of_listy +
oaz: oalist_abstract rep_key_order list_of_oalistz oalist_of_listz
for rep_key_order :: "'o \<Rightarrow> 'a key_order"
and list_of_oalistx :: "'x \<Rightarrow> ('a, 'b::zero, 'o) oalist_raw"
and oalist_of_listx :: "('a, 'b, 'o) oalist_raw \<Rightarrow> 'x"
and list_of_oalisty :: "'y \<Rightarrow> ('a, 'c::zero, 'o) oalist_raw"
and oalist_of_listy :: "('a, 'c, 'o) oalist_raw \<Rightarrow> 'y"
and list_of_oalistz :: "'z \<Rightarrow> ('a, 'd::zero, 'o) oalist_raw"
and oalist_of_listz :: "('a, 'd, 'o) oalist_raw \<Rightarrow> 'z"
begin
definition map_val :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'x \<Rightarrow> 'y"
where "map_val f xs = oalist_of_listy (map_val_raw f (list_of_oalistx xs))"
definition map2_val :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'x \<Rightarrow> 'y \<Rightarrow> 'z"
where "map2_val f xs ys =
oalist_of_listz (map2_val_raw f (map_val_raw (\<lambda>k b. f k b 0)) (map_val_raw (\<lambda>k. f k 0))
(list_of_oalistx xs) (list_of_oalisty ys))"
definition map2_val_rneutr :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> 'x \<Rightarrow> 'y \<Rightarrow> 'x"
where "map2_val_rneutr f xs ys =
oalist_of_listx (map2_val_raw f id (map_val_raw (\<lambda>k. f k 0)) (list_of_oalistx xs) (list_of_oalisty ys))"
definition lex_ord :: "'o \<Rightarrow> ('a \<Rightarrow> ('b, 'c) comp_opt) \<Rightarrow> ('x, 'y) comp_opt"
where "lex_ord ko f xs ys = lex_ord_raw ko f (list_of_oalistx xs) (list_of_oalisty ys)"
definition prod_ord :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> 'x \<Rightarrow> 'y \<Rightarrow> bool"
where "prod_ord f xs ys = prod_ord_raw f (list_of_oalistx xs) (list_of_oalisty ys)"
subsubsection \<open>@{const map_val}\<close>
lemma map_val_cong:
assumes "\<And>k v. (k, v) \<in> set (fst (list_of_oalistx xs)) \<Longrightarrow> f k v = g k v"
shows "map_val f xs = map_val g xs"
unfolding map_val_def by (rule arg_cong[where f=oalist_of_listy], rule map_val_raw_cong, elim assms)
lemma list_of_oalist_map_val [simp, code abstract]:
"list_of_oalisty (map_val f xs) = map_val_raw f (list_of_oalistx xs)"
unfolding map_val_def
by (rule oay.list_of_oalist_of_list_id, rule oalist_inv_map_val_raw, fact oalist_inv_list_of_oalist)
lemma lookup_map_val: "f k 0 = 0 \<Longrightarrow> oay.lookup (map_val f xs) k = f k (lookup xs k)"
by (simp add: oay.lookup_def lookup_def lookup_raw_map_val_raw oalist_inv_list_of_oalist)
subsubsection \<open>@{const map2_val} and @{const map2_val_rneutr}\<close>
lemma list_of_oalist_map2_val [simp, code abstract]:
"list_of_oalistz (map2_val f xs ys) =
map2_val_raw f (map_val_raw (\<lambda>k b. f k b 0)) (map_val_raw (\<lambda>k. f k 0)) (list_of_oalistx xs) (list_of_oalisty ys)"
unfolding map2_val_def
by (rule oaz.list_of_oalist_of_list_id, rule oalist_inv_map2_val_raw,
fact oalist_inv_list_of_oalist, fact oay.oalist_inv_list_of_oalist,
fact map2_val_compat'_map_val_raw, fact map2_val_compat'_map_val_raw)
lemma list_of_oalist_map2_val_rneutr [simp, code abstract]:
"list_of_oalistx (map2_val_rneutr f xs ys) =
map2_val_raw f id (map_val_raw (\<lambda>k c. f k 0 c)) (list_of_oalistx xs) (list_of_oalisty ys)"
unfolding map2_val_rneutr_def
by (rule list_of_oalist_of_list_id, rule oalist_inv_map2_val_raw,
fact oalist_inv_list_of_oalist, fact oay.oalist_inv_list_of_oalist,
fact map2_val_compat'_id, fact map2_val_compat'_map_val_raw)
lemma lookup_map2_val:
assumes "\<And>k. f k 0 0 = 0"
shows "oaz.lookup (map2_val f xs ys) k = f k (lookup xs k) (oay.lookup ys k)"
by (simp add: oaz.lookup_def oay.lookup_def lookup_def lookup_raw_map2_val_raw
map2_val_compat'_map_val_raw assms oalist_inv_list_of_oalist oay.oalist_inv_list_of_oalist)
lemma lookup_map2_val_rneutr:
assumes "\<And>k x. f k x 0 = x"
shows "lookup (map2_val_rneutr f xs ys) k = f k (lookup xs k) (oay.lookup ys k)"
proof (simp add: lookup_def oay.lookup_def, rule lookup_raw_map2_val_raw)
fix zs::"('a, 'b, 'o) oalist_raw"
assume "oalist_inv zs"
thus "id zs = map_val_raw (\<lambda>k v. f k v 0) zs" by (simp add: assms map_raw_id)
qed (fact oalist_inv_list_of_oalist, fact oay.oalist_inv_list_of_oalist,
- fact map2_val_compat'_id, fact map2_val_compat'_map_val_raw, rule refl, simp only: assms)
+ fact map2_val_compat'_id, fact map2_val_compat'_map_val_raw, rule HOL.refl, simp only: assms)
lemma map2_val_rneutr_singleton_eq_update_by_fun:
assumes "\<And>a x. f a x 0 = x" and "list_of_oalisty ys = ([(k, v)], oy)"
shows "map2_val_rneutr f xs ys = update_by_fun k (\<lambda>x. f k x v) xs"
by (simp add: map2_val_rneutr_def update_by_fun_def assms map2_val_raw_singleton_eq_update_by_fun_raw oalist_inv_list_of_oalist)
subsubsection \<open>@{const lex_ord} and @{const prod_ord}\<close>
lemma lex_ord_EqI:
"(\<And>k. k \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys)) \<Longrightarrow>
f k (lookup xs k) (oay.lookup ys k) = Some Eq) \<Longrightarrow>
lex_ord ko f xs ys = Some Eq"
by (simp add: lex_ord_def lookup_def oay.lookup_def, rule lex_ord_raw_EqI,
rule oalist_inv_list_of_oalist, rule oay.oalist_inv_list_of_oalist, blast)
lemma lex_ord_valI:
assumes "aux \<noteq> Some Eq" and "k \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys))"
shows "aux = f k (lookup xs k) (oay.lookup ys k) \<Longrightarrow>
(\<And>k'. k' \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys)) \<Longrightarrow>
lt ko k' k \<Longrightarrow> f k' (lookup xs k') (oay.lookup ys k') = Some Eq) \<Longrightarrow>
lex_ord ko f xs ys = aux"
by (simp (no_asm_use) add: lex_ord_def lookup_def oay.lookup_def, rule lex_ord_raw_valI,
rule oalist_inv_list_of_oalist, rule oay.oalist_inv_list_of_oalist, rule assms(1), rule assms(2), blast+)
lemma lex_ord_EqD:
"lex_ord ko f xs ys = Some Eq \<Longrightarrow>
k \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys)) \<Longrightarrow>
f k (lookup xs k) (oay.lookup ys k) = Some Eq"
by (simp add: lex_ord_def lookup_def oay.lookup_def, rule lex_ord_raw_EqD[where f=f],
rule oalist_inv_list_of_oalist, rule oay.oalist_inv_list_of_oalist, assumption, simp)
lemma lex_ord_valE:
assumes "lex_ord ko f xs ys = aux" and "aux \<noteq> Some Eq"
obtains k where "k \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys))"
and "aux = f k (lookup xs k) (oay.lookup ys k)"
and "\<And>k'. k' \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys)) \<Longrightarrow>
lt ko k' k \<Longrightarrow> f k' (lookup xs k') (oay.lookup ys k') = Some Eq"
proof -
note oalist_inv_list_of_oalist oay.oalist_inv_list_of_oalist
moreover from assms(1) have "lex_ord_raw ko f (list_of_oalistx xs) (list_of_oalisty ys) = aux"
by (simp only: lex_ord_def)
ultimately obtain k where 1: "k \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys))"
and "aux = f k (lookup_raw (list_of_oalistx xs) k) (lookup_raw (list_of_oalisty ys) k)"
and "\<And>k'. k' \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys)) \<Longrightarrow>
lt ko k' k \<Longrightarrow>
f k' (lookup_raw (list_of_oalistx xs) k') (lookup_raw (list_of_oalisty ys) k') = Some Eq"
using assms(2) by (rule lex_ord_raw_valE, blast)
from this(2, 3) have "aux = f k (lookup xs k) (oay.lookup ys k)"
and "\<And>k'. k' \<in> fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys)) \<Longrightarrow>
lt ko k' k \<Longrightarrow> f k' (lookup xs k') (oay.lookup ys k') = Some Eq"
by (simp_all only: lookup_def oay.lookup_def)
with 1 show ?thesis ..
qed
lemma prod_ord_alt:
"prod_ord P xs ys \<longleftrightarrow>
(\<forall>k\<in>fst ` set (fst (list_of_oalistx xs)) \<union> fst ` set (fst (list_of_oalisty ys)).
P k (lookup xs k) (oay.lookup ys k))"
by (simp add: prod_ord_def lookup_def oay.lookup_def prod_ord_raw_alt oalist_inv_list_of_oalist oay.oalist_inv_list_of_oalist)
end (* oalist_abstract3 *)
subsection \<open>Type \<open>oalist\<close>\<close>
global_interpretation ko: comparator "key_compare ko"
defines lookup_pair_ko = ko.lookup_pair
and update_by_pair_ko = ko.update_by_pair
and update_by_fun_pair_ko = ko.update_by_fun_pair
and update_by_fun_gr_pair_ko = ko.update_by_fun_gr_pair
and map2_val_pair_ko = ko.map2_val_pair
and lex_ord_pair_ko = ko.lex_ord_pair
and prod_ord_pair_ko = ko.prod_ord_pair
and sort_oalist_ko' = ko.sort_oalist
by (fact comparator_key_compare)
lemma ko_le: "ko.le = le_of_key_order"
by (intro ext, simp add: le_of_comp_def le_of_key_order_alt split: order.split)
global_interpretation ko: oalist_raw "\<lambda>x. x"
rewrites "comparator.lookup_pair (key_compare ko) = lookup_pair_ko ko"
and "comparator.update_by_pair (key_compare ko) = update_by_pair_ko ko"
and "comparator.update_by_fun_pair (key_compare ko) = update_by_fun_pair_ko ko"
and "comparator.update_by_fun_gr_pair (key_compare ko) = update_by_fun_gr_pair_ko ko"
and "comparator.map2_val_pair (key_compare ko) = map2_val_pair_ko ko"
and "comparator.lex_ord_pair (key_compare ko) = lex_ord_pair_ko ko"
and "comparator.prod_ord_pair (key_compare ko) = prod_ord_pair_ko ko"
and "comparator.sort_oalist (key_compare ko) = sort_oalist_ko' ko"
defines sort_oalist_aux_ko = ko.sort_oalist_aux
and lookup_ko = ko.lookup_raw
and sorted_domain_ko = ko.sorted_domain_raw
and tl_ko = ko.tl_raw
and min_key_val_ko = ko.min_key_val_raw
and update_by_ko = ko.update_by_raw
and update_by_fun_ko = ko.update_by_fun_raw
and update_by_fun_gr_ko = ko.update_by_fun_gr_raw
and map2_val_ko = ko.map2_val_raw
and lex_ord_ko = ko.lex_ord_raw
and prod_ord_ko = ko.prod_ord_raw
and oalist_eq_ko = ko.oalist_eq_raw
and sort_oalist_ko = ko.sort_oalist_raw
subgoal by (simp only: lookup_pair_ko_def)
subgoal by (simp only: update_by_pair_ko_def)
subgoal by (simp only: update_by_fun_pair_ko_def)
subgoal by (simp only: update_by_fun_gr_pair_ko_def)
subgoal by (simp only: map2_val_pair_ko_def)
subgoal by (simp only: lex_ord_pair_ko_def)
subgoal by (simp only: prod_ord_pair_ko_def)
subgoal by (simp only: sort_oalist_ko'_def)
done
typedef (overloaded) ('a, 'b) oalist = "{xs::('a, 'b::zero, 'a key_order) oalist_raw. ko.oalist_inv xs}"
morphisms list_of_oalist Abs_oalist
by (auto simp: ko.oalist_inv_def intro: ko.oalist_inv_raw_Nil)
lemma oalist_eq_iff: "xs = ys \<longleftrightarrow> list_of_oalist xs = list_of_oalist ys"
by (simp add: list_of_oalist_inject)
lemma oalist_eqI: "list_of_oalist xs = list_of_oalist ys \<Longrightarrow> xs = ys"
by (simp add: oalist_eq_iff)
text \<open>Formal, totalized constructor for @{typ "('a, 'b) oalist"}:\<close>
definition OAlist :: "('a \<times> 'b) list \<times> 'a key_order \<Rightarrow> ('a, 'b::zero) oalist" where
"OAlist xs = Abs_oalist (sort_oalist_ko xs)"
definition "oalist_of_list = OAlist"
lemma oalist_inv_list_of_oalist: "ko.oalist_inv (list_of_oalist xs)"
using list_of_oalist [of xs] by simp
lemma list_of_oalist_OAlist: "list_of_oalist (OAlist xs) = sort_oalist_ko xs"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
show ?thesis by (simp add: xs OAlist_def Abs_oalist_inverse ko.oalist_inv_raw_sort_oalist ko.oalist_inv_alt)
qed
lemma OAlist_list_of_oalist [code abstype]: "OAlist (list_of_oalist xs) = xs"
proof -
obtain xs' ox where xs: "list_of_oalist xs = (xs', ox)" by fastforce
have "ko.oalist_inv_raw ox xs'" by (simp add: xs[symmetric] ko.oalist_inv_alt[symmetric] oalist_inv_list_of_oalist)
thus ?thesis by (simp add: xs OAlist_def ko.sort_oalist_id, simp add: list_of_oalist_inverse xs[symmetric])
qed
lemma [code abstract]: "list_of_oalist (oalist_of_list xs) = sort_oalist_ko xs"
by (simp add: list_of_oalist_OAlist oalist_of_list_def)
global_interpretation oa: oalist_abstract "\<lambda>x. x" list_of_oalist OAlist
defines OAlist_lookup = oa.lookup
and OAlist_sorted_domain = oa.sorted_domain
and OAlist_empty = oa.empty
and OAlist_reorder = oa.reorder
and OAlist_tl = oa.tl
and OAlist_hd = oa.hd
and OAlist_except_min = oa.except_min
and OAlist_min_key_val = oa.min_key_val
and OAlist_insert = oa.insert
and OAlist_update_by_fun = oa.update_by_fun
and OAlist_update_by_fun_gr = oa.update_by_fun_gr
and OAlist_filter = oa.filter
and OAlist_map2_val_neutr = oa.map2_val_neutr
and OAlist_eq = oa.oalist_eq
apply standard
subgoal by (fact oalist_inv_list_of_oalist)
subgoal by (simp only: list_of_oalist_OAlist sort_oalist_ko_def)
subgoal by (fact OAlist_list_of_oalist)
done
global_interpretation oa: oalist_abstract3 "\<lambda>x. x"
"list_of_oalist::('a, 'b) oalist \<Rightarrow> ('a, 'b::zero, 'a key_order) oalist_raw" OAlist
"list_of_oalist::('a, 'c) oalist \<Rightarrow> ('a, 'c::zero, 'a key_order) oalist_raw" OAlist
"list_of_oalist::('a, 'd) oalist \<Rightarrow> ('a, 'd::zero, 'a key_order) oalist_raw" OAlist
defines OAlist_map_val = oa.map_val
and OAlist_map2_val = oa.map2_val
and OAlist_map2_val_rneutr = oa.map2_val_rneutr
and OAlist_lex_ord = oa.lex_ord
and OAlist_prod_ord = oa.prod_ord ..
lemmas OAlist_lookup_single = oa.lookup_oalist_of_list_single[folded oalist_of_list_def]
subsection \<open>Type \<open>oalist_tc\<close>\<close>
text \<open>``tc'' stands for ``type class''.\<close>
global_interpretation tc: comparator "comparator_of"
defines lookup_pair_tc = tc.lookup_pair
and update_by_pair_tc = tc.update_by_pair
and update_by_fun_pair_tc = tc.update_by_fun_pair
and update_by_fun_gr_pair_tc = tc.update_by_fun_gr_pair
and map2_val_pair_tc = tc.map2_val_pair
and lex_ord_pair_tc = tc.lex_ord_pair
and prod_ord_pair_tc = tc.prod_ord_pair
and sort_oalist_tc = tc.sort_oalist
by (fact comparator_of)
lemma tc_le_lt [simp]: "tc.le = (\<le>)" "tc.lt = (<)"
by (auto simp: le_of_comp_def lt_of_comp_def comparator_of_def intro!: ext split: order.split_asm if_split_asm)
typedef (overloaded) ('a, 'b) oalist_tc = "{xs::('a::linorder \<times> 'b::zero) list. tc.oalist_inv_raw xs}"
morphisms list_of_oalist_tc Abs_oalist_tc
by (auto intro: tc.oalist_inv_raw_Nil)
lemma oalist_tc_eq_iff: "xs = ys \<longleftrightarrow> list_of_oalist_tc xs = list_of_oalist_tc ys"
by (simp add: list_of_oalist_tc_inject)
lemma oalist_tc_eqI: "list_of_oalist_tc xs = list_of_oalist_tc ys \<Longrightarrow> xs = ys"
by (simp add: oalist_tc_eq_iff)
text \<open>Formal, totalized constructor for @{typ "('a, 'b) oalist_tc"}:\<close>
definition OAlist_tc :: "('a \<times> 'b) list \<Rightarrow> ('a::linorder, 'b::zero) oalist_tc" where
"OAlist_tc xs = Abs_oalist_tc (sort_oalist_tc xs)"
definition "oalist_tc_of_list = OAlist_tc"
lemma oalist_inv_list_of_oalist_tc: "tc.oalist_inv_raw (list_of_oalist_tc xs)"
using list_of_oalist_tc[of xs] by simp
lemma list_of_oalist_OAlist_tc: "list_of_oalist_tc (OAlist_tc xs) = sort_oalist_tc xs"
by (simp add: OAlist_tc_def Abs_oalist_tc_inverse tc.oalist_inv_raw_sort_oalist)
lemma OAlist_list_of_oalist_tc [code abstype]: "OAlist_tc (list_of_oalist_tc xs) = xs"
by (simp add: OAlist_tc_def tc.sort_oalist_id list_of_oalist_tc_inverse oalist_inv_list_of_oalist_tc)
lemma list_of_oalist_tc_of_list [code abstract]: "list_of_oalist_tc (oalist_tc_of_list xs) = sort_oalist_tc xs"
by (simp add: list_of_oalist_OAlist_tc oalist_tc_of_list_def)
lemma list_of_oalist_tc_of_list_id:
assumes "tc.oalist_inv_raw xs"
shows "list_of_oalist_tc (OAlist_tc xs) = xs"
using assms by (simp add: list_of_oalist_OAlist_tc tc.sort_oalist_id)
text \<open>It is better to define the following operations directly instead of interpreting
@{locale oalist_abstract}, because @{locale oalist_abstract} defines the operations via their
\<open>_raw\<close> analogues, whereas in this case we can define them directly via their \<open>_pair\<close> analogues.\<close>
definition OAlist_tc_lookup :: "('a::linorder, 'b::zero) oalist_tc \<Rightarrow> 'a \<Rightarrow> 'b"
where "OAlist_tc_lookup xs = lookup_pair_tc (list_of_oalist_tc xs)"
definition OAlist_tc_sorted_domain :: "('a::linorder, 'b::zero) oalist_tc \<Rightarrow> 'a list"
where "OAlist_tc_sorted_domain xs = map fst (list_of_oalist_tc xs)"
definition OAlist_tc_empty :: "('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_empty = OAlist_tc []"
definition OAlist_tc_except_min :: "('a, 'b) oalist_tc \<Rightarrow> ('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_except_min xs = OAlist_tc (tl (list_of_oalist_tc xs))"
definition OAlist_tc_min_key_val :: "('a::linorder, 'b::zero) oalist_tc \<Rightarrow> ('a \<times> 'b)"
where "OAlist_tc_min_key_val xs = hd (list_of_oalist_tc xs)"
definition OAlist_tc_insert :: "('a \<times> 'b) \<Rightarrow> ('a, 'b) oalist_tc \<Rightarrow> ('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_insert x xs = OAlist_tc (update_by_pair_tc x (list_of_oalist_tc xs))"
definition OAlist_tc_update_by_fun :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) oalist_tc \<Rightarrow> ('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_update_by_fun k f xs = OAlist_tc (update_by_fun_pair_tc k f (list_of_oalist_tc xs))"
definition OAlist_tc_update_by_fun_gr :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) oalist_tc \<Rightarrow> ('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_update_by_fun_gr k f xs = OAlist_tc (update_by_fun_gr_pair_tc k f (list_of_oalist_tc xs))"
definition OAlist_tc_filter :: "(('a \<times> 'b) \<Rightarrow> bool) \<Rightarrow> ('a, 'b) oalist_tc \<Rightarrow> ('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_filter P xs = OAlist_tc (filter P (list_of_oalist_tc xs))"
definition OAlist_tc_map_val :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a, 'b::zero) oalist_tc \<Rightarrow> ('a::linorder, 'c::zero) oalist_tc"
where "OAlist_tc_map_val f xs = OAlist_tc (map_val_pair f (list_of_oalist_tc xs))"
definition OAlist_tc_map2_val :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> ('a, 'b::zero) oalist_tc \<Rightarrow> ('a, 'c::zero) oalist_tc \<Rightarrow>
('a::linorder, 'd::zero) oalist_tc"
where "OAlist_tc_map2_val f xs ys =
OAlist_tc (map2_val_pair_tc f (map_val_pair (\<lambda>k b. f k b 0)) (map_val_pair (\<lambda>k. f k 0))
(list_of_oalist_tc xs) (list_of_oalist_tc ys))"
definition OAlist_tc_map2_val_rneutr :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) oalist_tc \<Rightarrow> ('a, 'c::zero) oalist_tc \<Rightarrow>
('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_map2_val_rneutr f xs ys =
OAlist_tc (map2_val_pair_tc f id (map_val_pair (\<lambda>k. f k 0)) (list_of_oalist_tc xs) (list_of_oalist_tc ys))"
definition OAlist_tc_map2_val_neutr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) oalist_tc \<Rightarrow>
('a, 'b) oalist_tc \<Rightarrow> ('a::linorder, 'b::zero) oalist_tc"
where "OAlist_tc_map2_val_neutr f xs ys = OAlist_tc (map2_val_pair_tc f id id (list_of_oalist_tc xs) (list_of_oalist_tc ys))"
definition OAlist_tc_lex_ord :: "('a \<Rightarrow> ('b, 'c) comp_opt) \<Rightarrow> (('a, 'b::zero) oalist_tc, ('a::linorder, 'c::zero) oalist_tc) comp_opt"
where "OAlist_tc_lex_ord f xs ys = lex_ord_pair_tc f (list_of_oalist_tc xs) (list_of_oalist_tc ys)"
definition OAlist_tc_prod_ord :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('a, 'b::zero) oalist_tc \<Rightarrow> ('a::linorder, 'c::zero) oalist_tc \<Rightarrow> bool"
where "OAlist_tc_prod_ord f xs ys = prod_ord_pair_tc f (list_of_oalist_tc xs) (list_of_oalist_tc ys)"
subsubsection \<open>@{const OAlist_tc_lookup}\<close>
lemma OAlist_tc_lookup_eq_valueI: "(k, v) \<in> set (list_of_oalist_tc xs) \<Longrightarrow> OAlist_tc_lookup xs k = v"
unfolding OAlist_tc_lookup_def using oalist_inv_list_of_oalist_tc by (rule tc.lookup_pair_eq_valueI)
lemma OAlist_tc_lookup_inj: "OAlist_tc_lookup xs = OAlist_tc_lookup ys \<Longrightarrow> xs = ys"
by (simp add: OAlist_tc_lookup_def oalist_inv_list_of_oalist_tc oalist_tc_eqI tc.lookup_pair_inj)
lemma OAlist_tc_lookup_oalist_of_list:
"distinct (map fst xs) \<Longrightarrow> OAlist_tc_lookup (oalist_tc_of_list xs) = lookup_dflt xs"
by (simp add: OAlist_tc_lookup_def list_of_oalist_tc_of_list tc.lookup_pair_sort_oalist')
subsubsection \<open>@{const OAlist_tc_sorted_domain}\<close>
lemma set_OAlist_tc_sorted_domain: "set (OAlist_tc_sorted_domain xs) = fst ` set (list_of_oalist_tc xs)"
unfolding OAlist_tc_sorted_domain_def by simp
lemma in_OAlist_tc_sorted_domain_iff_lookup: "k \<in> set (OAlist_tc_sorted_domain xs) \<longleftrightarrow> (OAlist_tc_lookup xs k \<noteq> 0)"
unfolding OAlist_tc_sorted_domain_def OAlist_tc_lookup_def using oalist_inv_list_of_oalist_tc tc.lookup_pair_eq_0
by fastforce
lemma sorted_OAlist_tc_sorted_domain: "sorted_wrt (<) (OAlist_tc_sorted_domain xs)"
unfolding OAlist_tc_sorted_domain_def tc_le_lt[symmetric] using oalist_inv_list_of_oalist_tc
by (rule tc.oalist_inv_rawD2)
subsubsection \<open>@{const OAlist_tc_empty} and Singletons\<close>
lemma list_of_oalist_OAlist_tc_empty [simp, code abstract]: "list_of_oalist_tc OAlist_tc_empty = []"
unfolding OAlist_tc_empty_def using tc.oalist_inv_raw_Nil by (rule list_of_oalist_tc_of_list_id)
lemma lookup_OAlist_tc_empty: "OAlist_tc_lookup OAlist_tc_empty k = 0"
by (simp add: OAlist_tc_lookup_def)
lemma OAlist_tc_lookup_single:
"OAlist_tc_lookup (oalist_tc_of_list [(k, v)]) k' = (if k = k' then v else 0)"
by (simp add: OAlist_tc_lookup_def list_of_oalist_tc_of_list tc.sort_oalist_def comparator_of_def split: order.split)
subsubsection \<open>@{const OAlist_tc_except_min}\<close>
lemma list_of_oalist_OAlist_tc_except_min [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_except_min xs) = tl (list_of_oalist_tc xs)"
unfolding OAlist_tc_except_min_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_tl, fact oalist_inv_list_of_oalist_tc)
lemma lookup_OAlist_tc_except_min:
"OAlist_tc_lookup (OAlist_tc_except_min xs) k =
(if (\<forall>k'\<in>fst ` set (list_of_oalist_tc xs). k \<le> k') then 0 else OAlist_tc_lookup xs k)"
by (simp add: OAlist_tc_lookup_def tc.lookup_pair_tl oalist_inv_list_of_oalist_tc split del: if_split cong: if_cong)
subsubsection \<open>@{const OAlist_tc_min_key_val}\<close>
lemma OAlist_tc_min_key_val_in:
assumes "list_of_oalist_tc xs \<noteq> []"
shows "OAlist_tc_min_key_val xs \<in> set (list_of_oalist_tc xs)"
unfolding OAlist_tc_min_key_val_def using assms by simp
lemma snd_OAlist_tc_min_key_val:
assumes "list_of_oalist_tc xs \<noteq> []"
shows "snd (OAlist_tc_min_key_val xs) = OAlist_tc_lookup xs (fst (OAlist_tc_min_key_val xs))"
proof -
let ?xs = "list_of_oalist_tc xs"
from assms have *: "OAlist_tc_min_key_val xs \<in> set ?xs" by (rule OAlist_tc_min_key_val_in)
show ?thesis unfolding OAlist_tc_lookup_def
by (rule HOL.sym, rule tc.lookup_pair_eq_valueI, fact oalist_inv_list_of_oalist_tc, simp add: *)
qed
lemma OAlist_tc_min_key_val_minimal:
assumes "z \<in> set (list_of_oalist_tc xs)"
shows "fst (OAlist_tc_min_key_val xs) \<le> fst z"
proof -
let ?xs = "list_of_oalist_tc xs"
from assms have "?xs \<noteq> []" by auto
hence "OAlist_tc_sorted_domain xs \<noteq> []" by (simp add: OAlist_tc_sorted_domain_def)
then obtain h xs' where eq: "OAlist_tc_sorted_domain xs = h # xs'" using list.exhaust by blast
with sorted_OAlist_tc_sorted_domain[of xs] have *: "\<forall>y\<in>set xs'. h < y" by simp
from assms have "fst z \<in> set (OAlist_tc_sorted_domain xs)" by (simp add: OAlist_tc_sorted_domain_def)
hence disj: "fst z = h \<or> fst z \<in> set xs'" by (simp add: eq)
from \<open>?xs \<noteq> []\<close> have "fst (OAlist_tc_min_key_val xs) = hd (OAlist_tc_sorted_domain xs)"
by (simp add: OAlist_tc_min_key_val_def OAlist_tc_sorted_domain_def hd_map)
also have "... = h" by (simp add: eq)
also from disj have "... \<le> fst z"
proof
assume "fst z = h"
thus ?thesis by simp
next
assume "fst z \<in> set xs'"
with * have "h < fst z" ..
thus ?thesis by simp
qed
finally show ?thesis .
qed
subsubsection \<open>@{const OAlist_tc_insert}\<close>
lemma list_of_oalist_OAlist_tc_insert [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_insert x xs) = update_by_pair_tc x (list_of_oalist_tc xs)"
unfolding OAlist_tc_insert_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_update_by_pair, fact oalist_inv_list_of_oalist_tc)
lemma lookup_OAlist_tc_insert: "OAlist_tc_lookup (OAlist_tc_insert (k, v) xs) k' = (if k = k' then v else OAlist_tc_lookup xs k')"
by (simp add: OAlist_tc_lookup_def tc.lookup_pair_update_by_pair oalist_inv_list_of_oalist_tc split del: if_split cong: if_cong)
subsubsection \<open>@{const OAlist_tc_update_by_fun} and @{const OAlist_tc_update_by_fun_gr}\<close>
lemma list_of_oalist_OAlist_tc_update_by_fun [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_update_by_fun k f xs) = update_by_fun_pair_tc k f (list_of_oalist_tc xs)"
unfolding OAlist_tc_update_by_fun_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_update_by_fun_pair, fact oalist_inv_list_of_oalist_tc)
lemma lookup_OAlist_tc_update_by_fun:
"OAlist_tc_lookup (OAlist_tc_update_by_fun k f xs) k' = (if k = k' then f else id) (OAlist_tc_lookup xs k')"
by (simp add: OAlist_tc_lookup_def tc.lookup_pair_update_by_fun_pair oalist_inv_list_of_oalist_tc split del: if_split cong: if_cong)
lemma list_of_oalist_OAlist_tc_update_by_fun_gr [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_update_by_fun_gr k f xs) = update_by_fun_gr_pair_tc k f (list_of_oalist_tc xs)"
unfolding OAlist_tc_update_by_fun_gr_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_update_by_fun_gr_pair, fact oalist_inv_list_of_oalist_tc)
lemma OAlist_tc_update_by_fun_gr_eq_OAlist_tc_update_by_fun: "OAlist_tc_update_by_fun_gr = OAlist_tc_update_by_fun"
by (rule, rule, rule,
simp add: OAlist_tc_update_by_fun_gr_def OAlist_tc_update_by_fun_def
tc.update_by_fun_gr_pair_eq_update_by_fun_pair oalist_inv_list_of_oalist_tc)
subsubsection \<open>@{const OAlist_tc_filter}\<close>
lemma list_of_oalist_OAlist_tc_filter [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_filter P xs) = filter P (list_of_oalist_tc xs)"
unfolding OAlist_tc_filter_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_filter, fact oalist_inv_list_of_oalist_tc)
lemma lookup_OAlist_tc_filter: "OAlist_tc_lookup (OAlist_tc_filter P xs) k = (let v = OAlist_tc_lookup xs k in if P (k, v) then v else 0)"
by (simp add: OAlist_tc_lookup_def tc.lookup_pair_filter oalist_inv_list_of_oalist_tc)
subsubsection \<open>@{const OAlist_tc_map_val}\<close>
lemma list_of_oalist_OAlist_tc_map_val [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_map_val f xs) = map_val_pair f (list_of_oalist_tc xs)"
unfolding OAlist_tc_map_val_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map_val_pair, fact oalist_inv_list_of_oalist_tc)
lemma OAlist_tc_map_val_cong:
assumes "\<And>k v. (k, v) \<in> set (list_of_oalist_tc xs) \<Longrightarrow> f k v = g k v"
shows "OAlist_tc_map_val f xs = OAlist_tc_map_val g xs"
unfolding OAlist_tc_map_val_def by (rule arg_cong[where f=OAlist_tc], rule tc.map_val_pair_cong, elim assms)
lemma lookup_OAlist_tc_map_val: "f k 0 = 0 \<Longrightarrow> OAlist_tc_lookup (OAlist_tc_map_val f xs) k = f k (OAlist_tc_lookup xs k)"
by (simp add: OAlist_tc_lookup_def tc.lookup_pair_map_val_pair oalist_inv_list_of_oalist_tc)
subsubsection \<open>@{const OAlist_tc_map2_val} @{const OAlist_tc_map2_val_rneutr} and @{const OAlist_tc_map2_val_neutr}\<close>
lemma list_of_oalist_map2_val [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_map2_val f xs ys) =
map2_val_pair_tc f (map_val_pair (\<lambda>k b. f k b 0)) (map_val_pair (\<lambda>k. f k 0)) (list_of_oalist_tc xs) (list_of_oalist_tc ys)"
unfolding OAlist_tc_map2_val_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map2_val_pair,
fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc,
fact tc.map2_val_compat_map_val_pair, fact tc.map2_val_compat_map_val_pair)
lemma list_of_oalist_OAlist_tc_map2_val_rneutr [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_map2_val_rneutr f xs ys) =
map2_val_pair_tc f id (map_val_pair (\<lambda>k c. f k 0 c)) (list_of_oalist_tc xs) (list_of_oalist_tc ys)"
unfolding OAlist_tc_map2_val_rneutr_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map2_val_pair,
fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc,
fact tc.map2_val_compat_id, fact tc.map2_val_compat_map_val_pair)
lemma list_of_oalist_OAlist_tc_map2_val_neutr [simp, code abstract]:
"list_of_oalist_tc (OAlist_tc_map2_val_neutr f xs ys) = map2_val_pair_tc f id id (list_of_oalist_tc xs) (list_of_oalist_tc ys)"
unfolding OAlist_tc_map2_val_neutr_def
by (rule list_of_oalist_tc_of_list_id, rule tc.oalist_inv_raw_map2_val_pair,
fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc,
fact tc.map2_val_compat_id, fact tc.map2_val_compat_id)
lemma lookup_OAlist_tc_map2_val:
assumes "\<And>k. f k 0 0 = 0"
shows "OAlist_tc_lookup (OAlist_tc_map2_val f xs ys) k = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)"
by (simp add: OAlist_tc_lookup_def tc.lookup_pair_map2_val_pair
tc.map2_val_compat_map_val_pair assms oalist_inv_list_of_oalist_tc)
lemma lookup_OAlist_tc_map2_val_rneutr:
assumes "\<And>k x. f k x 0 = x"
shows "OAlist_tc_lookup (OAlist_tc_map2_val_rneutr f xs ys) k = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)"
proof (simp add: OAlist_tc_lookup_def, rule tc.lookup_pair_map2_val_pair)
fix zs::"('a \<times> 'b) list"
assume "tc.oalist_inv_raw zs"
thus "id zs = map_val_pair (\<lambda>k v. f k v 0) zs" by (simp add: assms tc.map_pair_id)
qed (fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc,
fact tc.map2_val_compat_id, fact tc.map2_val_compat_map_val_pair, rule refl, simp only: assms)
lemma lookup_OAlist_tc_map2_val_neutr:
assumes "\<And>k x. f k x 0 = x" and "\<And>k x. f k 0 x = x"
shows "OAlist_tc_lookup (OAlist_tc_map2_val_neutr f xs ys) k = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)"
proof (simp add: OAlist_tc_lookup_def, rule tc.lookup_pair_map2_val_pair)
fix zs::"('a \<times> 'b) list"
assume "tc.oalist_inv_raw zs"
thus "id zs = map_val_pair (\<lambda>k v. f k v 0) zs" by (simp add: assms(1) tc.map_pair_id)
next
fix zs::"('a \<times> 'b) list"
assume "tc.oalist_inv_raw zs"
thus "id zs = map_val_pair (\<lambda>k. f k 0) zs" by (simp add: assms(2) tc.map_pair_id)
qed (fact oalist_inv_list_of_oalist_tc, fact oalist_inv_list_of_oalist_tc,
fact tc.map2_val_compat_id, fact tc.map2_val_compat_id, simp only: assms(1))
lemma OAlist_tc_map2_val_rneutr_singleton_eq_OAlist_tc_update_by_fun:
assumes "\<And>a x. f a x 0 = x" and "list_of_oalist_tc ys = [(k, v)]"
shows "OAlist_tc_map2_val_rneutr f xs ys = OAlist_tc_update_by_fun k (\<lambda>x. f k x v) xs"
by (simp add: OAlist_tc_map2_val_rneutr_def OAlist_tc_update_by_fun_def assms
tc.map2_val_pair_singleton_eq_update_by_fun_pair oalist_inv_list_of_oalist_tc)
subsubsection \<open>@{const OAlist_tc_lex_ord} and @{const OAlist_tc_prod_ord}\<close>
lemma OAlist_tc_lex_ord_EqI:
"(\<And>k. k \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys) \<Longrightarrow>
f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k) = Some Eq) \<Longrightarrow>
OAlist_tc_lex_ord f xs ys = Some Eq"
by (simp add: OAlist_tc_lex_ord_def OAlist_tc_lookup_def, rule tc.lex_ord_pair_EqI,
rule oalist_inv_list_of_oalist_tc, rule oalist_inv_list_of_oalist_tc, blast)
lemma OAlist_tc_lex_ord_valI:
assumes "aux \<noteq> Some Eq" and "k \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys)"
shows "aux = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k) \<Longrightarrow>
(\<And>k'. k' \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys) \<Longrightarrow>
k' < k \<Longrightarrow> f k' (OAlist_tc_lookup xs k') (OAlist_tc_lookup ys k') = Some Eq) \<Longrightarrow>
OAlist_tc_lex_ord f xs ys = aux"
by (simp (no_asm_use) add: OAlist_tc_lex_ord_def OAlist_tc_lookup_def, rule tc.lex_ord_pair_valI,
rule oalist_inv_list_of_oalist_tc, rule oalist_inv_list_of_oalist_tc, rule assms(1), rule assms(2), simp_all)
lemma OAlist_tc_lex_ord_EqD:
"OAlist_tc_lex_ord f xs ys = Some Eq \<Longrightarrow>
k \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys) \<Longrightarrow>
f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k) = Some Eq"
by (simp add: OAlist_tc_lex_ord_def OAlist_tc_lookup_def, rule tc.lex_ord_pair_EqD[where f=f],
rule oalist_inv_list_of_oalist_tc, rule oalist_inv_list_of_oalist_tc, assumption, simp)
lemma OAlist_tc_lex_ord_valE:
assumes "OAlist_tc_lex_ord f xs ys = aux" and "aux \<noteq> Some Eq"
obtains k where "k \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys)"
and "aux = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)"
and "\<And>k'. k' \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys) \<Longrightarrow>
k' < k \<Longrightarrow> f k' (OAlist_tc_lookup xs k') (OAlist_tc_lookup ys k') = Some Eq"
proof -
note oalist_inv_list_of_oalist_tc oalist_inv_list_of_oalist_tc
moreover from assms(1) have "lex_ord_pair_tc f (list_of_oalist_tc xs) (list_of_oalist_tc ys) = aux"
by (simp only: OAlist_tc_lex_ord_def)
ultimately obtain k where 1: "k \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys)"
and "aux = f k (lookup_pair_tc (list_of_oalist_tc xs) k) (lookup_pair_tc (list_of_oalist_tc ys) k)"
and "\<And>k'. k' \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys) \<Longrightarrow>
k' < k \<Longrightarrow>
f k' (lookup_pair_tc (list_of_oalist_tc xs) k') (lookup_pair_tc (list_of_oalist_tc ys) k') = Some Eq"
using assms(2) unfolding tc_le_lt[symmetric] by (rule tc.lex_ord_pair_valE, blast)
from this(2, 3) have "aux = f k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k)"
and "\<And>k'. k' \<in> fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys) \<Longrightarrow>
k' < k \<Longrightarrow> f k' (OAlist_tc_lookup xs k') (OAlist_tc_lookup ys k') = Some Eq"
by (simp_all only: OAlist_tc_lookup_def)
with 1 show ?thesis ..
qed
lemma OAlist_tc_prod_ord_alt:
"OAlist_tc_prod_ord P xs ys \<longleftrightarrow>
(\<forall>k\<in>fst ` set (list_of_oalist_tc xs) \<union> fst ` set (list_of_oalist_tc ys).
P k (OAlist_tc_lookup xs k) (OAlist_tc_lookup ys k))"
by (simp add: OAlist_tc_prod_ord_def OAlist_tc_lookup_def tc.prod_ord_pair_alt oalist_inv_list_of_oalist_tc)
subsubsection \<open>Instance of @{class equal}\<close>
instantiation oalist_tc :: (linorder, zero) equal
begin
definition equal_oalist_tc :: "('a, 'b) oalist_tc \<Rightarrow> ('a, 'b) oalist_tc \<Rightarrow> bool"
where "equal_oalist_tc xs ys = (list_of_oalist_tc xs = list_of_oalist_tc ys)"
instance by (intro_classes, simp add: equal_oalist_tc_def list_of_oalist_tc_inject)
end
subsection \<open>Experiment\<close>
lemma "oalist_tc_of_list [(0::nat, 4::nat), (1, 3), (0, 2), (1, 1)] = oalist_tc_of_list [(0, 4), (1, 3)]"
by eval
lemma "OAlist_tc_except_min (oalist_tc_of_list ([(1, 3), (0::nat, 4::nat), (0, 2), (1, 1)])) = oalist_tc_of_list [(1, 3)]"
by eval
lemma "OAlist_tc_min_key_val (oalist_tc_of_list [(1, 3), (0::nat, 4::nat), (0, 2), (1, 1)]) = (0, 4)"
by eval
lemma "OAlist_tc_lookup (oalist_tc_of_list [(0::nat, 4::nat), (1, 3), (0, 2), (1, 1)]) 1 = 3"
by eval
lemma "OAlist_tc_prod_ord (\<lambda>_. greater_eq)
(oalist_tc_of_list [(1, 4), (0::nat, 4::nat), (1, 3), (0, 2), (3, 1)])
(oalist_tc_of_list [(0, 4), (1, 3), (2, 2), (1, 1)]) = False"
by eval
lemma "OAlist_tc_map2_val_rneutr (\<lambda>_. minus)
(oalist_tc_of_list [(1, 4), (0::nat, 4::int), (1, 3), (0, 2), (3, 1)])
(oalist_tc_of_list [(0, 4), (1, 3), (2, 2), (1, 1)]) =
oalist_tc_of_list [(1, 1), (2, - 2), (3, 1)]"
by eval
end (* theory *)
diff --git a/thys/Polynomials/OAlist_Poly_Mapping.thy b/thys/Polynomials/OAlist_Poly_Mapping.thy
--- a/thys/Polynomials/OAlist_Poly_Mapping.thy
+++ b/thys/Polynomials/OAlist_Poly_Mapping.thy
@@ -1,494 +1,494 @@
(* Author: Alexander Maletzky *)
section \<open>Ordered Associative Lists for Polynomials\<close>
theory OAlist_Poly_Mapping
imports PP_Type MPoly_Type_Class_Ordered OAlist
begin
text \<open>We introduce a dedicated type for ordered associative lists (oalists) representing polynomials.
To that end, we require the order relation the oalists are sorted wrt. to be admissible term orders,
and furthermore sort the lists @{emph \<open>descending\<close>} rather than @{emph \<open>ascending\<close>}, because this
allows to implement various operations more efficiently.
For technical reasons, we must restrict the type of terms to types embeddable into
@{typ "(nat, nat) pp \<times> nat"}, though. All types we are interested in meet this requirement.\<close>
lemma comparator_lexicographic:
fixes f::"'a \<Rightarrow> 'b" and g::"'a \<Rightarrow> 'c"
assumes "comparator c1" and "comparator c2" and "\<And>x y. f x = f y \<Longrightarrow> g x = g y \<Longrightarrow> x = y"
shows "comparator (\<lambda>x y. case c1 (f x) (f y) of Eq \<Rightarrow> c2 (g x) (g y) | val \<Rightarrow> val)"
(is "comparator ?c3")
proof -
from assms(1) interpret c1: comparator c1 .
from assms(2) interpret c2: comparator c2 .
show ?thesis
proof
fix x y :: 'a
show "invert_order (?c3 x y) = ?c3 y x"
by (simp add: c1.eq c2.eq split: order.split,
metis invert_order.simps(1) invert_order.simps(2) c1.sym c2.sym order.distinct(5))
next
fix x y :: 'a
assume "?c3 x y = Eq"
hence "f x = f y" and "g x = g y" by (simp_all add: c1.eq c2.eq split: order.splits if_split_asm)
thus "x = y" by (rule assms(3))
next
fix x y z :: 'a
assume "?c3 x y = Lt"
hence d1: "c1 (f x) (f y) = Lt \<or> (c1 (f x) (f y) = Eq \<and> c2 (g x) (g y) = Lt)"
by (simp split: order.splits)
assume "?c3 y z = Lt"
hence d2: "c1 (f y) (f z) = Lt \<or> (c1 (f y) (f z) = Eq \<and> c2 (g y) (g z) = Lt)"
by (simp split: order.splits)
from d1 show "?c3 x z = Lt"
proof
assume 1: "c1 (f x) (f y) = Lt"
from d2 show ?thesis
proof
assume "c1 (f y) (f z) = Lt"
- with 1 have "c1 (f x) (f z) = Lt" by (rule c1.trans)
+ with 1 have "c1 (f x) (f z) = Lt" by (rule c1.comp_trans)
thus ?thesis by simp
next
assume "c1 (f y) (f z) = Eq \<and> c2 (g y) (g z) = Lt"
hence "f z = f y" and "c2 (g y) (g z) = Lt" by (simp_all add: c1.eq)
with 1 show ?thesis by simp
qed
next
assume "c1 (f x) (f y) = Eq \<and> c2 (g x) (g y) = Lt"
hence 1: "f x = f y" and 2: "c2 (g x) (g y) = Lt" by (simp_all add: c1.eq)
from d2 show ?thesis
proof
assume "c1 (f y) (f z) = Lt"
thus ?thesis by (simp add: 1)
next
assume "c1 (f y) (f z) = Eq \<and> c2 (g y) (g z) = Lt"
hence 3: "f y = f z" and "c2 (g y) (g z) = Lt" by (simp_all add: c1.eq)
- from 2 this(2) have "c2 (g x) (g z) = Lt" by (rule c2.trans)
+ from 2 this(2) have "c2 (g x) (g z) = Lt" by (rule c2.comp_trans)
thus ?thesis by (simp add: 1 3)
qed
qed
qed
qed
class nat_term =
fixes rep_nat_term :: "'a \<Rightarrow> ((nat, nat) pp \<times> nat)"
and splus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
assumes rep_nat_term_inj: "rep_nat_term x = rep_nat_term y \<Longrightarrow> x = y"
and full_component: "snd (rep_nat_term x) = i \<Longrightarrow> (\<exists>y. rep_nat_term y = (t, i))"
and splus_term: "rep_nat_term (splus x y) = pprod.splus (fst (rep_nat_term x)) (rep_nat_term y)"
begin
definition "lex_comp_aux = (\<lambda>x y. case comp_of_ord lex_pp (fst (rep_nat_term x)) (fst (rep_nat_term y)) of
Eq \<Rightarrow> comparator_of (snd (rep_nat_term x)) (snd (rep_nat_term y)) | val \<Rightarrow> val)"
lemma full_componentE:
assumes "snd (rep_nat_term x) = i"
obtains y where "rep_nat_term y = (t, i)"
proof -
from assms have "\<exists>y. rep_nat_term y = (t, i)" by (rule full_component)
then obtain y where "rep_nat_term y = (t, i)" ..
thus ?thesis ..
qed
end
class nat_pp_term = nat_term + zero + plus +
assumes rep_nat_term_zero: "rep_nat_term 0 = (0, 0)"
and splus_pp_term: "splus = (+)"
definition nat_term_comp :: "'a::nat_term comparator \<Rightarrow> bool"
where "nat_term_comp cmp \<longleftrightarrow>
(\<forall>u v. snd (rep_nat_term u) = snd (rep_nat_term v) \<longrightarrow> fst (rep_nat_term u) = 0 \<longrightarrow> cmp u v \<noteq> Gt) \<and>
(\<forall>u v. fst (rep_nat_term u) = fst (rep_nat_term v) \<longrightarrow> snd (rep_nat_term u) < snd (rep_nat_term v) \<longrightarrow> cmp u v = Lt) \<and>
(\<forall>t u v. cmp u v = Lt \<longrightarrow> cmp (splus t u) (splus t v) = Lt) \<and>
(\<forall>u v a b. fst (rep_nat_term u) = fst (rep_nat_term a) \<longrightarrow> fst (rep_nat_term v) = fst (rep_nat_term b) \<longrightarrow>
snd (rep_nat_term u) = snd (rep_nat_term v) \<longrightarrow> snd (rep_nat_term a) = snd (rep_nat_term b) \<longrightarrow>
cmp a b = Lt \<longrightarrow> cmp u v = Lt)"
lemma nat_term_compI:
assumes "\<And>u v. snd (rep_nat_term u) = snd (rep_nat_term v) \<Longrightarrow> fst (rep_nat_term u) = 0 \<Longrightarrow> cmp u v \<noteq> Gt"
and "\<And>u v. fst (rep_nat_term u) = fst (rep_nat_term v) \<Longrightarrow> snd (rep_nat_term u) < snd (rep_nat_term v) \<Longrightarrow> cmp u v = Lt"
and "\<And>t u v. cmp u v = Lt \<Longrightarrow> cmp (splus t u) (splus t v) = Lt"
and "\<And>u v a b. fst (rep_nat_term u) = fst (rep_nat_term a) \<Longrightarrow> fst (rep_nat_term v) = fst (rep_nat_term b) \<Longrightarrow>
snd (rep_nat_term u) = snd (rep_nat_term v) \<Longrightarrow> snd (rep_nat_term a) = snd (rep_nat_term b) \<Longrightarrow>
cmp a b = Lt \<Longrightarrow> cmp u v = Lt"
shows "nat_term_comp cmp"
unfolding nat_term_comp_def fst_conv snd_conv using assms by blast
lemma nat_term_compD1:
assumes "nat_term_comp cmp" and "snd (rep_nat_term u) = snd (rep_nat_term v)" and "fst (rep_nat_term u) = 0"
shows "cmp u v \<noteq> Gt"
using assms unfolding nat_term_comp_def fst_conv by blast
lemma nat_term_compD2:
assumes "nat_term_comp cmp" and "fst (rep_nat_term u) = fst (rep_nat_term v)" and "snd (rep_nat_term u) < snd (rep_nat_term v)"
shows "cmp u v = Lt"
using assms unfolding nat_term_comp_def fst_conv snd_conv by blast
lemma nat_term_compD3:
assumes "nat_term_comp cmp" and "cmp u v = Lt"
shows "cmp (splus t u) (splus t v) = Lt"
using assms unfolding nat_term_comp_def snd_conv by blast
lemma nat_term_compD4:
assumes "nat_term_comp cmp" and "fst (rep_nat_term u) = fst (rep_nat_term a)"
and "fst (rep_nat_term v) = fst (rep_nat_term b)" and "snd (rep_nat_term u) = snd (rep_nat_term v)"
and "snd (rep_nat_term a) = snd (rep_nat_term b)" and "cmp a b = Lt"
shows "cmp u v = Lt"
using assms unfolding nat_term_comp_def snd_conv by blast
lemma nat_term_compD1':
assumes "comparator cmp" and "nat_term_comp cmp" and "snd (rep_nat_term u) \<le> snd (rep_nat_term v)"
and "fst (rep_nat_term u) = 0"
shows "cmp u v \<noteq> Gt"
proof (cases "snd (rep_nat_term u) = snd (rep_nat_term v)")
case True
with assms(2) show ?thesis using assms(4) by (rule nat_term_compD1)
next
from assms(1) interpret cmp: comparator cmp .
case False
with assms(3) have a: "snd (rep_nat_term u) < snd (rep_nat_term v)" by simp
from refl obtain w::'a where eq: "rep_nat_term w = (0, snd (rep_nat_term v))" by (rule full_componentE)
have "cmp u w = Lt" by (rule nat_term_compD2, fact assms(2), simp_all add: eq assms(4) a)
moreover have "cmp w v \<noteq> Gt" by (rule nat_term_compD1, fact assms(2), simp_all add: eq)
ultimately show "cmp u v \<noteq> Gt" by (simp add: cmp.nGt_le_conv cmp.Lt_lt_conv)
qed
lemma nat_term_compD4':
assumes "comparator cmp" and "nat_term_comp cmp" and "fst (rep_nat_term u) = fst (rep_nat_term a)"
and "fst (rep_nat_term v) = fst (rep_nat_term b)" and "snd (rep_nat_term u) = snd (rep_nat_term v)"
and "snd (rep_nat_term a) = snd (rep_nat_term b)"
shows "cmp u v = cmp a b"
proof -
from assms(1) interpret cmp: comparator cmp .
show ?thesis
proof (cases "cmp a b")
case Eq
hence "fst (rep_nat_term u) = fst (rep_nat_term v)" by (simp add: cmp.eq assms(3, 4))
hence "rep_nat_term u = rep_nat_term v" using assms(5) by (rule prod_eqI)
hence "u = v" by (rule rep_nat_term_inj)
thus ?thesis by (simp add: Eq)
next
case Lt
with assms(2, 3, 4, 5, 6) have "cmp u v = Lt" by (rule nat_term_compD4)
thus ?thesis by (simp add: Lt)
next
case Gt
hence "cmp b a = Lt" by (simp only: cmp.Gt_lt_conv cmp.Lt_lt_conv)
with assms(2, 4, 3) assms(5, 6)[symmetric] have "cmp v u = Lt" by (rule nat_term_compD4)
hence "cmp u v = Gt" by (simp only: cmp.Gt_lt_conv cmp.Lt_lt_conv)
thus ?thesis by (simp add: Gt)
qed
qed
lemma nat_term_compD4'':
assumes "comparator cmp" and "nat_term_comp cmp" and "fst (rep_nat_term u) = fst (rep_nat_term a)"
and "fst (rep_nat_term v) = fst (rep_nat_term b)" and "snd (rep_nat_term u) \<le> snd (rep_nat_term v)"
and "snd (rep_nat_term a) = snd (rep_nat_term b)" and "cmp a b \<noteq> Gt"
shows "cmp u v \<noteq> Gt"
proof (cases "snd (rep_nat_term u) = snd (rep_nat_term v)")
case True
with assms(1, 2, 3, 4) have "cmp u v = cmp a b" using assms(6) by (rule nat_term_compD4')
thus ?thesis using assms(7) by simp
next
case False
from assms(1) interpret cmp: comparator cmp .
from refl obtain w::'a where w: "rep_nat_term w = (fst (rep_nat_term u), snd (rep_nat_term v))"
by (rule full_componentE)
have 1: "fst (rep_nat_term w) = fst (rep_nat_term a)" and 2: "snd (rep_nat_term w) = snd (rep_nat_term v)"
by (simp_all add: w assms(3))
from False assms(5) have *: "snd (rep_nat_term u) < snd (rep_nat_term v)" by simp
have "cmp u w = Lt" by (rule nat_term_compD2, fact assms(2), simp_all add: * w)
moreover from assms(1, 2) 1 assms(4) 2 assms(6) have "cmp w v = cmp a b" by (rule nat_term_compD4')
- ultimately show ?thesis using assms(7) by (metis cmp.nGt_le_conv cmp.nLt_le_conv cmp.trans)
+ ultimately show ?thesis using assms(7) by (metis cmp.nGt_le_conv cmp.nLt_le_conv cmp.comp_trans)
qed
lemma comparator_lex_comp_aux: "comparator (lex_comp_aux::'a::nat_term comparator)"
unfolding lex_comp_aux_def
proof (rule comparator_composition)
from lex_pp_antisym have as: "antisymp lex_pp" by (rule antisympI)
have "comparator (comp_of_ord (lex_pp::(nat, nat) pp \<Rightarrow> _))"
unfolding comp_of_ord_eq_comp_of_ords[OF as]
by (rule comp_of_ords, unfold_locales,
auto simp: lex_pp_refl intro: lex_pp_trans lex_pp_lin' elim!: lex_pp_antisym)
thus "comparator (\<lambda>x y::((nat, nat) pp \<times> nat). case comp_of_ord lex_pp (fst x) (fst y) of
Eq \<Rightarrow> comparator_of (snd x) (snd y) | val \<Rightarrow> val)"
using comparator_of prod_eqI by (rule comparator_lexicographic)
next
from rep_nat_term_inj show "inj rep_nat_term" by (rule injI)
qed
lemma nat_term_comp_lex_comp_aux: "nat_term_comp (lex_comp_aux::'a::nat_term comparator)"
proof -
from lex_pp_antisym have as: "antisymp lex_pp" by (rule antisympI)
interpret lex: comparator "comp_of_ord (lex_pp::(nat, nat) pp \<Rightarrow> _)"
unfolding comp_of_ord_eq_comp_of_ords[OF as]
by (rule comp_of_ords, unfold_locales,
auto simp: lex_pp_refl intro: lex_pp_trans lex_pp_lin' elim!: lex_pp_antisym)
show ?thesis
proof (rule nat_term_compI)
fix u v :: 'a
assume 1: "snd (rep_nat_term u) = snd (rep_nat_term v)" and 2: "fst (rep_nat_term u) = 0"
show "lex_comp_aux u v \<noteq> Gt"
by (simp add: lex_comp_aux_def 1 2 split: order.split, simp add: comp_of_ord_def lex_pp_zero_min)
next
fix u v :: 'a
assume 1: "fst (rep_nat_term u) = fst (rep_nat_term v)" and 2: "snd (rep_nat_term u) < snd (rep_nat_term v)"
show "lex_comp_aux u v = Lt"
by (simp add: lex_comp_aux_def 1 split: order.split, simp add: comparator_of_def 2)
next
fix t u v :: 'a
show "lex_comp_aux u v = Lt \<Longrightarrow> lex_comp_aux (splus t u) (splus t v) = Lt"
by (auto simp: lex_comp_aux_def splus_term pprod.splus_def comp_of_ord_def lex_pp_refl
split: order.splits if_splits intro: lex_pp_plus_monotone')
next
fix u v a b :: 'a
assume "fst (rep_nat_term u) = fst (rep_nat_term a)" and "fst (rep_nat_term v) = fst (rep_nat_term b)"
and "snd (rep_nat_term a) = snd (rep_nat_term b)" and "lex_comp_aux a b = Lt"
thus "lex_comp_aux u v = Lt" by (simp add: lex_comp_aux_def split: order.splits)
qed
qed
typedef (overloaded) 'a nat_term_order =
"{cmp::'a::nat_term comparator. comparator cmp \<and> nat_term_comp cmp}"
morphisms nat_term_compare Abs_nat_term_order
proof (rule, simp)
from comparator_lex_comp_aux nat_term_comp_lex_comp_aux
show "comparator lex_comp_aux \<and> nat_term_comp lex_comp_aux" ..
qed
lemma nat_term_compare_Abs_nat_term_order_id:
assumes "comparator cmp" and "nat_term_comp cmp"
shows "nat_term_compare (Abs_nat_term_order cmp) = cmp"
by (rule Abs_nat_term_order_inverse, simp add: assms)
instantiation nat_term_order :: (type) equal
begin
definition equal_nat_term_order :: "'a nat_term_order \<Rightarrow> 'a nat_term_order \<Rightarrow> bool" where "equal_nat_term_order = (=)"
instance by (standard, simp add: equal_nat_term_order_def)
end
definition nat_term_compare_inv :: "'a nat_term_order \<Rightarrow> 'a::nat_term comparator"
where "nat_term_compare_inv to = (\<lambda>x y. nat_term_compare to y x)"
definition key_order_of_nat_term_order :: "'a nat_term_order \<Rightarrow> 'a::nat_term key_order"
where key_order_of_nat_term_order_def [code del]:
"key_order_of_nat_term_order to = Abs_key_order (nat_term_compare to)"
definition key_order_of_nat_term_order_inv :: "'a nat_term_order \<Rightarrow> 'a::nat_term key_order"
where key_order_of_nat_term_order_inv_def [code del]:
"key_order_of_nat_term_order_inv to = Abs_key_order (nat_term_compare_inv to)"
definition le_of_nat_term_order :: "'a nat_term_order \<Rightarrow> 'a \<Rightarrow> 'a::nat_term \<Rightarrow> bool"
where "le_of_nat_term_order to = le_of_key_order (key_order_of_nat_term_order to)"
definition lt_of_nat_term_order :: "'a nat_term_order \<Rightarrow> 'a \<Rightarrow> 'a::nat_term \<Rightarrow> bool"
where "lt_of_nat_term_order to = lt_of_key_order (key_order_of_nat_term_order to)"
definition nat_term_order_of_le :: "'a::{linorder,nat_term} nat_term_order"
where "nat_term_order_of_le = Abs_nat_term_order (comparator_of)"
lemma comparator_nat_term_compare: "comparator (nat_term_compare to)"
using nat_term_compare by blast
lemma nat_term_comp_nat_term_compare: "nat_term_comp (nat_term_compare to)"
using nat_term_compare by blast
lemma nat_term_compare_splus: "nat_term_compare to (splus t u) (splus t v) = nat_term_compare to u v"
proof -
from comparator_nat_term_compare interpret cmp: comparator "nat_term_compare to" .
show ?thesis
proof (cases "nat_term_compare to u v")
case Eq
hence "splus t u = splus t v" by (simp add: cmp.eq)
thus ?thesis by (simp add: cmp.eq Eq)
next
case Lt
moreover from nat_term_comp_nat_term_compare this have "nat_term_compare to (splus t u) (splus t v) = Lt"
by (rule nat_term_compD3)
ultimately show ?thesis by simp
next
case Gt
hence "nat_term_compare to v u = Lt" using cmp.Gt_lt_conv cmp.Lt_lt_conv by auto
with nat_term_comp_nat_term_compare have "nat_term_compare to (splus t v) (splus t u) = Lt"
by (rule nat_term_compD3)
hence "nat_term_compare to (splus t u) (splus t v) = Gt" using cmp.Gt_lt_conv cmp.Lt_lt_conv by auto
with Gt show ?thesis by simp
qed
qed
lemma nat_term_compare_conv: "nat_term_compare to = key_compare (key_order_of_nat_term_order to)"
unfolding key_order_of_nat_term_order_def
by (rule sym, rule Abs_key_order_inverse, simp add: comparator_nat_term_compare)
lemma comparator_nat_term_compare_inv: "comparator (nat_term_compare_inv to)"
unfolding nat_term_compare_inv_def using comparator_nat_term_compare by (rule comparator_converse)
lemma nat_term_compare_inv_conv: "nat_term_compare_inv to = key_compare (key_order_of_nat_term_order_inv to)"
unfolding key_order_of_nat_term_order_inv_def
by (rule sym, rule Abs_key_order_inverse, simp add: comparator_nat_term_compare_inv)
lemma nat_term_compare_inv_alt [code_unfold]: "nat_term_compare_inv to x y = nat_term_compare to y x"
by (simp only: nat_term_compare_inv_def)
lemma le_of_nat_term_order [code]: "le_of_nat_term_order to x y = (nat_term_compare to x y \<noteq> Gt)"
by (simp add: le_of_key_order_alt le_of_nat_term_order_def nat_term_compare_conv)
lemma lt_of_nat_term_order [code]: "lt_of_nat_term_order to x y = (nat_term_compare to x y = Lt)"
by (simp add: lt_of_key_order_alt lt_of_nat_term_order_def nat_term_compare_conv)
lemma le_of_nat_term_order_alt:
"le_of_nat_term_order to = (\<lambda>u v. ko.le (key_order_of_nat_term_order_inv to) v u)"
by (intro ext, simp add: le_of_comp_def nat_term_compare_inv_conv[symmetric] le_of_nat_term_order_def
le_of_key_order_def nat_term_compare_conv[symmetric] nat_term_compare_inv_alt)
lemma lt_of_nat_term_order_alt:
"lt_of_nat_term_order to = (\<lambda>u v. ko.lt (key_order_of_nat_term_order_inv to) v u)"
by (intro ext, simp add: lt_of_comp_def nat_term_compare_inv_conv[symmetric] lt_of_nat_term_order_def
lt_of_key_order_def nat_term_compare_conv[symmetric] nat_term_compare_inv_alt)
lemma linorder_le_of_nat_term_order: "class.linorder (le_of_nat_term_order to) (lt_of_nat_term_order to)"
unfolding le_of_nat_term_order_alt lt_of_nat_term_order_alt using ko.linorder
by (rule linorder.dual_linorder)
lemma le_of_nat_term_order_zero_min: "le_of_nat_term_order to 0 (t::'a::nat_pp_term)"
unfolding le_of_nat_term_order
by (rule nat_term_compD1', fact comparator_nat_term_compare, fact nat_term_comp_nat_term_compare, simp_all add: rep_nat_term_zero)
lemma le_of_nat_term_order_plus_monotone:
assumes "le_of_nat_term_order to s (t::'a::nat_pp_term)"
shows "le_of_nat_term_order to (u + s) (u + t)"
using assms by (simp add: le_of_nat_term_order splus_pp_term[symmetric] nat_term_compare_splus)
global_interpretation ko_ntm: comparator "nat_term_compare_inv ko"
defines lookup_pair_ko_ntm = ko_ntm.lookup_pair
and update_by_pair_ko_ntm = ko_ntm.update_by_pair
and update_by_fun_pair_ko_ntm = ko_ntm.update_by_fun_pair
and update_by_fun_gr_pair_ko_ntm = ko_ntm.update_by_fun_gr_pair
and map2_val_pair_ko_ntm = ko_ntm.map2_val_pair
and lex_ord_pair_ko_ntm = ko_ntm.lex_ord_pair
and prod_ord_pair_ko_ntm = ko_ntm.prod_ord_pair
and sort_oalist_ko_ntm' = ko_ntm.sort_oalist
by (fact comparator_nat_term_compare_inv)
lemma ko_ntm_le: "ko_ntm.le to = (\<lambda>x y. le_of_nat_term_order to y x)"
by (intro ext, simp add: le_of_comp_def le_of_nat_term_order nat_term_compare_inv_def split: order.split)
global_interpretation ko_ntm: oalist_raw key_order_of_nat_term_order_inv
rewrites "comparator.lookup_pair (key_compare (key_order_of_nat_term_order_inv ko)) = lookup_pair_ko_ntm ko"
and "comparator.update_by_pair (key_compare (key_order_of_nat_term_order_inv ko)) = update_by_pair_ko_ntm ko"
and "comparator.update_by_fun_pair (key_compare (key_order_of_nat_term_order_inv ko)) = update_by_fun_pair_ko_ntm ko"
and "comparator.update_by_fun_gr_pair (key_compare (key_order_of_nat_term_order_inv ko)) = update_by_fun_gr_pair_ko_ntm ko"
and "comparator.map2_val_pair (key_compare (key_order_of_nat_term_order_inv ko)) = map2_val_pair_ko_ntm ko"
and "comparator.lex_ord_pair (key_compare (key_order_of_nat_term_order_inv ko)) = lex_ord_pair_ko_ntm ko"
and "comparator.prod_ord_pair (key_compare (key_order_of_nat_term_order_inv ko)) = prod_ord_pair_ko_ntm ko"
and "comparator.sort_oalist (key_compare (key_order_of_nat_term_order_inv ko)) = sort_oalist_ko_ntm' ko"
defines sort_oalist_aux_ko_ntm = ko_ntm.sort_oalist_aux
and lookup_ko_ntm = ko_ntm.lookup_raw
and sorted_domain_ko_ntm = ko_ntm.sorted_domain_raw
and tl_ko_ntm = ko_ntm.tl_raw
and min_key_val_ko_ntm = ko_ntm.min_key_val_raw
and update_by_ko_ntm = ko_ntm.update_by_raw
and update_by_fun_ko_ntm = ko_ntm.update_by_fun_raw
and update_by_fun_gr_ko_ntm = ko_ntm.update_by_fun_gr_raw
and map2_val_ko_ntm = ko_ntm.map2_val_raw
and lex_ord_ko_ntm = ko_ntm.lex_ord_raw
and prod_ord_ko_ntm = ko_ntm.prod_ord_raw
and oalist_eq_ko_ntm = ko_ntm.oalist_eq_raw
and sort_oalist_ko_ntm = ko_ntm.sort_oalist_raw
subgoal by (simp only: lookup_pair_ko_ntm_def nat_term_compare_inv_conv)
subgoal by (simp only: update_by_pair_ko_ntm_def nat_term_compare_inv_conv)
subgoal by (simp only: update_by_fun_pair_ko_ntm_def nat_term_compare_inv_conv)
subgoal by (simp only: update_by_fun_gr_pair_ko_ntm_def nat_term_compare_inv_conv)
subgoal by (simp only: map2_val_pair_ko_ntm_def nat_term_compare_inv_conv)
subgoal by (simp only: lex_ord_pair_ko_ntm_def nat_term_compare_inv_conv)
subgoal by (simp only: prod_ord_pair_ko_ntm_def nat_term_compare_inv_conv)
subgoal by (simp only: sort_oalist_ko_ntm'_def nat_term_compare_inv_conv)
done
lemma compute_min_key_val_ko_ntm [code]:
"min_key_val_ko_ntm ko (xs, ox) =
(if ko = ox then hd else min_list_param (\<lambda>x y. (le_of_nat_term_order ko) (fst y) (fst x))) xs"
proof -
have "ko.le (key_order_of_nat_term_order_inv ko) = (\<lambda>x y. le_of_nat_term_order ko y x)"
by (metis ko.nGt_le_conv le_of_nat_term_order nat_term_compare_inv_conv nat_term_compare_inv_def)
thus ?thesis by (simp only: min_key_val_ko_ntm_def oalist_raw.min_key_val_raw.simps)
qed
typedef (overloaded) ('a, 'b) oalist_ntm =
"{xs::('a, 'b::zero, 'a::nat_term nat_term_order) oalist_raw. ko_ntm.oalist_inv xs}"
morphisms list_of_oalist_ntm Abs_oalist_ntm
by (auto simp: ko_ntm.oalist_inv_def intro: ko.oalist_inv_raw_Nil)
lemma oalist_ntm_eq_iff: "xs = ys \<longleftrightarrow> list_of_oalist_ntm xs = list_of_oalist_ntm ys"
by (simp add: list_of_oalist_ntm_inject)
lemma oalist_ntm_eqI: "list_of_oalist_ntm xs = list_of_oalist_ntm ys \<Longrightarrow> xs = ys"
by (simp add: oalist_ntm_eq_iff)
text \<open>Formal, totalized constructor for @{typ "('a, 'b) oalist_ntm"}:\<close>
definition OAlist_ntm :: "('a \<times> 'b) list \<times> 'a nat_term_order \<Rightarrow> ('a::nat_term, 'b::zero) oalist_ntm"
where "OAlist_ntm xs = Abs_oalist_ntm (sort_oalist_ko_ntm xs)"
definition "oalist_of_list_ntm = OAlist_ntm"
lemma oalist_inv_list_of_oalist_ntm: "ko_ntm.oalist_inv (list_of_oalist_ntm xs)"
using list_of_oalist_ntm[of xs] by simp
lemma list_of_oalist_OAlist_ntm: "list_of_oalist_ntm (OAlist_ntm xs) = sort_oalist_ko_ntm xs"
proof -
obtain xs' ox where xs: "xs = (xs', ox)" by fastforce
have "ko_ntm.oalist_inv (sort_oalist_ko_ntm' ox xs', ox)"
using ko_ntm.oalist_inv_sort_oalist_raw by fastforce
thus ?thesis by (simp add: xs OAlist_ntm_def Abs_oalist_ntm_inverse)
qed
lemma OAlist_list_of_oalist_ntm [simp, code abstype]: "OAlist_ntm (list_of_oalist_ntm xs) = xs"
proof -
obtain xs' ox where xs: "list_of_oalist_ntm xs = (xs', ox)" by fastforce
have "ko_ntm.oalist_inv_raw ox xs'"
by (simp add: xs[symmetric] ko_ntm.oalist_inv_alt[symmetric] nat_term_compare_inv_conv oalist_inv_list_of_oalist_ntm)
thus ?thesis by (simp add: xs OAlist_ntm_def ko_ntm.sort_oalist_id, simp add: list_of_oalist_ntm_inverse xs[symmetric])
qed
lemma [code abstract]: "list_of_oalist_ntm (oalist_of_list_ntm xs) = sort_oalist_ko_ntm xs"
by (simp add: list_of_oalist_OAlist_ntm oalist_of_list_ntm_def)
global_interpretation oa_ntm: oalist_abstract key_order_of_nat_term_order_inv list_of_oalist_ntm OAlist_ntm
defines OAlist_lookup_ntm = oa_ntm.lookup
and OAlist_sorted_domain_ntm = oa_ntm.sorted_domain
and OAlist_empty_ntm = oa_ntm.empty
and OAlist_reorder_ntm = oa_ntm.reorder
and OAlist_tl_ntm = oa_ntm.tl
and OAlist_hd_ntm = oa_ntm.hd
and OAlist_except_min_ntm = oa_ntm.except_min
and OAlist_min_key_val_ntm = oa_ntm.min_key_val
and OAlist_insert_ntm = oa_ntm.insert
and OAlist_update_by_fun_ntm = oa_ntm.update_by_fun
and OAlist_update_by_fun_gr_ntm = oa_ntm.update_by_fun_gr
and OAlist_filter_ntm = oa_ntm.filter
and OAlist_map2_val_neutr_ntm = oa_ntm.map2_val_neutr
and OAlist_eq_ntm = oa_ntm.oalist_eq
apply unfold_locales
subgoal by (fact oalist_inv_list_of_oalist_ntm)
subgoal by (simp only: list_of_oalist_OAlist_ntm sort_oalist_ko_ntm_def)
subgoal by (fact OAlist_list_of_oalist_ntm)
done
global_interpretation oa_ntm: oalist_abstract3 key_order_of_nat_term_order_inv
"list_of_oalist_ntm::('a, 'b) oalist_ntm \<Rightarrow> ('a, 'b::zero, 'a::nat_term nat_term_order) oalist_raw" OAlist_ntm
"list_of_oalist_ntm::('a, 'c) oalist_ntm \<Rightarrow> ('a, 'c::zero, 'a nat_term_order) oalist_raw" OAlist_ntm
"list_of_oalist_ntm::('a, 'd) oalist_ntm \<Rightarrow> ('a, 'd::zero, 'a nat_term_order) oalist_raw" OAlist_ntm
defines OAlist_map_val_ntm = oa_ntm.map_val
and OAlist_map2_val_ntm = oa_ntm.map2_val
and OAlist_map2_val_rneutr_ntm = oa_ntm.map2_val_rneutr
and OAlist_lex_ord_ntm = oa_ntm.lex_ord
and OAlist_prod_ord_ntm = oa_ntm.prod_ord ..
lemmas OAlist_lookup_ntm_single = oa_ntm.lookup_oalist_of_list_single[folded oalist_of_list_ntm_def]
end (* theory *)
diff --git a/thys/Polynomials/Power_Products.thy b/thys/Polynomials/Power_Products.thy
--- a/thys/Polynomials/Power_Products.thy
+++ b/thys/Polynomials/Power_Products.thy
@@ -1,2698 +1,2698 @@
(* Author: Fabian Immler, Alexander Maletzky *)
section \<open>Abstract Power-Products\<close>
theory Power_Products
imports Complex_Main
"HOL-Library.Function_Algebras"
"HOL-Library.Countable"
"More_MPoly_Type"
"Utils"
Well_Quasi_Orders.Well_Quasi_Orders
begin
text \<open>This theory formalizes the concept of "power-products". A power-product can be thought of as
the product of some indeterminates, such as $x$, $x^2\,y$, $x\,y^3\,z^7$, etc., without any
scalar coefficient.
The approach in this theory is to capture the notion of "power-product" (also called "monomial") as
type class. A canonical instance for power-product is the type @{typ "'var \<Rightarrow>\<^sub>0 nat"}, which is
interpreted as mapping from variables in the power-product to exponents.
A slightly unintuitive (but fitting better with the standard type class instantiations of
@{typ "'a \<Rightarrow>\<^sub>0 'b"}) approach is to write addition to denote "multiplication" of power products.
For example, $x^2y$ would be represented as a function \<open>p = (X \<mapsto> 2, Y \<mapsto> 1)\<close>, $xz$ as a
function \<open>q = (X \<mapsto> 1, Z \<mapsto> 1)\<close>. With the (pointwise) instantiation of addition of @{typ "'a \<Rightarrow>\<^sub>0 'b"},
we will write \<open>p + q = (X \<mapsto> 3, Y \<mapsto> 1, Z \<mapsto> 1)\<close> for the product $x^2y \cdot xz = x^3yz$
\<close>
subsection \<open>Constant @{term Keys}\<close>
text \<open>Legacy:\<close>
lemmas keys_eq_empty_iff = keys_eq_empty
definition Keys :: "('a \<Rightarrow>\<^sub>0 'b::zero) set \<Rightarrow> 'a set"
where "Keys F = \<Union>(keys ` F)"
lemma in_Keys: "s \<in> Keys F \<longleftrightarrow> (\<exists>f\<in>F. s \<in> keys f)"
unfolding Keys_def by simp
lemma in_KeysI:
assumes "s \<in> keys f" and "f \<in> F"
shows "s \<in> Keys F"
unfolding in_Keys using assms ..
lemma in_KeysE:
assumes "s \<in> Keys F"
obtains f where "s \<in> keys f" and "f \<in> F"
using assms unfolding in_Keys ..
lemma Keys_mono:
assumes "A \<subseteq> B"
shows "Keys A \<subseteq> Keys B"
using assms by (auto simp add: Keys_def)
lemma Keys_insert: "Keys (insert a A) = keys a \<union> Keys A"
by (simp add: Keys_def)
lemma Keys_Un: "Keys (A \<union> B) = Keys A \<union> Keys B"
by (simp add: Keys_def)
lemma finite_Keys:
assumes "finite A"
shows "finite (Keys A)"
unfolding Keys_def by (rule, fact assms, rule finite_keys)
lemma Keys_not_empty:
assumes "a \<in> A" and "a \<noteq> 0"
shows "Keys A \<noteq> {}"
proof
assume "Keys A = {}"
from \<open>a \<noteq> 0\<close> have "keys a \<noteq> {}" using aux by fastforce
then obtain s where "s \<in> keys a" by blast
from this assms(1) have "s \<in> Keys A" by (rule in_KeysI)
with \<open>Keys A = {}\<close> show False by simp
qed
lemma Keys_empty [simp]: "Keys {} = {}"
by (simp add: Keys_def)
lemma Keys_zero [simp]: "Keys {0} = {}"
by (simp add: Keys_def)
lemma keys_subset_Keys:
assumes "f \<in> F"
shows "keys f \<subseteq> Keys F"
using in_KeysI[OF _ assms] by auto
lemma Keys_minus: "Keys (A - B) \<subseteq> Keys A"
by (auto simp add: Keys_def)
lemma Keys_minus_zero: "Keys (A - {0}) = Keys A"
proof (cases "0 \<in> A")
case True
hence "(A - {0}) \<union> {0} = A" by auto
hence "Keys A = Keys ((A - {0}) \<union> {0})" by simp
also have "... = Keys (A - {0}) \<union> Keys {0::('a \<Rightarrow>\<^sub>0 'b)}" by (fact Keys_Un)
also have "... = Keys (A - {0})" by simp
finally show ?thesis by simp
next
case False
hence "A - {0} = A" by simp
thus ?thesis by simp
qed
subsection \<open>Constant @{term except}\<close>
definition except_fun :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'b::zero)"
where "except_fun f S = (\<lambda>x. (f x when x \<notin> S))"
lift_definition except :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b::zero)" is except_fun
proof -
fix p::"'a \<Rightarrow> 'b" and S::"'a set"
assume "finite {t. p t \<noteq> 0}"
show "finite {t. except_fun p S t \<noteq> 0}"
proof (rule finite_subset[of _ "{t. p t \<noteq> 0}"], rule)
fix u
assume "u \<in> {t. except_fun p S t \<noteq> 0}"
hence "p u \<noteq> 0" by (simp add: except_fun_def)
thus "u \<in> {t. p t \<noteq> 0}" by simp
qed fact
qed
lemma lookup_except_when: "lookup (except p S) = (\<lambda>t. lookup p t when t \<notin> S)"
by (auto simp: except.rep_eq except_fun_def)
lemma lookup_except: "lookup (except p S) = (\<lambda>t. if t \<in> S then 0 else lookup p t)"
by (rule ext) (simp add: lookup_except_when)
lemma lookup_except_singleton: "lookup (except p {t}) t = 0"
by (simp add: lookup_except)
lemma except_zero [simp]: "except 0 S = 0"
by (rule poly_mapping_eqI) (simp add: lookup_except)
lemma lookup_except_eq_idI:
assumes "t \<notin> S"
shows "lookup (except p S) t = lookup p t"
using assms by (simp add: lookup_except)
lemma lookup_except_eq_zeroI:
assumes "t \<in> S"
shows "lookup (except p S) t = 0"
using assms by (simp add: lookup_except)
lemma except_empty [simp]: "except p {} = p"
by (rule poly_mapping_eqI) (simp add: lookup_except)
lemma except_eq_zeroI:
assumes "keys p \<subseteq> S"
shows "except p S = 0"
proof (rule poly_mapping_eqI, simp)
fix t
show "lookup (except p S) t = 0"
proof (cases "t \<in> S")
case True
thus ?thesis by (rule lookup_except_eq_zeroI)
next
case False then show ?thesis
by (metis assms in_keys_iff lookup_except_eq_idI subset_eq)
qed
qed
lemma except_eq_zeroE:
assumes "except p S = 0"
shows "keys p \<subseteq> S"
by (metis assms aux in_keys_iff lookup_except_eq_idI subset_iff)
lemma except_eq_zero_iff: "except p S = 0 \<longleftrightarrow> keys p \<subseteq> S"
by (rule, elim except_eq_zeroE, elim except_eq_zeroI)
lemma except_keys [simp]: "except p (keys p) = 0"
by (rule except_eq_zeroI, rule subset_refl)
lemma plus_except: "p = Poly_Mapping.single t (lookup p t) + except p {t}"
by (rule poly_mapping_eqI, simp add: lookup_add lookup_single lookup_except when_def split: if_split)
lemma keys_except: "keys (except p S) = keys p - S"
by (transfer, auto simp: except_fun_def)
lemma except_single: "except (Poly_Mapping.single u c) S = (Poly_Mapping.single u c when u \<notin> S)"
by (rule poly_mapping_eqI) (simp add: lookup_except lookup_single when_def)
lemma except_plus: "except (p + q) S = except p S + except q S"
by (rule poly_mapping_eqI) (simp add: lookup_except lookup_add)
lemma except_minus: "except (p - q) S = except p S - except q S"
by (rule poly_mapping_eqI) (simp add: lookup_except lookup_minus)
lemma except_uminus: "except (- p) S = - except p S"
by (rule poly_mapping_eqI) (simp add: lookup_except)
lemma except_except: "except (except p S) T = except p (S \<union> T)"
by (rule poly_mapping_eqI) (simp add: lookup_except)
lemma poly_mapping_keys_eqI:
assumes a1: "keys p = keys q" and a2: "\<And>t. t \<in> keys p \<Longrightarrow> lookup p t = lookup q t"
shows "p = q"
proof (rule poly_mapping_eqI)
fix t
show "lookup p t = lookup q t"
proof (cases "t \<in> keys p")
case True
thus ?thesis by (rule a2)
next
case False
moreover from this have "t \<notin> keys q" unfolding a1 .
ultimately have "lookup p t = 0" and "lookup q t = 0" unfolding in_keys_iff by simp_all
thus ?thesis by simp
qed
qed
lemma except_id_iff: "except p S = p \<longleftrightarrow> keys p \<inter> S = {}"
by (metis Diff_Diff_Int Diff_eq_empty_iff Diff_triv inf_le2 keys_except lookup_except_eq_idI
lookup_except_eq_zeroI not_in_keys_iff_lookup_eq_zero poly_mapping_keys_eqI)
lemma keys_subset_wf:
"wfP (\<lambda>p q::('a, 'b::zero) poly_mapping. keys p \<subset> keys q)"
unfolding wfP_def
proof (intro wfI_min)
fix x::"('a, 'b) poly_mapping" and Q
assume x_in: "x \<in> Q"
let ?Q0 = "card ` keys ` Q"
from x_in have "card (keys x) \<in> ?Q0" by simp
from wfE_min[OF wf this] obtain z0
where z0_in: "z0 \<in> ?Q0" and z0_min: "\<And>y. (y, z0) \<in> {(x, y). x < y} \<Longrightarrow> y \<notin> ?Q0" by auto
from z0_in obtain z where z0_def: "z0 = card (keys z)" and "z \<in> Q" by auto
show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> {(p, q). keys p \<subset> keys q} \<longrightarrow> y \<notin> Q"
proof (intro bexI[of _ z], rule, rule)
fix y::"('a, 'b) poly_mapping"
let ?y0 = "card (keys y)"
assume "(y, z) \<in> {(p, q). keys p \<subset> keys q}"
hence "keys y \<subset> keys z" by simp
hence "?y0 < z0" unfolding z0_def by (simp add: psubset_card_mono)
hence "(?y0, z0) \<in> {(x, y). x < y}" by simp
from z0_min[OF this] show "y \<notin> Q" by auto
qed (fact)
qed
lemma poly_mapping_except_induct:
assumes base: "P 0" and ind: "\<And>p t. p \<noteq> 0 \<Longrightarrow> t \<in> keys p \<Longrightarrow> P (except p {t}) \<Longrightarrow> P p"
shows "P p"
proof (induct rule: wfP_induct[OF keys_subset_wf])
fix p::"('a, 'b) poly_mapping"
assume "\<forall>q. keys q \<subset> keys p \<longrightarrow> P q"
hence IH: "\<And>q. keys q \<subset> keys p \<Longrightarrow> P q" by simp
show "P p"
proof (cases "p = 0")
case True
thus ?thesis using base by simp
next
case False
hence "keys p \<noteq> {}" by simp
then obtain t where "t \<in> keys p" by blast
show ?thesis
proof (rule ind, fact, fact, rule IH, simp only: keys_except, rule, rule Diff_subset, rule)
assume "keys p - {t} = keys p"
hence "t \<notin> keys p" by blast
from this \<open>t \<in> keys p\<close> show False ..
qed
qed
qed
lemma poly_mapping_except_induct':
assumes "\<And>p. (\<And>t. t \<in> keys p \<Longrightarrow> P (except p {t})) \<Longrightarrow> P p"
shows "P p"
proof (induct "card (keys p)" arbitrary: p)
case 0
with finite_keys[of p] have "keys p = {}" by simp
show ?case by (rule assms, simp add: \<open>keys p = {}\<close>)
next
case step: (Suc n)
show ?case
proof (rule assms)
fix t
assume "t \<in> keys p"
show "P (except p {t})"
proof (rule step(1), simp add: keys_except)
from step(2) \<open>t \<in> keys p\<close> finite_keys[of p] show "n = card (keys p - {t})" by simp
qed
qed
qed
lemma poly_mapping_plus_induct:
assumes "P 0" and "\<And>p c t. c \<noteq> 0 \<Longrightarrow> t \<notin> keys p \<Longrightarrow> P p \<Longrightarrow> P (Poly_Mapping.single t c + p)"
shows "P p"
proof (induct "card (keys p)" arbitrary: p)
case 0
with finite_keys[of p] have "keys p = {}" by simp
hence "p = 0" by simp
with assms(1) show ?case by simp
next
case step: (Suc n)
from step(2) obtain t where t: "t \<in> keys p" by (metis card_eq_SucD insert_iff)
define c where "c = lookup p t"
define q where "q = except p {t}"
have *: "p = Poly_Mapping.single t c + q"
by (rule poly_mapping_eqI, simp add: lookup_add lookup_single Poly_Mapping.when_def, intro conjI impI,
simp add: q_def lookup_except c_def, simp add: q_def lookup_except_eq_idI)
show ?case
proof (simp only: *, rule assms(2))
from t show "c \<noteq> 0"
using c_def by auto
next
show "t \<notin> keys q" by (simp add: q_def keys_except)
next
show "P q"
proof (rule step(1))
from step(2) \<open>t \<in> keys p\<close> show "n = card (keys q)" unfolding q_def keys_except
by (metis Suc_inject card.remove finite_keys)
qed
qed
qed
lemma except_Diff_singleton: "except p (keys p - {t}) = Poly_Mapping.single t (lookup p t)"
by (rule poly_mapping_eqI) (simp add: lookup_single in_keys_iff lookup_except when_def)
lemma except_Un_plus_Int: "except p (U \<union> V) + except p (U \<inter> V) = except p U + except p V"
by (rule poly_mapping_eqI) (simp add: lookup_except lookup_add)
corollary except_Int:
assumes "keys p \<subseteq> U \<union> V"
shows "except p (U \<inter> V) = except p U + except p V"
proof -
from assms have "except p (U \<union> V) = 0" by (rule except_eq_zeroI)
hence "except p (U \<inter> V) = except p (U \<union> V) + except p (U \<inter> V)" by simp
also have "\<dots> = except p U + except p V" by (fact except_Un_plus_Int)
finally show ?thesis .
qed
lemma except_keys_Int [simp]: "except p (keys p \<inter> U) = except p U"
by (rule poly_mapping_eqI) (simp add: in_keys_iff lookup_except)
lemma except_Int_keys [simp]: "except p (U \<inter> keys p) = except p U"
by (simp only: Int_commute[of U] except_keys_Int)
lemma except_keys_Diff: "except p (keys p - U) = except p (- U)"
proof -
have "except p (keys p - U) = except p (keys p \<inter> (- U))" by (simp only: Diff_eq)
also have "\<dots> = except p (- U)" by simp
finally show ?thesis .
qed
lemma except_decomp: "p = except p U + except p (- U)"
by (rule poly_mapping_eqI) (simp add: lookup_except lookup_add)
corollary except_Compl: "except p (- U) = p - except p U"
by (metis add_diff_cancel_left' except_decomp)
subsection \<open>'Divisibility' on Additive Structures\<close>
context plus begin
definition adds :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "adds" 50)
where "b adds a \<longleftrightarrow> (\<exists>k. a = b + k)"
lemma addsI [intro?]: "a = b + k \<Longrightarrow> b adds a"
unfolding adds_def ..
lemma addsE [elim?]: "b adds a \<Longrightarrow> (\<And>k. a = b + k \<Longrightarrow> P) \<Longrightarrow> P"
unfolding adds_def by blast
end
context comm_monoid_add
begin
lemma adds_refl [simp]: "a adds a"
proof
show "a = a + 0" by simp
qed
lemma adds_trans [trans]:
assumes "a adds b" and "b adds c"
shows "a adds c"
proof -
from assms obtain v where "b = a + v"
by (auto elim!: addsE)
moreover from assms obtain w where "c = b + w"
by (auto elim!: addsE)
ultimately have "c = a + (v + w)"
by (simp add: add.assoc)
then show ?thesis ..
qed
lemma subset_divisors_adds: "{c. c adds a} \<subseteq> {c. c adds b} \<longleftrightarrow> a adds b"
by (auto simp add: subset_iff intro: adds_trans)
lemma strict_subset_divisors_adds: "{c. c adds a} \<subset> {c. c adds b} \<longleftrightarrow> a adds b \<and> \<not> b adds a"
by (auto simp add: subset_iff intro: adds_trans)
lemma zero_adds [simp]: "0 adds a"
by (auto intro!: addsI)
lemma adds_plus_right [simp]: "a adds c \<Longrightarrow> a adds (b + c)"
by (auto intro!: add.left_commute addsI elim!: addsE)
lemma adds_plus_left [simp]: "a adds b \<Longrightarrow> a adds (b + c)"
using adds_plus_right [of a b c] by (simp add: ac_simps)
lemma adds_triv_right [simp]: "a adds b + a"
by (rule adds_plus_right) (rule adds_refl)
lemma adds_triv_left [simp]: "a adds a + b"
by (rule adds_plus_left) (rule adds_refl)
lemma plus_adds_mono:
assumes "a adds b"
and "c adds d"
shows "a + c adds b + d"
proof -
from \<open>a adds b\<close> obtain b' where "b = a + b'" ..
moreover from \<open>c adds d\<close> obtain d' where "d = c + d'" ..
ultimately have "b + d = (a + c) + (b' + d')"
by (simp add: ac_simps)
then show ?thesis ..
qed
lemma plus_adds_left: "a + b adds c \<Longrightarrow> a adds c"
by (simp add: adds_def add.assoc) blast
lemma plus_adds_right: "a + b adds c \<Longrightarrow> b adds c"
using plus_adds_left [of b a c] by (simp add: ac_simps)
end
class ninv_comm_monoid_add = comm_monoid_add +
assumes plus_eq_zero: "s + t = 0 \<Longrightarrow> s = 0"
begin
lemma plus_eq_zero_2: "t = 0" if "s + t = 0"
using that
by (simp only: add_commute[of s t] plus_eq_zero)
lemma adds_zero: "s adds 0 \<longleftrightarrow> (s = 0)"
proof
assume "s adds 0"
from this obtain k where "0 = s + k" unfolding adds_def ..
from this plus_eq_zero[of s k] show "s = 0"
by blast
next
assume "s = 0"
thus "s adds 0" by simp
qed
end
context canonically_ordered_monoid_add
begin
subclass ninv_comm_monoid_add by (standard, simp)
end
class comm_powerprod = cancel_comm_monoid_add
begin
lemma adds_canc: "s + u adds t + u \<longleftrightarrow> s adds t" for s t u::'a
unfolding adds_def
apply auto
apply (metis local.add.left_commute local.add_diff_cancel_left' local.add_diff_cancel_right')
using add_assoc add_commute by auto
lemma adds_canc_2: "u + s adds u + t \<longleftrightarrow> s adds t"
by (simp add: adds_canc ac_simps)
lemma add_minus_2: "(s + t) - s = t"
by simp
lemma adds_minus:
assumes "s adds t"
shows "(t - s) + s = t"
proof -
from assms adds_def[of s t] obtain u where u: "t = u + s" by (auto simp: ac_simps)
then have "t - s = u"
by simp
thus ?thesis using u by simp
qed
lemma plus_adds_0:
assumes "(s + t) adds u"
shows "s adds (u - t)"
proof -
from assms have "(s + t) adds ((u - t) + t)" using adds_minus local.plus_adds_right by presburger
thus ?thesis using adds_canc[of s t "u - t"] by simp
qed
lemma plus_adds_2:
assumes "t adds u" and "s adds (u - t)"
shows "(s + t) adds u"
by (metis adds_canc adds_minus assms)
lemma plus_adds:
shows "(s + t) adds u \<longleftrightarrow> (t adds u \<and> s adds (u - t))"
proof
assume a1: "(s + t) adds u"
show "t adds u \<and> s adds (u - t)"
proof
from plus_adds_right[OF a1] show "t adds u" .
next
from plus_adds_0[OF a1] show "s adds (u - t)" .
qed
next
assume "t adds u \<and> s adds (u - t)"
hence "t adds u" and "s adds (u - t)" by auto
from plus_adds_2[OF \<open>t adds u\<close> \<open>s adds (u - t)\<close>] show "(s + t) adds u" .
qed
lemma minus_plus:
assumes "s adds t"
shows "(t - s) + u = (t + u) - s"
proof -
from assms obtain k where k: "t = s + k" unfolding adds_def ..
hence "t - s = k" by simp
also from k have "(t + u) - s = k + u"
by (simp add: add_assoc)
finally show ?thesis by simp
qed
lemma minus_plus_minus:
assumes "s adds t" and "u adds v"
shows "(t - s) + (v - u) = (t + v) - (s + u)"
using add_commute assms(1) assms(2) diff_diff_add minus_plus by auto
lemma minus_plus_minus_cancel:
assumes "u adds t" and "s adds u"
shows "(t - u) + (u - s) = t - s"
by (metis assms(1) assms(2) local.add_diff_cancel_left' local.add_diff_cancel_right local.addsE minus_plus)
end
text \<open>Instances of class \<open>lcs_powerprod\<close> are types of commutative power-products admitting
(not necessarily unique) least common sums (inspired from least common multiplies).
Note that if the components of indeterminates are arbitrary integers (as for instance in Laurent
polynomials), then no unique lcss exist.\<close>
class lcs_powerprod = comm_powerprod +
fixes lcs::"'a \<Rightarrow> 'a \<Rightarrow> 'a"
assumes adds_lcs: "s adds (lcs s t)"
assumes lcs_adds: "s adds u \<Longrightarrow> t adds u \<Longrightarrow> (lcs s t) adds u"
assumes lcs_comm: "lcs s t = lcs t s"
begin
lemma adds_lcs_2: "t adds (lcs s t)"
by (simp only: lcs_comm[of s t], rule adds_lcs)
lemma lcs_adds_plus: "lcs s t adds s + t" by (simp add: lcs_adds)
text \<open>"gcs" stands for "greatest common summand".\<close>
definition gcs :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where "gcs s t = (s + t) - (lcs s t)"
lemma gcs_plus_lcs: "(gcs s t) + (lcs s t) = s + t"
unfolding gcs_def by (rule adds_minus, fact lcs_adds_plus)
lemma gcs_adds: "(gcs s t) adds s"
proof -
have "t adds (lcs s t)" (is "t adds ?l") unfolding lcs_comm[of s t] by (fact adds_lcs)
then obtain u where eq1: "?l = t + u" unfolding adds_def ..
from lcs_adds_plus[of s t] obtain v where eq2: "s + t = ?l + v" unfolding adds_def ..
hence "t + s = t + (u + v)" unfolding eq1 by (simp add: ac_simps)
hence s: "s = u + v" unfolding add_left_cancel .
show ?thesis unfolding eq2 gcs_def unfolding s by simp
qed
lemma gcs_comm: "gcs s t = gcs t s" unfolding gcs_def by (simp add: lcs_comm ac_simps)
lemma gcs_adds_2: "(gcs s t) adds t"
by (simp only: gcs_comm[of s t], rule gcs_adds)
end
class ulcs_powerprod = lcs_powerprod + ninv_comm_monoid_add
begin
lemma adds_antisym:
assumes "s adds t" "t adds s"
shows "s = t"
proof -
from \<open>s adds t\<close> obtain u where u_def: "t = s + u" unfolding adds_def ..
from \<open>t adds s\<close> obtain v where v_def: "s = t + v" unfolding adds_def ..
from u_def v_def have "s = (s + u) + v" by (simp add: ac_simps)
hence "s + 0 = s + (u + v)" by (simp add: ac_simps)
hence "u + v = 0" by simp
hence "u = 0" using plus_eq_zero[of u v] by simp
thus ?thesis using u_def by simp
qed
lemma lcs_unique:
assumes "s adds l" and "t adds l" and *: "\<And>u. s adds u \<Longrightarrow> t adds u \<Longrightarrow> l adds u"
shows "l = lcs s t"
by (rule adds_antisym, rule *, fact adds_lcs, fact adds_lcs_2, rule lcs_adds, fact+)
lemma lcs_zero: "lcs 0 t = t"
by (rule lcs_unique[symmetric], fact zero_adds, fact adds_refl)
lemma lcs_plus_left: "lcs (u + s) (u + t) = u + lcs s t"
proof (rule lcs_unique[symmetric], simp_all only: adds_canc_2, fact adds_lcs, fact adds_lcs_2,
simp add: add.commute[of u] plus_adds)
fix v
assume "u adds v \<and> s adds v - u"
hence "s adds v - u" ..
assume "t adds v - u"
with \<open>s adds v - u\<close> show "lcs s t adds v - u" by (rule lcs_adds)
qed
lemma lcs_plus_right: "lcs (s + u) (t + u) = (lcs s t) + u"
using lcs_plus_left[of u s t] by (simp add: ac_simps)
lemma adds_gcs:
assumes "u adds s" and "u adds t"
shows "u adds (gcs s t)"
proof -
from assms have "s + u adds s + t" and "t + u adds t + s"
by (simp_all add: plus_adds_mono)
hence "lcs (s + u) (t + u) adds s + t"
by (auto intro: lcs_adds simp add: ac_simps)
hence "u + (lcs s t) adds s + t" unfolding lcs_plus_right by (simp add: ac_simps)
hence "u adds (s + t) - (lcs s t)" unfolding plus_adds ..
thus ?thesis unfolding gcs_def .
qed
lemma gcs_unique:
assumes "g adds s" and "g adds t" and *: "\<And>u. u adds s \<Longrightarrow> u adds t \<Longrightarrow> u adds g"
shows "g = gcs s t"
by (rule adds_antisym, rule adds_gcs, fact, fact, rule *, fact gcs_adds, fact gcs_adds_2)
lemma gcs_plus_left: "gcs (u + s) (u + t) = u + gcs s t"
proof -
have "u + s + (u + t) - (u + lcs s t) = u + s + (u + t) - u - lcs s t" by (simp only: diff_diff_add)
also have "... = u + s + t + (u - u) - lcs s t" by (simp add: add.left_commute)
also have "... = u + s + t - lcs s t" by simp
also have "... = u + (s + t - lcs s t)"
using add_assoc add_commute local.lcs_adds_plus local.minus_plus by auto
finally show ?thesis unfolding gcs_def lcs_plus_left .
qed
lemma gcs_plus_right: "gcs (s + u) (t + u) = (gcs s t) + u"
using gcs_plus_left[of u s t] by (simp add: ac_simps)
lemma lcs_same [simp]: "lcs s s = s"
proof -
have "lcs s s adds s" by (rule lcs_adds, simp_all)
moreover have "s adds lcs s s" by (rule adds_lcs)
ultimately show ?thesis by (rule adds_antisym)
qed
lemma gcs_same [simp]: "gcs s s = s"
proof -
have "gcs s s adds s" by (rule gcs_adds)
moreover have "s adds gcs s s" by (rule adds_gcs, simp_all)
ultimately show ?thesis by (rule adds_antisym)
qed
end
subsection \<open>Dickson Classes\<close>
definition (in plus) dickson_grading :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
where "dickson_grading d \<longleftrightarrow>
((\<forall>s t. d (s + t) = max (d s) (d t)) \<and> (\<forall>n::nat. almost_full_on (adds) {x. d x \<le> n}))"
definition dgrad_set :: "('a \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> 'a set"
where "dgrad_set d m = {t. d t \<le> m}"
definition dgrad_set_le :: "('a \<Rightarrow> nat) \<Rightarrow> ('a set) \<Rightarrow> ('a set) \<Rightarrow> bool"
where "dgrad_set_le d S T \<longleftrightarrow> (\<forall>s\<in>S. \<exists>t\<in>T. d s \<le> d t)"
lemma dickson_gradingI:
assumes "\<And>s t. d (s + t) = max (d s) (d t)"
assumes "\<And>n::nat. almost_full_on (adds) {x. d x \<le> n}"
shows "dickson_grading d"
unfolding dickson_grading_def using assms by blast
lemma dickson_gradingD1: "dickson_grading d \<Longrightarrow> d (s + t) = max (d s) (d t)"
by (auto simp add: dickson_grading_def)
lemma dickson_gradingD2: "dickson_grading d \<Longrightarrow> almost_full_on (adds) {x. d x \<le> n}"
by (auto simp add: dickson_grading_def)
lemma dickson_gradingD2':
assumes "dickson_grading (d::'a::comm_monoid_add \<Rightarrow> nat)"
shows "wqo_on (adds) {x. d x \<le> n}"
proof (intro wqo_onI transp_onI)
fix x y z :: 'a
assume "x adds y" and "y adds z"
thus "x adds z" by (rule adds_trans)
next
from assms show "almost_full_on (adds) {x. d x \<le> n}" by (rule dickson_gradingD2)
qed
lemma dickson_gradingE:
assumes "dickson_grading d" and "\<And>i::nat. d ((seq::nat \<Rightarrow> 'a::plus) i) \<le> n"
obtains i j where "i < j" and "seq i adds seq j"
proof -
from assms(1) have "almost_full_on (adds) {x. d x \<le> n}" by (rule dickson_gradingD2)
moreover from assms(2) have "\<And>i. seq i \<in> {x. d x \<le> n}" by simp
ultimately obtain i j where "i < j" and "seq i adds seq j" by (rule almost_full_onD)
thus ?thesis ..
qed
lemma dickson_grading_adds_imp_le:
assumes "dickson_grading d" and "s adds t"
shows "d s \<le> d t"
proof -
from assms(2) obtain u where "t = s + u" ..
hence "d t = max (d s) (d u)" by (simp only: dickson_gradingD1[OF assms(1)])
thus ?thesis by simp
qed
lemma dickson_grading_minus:
assumes "dickson_grading d" and "s adds (t::'a::cancel_ab_semigroup_add)"
shows "d (t - s) \<le> d t"
proof -
from assms(2) obtain u where "t = s + u" ..
hence "t - s = u" by simp
from assms(1) have "d t = ord_class.max (d s) (d u)" unfolding \<open>t = s + u\<close> by (rule dickson_gradingD1)
thus ?thesis by (simp add: \<open>t - s = u\<close>)
qed
lemma dickson_grading_lcs:
assumes "dickson_grading d"
shows "d (lcs s t) \<le> max (d s) (d t)"
proof -
from assms have "d (lcs s t) \<le> d (s + t)" by (rule dickson_grading_adds_imp_le, intro lcs_adds_plus)
thus ?thesis by (simp only: dickson_gradingD1[OF assms])
qed
lemma dickson_grading_lcs_minus:
assumes "dickson_grading d"
shows "d (lcs s t - s) \<le> max (d s) (d t)"
proof -
from assms have "d (lcs s t - s) \<le> d (lcs s t)" by (rule dickson_grading_minus, intro adds_lcs)
also from assms have "... \<le> max (d s) (d t)" by (rule dickson_grading_lcs)
finally show ?thesis .
qed
lemma dgrad_set_leI:
assumes "\<And>s. s \<in> S \<Longrightarrow> \<exists>t\<in>T. d s \<le> d t"
shows "dgrad_set_le d S T"
using assms by (auto simp: dgrad_set_le_def)
lemma dgrad_set_leE:
assumes "dgrad_set_le d S T" and "s \<in> S"
obtains t where "t \<in> T" and "d s \<le> d t"
using assms by (auto simp: dgrad_set_le_def)
lemma dgrad_set_exhaust_expl:
assumes "finite F"
shows "F \<subseteq> dgrad_set d (Max (d ` F))"
proof
fix f
assume "f \<in> F"
hence "d f \<in> d ` F" by simp
with _ have "d f \<le> Max (d ` F)"
proof (rule Max_ge)
from assms show "finite (d ` F)" by auto
qed
hence "dgrad_set d (d f) \<subseteq> dgrad_set d (Max (d ` F))" by (auto simp: dgrad_set_def)
moreover have "f \<in> dgrad_set d (d f)" by (simp add: dgrad_set_def)
ultimately show "f \<in> dgrad_set d (Max (d ` F))" ..
qed
lemma dgrad_set_exhaust:
assumes "finite F"
obtains m where "F \<subseteq> dgrad_set d m"
proof
from assms show "F \<subseteq> dgrad_set d (Max (d ` F))" by (rule dgrad_set_exhaust_expl)
qed
lemma dgrad_set_le_trans [trans]:
assumes "dgrad_set_le d S T" and "dgrad_set_le d T U"
shows "dgrad_set_le d S U"
unfolding dgrad_set_le_def
proof
fix s
assume "s \<in> S"
with assms(1) obtain t where "t \<in> T" and 1: "d s \<le> d t" by (auto simp add: dgrad_set_le_def)
from assms(2) this(1) obtain u where "u \<in> U" and 2: "d t \<le> d u" by (auto simp add: dgrad_set_le_def)
from this(1) show "\<exists>u\<in>U. d s \<le> d u"
proof
from 1 2 show "d s \<le> d u" by (rule le_trans)
qed
qed
lemma dgrad_set_le_Un: "dgrad_set_le d (S \<union> T) U \<longleftrightarrow> (dgrad_set_le d S U \<and> dgrad_set_le d T U)"
by (auto simp add: dgrad_set_le_def)
lemma dgrad_set_le_subset:
assumes "S \<subseteq> T"
shows "dgrad_set_le d S T"
unfolding dgrad_set_le_def using assms by blast
lemma dgrad_set_le_refl: "dgrad_set_le d S S"
by (rule dgrad_set_le_subset, fact subset_refl)
lemma dgrad_set_le_dgrad_set:
assumes "dgrad_set_le d F G" and "G \<subseteq> dgrad_set d m"
shows "F \<subseteq> dgrad_set d m"
proof
fix f
assume "f \<in> F"
with assms(1) obtain g where "g \<in> G" and *: "d f \<le> d g" by (auto simp add: dgrad_set_le_def)
from assms(2) this(1) have "g \<in> dgrad_set d m" ..
hence "d g \<le> m" by (simp add: dgrad_set_def)
with * have "d f \<le> m" by (rule le_trans)
thus "f \<in> dgrad_set d m" by (simp add: dgrad_set_def)
qed
lemma dgrad_set_dgrad: "p \<in> dgrad_set d (d p)"
by (simp add: dgrad_set_def)
lemma dgrad_setI [intro]:
assumes "d t \<le> m"
shows "t \<in> dgrad_set d m"
using assms by (auto simp: dgrad_set_def)
lemma dgrad_setD:
assumes "t \<in> dgrad_set d m"
shows "d t \<le> m"
using assms by (simp add: dgrad_set_def)
lemma dgrad_set_zero [simp]: "dgrad_set (\<lambda>_. 0) m = UNIV"
by auto
lemma subset_dgrad_set_zero: "F \<subseteq> dgrad_set (\<lambda>_. 0) m"
by simp
lemma dgrad_set_subset:
assumes "m \<le> n"
shows "dgrad_set d m \<subseteq> dgrad_set d n"
using assms by (auto simp: dgrad_set_def)
lemma dgrad_set_closed_plus:
assumes "dickson_grading d" and "s \<in> dgrad_set d m" and "t \<in> dgrad_set d m"
shows "s + t \<in> dgrad_set d m"
proof -
from assms(1) have "d (s + t) = ord_class.max (d s) (d t)" by (rule dickson_gradingD1)
also from assms(2, 3) have "... \<le> m" by (simp add: dgrad_set_def)
finally show ?thesis by (simp add: dgrad_set_def)
qed
lemma dgrad_set_closed_minus:
assumes "dickson_grading d" and "s \<in> dgrad_set d m" and "t adds (s::'a::cancel_ab_semigroup_add)"
shows "s - t \<in> dgrad_set d m"
proof -
from assms(1, 3) have "d (s - t) \<le> d s" by (rule dickson_grading_minus)
also from assms(2) have "... \<le> m" by (simp add: dgrad_set_def)
finally show ?thesis by (simp add: dgrad_set_def)
qed
lemma dgrad_set_closed_lcs:
assumes "dickson_grading d" and "s \<in> dgrad_set d m" and "t \<in> dgrad_set d m"
shows "lcs s t \<in> dgrad_set d m"
proof -
from assms(1) have "d (lcs s t) \<le> ord_class.max (d s) (d t)" by (rule dickson_grading_lcs)
also from assms(2, 3) have "... \<le> m" by (simp add: dgrad_set_def)
finally show ?thesis by (simp add: dgrad_set_def)
qed
lemma dickson_gradingD_dgrad_set: "dickson_grading d \<Longrightarrow> almost_full_on (adds) (dgrad_set d m)"
by (auto dest: dickson_gradingD2 simp: dgrad_set_def)
lemma ex_finite_adds:
assumes "dickson_grading d" and "S \<subseteq> dgrad_set d m"
obtains T where "finite T" and "T \<subseteq> S" and "\<And>s. s \<in> S \<Longrightarrow> (\<exists>t\<in>T. t adds (s::'a::cancel_comm_monoid_add))"
proof -
have "reflp ((adds)::'a \<Rightarrow> _)" by (simp add: reflp_def)
moreover from assms(2) have "almost_full_on (adds) S"
proof (rule almost_full_on_subset)
from assms(1) show "almost_full_on (adds) (dgrad_set d m)" by (rule dickson_gradingD_dgrad_set)
qed
ultimately obtain T where "finite T" and "T \<subseteq> S" and "\<And>s. s \<in> S \<Longrightarrow> (\<exists>t\<in>T. t adds s)"
by (rule almost_full_on_finite_subsetE, blast)
thus ?thesis ..
qed
class graded_dickson_powerprod = ulcs_powerprod +
assumes ex_dgrad: "\<exists>d::'a \<Rightarrow> nat. dickson_grading d"
begin
definition dgrad_dummy where "dgrad_dummy = (SOME d. dickson_grading d)"
lemma dickson_grading_dgrad_dummy: "dickson_grading dgrad_dummy"
unfolding dgrad_dummy_def using ex_dgrad by (rule someI_ex)
end (* graded_dickson_powerprod *)
class dickson_powerprod = ulcs_powerprod +
assumes dickson: "almost_full_on (adds) UNIV"
begin
lemma dickson_grading_zero: "dickson_grading (\<lambda>_::'a. 0)"
by (simp add: dickson_grading_def dickson)
subclass graded_dickson_powerprod by (standard, rule, fact dickson_grading_zero)
end (* dickson_powerprod *)
text \<open>Class @{class graded_dickson_powerprod} is a slightly artificial construction. It is needed,
because type @{typ "nat \<Rightarrow>\<^sub>0 nat"} does not satisfy the usual conditions of a "Dickson domain" (as
formulated in class @{class dickson_powerprod}), but we still want to use that type as the type of
power-products in the computation of Gr\"obner bases. So, we exploit the fact that in a finite
set of polynomials (which is the input of Buchberger's algorithm) there is always some "highest"
indeterminate that occurs with non-zero exponent, and no "higher" indeterminates are generated
during the execution of the algorithm. This allows us to prove that the algorithm terminates, even
though there are in principle infinitely many indeterminates.\<close>
subsection \<open>Additive Linear Orderings\<close>
lemma group_eq_aux: "a + (b - a) = (b::'a::ab_group_add)"
proof -
have "a + (b - a) = b - a + a" by simp
also have "... = b" by simp
finally show ?thesis .
qed
class semi_canonically_ordered_monoid_add = ordered_comm_monoid_add +
assumes le_imp_add: "a \<le> b \<Longrightarrow> (\<exists>c. b = a + c)"
context canonically_ordered_monoid_add
begin
subclass semi_canonically_ordered_monoid_add
by (standard, simp only: le_iff_add)
end
class add_linorder_group = ordered_ab_semigroup_add_imp_le + ab_group_add + linorder
class add_linorder = ordered_ab_semigroup_add_imp_le + cancel_comm_monoid_add + semi_canonically_ordered_monoid_add + linorder
begin
subclass ordered_comm_monoid_add ..
subclass ordered_cancel_comm_monoid_add ..
lemma le_imp_inv:
assumes "a \<le> b"
shows "b = a + (b - a)"
using le_imp_add[OF assms] by auto
lemma max_eq_sum:
obtains y where "max a b = a + y"
unfolding max_def
proof (cases "a \<le> b")
case True
hence "b = a + (b - a)" by (rule le_imp_inv)
then obtain c where eq: "b = a + c" ..
show ?thesis
proof
from True show "max a b = a + c" unfolding max_def eq by simp
qed
next
case False
show ?thesis
proof
from False show "max a b = a + 0" unfolding max_def by simp
qed
qed
lemma min_plus_max:
shows "(min a b) + (max a b) = a + b"
proof (cases "a \<le> b")
case True
thus ?thesis unfolding min_def max_def by simp
next
case False
thus ?thesis unfolding min_def max_def by (simp add: ac_simps)
qed
end (* add_linorder *)
class add_linorder_min = add_linorder +
assumes zero_min: "0 \<le> x"
begin
subclass ninv_comm_monoid_add
proof
fix x y
assume *: "x + y = 0"
show "x = 0"
proof -
from zero_min[of x] have "0 = x \<or> x > 0" by auto
thus ?thesis
proof
assume "x > 0"
have "0 \<le> y" by (fact zero_min)
also have "... = 0 + y" by simp
also from \<open>x > 0\<close> have "... < x + y" by (rule add_strict_right_mono)
finally have "0 < x + y" .
hence "x + y \<noteq> 0" by simp
from this * show ?thesis ..
qed simp
qed
qed
lemma leq_add_right:
shows "x \<le> x + y"
using add_left_mono[OF zero_min[of y], of x] by simp
lemma leq_add_left:
shows "x \<le> y + x"
using add_right_mono[OF zero_min[of y], of x] by simp
subclass canonically_ordered_monoid_add
by (standard, rule, elim le_imp_add, elim exE, simp add: leq_add_right)
end (* add_linorder_min *)
class add_wellorder = add_linorder_min + wellorder
instantiation nat :: add_linorder
begin
instance by (standard, simp)
end (* add_linorder *)
instantiation nat :: add_linorder_min
begin
instance by (standard, simp)
end
instantiation nat :: add_wellorder
begin
instance ..
end
context add_linorder_group
begin
subclass add_linorder
proof (standard, intro exI)
fix a b
show "b = a + (b - a)" using add_commute local.diff_add_cancel by auto
qed
end (* add_linorder_group *)
instantiation int :: add_linorder_group
begin
instance ..
end
instantiation rat :: add_linorder_group
begin
instance ..
end
instantiation real :: add_linorder_group
begin
instance ..
end
subsection \<open>Ordered Power-Products\<close>
locale ordered_powerprod =
ordered_powerprod_lin: linorder ord ord_strict
for ord::"'a \<Rightarrow> 'a::comm_powerprod \<Rightarrow> bool" (infixl "\<preceq>" 50)
and ord_strict::"'a \<Rightarrow> 'a::comm_powerprod \<Rightarrow> bool" (infixl "\<prec>" 50) +
assumes zero_min: "0 \<preceq> t"
assumes plus_monotone: "s \<preceq> t \<Longrightarrow> s + u \<preceq> t + u"
begin
abbreviation ord_conv (infixl "\<succeq>" 50) where "ord_conv \<equiv> (\<preceq>)\<inverse>\<inverse>"
abbreviation ord_strict_conv (infixl "\<succ>" 50) where "ord_strict_conv \<equiv> (\<prec>)\<inverse>\<inverse>"
lemma ord_canc:
assumes "s + u \<preceq> t + u"
shows "s \<preceq> t"
proof (rule ordered_powerprod_lin.le_cases[of s t], simp)
assume "t \<preceq> s"
from assms plus_monotone[OF this, of u] have "t + u = s + u"
- using ordered_powerprod_lin.eq_iff by simp
+ using ordered_powerprod_lin.order.eq_iff by simp
hence "t = s" by simp
thus "s \<preceq> t" by simp
qed
lemma ord_adds:
assumes "s adds t"
shows "s \<preceq> t"
proof -
from assms have "\<exists>u. t = s + u" unfolding adds_def by simp
then obtain k where "t = s + k" ..
thus ?thesis using plus_monotone[OF zero_min[of k], of s] by (simp add: ac_simps)
qed
lemma ord_canc_left:
assumes "u + s \<preceq> u + t"
shows "s \<preceq> t"
using assms unfolding add.commute[of u] by (rule ord_canc)
lemma ord_strict_canc:
assumes "s + u \<prec> t + u"
shows "s \<prec> t"
using assms ord_canc[of s u t] add_right_cancel[of s u t]
ordered_powerprod_lin.le_imp_less_or_eq ordered_powerprod_lin.order.strict_implies_order by blast
lemma ord_strict_canc_left:
assumes "u + s \<prec> u + t"
shows "s \<prec> t"
using assms unfolding add.commute[of u] by (rule ord_strict_canc)
lemma plus_monotone_left:
assumes "s \<preceq> t"
shows "u + s \<preceq> u + t"
using assms
by (simp add: add.commute, rule plus_monotone)
lemma plus_monotone_strict:
assumes "s \<prec> t"
shows "s + u \<prec> t + u"
using assms
by (simp add: ordered_powerprod_lin.order.strict_iff_order plus_monotone)
lemma plus_monotone_strict_left:
assumes "s \<prec> t"
shows "u + s \<prec> u + t"
using assms
by (simp add: ordered_powerprod_lin.order.strict_iff_order plus_monotone_left)
end
locale gd_powerprod =
ordered_powerprod ord ord_strict
for ord::"'a \<Rightarrow> 'a::graded_dickson_powerprod \<Rightarrow> bool" (infixl "\<preceq>" 50)
and ord_strict (infixl "\<prec>" 50)
begin
definition dickson_le :: "('a \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
where "dickson_le d m s t \<longleftrightarrow> (d s \<le> m \<and> d t \<le> m \<and> s \<preceq> t)"
definition dickson_less :: "('a \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
where "dickson_less d m s t \<longleftrightarrow> (d s \<le> m \<and> d t \<le> m \<and> s \<prec> t)"
lemma dickson_leI:
assumes "d s \<le> m" and "d t \<le> m" and "s \<preceq> t"
shows "dickson_le d m s t"
using assms by (simp add: dickson_le_def)
lemma dickson_leD1:
assumes "dickson_le d m s t"
shows "d s \<le> m"
using assms by (simp add: dickson_le_def)
lemma dickson_leD2:
assumes "dickson_le d m s t"
shows "d t \<le> m"
using assms by (simp add: dickson_le_def)
lemma dickson_leD3:
assumes "dickson_le d m s t"
shows "s \<preceq> t"
using assms by (simp add: dickson_le_def)
lemma dickson_le_trans:
assumes "dickson_le d m s t" and "dickson_le d m t u"
shows "dickson_le d m s u"
using assms by (auto simp add: dickson_le_def)
lemma dickson_lessI:
assumes "d s \<le> m" and "d t \<le> m" and "s \<prec> t"
shows "dickson_less d m s t"
using assms by (simp add: dickson_less_def)
lemma dickson_lessD1:
assumes "dickson_less d m s t"
shows "d s \<le> m"
using assms by (simp add: dickson_less_def)
lemma dickson_lessD2:
assumes "dickson_less d m s t"
shows "d t \<le> m"
using assms by (simp add: dickson_less_def)
lemma dickson_lessD3:
assumes "dickson_less d m s t"
shows "s \<prec> t"
using assms by (simp add: dickson_less_def)
lemma dickson_less_irrefl: "\<not> dickson_less d m t t"
by (simp add: dickson_less_def)
lemma dickson_less_trans:
assumes "dickson_less d m s t" and "dickson_less d m t u"
shows "dickson_less d m s u"
using assms by (auto simp add: dickson_less_def)
lemma transp_dickson_less: "transp (dickson_less d m)"
by (rule transpI, fact dickson_less_trans)
lemma wfp_on_ord_strict:
assumes "dickson_grading d"
shows "wfp_on (\<prec>) {x. d x \<le> n}"
proof -
let ?A = "{x. d x \<le> n}"
have "strict (\<preceq>) = (\<prec>)" by (intro ext, simp only: ordered_powerprod_lin.less_le_not_le)
have "qo_on (adds) ?A" by (auto simp: qo_on_def reflp_on_def transp_on_def dest: adds_trans)
moreover from assms have "wqo_on (adds) ?A" by (rule dickson_gradingD2')
ultimately have "(\<forall>Q. (\<forall>x\<in>?A. \<forall>y\<in>?A. x adds y \<longrightarrow> Q x y) \<and> qo_on Q ?A \<longrightarrow> wfp_on (strict Q) ?A)"
by (simp only: wqo_extensions_wf_conv)
hence "(\<forall>x\<in>?A. \<forall>y\<in>?A. x adds y \<longrightarrow> x \<preceq> y) \<and> qo_on (\<preceq>) ?A \<longrightarrow> wfp_on (strict (\<preceq>)) ?A" ..
thus ?thesis unfolding \<open>strict (\<preceq>) = (\<prec>)\<close>
proof
show "(\<forall>x\<in>?A. \<forall>y\<in>?A. x adds y \<longrightarrow> x \<preceq> y) \<and> qo_on (\<preceq>) ?A"
proof (intro conjI ballI impI ord_adds)
show "qo_on (\<preceq>) ?A" by (auto simp: qo_on_def reflp_on_def transp_on_def)
qed
qed
qed
lemma wf_dickson_less:
assumes "dickson_grading d"
shows "wfP (dickson_less d m)"
proof (rule wfP_chain)
show "\<not> (\<exists>seq. \<forall>i. dickson_less d m (seq (Suc i)) (seq i))"
proof
assume "\<exists>seq. \<forall>i. dickson_less d m (seq (Suc i)) (seq i)"
then obtain seq::"nat \<Rightarrow> 'a" where "\<forall>i. dickson_less d m (seq (Suc i)) (seq i)" ..
hence *: "\<And>i. dickson_less d m (seq (Suc i)) (seq i)" ..
with transp_dickson_less have seq_decr: "\<And>i j. i < j \<Longrightarrow> dickson_less d m (seq j) (seq i)"
by (rule transp_sequence)
from assms obtain i j where "i < j" and i_adds_j: "seq i adds seq j"
proof (rule dickson_gradingE)
fix i
from * show "d (seq i) \<le> m" by (rule dickson_lessD2)
qed
from \<open>i < j\<close> have "dickson_less d m (seq j) (seq i)" by (rule seq_decr)
hence "seq j \<prec> seq i" by (rule dickson_lessD3)
moreover from i_adds_j have "seq i \<preceq> seq j" by (rule ord_adds)
ultimately show False by simp
qed
qed
end
text \<open>\<open>gd_powerprod\<close> stands for @{emph \<open>graded ordered Dickson power-products\<close>}.\<close>
locale od_powerprod =
ordered_powerprod ord ord_strict
for ord::"'a \<Rightarrow> 'a::dickson_powerprod \<Rightarrow> bool" (infixl "\<preceq>" 50)
and ord_strict (infixl "\<prec>" 50)
begin
sublocale gd_powerprod by standard
lemma wf_ord_strict: "wfP (\<prec>)"
proof (rule wfP_chain)
show "\<not> (\<exists>seq. \<forall>i. seq (Suc i) \<prec> seq i)"
proof
assume "\<exists>seq. \<forall>i. seq (Suc i) \<prec> seq i"
then obtain seq::"nat \<Rightarrow> 'a" where "\<forall>i. seq (Suc i) \<prec> seq i" ..
hence "\<And>i. seq (Suc i) \<prec> seq i" ..
with ordered_powerprod_lin.transp_less have seq_decr: "\<And>i j. i < j \<Longrightarrow> (seq j) \<prec> (seq i)"
by (rule transp_sequence)
from dickson obtain i j::nat where "i < j" and i_adds_j: "seq i adds seq j"
by (auto elim!: almost_full_onD)
from seq_decr[OF \<open>i < j\<close>] have "seq j \<preceq> seq i \<and> seq j \<noteq> seq i" by auto
hence "seq j \<preceq> seq i" and "seq j \<noteq> seq i" by simp_all
from \<open>seq j \<noteq> seq i\<close> \<open>seq j \<preceq> seq i\<close> ord_adds[OF i_adds_j]
- ordered_powerprod_lin.eq_iff[of "seq j" "seq i"]
+ ordered_powerprod_lin.order.eq_iff[of "seq j" "seq i"]
show False by simp
qed
qed
end
text \<open>\<open>od_powerprod\<close> stands for @{emph \<open>ordered Dickson power-products\<close>}.\<close>
subsection \<open>Functions as Power-Products\<close>
lemma finite_neq_0:
assumes fin_A: "finite {x. f x \<noteq> 0}" and fin_B: "finite {x. g x \<noteq> 0}" and "\<And>x. h x 0 0 = 0"
shows "finite {x. h x (f x) (g x) \<noteq> 0}"
proof -
from fin_A fin_B have "finite ({x. f x \<noteq> 0} \<union> {x. g x \<noteq> 0})" by (intro finite_UnI)
hence finite_union: "finite {x. (f x \<noteq> 0) \<or> (g x \<noteq> 0)}" by (simp only: Collect_disj_eq)
have "{x. h x (f x) (g x) \<noteq> 0} \<subseteq> {x. (f x \<noteq> 0) \<or> (g x \<noteq> 0)}"
proof (intro Collect_mono, rule)
fix x::'a
assume h_not_zero: "h x (f x) (g x) \<noteq> 0"
have "f x = 0 \<Longrightarrow> g x \<noteq> 0"
proof
assume "f x = 0" "g x = 0"
thus False using h_not_zero \<open>h x 0 0 = 0\<close> by simp
qed
thus "f x \<noteq> 0 \<or> g x \<noteq> 0" by auto
qed
from finite_subset[OF this] finite_union show "finite {x. h x (f x) (g x) \<noteq> 0}" .
qed
lemma finite_neq_0':
assumes "finite {x. f x \<noteq> 0}" and "finite {x. g x \<noteq> 0}" and "h 0 0 = 0"
shows "finite {x. h (f x) (g x) \<noteq> 0}"
using assms by (rule finite_neq_0)
lemma finite_neq_0_inv:
assumes fin_A: "finite {x. h x (f x) (g x) \<noteq> 0}" and fin_B: "finite {x. f x \<noteq> 0}" and "\<And>x y. h x 0 y = y"
shows "finite {x. g x \<noteq> 0}"
proof -
from fin_A and fin_B have "finite ({x. h x (f x) (g x) \<noteq> 0} \<union> {x. f x \<noteq> 0})" by (intro finite_UnI)
hence finite_union: "finite {x. (h x (f x) (g x) \<noteq> 0) \<or> f x \<noteq> 0}" by (simp only: Collect_disj_eq)
have "{x. g x \<noteq> 0} \<subseteq> {x. (h x (f x) (g x) \<noteq> 0) \<or> f x \<noteq> 0}"
by (intro Collect_mono, rule, rule disjCI, simp add: assms(3))
from this finite_union show "finite {x. g x \<noteq> 0}" by (rule finite_subset)
qed
lemma finite_neq_0_inv':
assumes inf_A: "finite {x. h (f x) (g x) \<noteq> 0}" and fin_B: "finite {x. f x \<noteq> 0}" and "\<And>x. h 0 x = x"
shows "finite {x. g x \<noteq> 0}"
using assms by (rule finite_neq_0_inv)
subsubsection \<open>@{typ "'a \<Rightarrow> 'b"} belongs to class @{class comm_powerprod}\<close>
instance "fun" :: (type, cancel_comm_monoid_add) comm_powerprod
by standard
subsubsection \<open>@{typ "'a \<Rightarrow> 'b"} belongs to class @{class ninv_comm_monoid_add}\<close>
instance "fun" :: (type, ninv_comm_monoid_add) ninv_comm_monoid_add
by (standard, simp only: plus_fun_def zero_fun_def fun_eq_iff, intro allI, rule plus_eq_zero, auto)
subsubsection \<open>@{typ "'a \<Rightarrow> 'b"} belongs to class @{class lcs_powerprod}\<close>
instantiation "fun" :: (type, add_linorder) lcs_powerprod
begin
definition lcs_fun::"('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)" where "lcs f g = (\<lambda>x. max (f x) (g x))"
lemma adds_funI:
assumes "s \<le> t"
shows "s adds (t::'a \<Rightarrow> 'b)"
proof (rule addsI, rule)
fix x
from assms have "s x \<le> t x" unfolding le_fun_def ..
hence "t x = s x + (t x - s x)" by (rule le_imp_inv)
thus "t x = (s + (t - s)) x" by simp
qed
lemma adds_fun_iff: "f adds (g::'a \<Rightarrow> 'b) \<longleftrightarrow> (\<forall>x. f x adds g x)"
unfolding adds_def plus_fun_def by metis
lemma adds_fun_iff': "f adds (g::'a \<Rightarrow> 'b) \<longleftrightarrow> (\<forall>x. \<exists>y. g x = f x + y)"
unfolding adds_fun_iff unfolding adds_def plus_fun_def ..
lemma adds_lcs_fun:
shows "s adds (lcs s (t::'a \<Rightarrow> 'b))"
by (rule adds_funI, simp only: le_fun_def lcs_fun_def, auto simp: max_def)
lemma lcs_comm_fun: "lcs s t = lcs t (s::'a \<Rightarrow> 'b)"
unfolding lcs_fun_def
by (auto simp: max_def intro!: ext)
lemma lcs_adds_fun:
assumes "s adds u" and "t adds (u::'a \<Rightarrow> 'b)"
shows "(lcs s t) adds u"
using assms unfolding lcs_fun_def adds_fun_iff'
proof -
assume a1: "\<forall>x. \<exists>y. u x = s x + y" and a2: "\<forall>x. \<exists>y. u x = t x + y"
show "\<forall>x. \<exists>y. u x = max (s x) (t x) + y"
proof
fix x
from a1 have b1: "\<exists>y. u x = s x + y" ..
from a2 have b2: "\<exists>y. u x = t x + y" ..
show "\<exists>y. u x = max (s x) (t x) + y" unfolding max_def
by (split if_split, intro conjI impI, rule b2, rule b1)
qed
qed
instance
apply standard
subgoal by (rule adds_lcs_fun)
subgoal by (rule lcs_adds_fun)
subgoal by (rule lcs_comm_fun)
done
end
lemma leq_lcs_fun_1: "s \<le> (lcs s (t::'a \<Rightarrow> 'b::add_linorder))"
by (simp add: lcs_fun_def le_fun_def)
lemma leq_lcs_fun_2: "t \<le> (lcs s (t::'a \<Rightarrow> 'b::add_linorder))"
by (simp add: lcs_fun_def le_fun_def)
lemma lcs_leq_fun:
assumes "s \<le> u" and "t \<le> (u::'a \<Rightarrow> 'b::add_linorder)"
shows "(lcs s t) \<le> u"
using assms by (simp add: lcs_fun_def le_fun_def)
lemma adds_fun: "s adds t \<longleftrightarrow> s \<le> t"
for s t::"'a \<Rightarrow> 'b::add_linorder_min"
proof
assume "s adds t"
from this obtain k where "t = s + k" ..
show "s \<le> t" unfolding \<open>t = s + k\<close> le_fun_def plus_fun_def le_iff_add by (simp add: leq_add_right)
qed (rule adds_funI)
lemma gcs_fun: "gcs s (t::'a \<Rightarrow> ('b::add_linorder)) = (\<lambda>x. min (s x) (t x))"
proof -
show ?thesis unfolding gcs_def lcs_fun_def fun_diff_def
proof (simp, rule)
fix x
have eq: "s x + t x = max (s x) (t x) + min (s x) (t x)" by (metis add.commute min_def max_def)
thus "s x + t x - max (s x) (t x) = min (s x) (t x)" by simp
qed
qed
lemma gcs_leq_fun_1: "(gcs s (t::'a \<Rightarrow> 'b::add_linorder)) \<le> s"
by (simp add: gcs_fun le_fun_def)
lemma gcs_leq_fun_2: "(gcs s (t::'a \<Rightarrow> 'b::add_linorder)) \<le> t"
by (simp add: gcs_fun le_fun_def)
lemma leq_gcs_fun:
assumes "u \<le> s" and "u \<le> (t::'a \<Rightarrow> 'b::add_linorder)"
shows "u \<le> (gcs s t)"
using assms by (simp add: gcs_fun le_fun_def)
subsubsection \<open>@{typ "'a \<Rightarrow> 'b"} belongs to class @{class ulcs_powerprod}\<close>
instance "fun" :: (type, add_linorder_min) ulcs_powerprod ..
subsubsection \<open>Power-products in a given set of indeterminates\<close>
definition supp_fun::"('a \<Rightarrow> 'b::zero) \<Rightarrow> 'a set" where "supp_fun f = {x. f x \<noteq> 0}"
text \<open>@{term supp_fun} for general functions is like @{term keys} for @{type poly_mapping},
but does not need to be finite.\<close>
lemma keys_eq_supp: "keys s = supp_fun (lookup s)"
unfolding supp_fun_def by (transfer, rule)
lemma supp_fun_zero [simp]: "supp_fun 0 = {}"
by (auto simp: supp_fun_def)
lemma supp_fun_eq_zero_iff: "supp_fun f = {} \<longleftrightarrow> f = 0"
by (auto simp: supp_fun_def)
lemma sub_supp_empty: "supp_fun s \<subseteq> {} \<longleftrightarrow> (s = 0)"
by (auto simp: supp_fun_def)
lemma except_fun_idI: "supp_fun f \<inter> V = {} \<Longrightarrow> except_fun f V = f"
by (auto simp: except_fun_def supp_fun_def when_def intro!: ext)
lemma supp_except_fun: "supp_fun (except_fun s V) = supp_fun s - V"
by (auto simp: except_fun_def supp_fun_def)
lemma supp_fun_plus_subset: "supp_fun (s + t) \<subseteq> supp_fun s \<union> supp_fun (t::'a \<Rightarrow> 'b::monoid_add)"
unfolding supp_fun_def by force
lemma fun_eq_zeroI:
assumes "\<And>x. x \<in> supp_fun f \<Longrightarrow> f x = 0"
shows "f = 0"
proof (rule, simp)
fix x
show "f x = 0"
proof (cases "x \<in> supp_fun f")
case True
then show ?thesis by (rule assms)
next
case False
then show ?thesis by (simp add: supp_fun_def)
qed
qed
lemma except_fun_cong1:
"supp_fun s \<inter> ((V - U) \<union> (U - V)) \<subseteq> {} \<Longrightarrow> except_fun s V = except_fun s U"
by (auto simp: except_fun_def when_def supp_fun_def intro!: ext)
lemma adds_except_fun:
"s adds t = (except_fun s V adds except_fun t V \<and> except_fun s (- V) adds except_fun t (- V))"
for s t :: "'a \<Rightarrow> 'b::add_linorder"
by (auto simp: supp_fun_def except_fun_def adds_fun_iff when_def)
lemma adds_except_fun_singleton: "s adds t = (except_fun s {v} adds except_fun t {v} \<and> s v adds t v)"
for s t :: "'a \<Rightarrow> 'b::add_linorder"
by (auto simp: supp_fun_def except_fun_def adds_fun_iff when_def)
subsubsection \<open>Dickson's lemma for power-products in finitely many indeterminates\<close>
lemma Dickson_fun:
assumes "finite V"
shows "almost_full_on (adds) {x::'a \<Rightarrow> 'b::add_wellorder. supp_fun x \<subseteq> V}"
using assms
proof (induct V)
case empty
have "finite {0}" by simp
moreover have "reflp_on (adds) {0::'a \<Rightarrow> 'b}" by (simp add: reflp_on_def)
ultimately have "almost_full_on (adds) {0::'a \<Rightarrow> 'b}" by (rule finite_almost_full_on)
thus ?case by (simp add: supp_fun_eq_zero_iff)
next
case (insert v V)
show ?case
proof (rule almost_full_onI)
fix seq::"nat \<Rightarrow> 'a \<Rightarrow> 'b"
assume "\<forall>i. seq i \<in> {x. supp_fun x \<subseteq> insert v V}"
hence a: "supp_fun (seq i) \<subseteq> insert v V" for i by simp
define seq' where "seq' = (\<lambda>i. (except_fun (seq i) {v}, except_fun (seq i) V))"
have "almost_full_on (adds) {x::'a \<Rightarrow> 'b. supp_fun x \<subseteq> {v}}"
proof (rule almost_full_onI)
fix f::"nat \<Rightarrow> 'a \<Rightarrow> 'b"
assume "\<forall>i. f i \<in> {x. supp_fun x \<subseteq> {v}}"
hence b: "supp_fun (f i) \<subseteq> {v}" for i by simp
let ?f = "\<lambda>i. f i v"
have "wfP ((<)::'b \<Rightarrow> _)" by (simp add: wf wfP_def)
hence "\<forall>f::nat \<Rightarrow> 'b. \<exists>i. f i \<le> f (Suc i)"
by (simp add: wf_iff_no_infinite_down_chain[to_pred] not_less)
hence "\<exists>i. ?f i \<le> ?f (Suc i)" ..
then obtain i where "?f i \<le> ?f (Suc i)" ..
have "i < Suc i" by simp
moreover have "f i adds f (Suc i)" unfolding adds_fun_iff
proof
fix x
show "f i x adds f (Suc i) x"
proof (cases "x = v")
case True
with \<open>?f i \<le> ?f (Suc i)\<close> show ?thesis by (simp add: adds_def le_iff_add)
next
case False
with b have "x \<notin> supp_fun (f i)" and "x \<notin> supp_fun (f (Suc i))" by blast+
thus ?thesis by (simp add: supp_fun_def)
qed
qed
ultimately show "good (adds) f" by (meson goodI)
qed
with insert(3) have
"almost_full_on (prod_le (adds) (adds)) ({x::'a \<Rightarrow> 'b. supp_fun x \<subseteq> V} \<times> {x::'a \<Rightarrow> 'b. supp_fun x \<subseteq> {v}})"
(is "almost_full_on ?P ?A") by (rule almost_full_on_Sigma)
moreover from a have "seq' i \<in> ?A" for i by (auto simp add: seq'_def supp_except_fun)
ultimately obtain i j where "i < j" and "?P (seq' i) (seq' j)" by (rule almost_full_onD)
have "seq i adds seq j" unfolding adds_except_fun[where s="seq i" and V=V]
proof
from \<open>?P (seq' i) (seq' j)\<close> show "except_fun (seq i) V adds except_fun (seq j) V"
by (simp add: prod_le_def seq'_def)
next
from \<open>?P (seq' i) (seq' j)\<close> have "except_fun (seq i) {v} adds except_fun (seq j) {v}"
by (simp add: prod_le_def seq'_def)
moreover have "except_fun (seq i) (- V) = except_fun (seq i) {v}"
by (rule except_fun_cong1; use a[of i] insert.hyps(2) in blast)
moreover have "except_fun (seq j) (- V) = except_fun (seq j) {v}"
by (rule except_fun_cong1; use a[of j] insert.hyps(2) in blast)
ultimately show "except_fun (seq i) (- V) adds except_fun (seq j) (- V)" by simp
qed
with \<open>i < j\<close> show "good (adds) seq" by (meson goodI)
qed
qed
instance "fun" :: (finite, add_wellorder) dickson_powerprod
proof
have "finite (UNIV::'a set)" by simp
hence "almost_full_on (adds) {x::'a \<Rightarrow> 'b. supp_fun x \<subseteq> UNIV}" by (rule Dickson_fun)
thus "almost_full_on (adds) (UNIV::('a \<Rightarrow> 'b) set)" by simp
qed
subsubsection \<open>Lexicographic Term Order\<close>
text \<open>Term orders are certain linear orders on power-products, satisfying additional requirements.
Further information on term orders can be found, e.\,g., in @{cite Robbiano1985}.\<close>
context wellorder
begin
lemma neq_fun_alt:
assumes "s \<noteq> (t::'a \<Rightarrow> 'b)"
obtains x where "s x \<noteq> t x" and "\<And>y. s y \<noteq> t y \<Longrightarrow> x \<le> y"
proof -
from assms ext[of s t] have "\<exists>x. s x \<noteq> t x" by auto
with exists_least_iff[of "\<lambda>x. s x \<noteq> t x"]
obtain x where x1: "s x \<noteq> t x" and x2: "\<And>y. y < x \<Longrightarrow> s y = t y"
by auto
show ?thesis
proof
from x1 show "s x \<noteq> t x" .
next
fix y
assume "s y \<noteq> t y"
with x2[of y] have "\<not> y < x" by auto
thus "x \<le> y" by simp
qed
qed
definition lex_fun::"('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> bool" where
"lex_fun s t \<equiv> (\<forall>x. s x \<le> t x \<or> (\<exists>y<x. s y \<noteq> t y))"
definition "lex_fun_strict s t \<longleftrightarrow> lex_fun s t \<and> \<not> lex_fun t s"
text \<open>Attention! @{term lex_fun} reverses the order of the indeterminates: if @{term x} is smaller than
@{term y} w.r.t. the order on @{typ 'a}, then the @{emph \<open>power-product\<close>} @{term x} is
@{emph \<open>greater\<close>} than the @{emph \<open>power-product\<close>} @{term y}.\<close>
lemma lex_fun_alt:
shows "lex_fun s t = (s = t \<or> (\<exists>x. s x < t x \<and> (\<forall>y<x. s y = t y)))" (is "?L = ?R")
proof
assume ?L
show ?R
proof (cases "s = t")
assume "s = t"
thus ?R by simp
next
assume "s \<noteq> t"
from neq_fun_alt[OF this] obtain x0
where x0_neq: "s x0 \<noteq> t x0" and x0_min: "\<And>z. s z \<noteq> t z \<Longrightarrow> x0 \<le> z" by auto
show ?R
proof (intro disjI2, rule exI[of _ x0], intro conjI)
from \<open>?L\<close> have "s x0 \<le> t x0 \<or> (\<exists>y. y < x0 \<and> s y \<noteq> t y)" unfolding lex_fun_def ..
thus "s x0 < t x0"
proof
assume "s x0 \<le> t x0"
from this x0_neq show ?thesis by simp
next
assume "\<exists>y. y < x0 \<and> s y \<noteq> t y"
then obtain y where "y < x0" and y_neq: "s y \<noteq> t y" by auto
from \<open>y < x0\<close> x0_min[OF y_neq] show ?thesis by simp
qed
next
show "\<forall>y<x0. s y = t y"
proof (rule, rule)
fix y
assume "y < x0"
hence "\<not> x0 \<le> y" by simp
from this x0_min[of y] show "s y = t y" by auto
qed
qed
qed
next
assume ?R
thus ?L
proof
assume "s = t"
thus ?thesis by (simp add: lex_fun_def)
next
assume "\<exists>x. s x < t x \<and> (\<forall>y<x. s y = t y)"
then obtain y where y: "s y < t y" and y_min: "\<forall>z<y. s z = t z" by auto
show ?thesis unfolding lex_fun_def
proof
fix x
show "s x \<le> t x \<or> (\<exists>y<x. s y \<noteq> t y)"
proof (cases "s x \<le> t x")
assume "s x \<le> t x"
thus ?thesis by simp
next
assume x: "\<not> s x \<le> t x"
show ?thesis
proof (intro disjI2, rule exI[of _ y], intro conjI)
have "\<not> x \<le> y"
proof
assume "x \<le> y"
hence "x < y \<or> y = x" by auto
thus False
proof
assume "x < y"
from x y_min[rule_format, OF this] show ?thesis by simp
next
assume "y = x"
from this x y show ?thesis
by (auto simp: preorder_class.less_le_not_le)
qed
qed
thus "y < x" by simp
next
from y show "s y \<noteq> t y" by simp
qed
qed
qed
qed
qed
lemma lex_fun_refl: "lex_fun s s"
unfolding lex_fun_alt by simp
lemma lex_fun_antisym:
assumes "lex_fun s t" and "lex_fun t s"
shows "s = t"
proof
fix x
from assms(1) have "s = t \<or> (\<exists>x. s x < t x \<and> (\<forall>y<x. s y = t y))"
unfolding lex_fun_alt .
thus "s x = t x"
proof
assume "s = t"
thus ?thesis by simp
next
assume "\<exists>x. s x < t x \<and> (\<forall>y<x. s y = t y)"
then obtain x0 where x0: "s x0 < t x0" and x0_min: "\<forall>y<x0. s y = t y" by auto
from assms(2) have "t = s \<or> (\<exists>x. t x < s x \<and> (\<forall>y<x. t y = s y))" unfolding lex_fun_alt .
thus ?thesis
proof
assume "t = s"
thus ?thesis by simp
next
assume "\<exists>x. t x < s x \<and> (\<forall>y<x. t y = s y)"
then obtain x1 where x1: "t x1 < s x1" and x1_min: "\<forall>y<x1. t y = s y" by auto
have "x0 < x1 \<or> x1 < x0 \<or> x1 = x0" using local.antisym_conv3 by auto
show ?thesis
proof (rule linorder_cases[of x0 x1])
assume "x1 < x0"
from x0_min[rule_format, OF this] x1 show ?thesis by simp
next
assume "x0 = x1"
from this x0 x1 show ?thesis by simp
next
assume "x0 < x1"
from x1_min[rule_format, OF this] x0 show ?thesis by simp
qed
qed
qed
qed
lemma lex_fun_trans:
assumes "lex_fun s t" and "lex_fun t u"
shows "lex_fun s u"
proof -
from assms(1) have "s = t \<or> (\<exists>x. s x < t x \<and> (\<forall>y<x. s y = t y))" unfolding lex_fun_alt .
thus ?thesis
proof
assume "s = t"
from this assms(2) show ?thesis by simp
next
assume "\<exists>x. s x < t x \<and> (\<forall>y<x. s y = t y)"
then obtain x0 where x0: "s x0 < t x0" and x0_min: "\<forall>y<x0. s y = t y"
by auto
from assms(2) have "t = u \<or> (\<exists>x. t x < u x \<and> (\<forall>y<x. t y = u y))" unfolding lex_fun_alt .
thus ?thesis
proof
assume "t = u"
from this assms(1) show ?thesis by simp
next
assume "\<exists>x. t x < u x \<and> (\<forall>y<x. t y = u y)"
then obtain x1 where x1: "t x1 < u x1" and x1_min: "\<forall>y<x1. t y = u y" by auto
show ?thesis unfolding lex_fun_alt
proof (intro disjI2)
show "\<exists>x. s x < u x \<and> (\<forall>y<x. s y = u y)"
proof (rule linorder_cases[of x0 x1])
assume "x1 < x0"
show ?thesis
proof (rule exI[of _ x1], intro conjI)
from x0_min[rule_format, OF \<open>x1 < x0\<close>] x1 show "s x1 < u x1" by simp
next
show "\<forall>y<x1. s y = u y"
proof (intro allI, intro impI)
fix y
assume "y < x1"
from this \<open>x1 < x0\<close> have "y < x0" by simp
from x0_min[rule_format, OF this] x1_min[rule_format, OF \<open>y < x1\<close>]
show "s y = u y" by simp
qed
qed
next
assume "x0 < x1"
show ?thesis
proof (rule exI[of _ x0], intro conjI)
from x1_min[rule_format, OF \<open>x0 < x1\<close>] x0 show "s x0 < u x0" by simp
next
show "\<forall>y<x0. s y = u y"
proof (intro allI, intro impI)
fix y
assume "y < x0"
from this \<open>x0 < x1\<close> have "y < x1" by simp
from x0_min[rule_format, OF \<open>y < x0\<close>] x1_min[rule_format, OF this]
show "s y = u y" by simp
qed
qed
next
assume "x0 = x1"
show ?thesis
proof (rule exI[of _ x1], intro conjI)
from \<open>x0 = x1\<close> x0 x1 show "s x1 < u x1" by simp
next
show "\<forall>y<x1. s y = u y"
proof (intro allI, intro impI)
fix y
assume "y < x1"
hence "y < x0" using \<open>x0 = x1\<close> by simp
from x0_min[rule_format, OF this] x1_min[rule_format, OF \<open>y < x1\<close>]
show "s y = u y" by simp
qed
qed
qed
qed
qed
qed
qed
lemma lex_fun_lin: "lex_fun s t \<or> lex_fun t s" for s t::"'a \<Rightarrow> 'b::{ordered_comm_monoid_add, linorder}"
proof (intro disjCI)
assume "\<not> lex_fun t s"
hence a: "\<forall>x. \<not> (t x < s x) \<or> (\<exists>y<x. t y \<noteq> s y)" unfolding lex_fun_alt by auto
show "lex_fun s t" unfolding lex_fun_def
proof
fix x
from a have "\<not> (t x < s x) \<or> (\<exists>y<x. t y \<noteq> s y)" ..
thus "s x \<le> t x \<or> (\<exists>y<x. s y \<noteq> t y)" by auto
qed
qed
corollary lex_fun_strict_alt [code]:
"lex_fun_strict s t = (\<not> lex_fun t s)" for s t::"'a \<Rightarrow> 'b::{ordered_comm_monoid_add, linorder}"
unfolding lex_fun_strict_def using lex_fun_lin[of s t] by auto
lemma lex_fun_zero_min: "lex_fun 0 s" for s::"'a \<Rightarrow> 'b::add_linorder_min"
by (simp add: lex_fun_def zero_min)
lemma lex_fun_plus_monotone:
"lex_fun (s + u) (t + u)" if "lex_fun s t"
for s t::"'a \<Rightarrow> 'b::ordered_cancel_comm_monoid_add"
unfolding lex_fun_def
proof
fix x
from that have "s x \<le> t x \<or> (\<exists>y<x. s y \<noteq> t y)" unfolding lex_fun_def ..
thus "(s + u) x \<le> (t + u) x \<or> (\<exists>y<x. (s + u) y \<noteq> (t + u) y)"
proof
assume a1: "s x \<le> t x"
show ?thesis
proof (intro disjI1)
from a1 show "(s + u) x \<le> (t + u) x" by (auto simp: add_right_mono)
qed
next
assume "\<exists>y<x. s y \<noteq> t y"
then obtain y where "y < x" and a2: "s y \<noteq> t y" by auto
show ?thesis
proof (intro disjI2, rule exI[of _ y], intro conjI, fact)
from a2 show "(s + u) y \<noteq> (t + u) y" by (auto simp: add_right_mono)
qed
qed
qed
end (* wellorder *)
subsubsection \<open>Degree\<close>
definition deg_fun::"('a \<Rightarrow> 'b::comm_monoid_add) \<Rightarrow> 'b" where "deg_fun s \<equiv> \<Sum>x\<in>(supp_fun s). s x"
lemma deg_fun_zero[simp]: "deg_fun 0 = 0"
by (auto simp: deg_fun_def)
lemma deg_fun_eq_0_iff:
assumes "finite (supp_fun (s::'a \<Rightarrow> 'b::add_linorder_min))"
shows "deg_fun s = 0 \<longleftrightarrow> s = 0"
proof
assume "deg_fun s = 0"
hence *: "(\<Sum>x\<in>(supp_fun s). s x) = 0" by (simp only: deg_fun_def)
have **: "\<And>x. 0 \<le> s x" by (rule zero_min)
from * have "\<And>x. x \<in> supp_fun s \<Longrightarrow> s x = 0" by (simp only: sum_nonneg_eq_0_iff[OF assms **])
thus "s = 0" by (rule fun_eq_zeroI)
qed simp
lemma deg_fun_superset:
fixes A::"'a set"
assumes "supp_fun s \<subseteq> A" and "finite A"
shows "deg_fun s = (\<Sum>x\<in>A. s x)"
unfolding deg_fun_def
proof (rule sum.mono_neutral_cong_left, fact, fact, rule)
fix x
assume "x \<in> A - supp_fun s"
hence "x \<notin> supp_fun s" by simp
thus "s x = 0" by (simp add: supp_fun_def)
qed rule
lemma deg_fun_plus:
assumes "finite (supp_fun s)" and "finite (supp_fun t)"
shows "deg_fun (s + t) = deg_fun s + deg_fun (t::'a \<Rightarrow> 'b::comm_monoid_add)"
proof -
from assms have fin: "finite (supp_fun s \<union> supp_fun t)" by simp
have "deg_fun (s + t) = (\<Sum>x\<in>(supp_fun (s + t)). s x + t x)" by (simp add: deg_fun_def)
also from fin have "... = (\<Sum>x\<in>(supp_fun s \<union> supp_fun t). s x + t x)"
proof (rule sum.mono_neutral_cong_left)
show "\<forall>x\<in>supp_fun s \<union> supp_fun t - supp_fun (s + t). s x + t x = 0"
proof
fix x
assume "x \<in> supp_fun s \<union> supp_fun t - supp_fun (s + t)"
hence "x \<notin> supp_fun (s + t)" by simp
thus "s x + t x = 0" by (simp add: supp_fun_def)
qed
qed (rule supp_fun_plus_subset, rule)
also have "\<dots> = (\<Sum>x\<in>(supp_fun s \<union> supp_fun t). s x) + (\<Sum>x\<in>(supp_fun s \<union> supp_fun t). t x)"
by (rule sum.distrib)
also from fin have "(\<Sum>x\<in>(supp_fun s \<union> supp_fun t). s x) = deg_fun s" unfolding deg_fun_def
proof (rule sum.mono_neutral_cong_right)
show "\<forall>x\<in>supp_fun s \<union> supp_fun t - supp_fun s. s x = 0"
proof
fix x
assume "x \<in> supp_fun s \<union> supp_fun t - supp_fun s"
hence "x \<notin> supp_fun s" by simp
thus "s x = 0" by (simp add: supp_fun_def)
qed
qed simp_all
also from fin have "(\<Sum>x\<in>(supp_fun s \<union> supp_fun t). t x) = deg_fun t" unfolding deg_fun_def
proof (rule sum.mono_neutral_cong_right)
show "\<forall>x\<in>supp_fun s \<union> supp_fun t - supp_fun t. t x = 0"
proof
fix x
assume "x \<in> supp_fun s \<union> supp_fun t - supp_fun t"
hence "x \<notin> supp_fun t" by simp
thus "t x = 0" by (simp add: supp_fun_def)
qed
qed simp_all
finally show ?thesis .
qed
lemma deg_fun_leq:
assumes "finite (supp_fun s)" and "finite (supp_fun t)" and "s \<le> (t::'a \<Rightarrow> 'b::ordered_comm_monoid_add)"
shows "deg_fun s \<le> deg_fun t"
proof -
let ?A = "supp_fun s \<union> supp_fun t"
from assms(1) assms(2) have 1: "finite ?A" by simp
have s: "supp_fun s \<subseteq> ?A" and t: "supp_fun t \<subseteq> ?A" by simp_all
show ?thesis unfolding deg_fun_superset[OF s 1] deg_fun_superset[OF t 1]
proof (rule sum_mono)
fix i
from assms(3) show "s i \<le> t i" unfolding le_fun_def ..
qed
qed
subsubsection \<open>General Degree-Orders\<close>
context linorder
begin
lemma ex_min:
assumes "finite (A::'a set)" and "A \<noteq> {}"
shows "\<exists>y\<in>A. (\<forall>z\<in>A. y \<le> z)"
using assms
proof (induct rule: finite_induct)
assume "{} \<noteq> {}"
thus "\<exists>y\<in>{}. \<forall>z\<in>{}. y \<le> z" by simp
next
fix a::'a and A::"'a set"
assume "a \<notin> A" and IH: "A \<noteq> {} \<Longrightarrow> \<exists>y\<in>A. (\<forall>z\<in>A. y \<le> z)"
show "\<exists>y\<in>insert a A. (\<forall>z\<in>insert a A. y \<le> z)"
proof (cases "A = {}")
case True
show ?thesis
proof (rule bexI[of _ a], intro ballI)
fix z
assume "z \<in> insert a A"
from this True have "z = a" by simp
thus "a \<le> z" by simp
qed (simp)
next
case False
from IH[OF False] obtain y where "y \<in> A" and y_min: "\<forall>z\<in>A. y \<le> z" by auto
from linear[of a y] show ?thesis
proof
assume "y \<le> a"
show ?thesis
proof (rule bexI[of _ y], intro ballI)
fix z
assume "z \<in> insert a A"
hence "z = a \<or> z \<in> A" by simp
thus "y \<le> z"
proof
assume "z = a"
from this \<open>y \<le> a\<close> show "y \<le> z" by simp
next
assume "z \<in> A"
from y_min[rule_format, OF this] show "y \<le> z" .
qed
next
from \<open>y \<in> A\<close> show "y \<in> insert a A" by simp
qed
next
assume "a \<le> y"
show ?thesis
proof (rule bexI[of _ a], intro ballI)
fix z
assume "z \<in> insert a A"
hence "z = a \<or> z \<in> A" by simp
thus "a \<le> z"
proof
assume "z = a"
from this show "a \<le> z" by simp
next
assume "z \<in> A"
from y_min[rule_format, OF this] \<open>a \<le> y\<close> show "a \<le> z" by simp
qed
qed (simp)
qed
qed
qed
definition dord_fun::"(('a \<Rightarrow> 'b::ordered_comm_monoid_add) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "dord_fun ord s t \<equiv> (let d1 = deg_fun s; d2 = deg_fun t in (d1 < d2 \<or> (d1 = d2 \<and> ord s t)))"
lemma dord_fun_degD:
assumes "dord_fun ord s t"
shows "deg_fun s \<le> deg_fun t"
using assms unfolding dord_fun_def Let_def by auto
lemma dord_fun_refl:
assumes "ord s s"
shows "dord_fun ord s s"
using assms unfolding dord_fun_def by simp
lemma dord_fun_antisym:
assumes ord_antisym: "ord s t \<Longrightarrow> ord t s \<Longrightarrow> s = t" and "dord_fun ord s t" and "dord_fun ord t s"
shows "s = t"
proof -
from assms(3) have ts: "deg_fun t < deg_fun s \<or> (deg_fun t = deg_fun s \<and> ord t s)"
unfolding dord_fun_def Let_def .
from assms(2) have st: "deg_fun s < deg_fun t \<or> (deg_fun s = deg_fun t \<and> ord s t)"
unfolding dord_fun_def Let_def .
thus ?thesis
proof
assume "deg_fun s < deg_fun t"
thus ?thesis using ts by auto
next
assume "deg_fun s = deg_fun t \<and> ord s t"
hence "deg_fun s = deg_fun t" and "ord s t" by simp_all
from \<open>deg_fun s = deg_fun t\<close> ts have "ord t s" by simp
with \<open>ord s t\<close> show ?thesis by (rule ord_antisym)
qed
qed
lemma dord_fun_trans:
assumes ord_trans: "ord s t \<Longrightarrow> ord t u \<Longrightarrow> ord s u" and "dord_fun ord s t" and "dord_fun ord t u"
shows "dord_fun ord s u"
proof -
from assms(3) have ts: "deg_fun t < deg_fun u \<or> (deg_fun t = deg_fun u \<and> ord t u)"
unfolding dord_fun_def Let_def .
from assms(2) have st: "deg_fun s < deg_fun t \<or> (deg_fun s = deg_fun t \<and> ord s t)"
unfolding dord_fun_def Let_def .
thus ?thesis
proof
assume "deg_fun s < deg_fun t"
from this dord_fun_degD[OF assms(3)] have "deg_fun s < deg_fun u" by simp
thus ?thesis by (simp add: dord_fun_def Let_def)
next
assume "deg_fun s = deg_fun t \<and> ord s t"
hence "deg_fun s = deg_fun t" and "ord s t" by simp_all
from ts show ?thesis
proof
assume "deg_fun t < deg_fun u"
hence "deg_fun s < deg_fun u" using \<open>deg_fun s = deg_fun t\<close> by simp
thus ?thesis by (simp add: dord_fun_def Let_def)
next
assume "deg_fun t = deg_fun u \<and> ord t u"
hence "deg_fun t = deg_fun u" and "ord t u" by simp_all
from ord_trans[OF \<open>ord s t\<close> \<open>ord t u\<close>] \<open>deg_fun s = deg_fun t\<close> \<open>deg_fun t = deg_fun u\<close> show ?thesis
by (simp add: dord_fun_def Let_def)
qed
qed
qed
lemma dord_fun_lin:
"dord_fun ord s t \<or> dord_fun ord t s"
if "ord s t \<or> ord t s"
for s t::"'a \<Rightarrow> 'b::{ordered_comm_monoid_add, linorder}"
proof (intro disjCI)
assume "\<not> dord_fun ord t s"
hence "deg_fun s \<le> deg_fun t \<and> (deg_fun t \<noteq> deg_fun s \<or> \<not> ord t s)"
unfolding dord_fun_def Let_def by auto
hence "deg_fun s \<le> deg_fun t" and dis1: "deg_fun t \<noteq> deg_fun s \<or> \<not> ord t s" by simp_all
show "dord_fun ord s t" unfolding dord_fun_def Let_def
proof (intro disjCI)
assume "\<not> (deg_fun s = deg_fun t \<and> ord s t)"
hence dis2: "deg_fun s \<noteq> deg_fun t \<or> \<not> ord s t" by simp
show "deg_fun s < deg_fun t"
proof (cases "deg_fun s = deg_fun t")
case True
from True dis1 have "\<not> ord t s" by simp
from True dis2 have "\<not> ord s t" by simp
from \<open>\<not> ord s t\<close> \<open>\<not> ord t s\<close> that show ?thesis by simp
next
case False
from this \<open>deg_fun s \<le> deg_fun t\<close> show ?thesis by simp
qed
qed
qed
lemma dord_fun_zero_min:
fixes s t::"'a \<Rightarrow> 'b::add_linorder_min"
assumes ord_refl: "\<And>t. ord t t" and "finite (supp_fun s)"
shows "dord_fun ord 0 s"
unfolding dord_fun_def Let_def deg_fun_zero
proof (rule disjCI)
assume "\<not> (0 = deg_fun s \<and> ord 0 s)"
hence dis: "deg_fun s \<noteq> 0 \<or> \<not> ord 0 s" by simp
show "0 < deg_fun s"
proof (cases "deg_fun s = 0")
case True
hence "s = 0" using deg_fun_eq_0_iff[OF assms(2)] by auto
hence "ord 0 s" using ord_refl by simp
with True dis show ?thesis by simp
next
case False
thus ?thesis by (auto simp: zero_less_iff_neq_zero)
qed
qed
lemma dord_fun_plus_monotone:
fixes s t u ::"'a \<Rightarrow> 'b::{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le}"
assumes ord_monotone: "ord s t \<Longrightarrow> ord (s + u) (t + u)" and "finite (supp_fun s)"
and "finite (supp_fun t)" and "finite (supp_fun u)" and "dord_fun ord s t"
shows "dord_fun ord (s + u) (t + u)"
proof -
from assms(5) have "deg_fun s < deg_fun t \<or> (deg_fun s = deg_fun t \<and> ord s t)"
unfolding dord_fun_def Let_def .
thus ?thesis
proof
assume "deg_fun s < deg_fun t"
hence "deg_fun (s + u) < deg_fun (t + u)" by (auto simp: deg_fun_plus[OF _ assms(4)] assms(2) assms(3))
thus ?thesis unfolding dord_fun_def Let_def by simp
next
assume "deg_fun s = deg_fun t \<and> ord s t"
hence "deg_fun s = deg_fun t" and "ord s t" by simp_all
from \<open>deg_fun s = deg_fun t\<close> have "deg_fun (s + u) = deg_fun (t + u)"
by (auto simp: deg_fun_plus[OF _ assms(4)] assms(2) assms(3))
from this ord_monotone[OF \<open>ord s t\<close>] show ?thesis unfolding dord_fun_def Let_def by simp
qed
qed
end (* linorder *)
context wellorder
begin
subsubsection \<open>Degree-Lexicographic Term Order\<close>
definition dlex_fun::"('a \<Rightarrow> 'b::ordered_comm_monoid_add) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "dlex_fun \<equiv> dord_fun lex_fun"
definition "dlex_fun_strict s t \<longleftrightarrow> dlex_fun s t \<and> \<not> dlex_fun t s"
lemma dlex_fun_refl:
shows "dlex_fun s s"
unfolding dlex_fun_def by (rule dord_fun_refl, rule lex_fun_refl)
lemma dlex_fun_antisym:
assumes "dlex_fun s t" and "dlex_fun t s"
shows "s = t"
by (rule dord_fun_antisym, erule lex_fun_antisym, assumption,
simp_all only: dlex_fun_def[symmetric], fact+)
lemma dlex_fun_trans:
assumes "dlex_fun s t" and "dlex_fun t u"
shows "dlex_fun s u"
by (simp only: dlex_fun_def, rule dord_fun_trans, erule lex_fun_trans, assumption,
simp_all only: dlex_fun_def[symmetric], fact+)
lemma dlex_fun_lin: "dlex_fun s t \<or> dlex_fun t s"
for s t::"('a \<Rightarrow> 'b::{ordered_comm_monoid_add, linorder})"
unfolding dlex_fun_def by (rule dord_fun_lin, rule lex_fun_lin)
corollary dlex_fun_strict_alt [code]:
"dlex_fun_strict s t = (\<not> dlex_fun t s)" for s t::"'a \<Rightarrow> 'b::{ordered_comm_monoid_add, linorder}"
unfolding dlex_fun_strict_def using dlex_fun_lin by auto
lemma dlex_fun_zero_min:
fixes s t::"('a \<Rightarrow> 'b::add_linorder_min)"
assumes "finite (supp_fun s)"
shows "dlex_fun 0 s"
unfolding dlex_fun_def by (rule dord_fun_zero_min, rule lex_fun_refl, fact)
lemma dlex_fun_plus_monotone:
fixes s t u::"'a \<Rightarrow> 'b::{ordered_cancel_comm_monoid_add, ordered_ab_semigroup_add_imp_le}"
assumes "finite (supp_fun s)" and "finite (supp_fun t)" and "finite (supp_fun u)" and "dlex_fun s t"
shows "dlex_fun (s + u) (t + u)"
using lex_fun_plus_monotone[of s t u] assms unfolding dlex_fun_def
by (rule dord_fun_plus_monotone)
subsubsection \<open>Degree-Reverse-Lexicographic Term Order\<close>
abbreviation rlex_fun::"('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> bool" where
"rlex_fun s t \<equiv> lex_fun t s"
text \<open>Note that @{const rlex_fun} is not precisely the reverse-lexicographic order relation on
power-products. Normally, the @{emph \<open>last\<close>} (i.\,e. highest) indeterminate whose exponent differs
in the two power-products to be compared is taken, but since we do not require the domain to be finite,
there might not be such a last indeterminate. Therefore, we simply take the converse of
@{const lex_fun}.\<close>
definition drlex_fun::"('a \<Rightarrow> 'b::ordered_comm_monoid_add) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where "drlex_fun \<equiv> dord_fun rlex_fun"
definition "drlex_fun_strict s t \<longleftrightarrow> drlex_fun s t \<and> \<not> drlex_fun t s"
lemma drlex_fun_refl:
shows "drlex_fun s s"
unfolding drlex_fun_def by (rule dord_fun_refl, fact lex_fun_refl)
lemma drlex_fun_antisym:
assumes "drlex_fun s t" and "drlex_fun t s"
shows "s = t"
by (rule dord_fun_antisym, erule lex_fun_antisym, assumption,
simp_all only: drlex_fun_def[symmetric], fact+)
lemma drlex_fun_trans:
assumes "drlex_fun s t" and "drlex_fun t u"
shows "drlex_fun s u"
by (simp only: drlex_fun_def, rule dord_fun_trans, erule lex_fun_trans, assumption,
simp_all only: drlex_fun_def[symmetric], fact+)
lemma drlex_fun_lin: "drlex_fun s t \<or> drlex_fun t s"
for s t::"('a \<Rightarrow> 'b::{ordered_comm_monoid_add, linorder})"
unfolding drlex_fun_def by (rule dord_fun_lin, rule lex_fun_lin)
corollary drlex_fun_strict_alt [code]:
"drlex_fun_strict s t = (\<not> drlex_fun t s)" for s t::"'a \<Rightarrow> 'b::{ordered_comm_monoid_add, linorder}"
unfolding drlex_fun_strict_def using drlex_fun_lin by auto
lemma drlex_fun_zero_min:
fixes s t::"('a \<Rightarrow> 'b::add_linorder_min)"
assumes "finite (supp_fun s)"
shows "drlex_fun 0 s"
unfolding drlex_fun_def by (rule dord_fun_zero_min, rule lex_fun_refl, fact)
lemma drlex_fun_plus_monotone:
fixes s t u::"'a \<Rightarrow> 'b::{ordered_cancel_comm_monoid_add, ordered_ab_semigroup_add_imp_le}"
assumes "finite (supp_fun s)" and "finite (supp_fun t)" and "finite (supp_fun u)" and "drlex_fun s t"
shows "drlex_fun (s + u) (t + u)"
using lex_fun_plus_monotone[of t s u] assms unfolding drlex_fun_def
by (rule dord_fun_plus_monotone)
end (* wellorder *)
text\<open>Every finite linear ordering is also a well-ordering. This fact is particularly useful when
working with fixed finite sets of indeterminates.\<close>
class finite_linorder = finite + linorder
begin
subclass wellorder
proof
fix P::"'a \<Rightarrow> bool" and a
assume hyp: "\<And>x. (\<And>y. (y < x) \<Longrightarrow> P y) \<Longrightarrow> P x"
show "P a"
proof (rule ccontr)
assume "\<not> P a"
have "finite {x. \<not> P x}" (is "finite ?A") by simp
from \<open>\<not> P a\<close> have "a \<in> ?A" by simp
hence "?A \<noteq> {}" by auto
from ex_min[OF \<open>finite ?A\<close> this] obtain b where "b \<in> ?A" and b_min: "\<forall>y\<in>?A. b \<le> y" by auto
from \<open>b \<in> ?A\<close> have "\<not> P b" by simp
with hyp[of b] obtain y where "y < b" and "\<not> P y" by auto
from \<open>\<not> P y\<close> have "y \<in> ?A" by simp
with b_min have "b \<le> y" by simp
with \<open>y < b\<close> show False by simp
qed
qed
end
subsection \<open>Type @{type poly_mapping}\<close>
lemma poly_mapping_eq_zeroI:
assumes "keys s = {}"
shows "s = (0::('a, 'b::zero) poly_mapping)"
proof (rule poly_mapping_eqI, simp)
fix x
from assms show "lookup s x = 0" by auto
qed
lemma keys_plus_ninv_comm_monoid_add: "keys (s + t) = keys s \<union> keys (t::'a \<Rightarrow>\<^sub>0 'b::ninv_comm_monoid_add)"
proof (rule, fact Poly_Mapping.keys_add, rule)
fix x
assume "x \<in> keys s \<union> keys t"
thus "x \<in> keys (s + t)"
proof
assume "x \<in> keys s"
thus ?thesis
by (metis in_keys_iff lookup_add plus_eq_zero)
next
assume "x \<in> keys t"
thus ?thesis
by (metis in_keys_iff lookup_add plus_eq_zero_2)
qed
qed
lemma lookup_zero_fun: "lookup 0 = 0"
by (simp only: zero_poly_mapping.rep_eq zero_fun_def)
lemma lookup_plus_fun: "lookup (s + t) = lookup s + lookup t"
by (simp only: plus_poly_mapping.rep_eq plus_fun_def)
lemma lookup_uminus_fun: "lookup (- s) = - lookup s"
by (fact uminus_poly_mapping.rep_eq)
lemma lookup_minus_fun: "lookup (s - t) = lookup s - lookup t"
by (simp only: minus_poly_mapping.rep_eq, rule, simp only: minus_apply)
lemma poly_mapping_adds_iff: "s adds t \<longleftrightarrow> lookup s adds lookup t"
unfolding adds_def
proof
assume "\<exists>k. t = s + k"
then obtain k where *: "t = s + k" ..
show "\<exists>k. lookup t = lookup s + k"
proof
from * show "lookup t = lookup s + lookup k" by (simp only: lookup_plus_fun)
qed
next
assume "\<exists>k. lookup t = lookup s + k"
then obtain k where *: "lookup t = lookup s + k" ..
have **: "k \<in> {f. finite {x. f x \<noteq> 0}}"
proof
have "finite {x. lookup t x \<noteq> 0}" by transfer
hence "finite {x. lookup s x + k x \<noteq> 0}" by (simp only: * plus_fun_def)
moreover have "finite {x. lookup s x \<noteq> 0}" by transfer
ultimately show "finite {x. k x \<noteq> 0}" by (rule finite_neq_0_inv', simp)
qed
show "\<exists>k. t = s + k"
proof
show "t = s + Abs_poly_mapping k"
by (rule poly_mapping_eqI, simp add: * lookup_add Abs_poly_mapping_inverse[OF **])
qed
qed
subsubsection \<open>@{typ "('a, 'b) poly_mapping"} belongs to class @{class comm_powerprod}\<close>
instance poly_mapping :: (type, cancel_comm_monoid_add) comm_powerprod
by standard
subsubsection \<open>@{typ "('a, 'b) poly_mapping"} belongs to class @{class ninv_comm_monoid_add}\<close>
instance poly_mapping :: (type, ninv_comm_monoid_add) ninv_comm_monoid_add
proof (standard, transfer)
fix s t::"'a \<Rightarrow> 'b"
assume "(\<lambda>k. s k + t k) = (\<lambda>_. 0)"
hence "s + t = 0" by (simp only: plus_fun_def zero_fun_def)
hence "s = 0" by (rule plus_eq_zero)
thus "s = (\<lambda>_. 0)" by (simp only: zero_fun_def)
qed
subsubsection \<open>@{typ "('a, 'b) poly_mapping"} belongs to class @{class lcs_powerprod}\<close>
instantiation poly_mapping :: (type, add_linorder) lcs_powerprod
begin
lift_definition lcs_poly_mapping::"('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b)" is "\<lambda>s t. \<lambda>x. max (s x) (t x)"
proof -
fix fun1 fun2::"'a \<Rightarrow> 'b"
assume "finite {t. fun1 t \<noteq> 0}" and "finite {t. fun2 t \<noteq> 0}"
from finite_neq_0'[OF this, of max] show "finite {t. max (fun1 t) (fun2 t) \<noteq> 0}"
by (auto simp: max_def)
qed
lemma adds_poly_mappingI:
assumes "lookup s \<le> lookup (t::'a \<Rightarrow>\<^sub>0 'b)"
shows "s adds t"
unfolding poly_mapping_adds_iff using assms by (rule adds_funI)
lemma lookup_lcs_fun: "lookup (lcs s t) = lcs (lookup s) (lookup (t:: 'a \<Rightarrow>\<^sub>0 'b))"
by (simp only: lcs_poly_mapping.rep_eq lcs_fun_def)
instance
by (standard, simp_all only: poly_mapping_adds_iff lookup_lcs_fun, rule adds_lcs, elim lcs_adds,
assumption, rule poly_mapping_eqI, simp only: lookup_lcs_fun lcs_comm)
end
lemma adds_poly_mapping: "s adds t \<longleftrightarrow> lookup s \<le> lookup t"
for s t::"'a \<Rightarrow>\<^sub>0 'b::add_linorder_min"
by (simp only: poly_mapping_adds_iff adds_fun)
lemma lookup_gcs_fun: "lookup (gcs s (t::'a \<Rightarrow>\<^sub>0 ('b::add_linorder))) = gcs (lookup s) (lookup t)"
proof
fix x
show "lookup (gcs s t) x = gcs (lookup s) (lookup t) x"
by (simp add: gcs_def lookup_minus lookup_add lookup_lcs_fun)
qed
subsubsection \<open>@{typ "('a, 'b) poly_mapping"} belongs to class @{class ulcs_powerprod}\<close>
instance poly_mapping :: (type, add_linorder_min) ulcs_powerprod ..
subsubsection \<open>Power-products in a given set of indeterminates.\<close>
lemma adds_except:
"s adds t = (except s V adds except t V \<and> except s (- V) adds except t (- V))"
for s t :: "'a \<Rightarrow>\<^sub>0 'b::add_linorder"
by (simp add: poly_mapping_adds_iff adds_except_fun[of "lookup s", where V=V] except.rep_eq)
lemma adds_except_singleton:
"s adds t \<longleftrightarrow> (except s {v} adds except t {v} \<and> lookup s v adds lookup t v)"
for s t :: "'a \<Rightarrow>\<^sub>0 'b::add_linorder"
by (simp add: poly_mapping_adds_iff adds_except_fun_singleton[of "lookup s", where v=v] except.rep_eq)
subsubsection \<open>Dickson's lemma for power-products in finitely many indeterminates\<close>
context countable
begin
definition elem_index :: "'a \<Rightarrow> nat" where "elem_index = (SOME f. inj f)"
lemma inj_elem_index: "inj elem_index"
unfolding elem_index_def using ex_inj by (rule someI_ex)
lemma elem_index_inj:
assumes "elem_index x = elem_index y"
shows "x = y"
using inj_elem_index assms by (rule injD)
lemma finite_nat_seg: "finite {x. elem_index x < n}"
proof (rule finite_imageD)
have "elem_index ` {x. elem_index x < n} \<subseteq> {0..<n}" by auto
moreover have "finite ..." ..
ultimately show "finite (elem_index ` {x. elem_index x < n})" by (rule finite_subset)
next
from inj_elem_index show "inj_on elem_index {x. elem_index x < n}" using inj_on_subset by blast
qed
end (* countable *)
lemma Dickson_poly_mapping:
assumes "finite V"
shows "almost_full_on (adds) {x::'a \<Rightarrow>\<^sub>0 'b::add_wellorder. keys x \<subseteq> V}"
proof (rule almost_full_onI)
fix seq::"nat \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b"
assume a: "\<forall>i. seq i \<in> {x::'a \<Rightarrow>\<^sub>0 'b. keys x \<subseteq> V}"
define seq' where "seq' = (\<lambda>i. lookup (seq i))"
from assms have "almost_full_on (adds) {x::'a \<Rightarrow> 'b. supp_fun x \<subseteq> V}" by (rule Dickson_fun)
moreover from a have "\<And>i. seq' i \<in> {x::'a \<Rightarrow> 'b. supp_fun x \<subseteq> V}"
by (auto simp: seq'_def keys_eq_supp)
ultimately obtain i j where "i < j" and "seq' i adds seq' j" by (rule almost_full_onD)
from this(2) have "seq i adds seq j" by (simp add: seq'_def poly_mapping_adds_iff)
with \<open>i < j\<close> show "good (adds) seq" by (rule goodI)
qed
definition varnum :: "'x set \<Rightarrow> ('x::countable \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> nat"
where "varnum X t = (if keys t - X = {} then 0 else Suc (Max (elem_index ` (keys t - X))))"
lemma elem_index_less_varnum:
assumes "x \<in> keys t"
obtains "x \<in> X" | "elem_index x < varnum X t"
proof (cases "x \<in> X")
case True
thus ?thesis ..
next
case False
with assms have 1: "x \<in> keys t - X" by simp
hence "keys t - X \<noteq> {}" by blast
hence eq: "varnum X t = Suc (Max (elem_index ` (keys t - X)))" by (simp add: varnum_def)
hence "elem_index x < varnum X t" using 1 by (simp add: less_Suc_eq_le)
thus ?thesis ..
qed
lemma varnum_plus:
"varnum X (s + t) = max (varnum X s) (varnum X (t::'x::countable \<Rightarrow>\<^sub>0 'b::ninv_comm_monoid_add))"
proof (simp add: varnum_def keys_plus_ninv_comm_monoid_add image_Un Un_Diff del: Diff_eq_empty_iff, intro impI)
assume 1: "keys s - X \<noteq> {}" and 2: "keys t - X \<noteq> {}"
have "finite (elem_index ` (keys s - X))" by simp
moreover from 1 have "elem_index ` (keys s - X) \<noteq> {}" by simp
moreover have "finite (elem_index ` (keys t - X))" by simp
moreover from 2 have "elem_index ` (keys t - X) \<noteq> {}" by simp
ultimately show "Max (elem_index ` (keys s - X) \<union> elem_index ` (keys t - X)) =
max (Max (elem_index ` (keys s - X))) (Max (elem_index ` (keys t - X)))"
by (rule Max_Un)
qed
lemma dickson_grading_varnum:
assumes "finite X"
shows "dickson_grading ((varnum X)::('x::countable \<Rightarrow>\<^sub>0 'b::add_wellorder) \<Rightarrow> nat)"
using varnum_plus
proof (rule dickson_gradingI)
fix m::nat
let ?V = "X \<union> {x. elem_index x < m}"
have "{t::'x \<Rightarrow>\<^sub>0 'b. varnum X t \<le> m} \<subseteq> {t. keys t \<subseteq> ?V}"
proof (rule, simp, intro subsetI, simp)
fix t::"'x \<Rightarrow>\<^sub>0 'b" and x::'x
assume "varnum X t \<le> m"
assume "x \<in> keys t"
thus "x \<in> X \<or> elem_index x < m"
proof (rule elem_index_less_varnum)
assume "x \<in> X"
thus ?thesis ..
next
assume "elem_index x < varnum X t"
hence "elem_index x < m" using \<open>varnum X t \<le> m\<close> by (rule less_le_trans)
thus ?thesis ..
qed
qed
thus "almost_full_on (adds) {t::'x \<Rightarrow>\<^sub>0 'b. varnum X t \<le> m}"
proof (rule almost_full_on_subset)
from assms finite_nat_seg have "finite ?V" by (rule finite_UnI)
thus "almost_full_on (adds) {t::'x \<Rightarrow>\<^sub>0 'b. keys t \<subseteq> ?V}" by (rule Dickson_poly_mapping)
qed
qed
corollary dickson_grading_varnum_empty:
"dickson_grading ((varnum {})::(_ \<Rightarrow>\<^sub>0 _::add_wellorder) \<Rightarrow> nat)"
using finite.emptyI by (rule dickson_grading_varnum)
lemma varnum_le_iff: "varnum X t \<le> n \<longleftrightarrow> keys t \<subseteq> X \<union> {x. elem_index x < n}"
by (auto simp: varnum_def Suc_le_eq)
lemma varnum_zero [simp]: "varnum X 0 = 0"
by (simp add: varnum_def)
lemma varnum_empty_eq_zero_iff: "varnum {} t = 0 \<longleftrightarrow> t = 0"
proof
assume "varnum {} t = 0"
hence "keys t = {}" by (simp add: varnum_def split: if_splits)
thus "t = 0" by (rule poly_mapping_eq_zeroI)
qed simp
instance poly_mapping :: (countable, add_wellorder) graded_dickson_powerprod
by standard (rule, fact dickson_grading_varnum_empty)
instance poly_mapping :: (finite, add_wellorder) dickson_powerprod
proof
have "finite (UNIV::'a set)" by simp
hence "almost_full_on (adds) {x::'a \<Rightarrow>\<^sub>0 'b. keys x \<subseteq> UNIV}" by (rule Dickson_poly_mapping)
thus "almost_full_on (adds) (UNIV::('a \<Rightarrow>\<^sub>0 'b) set)" by simp
qed
subsubsection \<open>Lexicographic Term Order\<close>
definition lex_pm :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a::linorder \<Rightarrow>\<^sub>0 'b::{zero,linorder}) \<Rightarrow> bool"
where "lex_pm = (\<le>)"
definition lex_pm_strict :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a::linorder \<Rightarrow>\<^sub>0 'b::{zero,linorder}) \<Rightarrow> bool"
where "lex_pm_strict = (<)"
lemma lex_pm_alt: "lex_pm s t = (s = t \<or> (\<exists>x. lookup s x < lookup t x \<and> (\<forall>y<x. lookup s y = lookup t y)))"
unfolding lex_pm_def by (metis less_eq_poly_mapping.rep_eq less_funE less_funI poly_mapping_eq_iff)
lemma lex_pm_refl: "lex_pm s s"
by (simp add: lex_pm_def)
lemma lex_pm_antisym: "lex_pm s t \<Longrightarrow> lex_pm t s \<Longrightarrow> s = t"
by (simp add: lex_pm_def)
lemma lex_pm_trans: "lex_pm s t \<Longrightarrow> lex_pm t u \<Longrightarrow> lex_pm s u"
by (simp add: lex_pm_def)
lemma lex_pm_lin: "lex_pm s t \<or> lex_pm t s"
by (simp add: lex_pm_def linear)
corollary lex_pm_strict_alt [code]: "lex_pm_strict s t = (\<not> lex_pm t s)"
by (auto simp: lex_pm_strict_def lex_pm_def)
lemma lex_pm_zero_min: "lex_pm 0 s" for s::"_ \<Rightarrow>\<^sub>0 _::add_linorder_min"
proof (rule ccontr)
assume "\<not> lex_pm 0 s"
hence "lex_pm_strict s 0" by (simp add: lex_pm_strict_alt)
thus False by (simp add: lex_pm_strict_def less_poly_mapping.rep_eq less_fun_def)
qed
lemma lex_pm_plus_monotone: "lex_pm s t \<Longrightarrow> lex_pm (s + u) (t + u)"
for s t::"_ \<Rightarrow>\<^sub>0 _::{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le}"
by (simp add: lex_pm_def add_right_mono)
subsubsection \<open>Degree\<close>
lift_definition deg_pm::"('a \<Rightarrow>\<^sub>0 'b::comm_monoid_add) \<Rightarrow> 'b" is deg_fun .
lemma deg_pm_zero[simp]: "deg_pm 0 = 0"
by (simp add: deg_pm.rep_eq lookup_zero_fun)
lemma deg_pm_eq_0_iff[simp]: "deg_pm s = 0 \<longleftrightarrow> s = 0" for s::"'a \<Rightarrow>\<^sub>0 'b::add_linorder_min"
by (simp only: deg_pm.rep_eq poly_mapping_eq_iff lookup_zero_fun, rule deg_fun_eq_0_iff,
simp add: keys_eq_supp[symmetric])
lemma deg_pm_superset:
assumes "keys s \<subseteq> A" and "finite A"
shows "deg_pm s = (\<Sum>x\<in>A. lookup s x)"
using assms by (simp only: deg_pm.rep_eq keys_eq_supp, elim deg_fun_superset)
lemma deg_pm_plus: "deg_pm (s + t) = deg_pm s + deg_pm (t::'a \<Rightarrow>\<^sub>0 'b::comm_monoid_add)"
by (simp only: deg_pm.rep_eq lookup_plus_fun, rule deg_fun_plus, simp_all add: keys_eq_supp[symmetric])
lemma deg_pm_single: "deg_pm (Poly_Mapping.single x k) = k"
proof -
have "keys (Poly_Mapping.single x k) \<subseteq> {x}" by simp
moreover have "finite {x}" by simp
ultimately have "deg_pm (Poly_Mapping.single x k) = (\<Sum>y\<in>{x}. lookup (Poly_Mapping.single x k) y)"
by (rule deg_pm_superset)
also have "... = k" by simp
finally show ?thesis .
qed
subsubsection \<open>General Degree-Orders\<close>
context linorder
begin
lift_definition dord_pm::"(('a \<Rightarrow>\<^sub>0 'b::ordered_comm_monoid_add) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
is dord_fun by (metis local.dord_fun_def)
lemma dord_pm_alt: "dord_pm ord = (\<lambda>x y. deg_pm x < deg_pm y \<or> (deg_pm x = deg_pm y \<and> ord x y))"
by (intro ext) (transfer, simp add: dord_fun_def Let_def)
lemma dord_pm_degD:
assumes "dord_pm ord s t"
shows "deg_pm s \<le> deg_pm t"
using assms by (simp only: dord_pm.rep_eq deg_pm.rep_eq, elim dord_fun_degD)
lemma dord_pm_refl:
assumes "ord s s"
shows "dord_pm ord s s"
using assms by (simp only: dord_pm.rep_eq, intro dord_fun_refl, simp add: lookup_inverse)
lemma dord_pm_antisym:
assumes "ord s t \<Longrightarrow> ord t s \<Longrightarrow> s = t" and "dord_pm ord s t" and "dord_pm ord t s"
shows "s = t"
using assms
proof (simp only: dord_pm.rep_eq poly_mapping_eq_iff)
assume 1: "(ord s t \<Longrightarrow> ord t s \<Longrightarrow> lookup s = lookup t)"
assume 2: "dord_fun (map_fun Abs_poly_mapping id \<circ> ord \<circ> Abs_poly_mapping) (lookup s) (lookup t)"
assume 3: "dord_fun (map_fun Abs_poly_mapping id \<circ> ord \<circ> Abs_poly_mapping) (lookup t) (lookup s)"
from _ 2 3 show "lookup s = lookup t" by (rule dord_fun_antisym, simp add: lookup_inverse 1)
qed
lemma dord_pm_trans:
assumes "ord s t \<Longrightarrow> ord t u \<Longrightarrow> ord s u" and "dord_pm ord s t" and "dord_pm ord t u"
shows "dord_pm ord s u"
using assms
proof (simp only: dord_pm.rep_eq poly_mapping_eq_iff)
assume 1: "(ord s t \<Longrightarrow> ord t u \<Longrightarrow> ord s u)"
assume 2: "dord_fun (map_fun Abs_poly_mapping id \<circ> ord \<circ> Abs_poly_mapping) (lookup s) (lookup t)"
assume 3: "dord_fun (map_fun Abs_poly_mapping id \<circ> ord \<circ> Abs_poly_mapping) (lookup t) (lookup u)"
from _ 2 3 show "dord_fun (map_fun Abs_poly_mapping id \<circ> ord \<circ> Abs_poly_mapping) (lookup s) (lookup u)"
by (rule dord_fun_trans, simp add: lookup_inverse 1)
qed
lemma dord_pm_lin:
"dord_pm ord s t \<or> dord_pm ord t s"
if "ord s t \<or> ord t s"
for s t::"'a \<Rightarrow>\<^sub>0 'b::{ordered_comm_monoid_add, linorder}"
using that by (simp only: dord_pm.rep_eq, intro dord_fun_lin, simp add: lookup_inverse)
lemma dord_pm_zero_min: "dord_pm ord 0 s"
if ord_refl: "\<And>t. ord t t"
for s t::"'a \<Rightarrow>\<^sub>0 'b::add_linorder_min"
using that
by (simp only: dord_pm.rep_eq lookup_zero_fun, intro dord_fun_zero_min,
simp add: lookup_inverse, simp add: keys_eq_supp[symmetric])
lemma dord_pm_plus_monotone:
fixes s t u ::"'a \<Rightarrow>\<^sub>0 'b::{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le}"
assumes "ord s t \<Longrightarrow> ord (s + u) (t + u)" and "dord_pm ord s t"
shows "dord_pm ord (s + u) (t + u)"
using assms
by (simp only: dord_pm.rep_eq lookup_plus_fun, intro dord_fun_plus_monotone,
simp add: lookup_inverse lookup_plus_fun[symmetric],
simp add: keys_eq_supp[symmetric],
simp add: keys_eq_supp[symmetric],
simp add: keys_eq_supp[symmetric],
simp add: lookup_inverse)
end (* linorder *)
subsubsection \<open>Degree-Lexicographic Term Order\<close>
definition dlex_pm::"('a::linorder \<Rightarrow>\<^sub>0 'b::{ordered_comm_monoid_add,linorder}) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "dlex_pm \<equiv> dord_pm lex_pm"
definition "dlex_pm_strict s t \<longleftrightarrow> dlex_pm s t \<and> \<not> dlex_pm t s"
lemma dlex_pm_refl: "dlex_pm s s"
unfolding dlex_pm_def using lex_pm_refl by (rule dord_pm_refl)
lemma dlex_pm_antisym: "dlex_pm s t \<Longrightarrow> dlex_pm t s \<Longrightarrow> s = t"
unfolding dlex_pm_def using lex_pm_antisym by (rule dord_pm_antisym)
lemma dlex_pm_trans: "dlex_pm s t \<Longrightarrow> dlex_pm t u \<Longrightarrow> dlex_pm s u"
unfolding dlex_pm_def using lex_pm_trans by (rule dord_pm_trans)
lemma dlex_pm_lin: "dlex_pm s t \<or> dlex_pm t s"
unfolding dlex_pm_def using lex_pm_lin by (rule dord_pm_lin)
corollary dlex_pm_strict_alt [code]: "dlex_pm_strict s t = (\<not> dlex_pm t s)"
unfolding dlex_pm_strict_def using dlex_pm_lin by auto
lemma dlex_pm_zero_min: "dlex_pm 0 s"
for s t::"(_ \<Rightarrow>\<^sub>0 _::add_linorder_min)"
unfolding dlex_pm_def using lex_pm_refl by (rule dord_pm_zero_min)
lemma dlex_pm_plus_monotone: "dlex_pm s t \<Longrightarrow> dlex_pm (s + u) (t + u)"
for s t::"_ \<Rightarrow>\<^sub>0 _::{ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add}"
unfolding dlex_pm_def using lex_pm_plus_monotone by (rule dord_pm_plus_monotone)
subsubsection \<open>Degree-Reverse-Lexicographic Term Order\<close>
definition drlex_pm::"('a::linorder \<Rightarrow>\<^sub>0 'b::{ordered_comm_monoid_add,linorder}) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "drlex_pm \<equiv> dord_pm (\<lambda>s t. lex_pm t s)"
definition "drlex_pm_strict s t \<longleftrightarrow> drlex_pm s t \<and> \<not> drlex_pm t s"
lemma drlex_pm_refl: "drlex_pm s s"
unfolding drlex_pm_def using lex_pm_refl by (rule dord_pm_refl)
lemma drlex_pm_antisym: "drlex_pm s t \<Longrightarrow> drlex_pm t s \<Longrightarrow> s = t"
unfolding drlex_pm_def using lex_pm_antisym by (rule dord_pm_antisym)
lemma drlex_pm_trans: "drlex_pm s t \<Longrightarrow> drlex_pm t u \<Longrightarrow> drlex_pm s u"
unfolding drlex_pm_def using lex_pm_trans by (rule dord_pm_trans)
lemma drlex_pm_lin: "drlex_pm s t \<or> drlex_pm t s"
unfolding drlex_pm_def using lex_pm_lin by (rule dord_pm_lin)
corollary drlex_pm_strict_alt [code]: "drlex_pm_strict s t = (\<not> drlex_pm t s)"
unfolding drlex_pm_strict_def using drlex_pm_lin by auto
lemma drlex_pm_zero_min: "drlex_pm 0 s"
for s t::"(_ \<Rightarrow>\<^sub>0 _::add_linorder_min)"
unfolding drlex_pm_def using lex_pm_refl by (rule dord_pm_zero_min)
lemma drlex_pm_plus_monotone: "drlex_pm s t \<Longrightarrow> drlex_pm (s + u) (t + u)"
for s t::"_ \<Rightarrow>\<^sub>0 _::{ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add}"
unfolding drlex_pm_def using lex_pm_plus_monotone by (rule dord_pm_plus_monotone)
end (* theory *)
diff --git a/thys/Propositional_Proof_Systems/SC_Cut.thy b/thys/Propositional_Proof_Systems/SC_Cut.thy
--- a/thys/Propositional_Proof_Systems/SC_Cut.thy
+++ b/thys/Propositional_Proof_Systems/SC_Cut.thy
@@ -1,166 +1,166 @@
theory SC_Cut
imports SC
begin
subsubsection\<open>Contraction\<close>
text\<open>First, we need contraction:\<close>
lemma contract:
"(F,F,\<Gamma> \<Rightarrow> \<Delta> \<longrightarrow> F,\<Gamma> \<Rightarrow> \<Delta>) \<and> (\<Gamma> \<Rightarrow> F,F,\<Delta> \<longrightarrow> \<Gamma> \<Rightarrow> F,\<Delta>)"
proof(induction F arbitrary: \<Gamma> \<Delta>)
case (Atom k)
have "Atom k, Atom k, \<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> Atom k, \<Gamma> \<Rightarrow> \<Delta>"
by(induction "Atom k, Atom k, \<Gamma>" \<Delta> arbitrary: \<Gamma> rule: SCp.induct)
(auto dest!: multi_member_split intro!: SCp.intros(3-) simp add: lem2 lem1 SCp.intros)
moreover have "\<Gamma> \<Rightarrow> Atom k, Atom k, \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> Atom k, \<Delta>"
by(induction "\<Gamma>" "Atom k, Atom k, \<Delta>" arbitrary: \<Delta> rule: SCp.induct)
(auto dest!: multi_member_split intro!: SCp.intros(3-) simp add: lem2 lem1 SCp.intros)
ultimately show ?case by blast
next
case Bot
have "\<bottom>, \<bottom>, \<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> \<bottom>, \<Gamma> \<Rightarrow> \<Delta>" by(induction "\<bottom>, \<bottom>, \<Gamma>" \<Delta> rule: SCp.induct; blast)
moreover have "(\<Gamma> \<Rightarrow> \<bottom>, \<bottom>, \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> \<bottom>, \<Delta>)"
using Bot_delR by fastforce
ultimately show ?case by blast
next
case (Not F)
then show ?case by (meson NotL_inv NotR_inv SCp.NotL SCp.NotR)
next
case (And F1 F2) then show ?case by(auto intro!: SCp.intros(3-) dest!: AndR_inv AndL_inv) (metis add_mset_commute)
next
case (Or F1 F2) then show ?case by(auto intro!: SCp.intros(3-) dest!: OrR_inv OrL_inv) (metis add_mset_commute)
next
(* Okay, so the induction hypothesis is poison for the simplifier.
For some reason, this didn't cause trouble for the other two cases, but here, it does. *)
case (Imp F1 F2) show ?case by(auto dest!: ImpR_inv ImpL_inv intro!: SCp.intros(3-)) (insert Imp.IH; blast)+
qed
corollary
shows contractL: "F, F, \<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> F, \<Gamma> \<Rightarrow> \<Delta>"
and contractR: "\<Gamma> \<Rightarrow> F, F, \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> F, \<Delta>"
using contract by blast+
subsubsection\<open>Cut\<close>
text\<open>Which cut rule we show is up to us:\<close>
lemma cut_cs_cf:
assumes context_sharing_Cut: "\<And>(A::'a formula) \<Gamma> \<Delta>. \<Gamma> \<Rightarrow> A,\<Delta> \<Longrightarrow> A,\<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> \<Delta>"
shows context_free_Cut: "\<Gamma> \<Rightarrow> (A::'a formula),\<Delta> \<Longrightarrow> A,\<Gamma>' \<Rightarrow> \<Delta>' \<Longrightarrow> \<Gamma> + \<Gamma>' \<Rightarrow> \<Delta> + \<Delta>'"
by(intro context_sharing_Cut[of "\<Gamma> + \<Gamma>'" A "\<Delta> + \<Delta>'"])
(metis add.commute union_mset_add_mset_left weakenL_set weakenR_set)+
lemma cut_cf_cs:
assumes context_free_Cut: "\<And>(A::'a formula) \<Gamma> \<Gamma>' \<Delta> \<Delta>'. \<Gamma> \<Rightarrow> A,\<Delta> \<Longrightarrow> A,\<Gamma>' \<Rightarrow> \<Delta>' \<Longrightarrow> \<Gamma> + \<Gamma>' \<Rightarrow> \<Delta> + \<Delta>'"
shows context_sharing_Cut: "\<Gamma> \<Rightarrow> (A::'a formula),\<Delta> \<Longrightarrow> A,\<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> \<Delta>"
proof -
have contract\<Gamma>\<Gamma>: "\<Gamma>+\<Gamma>' \<Rightarrow> \<Delta> \<Longrightarrow> \<Gamma>' \<subseteq># \<Gamma> \<Longrightarrow> \<Gamma> \<Rightarrow> \<Delta>" for \<Gamma> \<Gamma>' \<Delta>
proof(induction "\<Gamma>'" arbitrary: \<Gamma>)
case empty thus ?case using weakenL_set by force
next
case (add x \<Gamma>' \<Gamma>)
from add.prems(2) have "x \<in># \<Gamma>" by (simp add: insert_subset_eq_iff)
then obtain \<Gamma>'' where \<Gamma>[simp]: "\<Gamma> = x,\<Gamma>''" by (meson multi_member_split)
from add.prems(1) have "x,x,\<Gamma>'' + \<Gamma>' \<Rightarrow> \<Delta>" by simp
with contractL have "x, \<Gamma>'' + \<Gamma>' \<Rightarrow> \<Delta>" .
- with add.IH[of \<Gamma>] show ?case using \<Gamma> add.prems(2) subset_mset.order.trans by force
+ with add.IH[of \<Gamma>] show ?case using \<Gamma> add.prems(2) subset_mset.trans by force
qed
have contract\<Delta>\<Delta>: "\<Gamma> \<Rightarrow> \<Delta>+\<Delta>' \<Longrightarrow> \<Delta>' \<subseteq># \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> \<Delta>" for \<Gamma> \<Delta> \<Delta>'
proof(induction "\<Delta>'" arbitrary: \<Delta>)
case empty thus ?case using weakenL_set by force
next
case (add x \<Delta>' \<Delta>)
from add.prems(2) have "x \<in># \<Delta>" by (simp add: insert_subset_eq_iff)
then obtain \<Delta>'' where \<Delta>[simp]: "\<Delta> = x,\<Delta>''" by (metis multi_member_split)
from add.prems(1) have "\<Gamma> \<Rightarrow> x,x,\<Delta>'' + \<Delta>'" by simp
with contractR have "\<Gamma> \<Rightarrow> x, \<Delta>'' + \<Delta>'" .
- with add.IH[of \<Delta>] show ?case using \<Delta> add.prems(2) subset_mset.order.trans by force
+ with add.IH[of \<Delta>] show ?case using \<Delta> add.prems(2) subset_mset.trans by force
qed
show "\<Gamma> \<Rightarrow> A,\<Delta> \<Longrightarrow> A,\<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> \<Delta>"
using context_free_Cut[of \<Gamma> A \<Delta> \<Gamma> \<Delta>] contract\<Gamma>\<Gamma> contract\<Delta>\<Delta>
by blast
qed
(* since these are the only lemmas that need contraction on sets, I've decided to transfer those in place *)
text\<open>According to Troelstra and Schwichtenberg~\cite{troelstra2000basic}, the sharing variant is the one we want to prove.\<close>
lemma Cut_Atom_pre: "Atom k,\<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> Atom k,\<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> \<Delta>"
proof(induction "Atom k,\<Gamma>" "\<Delta>" arbitrary: \<Gamma> rule: SCp.induct)
case BotL
hence "\<bottom> \<in># \<Gamma>" by simp
thus ?case using SCp.BotL by blast
next
case (Ax l \<Delta>)
show ?case proof cases
assume "l = k"
with \<open>Atom l \<in># \<Delta>\<close> obtain \<Delta>' where "\<Delta> = Atom k, \<Delta>'" by (meson multi_member_split)
with \<open>\<Gamma> \<Rightarrow> Atom k, \<Delta>\<close> have "\<Gamma> \<Rightarrow> \<Delta>" using contractR by blast
thus ?thesis by auto
next
assume "l \<noteq> k"
with \<open>Atom l \<in># Atom k, \<Gamma>\<close> have "Atom l \<in># \<Gamma>" by simp
with \<open>Atom l \<in># \<Delta>\<close> show ?thesis using SCp.Ax[of l] by blast
qed
next
case NotL
thus ?case by(auto simp: add_eq_conv_ex intro!: SCp.NotL dest!: NotL_inv)
next
case NotR
then show ?case by(auto intro!: SCp.NotR dest!: NotR_inv)
next
case AndL
thus ?case by(auto simp: add_eq_conv_ex intro!: SCp.AndL dest!: AndL_inv)
next
case AndR
then show ?case by(auto intro!: SCp.AndR dest!: AndR_inv)
next
case OrL
thus ?case by(auto simp: add_eq_conv_ex intro!: SCp.OrL dest!: OrL_inv)
next
case OrR
thus ?case by(auto intro!: SCp.OrR dest!: OrR_inv)
next
case ImpL
thus ?case by(auto simp: add_eq_conv_ex intro!: SCp.ImpL dest!: ImpL_inv)
next
case ImpR
then show ?case by (auto dest!: ImpR_inv intro!: SCp.ImpR)
qed
text\<open>We can show the admissibility of the cut rule by induction on the cut formula
(or, if you will, as a procedure that splits the cut into smaller formulas that get cut).
The only mildly complicated case is that of cutting in an @{term "Atom k"}.
It is, contrary to the general case, only mildly complicated, since the cut formula can only appear principal in the axiom rules.\<close>
theorem cut: "\<Gamma> \<Rightarrow> F,\<Delta> \<Longrightarrow> F,\<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> \<Gamma> \<Rightarrow> \<Delta>"
proof(induction F arbitrary: \<Gamma> \<Delta>)
case Atom thus ?case using Cut_Atom_pre by metis
next
case Bot thus ?case using Bot_delR by fastforce
next
case Not with NotL_inv NotR_inv show ?case by blast
next
case And thus ?case by (meson AndL_inv AndR_inv weakenL)
next
case Or thus ?case by (metis OrL_inv OrR_inv weakenR)
next
case (Imp F G)
(* an automatic proof:
thus ?case by (meson ImpL_inv ImpR_inv weakenR)
a readable one: *)
from ImpR_inv \<open>\<Gamma> \<Rightarrow> F \<^bold>\<rightarrow> G, \<Delta>\<close> have R: "F, \<Gamma> \<Rightarrow> G, \<Delta>" by blast
from ImpL_inv \<open>F \<^bold>\<rightarrow> G, \<Gamma> \<Rightarrow> \<Delta>\<close> have L: "\<Gamma> \<Rightarrow> F, \<Delta>" "G, \<Gamma> \<Rightarrow> \<Delta>" by blast+
from L(1) have "\<Gamma> \<Rightarrow> F, G, \<Delta>" using weakenR add_ac(3) by blast
with R have "\<Gamma> \<Rightarrow> G, \<Delta>" using Imp.IH(1) by simp
with L(2) show "\<Gamma> \<Rightarrow> \<Delta>" using Imp.IH(2) by clarsimp
qed
(* For this proof to work with FOL, I think we would need very special inversion rules.
For example, for \<forall>R, we would need that there's a (finite!) multiset of substitutions S, such that
instead of having \<forall>x.F,\<Delta>, having {F[s/x] | s \<in> S} + \<Delta> is enough. I don't think that holds,
but we may be able to cheat ourselves around it\<dots> *)
corollary cut_cf: "\<lbrakk> \<Gamma> \<Rightarrow> A, \<Delta>; A, \<Gamma>' \<Rightarrow> \<Delta>'\<rbrakk> \<Longrightarrow> \<Gamma> + \<Gamma>' \<Rightarrow> \<Delta> + \<Delta>'"
using cut_cs_cf cut by metis
lemma assumes cut: "\<And> \<Gamma>' \<Delta>' (A::'a formula). \<lbrakk>\<Gamma>' \<Rightarrow> A, \<Delta>'; A, \<Gamma>' \<Rightarrow> \<Delta>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<Rightarrow> \<Delta>'"
shows contraction_admissibility: "F,F,\<Gamma> \<Rightarrow> \<Delta> \<Longrightarrow> (F::'a formula),\<Gamma> \<Rightarrow> \<Delta>"
by(rule cut[of "F,\<Gamma>" F \<Delta>, OF extended_Ax]) simp_all
(* yes, unpublished Chapman p 2/5, second to last paragraph. that's right. we can prove contraction with cut. but what's that good for? *)
end
diff --git a/thys/PseudoHoops/LeftComplementedMonoid.thy b/thys/PseudoHoops/LeftComplementedMonoid.thy
--- a/thys/PseudoHoops/LeftComplementedMonoid.thy
+++ b/thys/PseudoHoops/LeftComplementedMonoid.thy
@@ -1,183 +1,183 @@
section\<open>Left Complemented Monoid\<close>
theory LeftComplementedMonoid
imports Operations LatticeProperties.Lattice_Prop
begin
class right_pordered_monoid_mult = order + monoid_mult +
assumes mult_right_mono: "a \<le> b \<Longrightarrow> a * c \<le> b * c"
class one_greatest = one + ord +
assumes one_greatest [simp]: "a \<le> 1"
class integral_right_pordered_monoid_mult = right_pordered_monoid_mult + one_greatest
class left_residuated = ord + times + left_imp +
assumes left_residual: "(x * a \<le> b) = (x \<le> a l\<rightarrow> b)"
class left_residuated_pordered_monoid = integral_right_pordered_monoid_mult + left_residuated
class left_inf = inf + times + left_imp +
assumes inf_l_def: "(a \<sqinter> b) = (a l\<rightarrow> b) * a"
class left_complemented_monoid = left_residuated_pordered_monoid + left_inf +
assumes right_divisibility: "(a \<le> b) = (\<exists> c . a = c * b)"
begin
lemma lcm_D: "a l\<rightarrow> a = 1"
- apply (rule antisym, simp)
+ apply (rule order.antisym, simp)
by (unfold left_residual [THEN sym], simp)
subclass semilattice_inf
apply unfold_locales
apply (metis inf_l_def right_divisibility)
apply (metis inf_l_def left_residual order_refl)
by (metis inf_l_def left_residual mult_right_mono right_divisibility)
lemma left_one_inf [simp]: "1 \<sqinter> a = a"
- by (rule antisym, simp_all)
+ by (rule order.antisym, simp_all)
lemma left_one_impl [simp]: "1 l\<rightarrow> a = a"
(*by (metis inf_l_def left_one_inf mult.right_neutral)*)
proof -
have "1 l\<rightarrow> a = 1 \<sqinter> a" by (unfold inf_l_def, simp)
also have "1 \<sqinter> a = a" by simp
finally show ?thesis .
qed
lemma lcm_A: "(a l\<rightarrow> b) * a = (b l\<rightarrow> a) * b"
apply (unfold inf_l_def [THEN sym])
by (simp add: inf_commute)
lemma lcm_B: "((a * b) l\<rightarrow> c) = (a l\<rightarrow> (b l\<rightarrow> c))"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: left_residual [THEN sym] mult.assoc)
apply (simp add: left_residual)
apply (simp add: left_residual [THEN sym])
apply (rule_tac y="(b l\<rightarrow> c) * b" in order_trans)
apply (simp add: mult.assoc [THEN sym])
apply (rule mult_right_mono)
by (simp_all add: left_residual)
lemma lcm_C: "(a \<le> b) = ((a l\<rightarrow> b) = 1)"
(*by (metis eq_iff left_residual mult.left_neutral one_greatest)*)
proof -
have "(a \<le> b) = (1 * a \<le> b)" by simp
also have "\<dots> = (1 \<le> a l\<rightarrow> b)" by (unfold left_residual, simp)
- also have "\<dots> = (a l\<rightarrow> b = 1)" apply safe by(rule antisym, simp_all)
+ also have "\<dots> = (a l\<rightarrow> b = 1)" apply safe by(rule order.antisym, simp_all)
finally show ?thesis .
qed
end
class less_def = ord +
assumes less_def: "(a < b) = ((a \<le> b) \<and> \<not> (b \<le> a))"
class one_times = one + times +
assumes one_left [simp]: "1 * a = a"
and one_right [simp]: "a * 1 = a"
class left_complemented_monoid_algebra = left_imp + one_times + left_inf + less_def +
assumes left_impl_one [simp]: "a l\<rightarrow> a = 1"
and left_impl_times: "(a l\<rightarrow> b) * a = (b l\<rightarrow> a) * b"
and left_impl_ded: "((a * b) l\<rightarrow> c) = (a l\<rightarrow> (b l\<rightarrow> c))"
and left_lesseq: "(a \<le> b) = ((a l\<rightarrow> b) = 1)"
begin
lemma A: "(1 l\<rightarrow> a) l\<rightarrow> 1 = 1"
proof -
have "(1 l\<rightarrow> a) l\<rightarrow> 1 = (1 l\<rightarrow> a) * 1 l\<rightarrow> 1" by simp
also have "... = (a l\<rightarrow> 1) * a l\<rightarrow> 1" by (subst left_impl_times, simp)
also have "... = 1" by (simp add: left_impl_ded)
finally show ?thesis .
qed
subclass order
proof
fix x y show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)" by (unfold less_def, simp)
next
fix x show "x \<le> x" by (unfold left_lesseq, simp)
next
fix x y z assume a: "x \<le> y" assume b: "y \<le> z"
have "x l\<rightarrow> z = 1 * x l\<rightarrow> z" by simp
also have "... = (x l\<rightarrow> y) * x l\<rightarrow> z" by (cut_tac a, simp add: left_lesseq)
also have "... = (y l\<rightarrow> x) * y l\<rightarrow> z" by (simp add: left_impl_times)
also have "... = (y l\<rightarrow> x) l\<rightarrow> (y l\<rightarrow> z)" by (simp add: left_impl_ded)
also have "... = (y l\<rightarrow> x) l\<rightarrow> 1" by (cut_tac b, simp add: left_lesseq)
also have "... = (1 * y l\<rightarrow> x) l\<rightarrow> 1" by simp
also have "... = (1 l\<rightarrow> (y l\<rightarrow> x)) l\<rightarrow> 1" by (subst left_impl_ded, simp)
also have "... = 1" by (simp add: A)
finally show "x \<le> z" by (simp add: left_lesseq)
next
fix x y z assume a: "x \<le> y" assume b: "y \<le> x"
have "x = (x l\<rightarrow> y) * x" by (cut_tac a, simp add: left_lesseq)
also have "... = (y l\<rightarrow> x) * y" by (simp add: left_impl_times)
also have "... = y" by (cut_tac b, simp add: left_lesseq)
finally show "x = y" .
qed
lemma B: "(1 l\<rightarrow> a) \<le> 1"
apply (unfold left_lesseq)
by (rule A)
lemma C: "a \<le> (1 l\<rightarrow> a)"
by (simp add: left_lesseq left_impl_ded [THEN sym])
lemma D: "a \<le> 1"
apply (rule_tac y = "1 l\<rightarrow> a" in order_trans)
by (simp_all add: C B)
lemma less_eq: "(a \<le> b) = (\<exists> c . a = (c * b))"
(*by (metis left_impl_ded left_impl_one left_impl_times left_lesseq one_left)*)
apply safe
apply (unfold left_lesseq)
apply (rule_tac x = "b l\<rightarrow> a" in exI)
apply (simp add: left_impl_times)
apply (simp add: left_impl_ded)
apply (case_tac "c \<le> 1")
apply (simp add: left_lesseq)
by (simp add: D)
lemma F: "(a * b) * c l\<rightarrow> z = a * (b * c) l\<rightarrow> z"
by (simp add: left_impl_ded)
lemma associativity: "(a * b) * c = a * (b * c)"
(*by (metis F left_impl_one left_impl_times one_left)*)
- apply (rule antisym)
+ apply (rule order.antisym)
apply (unfold left_lesseq)
apply (simp add: F)
by (simp add: F [THEN sym])
lemma H: "a * b \<le> b"
apply (simp add: less_eq)
by auto
lemma I: "a * b l\<rightarrow> b = 1"
apply (case_tac "a * b \<le> b")
apply (simp add: left_lesseq)
by (simp add: H)
lemma K: "a \<le> b \<Longrightarrow> a * c \<le> b * c"
apply (unfold less_eq)
apply clarify
apply (rule_tac x = "ca" in exI)
by (simp add: associativity)
lemma L: "(x * a \<le> b) = (x \<le> a l\<rightarrow> b)"
by (simp add: left_lesseq left_impl_ded)
subclass left_complemented_monoid
apply unfold_locales
apply (simp_all add: less_def D associativity K)
apply (simp add: L)
by (simp add: less_eq)
end
lemma (in left_complemented_monoid) left_complemented_monoid:
"class.left_complemented_monoid_algebra (*) inf (l\<rightarrow>) (\<le>) (<) 1"
by (unfold_locales, simp_all add: less_le_not_le lcm_A lcm_B lcm_C lcm_D)
end
diff --git a/thys/PseudoHoops/PseudoHoopFilters.thy b/thys/PseudoHoops/PseudoHoopFilters.thy
--- a/thys/PseudoHoops/PseudoHoopFilters.thy
+++ b/thys/PseudoHoops/PseudoHoopFilters.thy
@@ -1,1093 +1,1093 @@
section\<open>Filters and Congruences\<close>
theory PseudoHoopFilters
imports PseudoHoops
begin
context pseudo_hoop_algebra
begin
definition
"filters = {F . F \<noteq> {} \<and> (\<forall> a b . a \<in> F \<and> b \<in> F \<longrightarrow> a * b \<in> F) \<and> (\<forall> a b . a \<in> F \<and> a \<le> b \<longrightarrow> b \<in> F)}"
definition
"properfilters = {F . F \<in> filters \<and> F \<noteq> UNIV}"
definition
"maximalfilters = {F . F \<in> filters \<and> (\<forall> A . A \<in> filters \<and> F \<subseteq> A \<longrightarrow> A = F \<or> A = UNIV)}"
definition
"ultrafilters = properfilters \<inter> maximalfilters"
lemma filter_i: "F \<in> filters \<Longrightarrow> a \<in> F \<Longrightarrow> b \<in> F \<Longrightarrow> a * b \<in> F"
by (simp add: filters_def)
lemma filter_ii: "F \<in> filters \<Longrightarrow> a \<in> F \<Longrightarrow> a \<le> b \<Longrightarrow> b \<in> F"
by (simp add: filters_def)
lemma filter_iii [simp]: "F \<in> filters \<Longrightarrow> 1 \<in> F"
by (auto simp add: filters_def)
lemma filter_left_impl:
"(F \<in> filters) = ((1 \<in> F) \<and> (\<forall> a b . a \<in> F \<and> a l\<rightarrow> b \<in> F \<longrightarrow> b \<in> F))"
apply safe
apply simp
apply (frule_tac a = "a l\<rightarrow> b" and b = a in filter_i)
apply simp
apply simp
apply (rule_tac a = "(a l\<rightarrow> b) * a" in filter_ii)
apply simp
apply simp
apply (simp add: inf_l_def [THEN sym])
apply (subst filters_def)
apply safe
apply (subgoal_tac "a l\<rightarrow> (b l\<rightarrow> a * b) \<in> F")
apply blast
apply (subst left_impl_ded [THEN sym])
apply (subst left_impl_one)
apply safe
apply (subst (asm) left_lesseq)
by blast
lemma filter_right_impl:
"(F \<in> filters) = ((1 \<in> F) \<and> (\<forall> a b . a \<in> F \<and> a r\<rightarrow> b \<in> F \<longrightarrow> b \<in> F))"
apply safe
apply simp
apply (frule_tac a = a and b = "a r\<rightarrow> b" in filter_i)
apply simp
apply simp
apply (rule_tac a = "a * (a r\<rightarrow> b)" in filter_ii)
apply simp
apply simp
apply (simp add: inf_r_def [THEN sym])
apply (subst filters_def)
apply safe
apply (subgoal_tac "b r\<rightarrow> (a r\<rightarrow> a * b) \<in> F")
apply blast
apply (subst right_impl_ded [THEN sym])
apply (subst right_impl_one)
apply safe
apply (subst (asm) right_lesseq)
by blast
lemma [simp]: "A \<subseteq> filters \<Longrightarrow> \<Inter> A \<in> filters"
apply (simp add: filters_def)
apply safe
apply (simp add: Inter_eq)
apply (drule_tac x = "1" in spec)
apply safe
apply (erule notE)
apply (subgoal_tac "x \<in> filters")
apply simp
apply (simp add: filters_def)
apply blast
apply (frule rev_subsetD)
apply simp
apply simp
apply (frule rev_subsetD)
apply simp
apply (subgoal_tac "a \<in> X")
apply blast
by blast
definition
"filterof X = \<Inter> {F . F \<in> filters \<and> X \<subseteq> F}"
lemma [simp]: "filterof X \<in> filters"
by (auto simp add: filterof_def)
lemma times_le_mono [simp]: "x \<le> y \<Longrightarrow> u \<le> v \<Longrightarrow> x * u \<le> y * v"
apply (rule_tac y = "x * v" in order_trans)
by (simp_all add: mult_left_mono mult_right_mono)
lemma prop_3_2_i:
"filterof X = {a . \<exists> n x . x \<in> X *^ n \<and> x \<le> a}"
apply safe
apply (subgoal_tac "{a . \<exists> n x . x \<in> X *^ n \<and> x \<le> a} \<in> filters")
apply (simp add: filterof_def)
apply (drule_tac x = "{a::'a. \<exists>(n::nat) x::'a. x \<in> X *^ n \<and> x \<le> a}" in spec)
apply safe
apply (rule_tac x = "1::nat" in exI)
apply (rule_tac x = "xa" in exI)
apply (simp add: times_set_def)
apply (drule drop_assumption)
apply (simp add: filters_def)
apply safe
apply (rule_tac x = "1" in exI)
apply (rule_tac x = "0" in exI)
apply (rule_tac x = "1" in exI)
apply simp
apply (rule_tac x = "n + na" in exI)
apply (rule_tac x = "x * xa" in exI)
apply safe
apply (simp add: power_set_add times_set_def)
apply blast
apply simp
apply (rule_tac x = "n" in exI)
apply (rule_tac x = "x" in exI)
apply simp
apply (simp add: filterof_def)
apply safe
apply (rule filter_ii)
apply simp_all
apply (subgoal_tac "\<forall>x . x \<in> X *^ n \<longrightarrow> x \<in> xb")
apply simp
apply (induct_tac n)
apply (simp add: power_set_0)
apply (simp add: power_set_Suc times_set_def)
apply safe
apply (rule filter_i)
apply simp_all
by blast
lemma ultrafilter_union:
"ultrafilters = {F . F \<in> filters \<and> F \<noteq> UNIV \<and> (\<forall> x . x \<notin> F \<longrightarrow> filterof (F \<union> {x}) = UNIV)}"
apply (simp add: ultrafilters_def maximalfilters_def properfilters_def filterof_def)
by auto
lemma filterof_sub: "F \<in> filters \<Longrightarrow> X \<subseteq> F \<Longrightarrow> filterof X \<subseteq> F"
apply (simp add: filterof_def)
by blast
lemma filterof_elem [simp]: "x \<in> X \<Longrightarrow> x \<in> filterof X"
apply (simp add: filterof_def)
by blast
lemma [simp]: "filterof X \<in> filters"
apply (simp add: filters_def prop_3_2_i)
apply safe
apply (rule_tac x = 1 in exI)
apply (rule_tac x = 0 in exI)
apply (rule_tac x = 1 in exI)
apply auto [1]
apply (rule_tac x = "n + na" in exI)
apply (rule_tac x = "x * xa" in exI)
apply safe
apply (unfold power_set_add)
apply (simp add: times_set_def)
apply auto [1]
apply (rule_tac y = "x * b" in order_trans)
apply (rule mult_left_mono)
apply simp
apply (simp add: mult_right_mono)
apply (rule_tac x = n in exI)
apply (rule_tac x = x in exI)
by simp
lemma singleton_power [simp]: "{a} *^ n = {b . b = a ^ n}"
apply (induct_tac n)
apply auto [1]
by (simp add: times_set_def)
lemma power_pair: "x \<in> {a, b} *^ n \<Longrightarrow> \<exists> i j . i + j = n \<and> x \<le> a ^ i \<and> x \<le> b ^ j"
apply (subgoal_tac "\<forall> x . x \<in> {a, b} *^ n \<longrightarrow> (\<exists> i j . i + j = n \<and> x \<le> a ^ i \<and> x \<le> b ^ j)")
apply auto[1]
apply (drule drop_assumption)
apply (induct_tac n)
apply auto [1]
apply safe
apply (simp add: times_set_def)
apply safe
apply (drule_tac x = y in spec)
apply safe
apply (rule_tac x = "i + 1" in exI)
apply (rule_tac x = "j" in exI)
apply simp
apply (rule_tac y = y in order_trans)
apply simp_all
apply (drule_tac x = y in spec)
apply safe
apply (rule_tac x = "i" in exI)
apply (rule_tac x = "j+1" in exI)
apply simp
apply (rule_tac y = y in order_trans)
by simp_all
lemma filterof_supremum:
"c \<in> supremum {a, b} \<Longrightarrow> filterof {c} = filterof {a} \<inter> filterof {b}"
apply safe
apply (cut_tac X = "{c}" and F = "filterof {a}" in filterof_sub)
apply simp_all
apply (simp add: supremum_def upper_bound_def)
apply safe
apply (rule_tac a = a in filter_ii)
apply simp_all
apply blast
apply (cut_tac X = "{c}" and F = "filterof {b}" in filterof_sub)
apply simp_all
apply (simp add: supremum_def upper_bound_def)
apply safe
apply (rule_tac a = b in filter_ii)
apply simp_all
apply blast
apply (subst (asm) prop_3_2_i)
apply simp
apply (subst (asm) prop_3_2_i)
apply simp
apply safe
apply (cut_tac A = "{a, b}" and a = c and b = x and n = "n + na" in lemma_2_8_ii1)
apply simp
apply (subst prop_3_2_i)
apply simp
apply (rule_tac x = "n + na" in exI)
apply (subgoal_tac "infimum ((\<lambda>xa::'a. xa r\<rightarrow> x) ` ({a, b} *^ (n + na))) = {1}")
apply simp
apply (simp add: right_lesseq)
apply (subst infimum_unique)
apply (subst infimum_def lower_bound_def)
apply (subst lower_bound_def)
apply safe
apply simp_all
apply (drule power_pair)
apply safe
apply (subst right_residual [THEN sym])
apply simp
apply (case_tac "n \<le> i")
apply (rule_tac y = "a ^ n" in order_trans)
apply (rule_tac y = "a ^ i" in order_trans)
apply simp_all
apply (subgoal_tac "na \<le> j")
apply (rule_tac y = "b ^ na" in order_trans)
apply (rule_tac y = "b ^ j" in order_trans)
by simp_all
definition "d1 a b = (a l\<rightarrow> b) * (b l\<rightarrow> a)"
definition "d2 a b = (a r\<rightarrow> b) * (b r\<rightarrow> a)"
definition "d3 a b = d1 b a"
definition "d4 a b = d2 b a"
lemma [simp]: "(a * b = 1) = (a = 1 \<and> b = 1)"
apply (rule iffI)
apply (rule conjI)
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (rule_tac y = "a*b" in order_trans)
apply simp
apply (drule drop_assumption)
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (rule_tac y = "a*b" in order_trans)
apply simp
apply (drule drop_assumption)
apply simp
by simp
lemma lemma_3_5_i_1: "(d1 a b = 1) = (a = b)"
apply (simp add: d1_def left_lesseq [THEN sym])
by auto
lemma lemma_3_5_i_2: "(d2 a b = 1) = (a = b)"
apply (simp add: d2_def right_lesseq [THEN sym])
by auto
lemma lemma_3_5_i_3: "(d3 a b = 1) = (a = b)"
apply (simp add: d3_def lemma_3_5_i_1)
by auto
lemma lemma_3_5_i_4: "(d4 a b = 1) = (a = b)"
apply (simp add: d4_def lemma_3_5_i_2)
by auto
lemma lemma_3_5_ii_1 [simp]: "d1 a a = 1"
apply (subst lemma_3_5_i_1)
by simp
lemma lemma_3_5_ii_2 [simp]: "d2 a a = 1"
apply (subst lemma_3_5_i_2)
by simp
lemma lemma_3_5_ii_3 [simp]: "d3 a a = 1"
apply (subst lemma_3_5_i_3)
by simp
lemma lemma_3_5_ii_4 [simp]: "d4 a a = 1"
apply (subst lemma_3_5_i_4)
by simp
lemma [simp]: "(a l\<rightarrow> 1) = 1"
by (simp add: left_lesseq [THEN sym])
end
context pseudo_hoop_algebra begin
lemma [simp]: "(a r\<rightarrow> 1) = 1"
by simp
lemma lemma_3_5_iii_1 [simp]: "d1 a 1 = a"
by (simp add: d1_def)
lemma lemma_3_5_iii_2 [simp]: "d2 a 1 = a"
by (simp add: d2_def)
lemma lemma_3_5_iii_3 [simp]: "d3 a 1 = a"
by (simp add: d3_def d1_def)
lemma lemma_3_5_iii_4 [simp]: "d4 a 1 = a"
by (simp add: d4_def d2_def)
lemma lemma_3_5_iv_1: "(d1 b c) * (d1 a b) * (d1 b c) \<le> d1 a c"
apply (simp add: d1_def)
apply (subgoal_tac "(b l\<rightarrow> c) * (c l\<rightarrow> b) * ((a l\<rightarrow> b) * (b l\<rightarrow> a)) * ((b l\<rightarrow> c) * (c l\<rightarrow> b)) =
((b l\<rightarrow> c) * (c l\<rightarrow> b) * (a l\<rightarrow> b)) * ((b l\<rightarrow> a) * (b l\<rightarrow> c) * (c l\<rightarrow> b))")
apply simp
apply (rule mult_mono)
apply (rule_tac y = "(b l\<rightarrow> c) * (a l\<rightarrow> b)" in order_trans)
apply (rule mult_right_mono)
apply simp
apply (simp add: lemma_2_5_14)
apply (rule_tac y = "(b l\<rightarrow> a) * (c l\<rightarrow> b)" in order_trans)
apply (rule mult_right_mono)
apply simp
apply (simp add: lemma_2_5_14)
by (simp add: mult.assoc)
lemma lemma_3_5_iv_2: "(d2 a b) * (d2 b c) * (d2 a b) \<le> d2 a c"
apply (simp add: d2_def)
apply (subgoal_tac "(a r\<rightarrow> b) * (b r\<rightarrow> a) * ((b r\<rightarrow> c) * (c r\<rightarrow> b)) * ((a r\<rightarrow> b) * (b r\<rightarrow> a)) =
((a r\<rightarrow> b) * (b r\<rightarrow> a) * (b r\<rightarrow> c)) * ((c r\<rightarrow> b) * (a r\<rightarrow> b) * (b r\<rightarrow> a))")
apply simp
apply (rule mult_mono)
apply (rule_tac y = "(a r\<rightarrow> b) * (b r\<rightarrow> c)" in order_trans)
apply (rule mult_right_mono)
apply simp
apply (simp add: lemma_2_5_15)
apply (rule_tac y = "(c r\<rightarrow> b) * (b r\<rightarrow> a)" in order_trans)
apply (rule mult_right_mono)
apply simp
apply (simp add: lemma_2_5_15)
by (simp add: mult.assoc)
lemma lemma_3_5_iv_3: "(d3 a b) * (d3 b c) * (d3 a b) \<le> d3 a c"
by (simp add: d3_def lemma_3_5_iv_1)
lemma lemma_3_5_iv_4: "(d4 b c) * (d4 a b) * (d4 b c) \<le> d4 a c"
by (simp add: d4_def lemma_3_5_iv_2)
definition
"cong_r F a b \<equiv> d1 a b \<in> F"
definition
"cong_l F a b \<equiv> d2 a b \<in> F"
lemma cong_r_filter: "F \<in> filters \<Longrightarrow> (cong_r F a b) = (a l\<rightarrow> b \<in> F \<and> b l\<rightarrow> a \<in> F)"
apply (simp add: cong_r_def d1_def)
apply safe
apply (rule filter_ii)
apply simp_all
apply simp
apply (rule filter_ii)
apply simp_all
apply simp
by (simp add: filter_i)
lemma cong_r_symmetric: "F \<in> filters \<Longrightarrow> (cong_r F a b) = (cong_r F b a)"
apply (simp add: cong_r_filter)
by blast
lemma cong_r_d3: "F \<in> filters \<Longrightarrow> (cong_r F a b) = (d3 a b \<in> F)"
apply (simp add: d3_def)
apply (subst cong_r_symmetric)
by (simp_all add: cong_r_def)
lemma cong_l_filter: "F \<in> filters \<Longrightarrow> (cong_l F a b) = (a r\<rightarrow> b \<in> F \<and> b r\<rightarrow> a \<in> F)"
apply (simp add: cong_l_def d2_def)
apply safe
apply (rule filter_ii)
apply simp_all
apply simp
apply (rule filter_ii)
apply simp_all
apply simp
by (simp add: filter_i)
lemma cong_l_symmetric: "F \<in> filters \<Longrightarrow> (cong_l F a b) = (cong_l F b a)"
apply (simp add: cong_l_filter)
by blast
lemma cong_l_d4: "F \<in> filters \<Longrightarrow> (cong_l F a b) = (d4 a b \<in> F)"
apply (simp add: d4_def)
apply (subst cong_l_symmetric)
by (simp_all add: cong_l_def)
lemma cong_r_reflexive: "F \<in> filters \<Longrightarrow> cong_r F a a"
by (simp add: cong_r_def)
lemma cong_r_transitive: "F \<in> filters \<Longrightarrow> cong_r F a b \<Longrightarrow> cong_r F b c \<Longrightarrow> cong_r F a c"
apply (simp add: cong_r_filter)
apply safe
apply (rule_tac a = "(b l\<rightarrow> c) * (a l\<rightarrow> b)" in filter_ii)
apply simp_all
apply (rule filter_i)
apply simp_all
apply (simp add: lemma_2_5_14)
apply (rule_tac a = "(b l\<rightarrow> a) * (c l\<rightarrow> b)" in filter_ii)
apply simp_all
apply (rule filter_i)
apply simp_all
by (simp add: lemma_2_5_14)
lemma cong_l_reflexive: "F \<in> filters \<Longrightarrow> cong_l F a a"
by (simp add: cong_l_def)
lemma cong_l_transitive: "F \<in> filters \<Longrightarrow> cong_l F a b \<Longrightarrow> cong_l F b c \<Longrightarrow> cong_l F a c"
apply (simp add: cong_l_filter)
apply safe
apply (rule_tac a = "(a r\<rightarrow> b) * (b r\<rightarrow> c)" in filter_ii)
apply simp_all
apply (rule filter_i)
apply simp_all
apply (simp add: lemma_2_5_15)
apply (rule_tac a = "(c r\<rightarrow> b) * (b r\<rightarrow> a)" in filter_ii)
apply simp_all
apply (rule filter_i)
apply simp_all
by (simp add: lemma_2_5_15)
lemma lemma_3_7_i: "F \<in> filters \<Longrightarrow> F = {a . cong_r F a 1}"
by (simp add: cong_r_def)
lemma lemma_3_7_ii: "F \<in> filters \<Longrightarrow> F = {a . cong_l F a 1}"
by (simp add: cong_l_def)
lemma lemma_3_8_i: "F \<in> filters \<Longrightarrow> (cong_r F a b) = (\<exists> x y . x \<in> F \<and> y \<in> F \<and> x * a = y * b)"
apply (subst cong_r_filter)
apply safe
apply (rule_tac x = "a l\<rightarrow> b" in exI)
apply (rule_tac x = "b l\<rightarrow> a" in exI)
apply (simp add: left_impl_times)
apply (subgoal_tac "x \<le> a l\<rightarrow> b")
apply (simp add: filter_ii)
apply (simp add: left_residual [THEN sym])
apply (subgoal_tac "y \<le> b l\<rightarrow> a")
apply (simp add: filter_ii)
apply (simp add: left_residual [THEN sym])
apply (subgoal_tac "y * b = x * a")
by simp_all
lemma lemma_3_8_ii: "F \<in> filters \<Longrightarrow> (cong_l F a b) = (\<exists> x y . x \<in> F \<and> y \<in> F \<and> a * x = b * y)"
apply (subst cong_l_filter)
apply safe
apply (rule_tac x = "a r\<rightarrow> b" in exI)
apply (rule_tac x = "b r\<rightarrow> a" in exI)
apply (simp add: right_impl_times)
apply (subgoal_tac "x \<le> a r\<rightarrow> b")
apply (simp add: filter_ii)
apply (simp add: right_residual [THEN sym])
apply (subgoal_tac "y \<le> b r\<rightarrow> a")
apply (simp add: filter_ii)
apply (simp add: right_residual [THEN sym])
apply (subgoal_tac "b * y = a * x")
by simp_all
lemma lemma_3_9_i: "F \<in> filters \<Longrightarrow> cong_r F a b \<Longrightarrow> cong_r F c d \<Longrightarrow> (a l\<rightarrow> c \<in> F) = (b l\<rightarrow> d \<in> F)"
apply (simp add: cong_r_filter)
apply safe
apply (rule_tac a = "(a l\<rightarrow> d) * (b l\<rightarrow> a)" in filter_ii)
apply (simp_all add: lemma_2_5_14)
apply (rule_tac a = "((c l\<rightarrow> d) * (a l\<rightarrow> c)) * (b l\<rightarrow> a)" in filter_ii)
apply simp_all
apply (simp add: filter_i)
apply (rule mult_right_mono)
apply (simp_all add: lemma_2_5_14)
apply (rule_tac a = "(b l\<rightarrow> c) * (a l\<rightarrow> b)" in filter_ii)
apply (simp_all add: lemma_2_5_14)
apply (rule_tac a = "((d l\<rightarrow> c) * (b l\<rightarrow> d)) * (a l\<rightarrow> b)" in filter_ii)
apply simp_all
apply (simp add: filter_i)
apply (rule mult_right_mono)
by (simp_all add: lemma_2_5_14)
lemma lemma_3_9_ii: "F \<in> filters \<Longrightarrow> cong_l F a b \<Longrightarrow> cong_l F c d \<Longrightarrow> (a r\<rightarrow> c \<in> F) = (b r\<rightarrow> d \<in> F)"
apply (simp add: cong_l_filter)
apply safe
apply (rule_tac a = "(b r\<rightarrow> a) * (a r\<rightarrow> d)" in filter_ii)
apply (simp_all add: lemma_2_5_15)
apply (rule_tac a = "(b r\<rightarrow> a) * ((a r\<rightarrow> c) * (c r\<rightarrow> d))" in filter_ii)
apply simp_all
apply (simp add: filter_i)
apply (rule mult_left_mono)
apply (simp_all add: lemma_2_5_15)
apply (rule_tac a = "(a r\<rightarrow> b) * (b r\<rightarrow> c)" in filter_ii)
apply (simp_all add: lemma_2_5_15)
apply (rule_tac a = "(a r\<rightarrow> b) * ((b r\<rightarrow> d) * (d r\<rightarrow> c))" in filter_ii)
apply simp_all
apply (simp add: filter_i)
apply (rule mult_left_mono)
by (simp_all add: lemma_2_5_15)
definition
"normalfilters = {F . F \<in> filters \<and> (\<forall> a b . (a l\<rightarrow> b \<in> F) = (a r\<rightarrow> b \<in> F))}"
lemma normalfilter_i:
"H \<in> normalfilters \<Longrightarrow> a l\<rightarrow> b \<in> H \<Longrightarrow> a r\<rightarrow> b \<in> H"
by (simp add: normalfilters_def)
lemma normalfilter_ii:
"H \<in> normalfilters \<Longrightarrow> a r\<rightarrow> b \<in> H \<Longrightarrow> a l\<rightarrow> b \<in> H"
by (simp add: normalfilters_def)
lemma [simp]: "H \<in> normalfilters \<Longrightarrow> H \<in> filters"
by (simp add: normalfilters_def)
lemma lemma_3_10_i_ii:
"H \<in> normalfilters \<Longrightarrow> {a} ** H = H ** {a}"
apply (simp add: times_set_def)
apply safe
apply simp
apply (rule_tac x = "a l\<rightarrow> a * y" in bexI)
apply (simp add: inf_l_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply simp
apply (rule normalfilter_ii)
apply simp_all
apply (rule_tac a = "y" in filter_ii)
apply simp_all
apply (simp add: right_residual [THEN sym])
apply (rule_tac x = "a r\<rightarrow> xa * a" in bexI)
apply (simp add: inf_r_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply simp
apply (rule normalfilter_i)
apply simp_all
apply (rule_tac a = "xa" in filter_ii)
apply simp_all
by (simp add: left_residual [THEN sym])
lemma lemma_3_10_ii_iii:
"H \<in> filters \<Longrightarrow> (\<forall> a . {a} ** H = H ** {a}) \<Longrightarrow> cong_r H = cong_l H"
apply (subst fun_eq_iff)
apply (subst fun_eq_iff)
apply safe
apply (subst (asm) lemma_3_8_i)
apply simp_all
apply safe
apply (subst lemma_3_8_ii)
apply simp_all
apply (subgoal_tac "xb * x \<in> {x} ** H")
apply (subgoal_tac "y * xa \<in> {xa} ** H")
apply (drule drop_assumption)
apply (drule drop_assumption)
apply (simp add: times_set_def)
apply safe
apply (rule_tac x = ya in exI)
apply simp
apply (rule_tac x = yb in exI)
apply simp
apply (drule_tac x = xa in spec)
apply (simp add: times_set_def)
apply auto[1]
apply (drule_tac x = x in spec)
apply simp
apply (simp add: times_set_def)
apply (rule_tac x = xb in bexI)
apply simp_all
apply (subst (asm) lemma_3_8_ii)
apply simp_all
apply safe
apply (subst lemma_3_8_i)
apply simp_all
apply (subgoal_tac "x * xb \<in> H ** {x}")
apply (subgoal_tac "xa * y \<in> H ** {xa}")
apply (drule drop_assumption)
apply (drule drop_assumption)
apply (simp add: times_set_def)
apply safe
apply (rule_tac x = xc in exI)
apply simp
apply (rule_tac x = xd in exI)
apply simp
apply (drule_tac x = xa in spec)
apply (simp add: times_set_def)
apply auto[1]
apply (drule_tac x = x in spec)
apply (subgoal_tac "x * xb \<in> {x} ** H")
apply simp
apply (subst times_set_def)
by blast
lemma lemma_3_10_i_iii:
"H \<in> normalfilters \<Longrightarrow> cong_r H = cong_l H"
by (simp add: lemma_3_10_i_ii lemma_3_10_ii_iii)
lemma order_impl_l [simp]: "a \<le> b \<Longrightarrow> a l\<rightarrow> b = 1"
by (simp add: left_lesseq)
end
context pseudo_hoop_algebra begin
lemma impl_l_d1: "(a l\<rightarrow> b) = d1 a (a \<sqinter> b)"
by (simp add: d1_def lemma_2_6_20_a )
lemma impl_r_d2: "(a r\<rightarrow> b) = d2 a (a \<sqinter> b)"
by (simp add: d2_def lemma_2_6_21_a)
lemma lemma_3_10_iii_i:
"H \<in> filters \<Longrightarrow> cong_r H = cong_l H \<Longrightarrow> H \<in> normalfilters"
apply (unfold normalfilters_def)
apply (simp add: impl_l_d1 impl_r_d2)
apply safe
apply (subgoal_tac "cong_r H a (a \<sqinter> b)")
apply simp
apply (subst (asm) cong_l_def)
apply simp
apply (subst cong_r_def)
apply simp
apply (subgoal_tac "cong_r H a (a \<sqinter> b)")
apply (subst (asm) cong_r_def)
apply simp
apply simp
apply (subst cong_l_def)
by simp
lemma lemma_3_10_ii_i:
"H \<in> filters \<Longrightarrow> (\<forall> a . {a} ** H = H ** {a}) \<Longrightarrow> H \<in> normalfilters"
apply (rule lemma_3_10_iii_i)
apply simp
apply (rule lemma_3_10_ii_iii)
by simp_all
definition
"allpowers x n = {y . \<exists> i. i < n \<and> y = x ^ i}"
lemma times_set_in: "a \<in> A \<Longrightarrow> b \<in> B \<Longrightarrow> c = a * b \<Longrightarrow> c \<in> A ** B"
apply (simp add: times_set_def)
by auto
lemma power_set_power: "a \<in> A \<Longrightarrow> a ^ n \<in> A *^ n"
apply (induct_tac n)
apply simp
apply simp
apply (rule_tac a = a and b = "a ^ n" in times_set_in)
by simp_all
lemma normal_filter_union: "H \<in> normalfilters \<Longrightarrow> (H \<union> {x}) *^ n = (H ** (allpowers x n)) \<union> {x ^ n} "
apply (induct_tac n)
apply (simp add: times_set_def allpowers_def)
apply safe
apply simp
apply (simp add: times_set_def)
apply safe
apply (simp add: allpowers_def)
apply safe
apply (subgoal_tac "x * xa \<in> H ** {x}")
apply (simp add: times_set_def)
apply safe
apply (drule_tac x = "xb" in bspec)
apply simp
apply (drule_tac x = "x ^ (i + 1)" in spec)
apply simp
apply safe
apply (erule notE)
apply (rule_tac x = "i + 1" in exI)
apply simp
apply (erule notE)
apply (simp add: mult.assoc [THEN sym])
apply (drule_tac a = x in lemma_3_10_i_ii)
apply (subgoal_tac "H ** {x} = {x} ** H")
apply simp
apply (simp add: times_set_def)
apply auto[1]
apply simp
apply (drule_tac x = "xaa" in bspec)
apply simp
apply (drule_tac x = "x ^ n" in bspec)
apply (simp add: allpowers_def)
apply blast
apply simp
apply (drule_tac x = "xaa * xb" in bspec)
apply (simp add: filter_i)
apply (simp add: mult.assoc)
apply (drule_tac x = "ya" in bspec)
apply (simp add: allpowers_def)
apply safe
apply (rule_tac x = i in exI)
apply simp
apply simp
apply (subst (asm) times_set_def)
apply (subst (asm) times_set_def)
apply simp
apply safe
apply (subst (asm) allpowers_def)
apply (subst (asm) allpowers_def)
apply safe
apply (case_tac "i = 0")
apply simp
apply (rule_tac a = xa and b = 1 in times_set_in)
apply blast
apply (simp add: allpowers_def times_set_def)
apply safe
apply simp
apply (drule_tac x = 1 in bspec)
apply simp
apply (drule_tac x = 1 in spec)
apply simp
apply (drule_tac x = 0 in spec)
apply auto[1]
apply simp
apply (rule_tac a = xaa and b = "x ^ i" in times_set_in)
apply blast
apply (case_tac "i = n")
apply simp
apply (simp add: allpowers_def)
apply safe
apply (subgoal_tac "x ^ i \<in> H ** {y . \<exists>i<n. y = x ^ i}")
apply simp
apply (rule_tac a = 1 and b = "x ^ i" in times_set_in)
apply simp
apply simp
apply (rule_tac x = i in exI)
apply simp
apply simp
apply (rule power_set_power)
by simp
lemma lemma_3_11_i: "H \<in> normalfilters \<Longrightarrow> filterof (H \<union> {x}) = {a . \<exists> n h . h \<in> H \<and> h * x ^ n \<le> a}"
apply (subst prop_3_2_i)
apply (subst normal_filter_union)
apply simp_all
apply safe
apply (rule_tac x = n in exI)
apply (rule_tac x = 1 in exI)
apply simp
apply (simp_all add: allpowers_def times_set_def)
apply safe
apply (rule_tac x = i in exI)
apply (rule_tac x = xb in exI)
apply simp
apply (rule_tac x = "n + 1" in exI)
apply (rule_tac x = "h * x ^ n" in exI)
apply safe
apply (erule notE)
apply (rule_tac x = h in bexI)
apply (rule_tac x = "x ^ n" in exI)
by auto
lemma lemma_3_11_ii: "H \<in> normalfilters \<Longrightarrow> filterof (H \<union> {x}) = {a . \<exists> n h . h \<in> H \<and> (x ^ n) * h \<le> a}"
apply (subst lemma_3_11_i)
apply simp_all
apply safe
apply (rule_tac x = n in exI)
apply (subgoal_tac "h * x ^ n \<in> {x ^ n} ** H")
apply (simp add: times_set_def)
apply auto[1]
apply (drule_tac a = "x ^ n" in lemma_3_10_i_ii)
apply simp
apply (simp add: times_set_def)
apply auto[1]
apply (rule_tac x = n in exI)
apply (subgoal_tac "(x ^ n) * h \<in> H ** {x ^ n}")
apply (simp add: times_set_def)
apply auto[1]
apply (drule_tac a = "x ^ n" in lemma_3_10_i_ii)
apply (drule_tac sym)
apply simp
apply (simp add: times_set_def)
by auto
lemma lemma_3_12_i_ii:
"H \<in> normalfilters \<Longrightarrow> H \<in> ultrafilters \<Longrightarrow> x \<notin> H \<Longrightarrow> (\<exists> n . x ^ n l\<rightarrow> a \<in> H)"
apply (subst (asm) ultrafilter_union)
apply clarify
apply (drule_tac x = x in spec)
apply clarify
apply (subst (asm) lemma_3_11_i)
apply assumption
apply (subgoal_tac "a \<in> {a::'a. \<exists>(n::nat) h::'a. h \<in> H \<and> h * x ^ n \<le> a}")
apply clarify
apply (rule_tac x = n in exI)
apply (simp add: left_residual)
apply (rule filter_ii)
by simp_all
lemma lemma_3_12_ii_i:
"H \<in> normalfilters \<Longrightarrow> H \<in> properfilters \<Longrightarrow> (\<forall> x a . x \<notin> H \<longrightarrow> (\<exists> n . x ^ n l\<rightarrow> a \<in> H)) \<Longrightarrow> H \<in> maximalfilters"
apply (subgoal_tac "H \<in> ultrafilters")
apply (simp add: ultrafilters_def)
apply (subst ultrafilter_union)
apply clarify
apply (subst (asm) properfilters_def)
apply clarify
apply (subst lemma_3_11_i)
apply simp_all
apply safe
apply simp_all
apply (drule_tac x = x in spec)
apply clarify
apply (drule_tac x = xb in spec)
apply clarify
apply (rule_tac x = n in exI)
apply (rule_tac x = "x ^ n l\<rightarrow> xb" in exI)
apply clarify
apply (subst inf_l_def [THEN sym])
by simp
lemma lemma_3_12_i_iii:
"H \<in> normalfilters \<Longrightarrow> H \<in> ultrafilters \<Longrightarrow> x \<notin> H \<Longrightarrow> (\<exists> n . x ^ n r\<rightarrow> a \<in> H)"
apply (subst (asm) ultrafilter_union)
apply clarify
apply (drule_tac x = x in spec)
apply clarify
apply (subst (asm) lemma_3_11_ii)
apply assumption
apply (subgoal_tac "a \<in> {a::'a. \<exists>(n::nat) h::'a. h \<in> H \<and> (x ^ n) * h \<le> a}")
apply clarify
apply (rule_tac x = n in exI)
apply (simp add: right_residual)
apply (rule filter_ii)
by simp_all
lemma lemma_3_12_iii_i:
"H \<in> normalfilters \<Longrightarrow> H \<in> properfilters \<Longrightarrow> (\<forall> x a . x \<notin> H \<longrightarrow> (\<exists> n . x ^ n r\<rightarrow> a \<in> H)) \<Longrightarrow> H \<in> maximalfilters"
apply (subgoal_tac "H \<in> ultrafilters")
apply (simp add: ultrafilters_def)
apply (subst ultrafilter_union)
apply clarify
apply (subst (asm) properfilters_def)
apply clarify
apply (subst lemma_3_11_ii)
apply simp_all
apply safe
apply simp_all
apply (drule_tac x = x in spec)
apply clarify
apply (drule_tac x = xb in spec)
apply clarify
apply (rule_tac x = n in exI)
apply (rule_tac x = "x ^ n r\<rightarrow> xb" in exI)
apply clarify
apply (subst inf_r_def [THEN sym])
by simp
definition
"cong H = (\<lambda> x y . cong_l H x y \<and> cong_r H x y)"
definition
"congruences = {R . equivp R \<and> (\<forall> a b c d . R a b \<and> R c d \<longrightarrow> R (a * c) (b * d) \<and> R (a l\<rightarrow> c) (b l\<rightarrow> d) \<and> R (a r\<rightarrow> c) (b r\<rightarrow> d))}"
lemma cong_l: "H \<in> normalfilters \<Longrightarrow> cong H = cong_l H"
by (simp add: cong_def lemma_3_10_i_iii)
lemma cong_r: "H \<in> normalfilters \<Longrightarrow> cong H = cong_r H"
by (simp add: cong_def lemma_3_10_i_iii)
lemma cong_equiv: "H \<in> normalfilters \<Longrightarrow> equivp (cong H)"
apply (simp add: cong_l)
apply (simp add: equivp_reflp_symp_transp reflp_def refl_on_def cong_l_reflexive cong_l_symmetric symp_def sym_def transp_def trans_def)
apply safe
apply (rule cong_l_transitive)
by simp_all
lemma cong_trans: "H \<in> normalfilters \<Longrightarrow> cong H x y \<Longrightarrow> cong H y z \<Longrightarrow> cong H x z"
apply (drule cong_equiv)
apply (drule equivp_transp)
by simp_all
lemma lemma_3_13 [simp]:
"H \<in> normalfilters \<Longrightarrow> cong H \<in> congruences"
apply (subst congruences_def)
apply safe
apply (simp add: cong_equiv)
apply (rule_tac y = "b * c" in cong_trans)
apply simp_all
apply (simp add: cong_r lemma_3_8_i)
apply safe
apply (rule_tac x = x in exI)
apply simp
apply (rule_tac x = y in exI)
apply (simp add: mult.assoc [THEN sym])
apply (simp add: cong_l lemma_3_8_ii)
apply safe
apply (rule_tac x = xa in exI)
apply simp
apply (rule_tac x = ya in exI)
apply (simp add: mult.assoc)
apply (rule_tac y = "a l\<rightarrow> d" in cong_trans)
apply simp
apply (simp add: cong_r cong_r_filter)
apply safe
apply (rule_tac a = "c l\<rightarrow> d" in filter_ii)
apply simp_all
apply (subst left_residual [THEN sym])
apply (simp add: lemma_2_5_14)
apply (rule_tac a = "d l\<rightarrow> c" in filter_ii)
apply simp_all
apply (subst left_residual [THEN sym])
apply (simp add: lemma_2_5_14)
apply (subst cong_l)
apply simp
apply (simp add: cong_r cong_r_filter cong_l_filter)
apply safe
apply (rule_tac a = "b l\<rightarrow> a" in filter_ii)
apply simp_all
apply (subst right_residual [THEN sym])
apply (simp add: lemma_2_5_14)
apply (rule_tac a = "a l\<rightarrow> b" in filter_ii)
apply simp_all
apply (subst right_residual [THEN sym])
apply (simp add: lemma_2_5_14)
apply (rule_tac y = "a r\<rightarrow> d" in cong_trans)
apply simp
apply (simp add: cong_l cong_l_filter)
apply safe
apply (rule_tac a = "c r\<rightarrow> d" in filter_ii)
apply simp_all
apply (subst right_residual [THEN sym])
apply (simp add: lemma_2_5_15)
apply (rule_tac a = "d r\<rightarrow> c" in filter_ii)
apply simp_all
apply (subst right_residual [THEN sym])
apply (simp add: lemma_2_5_15)
apply (subst cong_r)
apply simp
apply (simp add: cong_l cong_l_filter cong_r_filter)
apply safe
apply (rule_tac a = "b r\<rightarrow> a" in filter_ii)
apply simp_all
apply (subst left_residual [THEN sym])
apply (simp add: lemma_2_5_15)
apply (rule_tac a = "a r\<rightarrow> b" in filter_ii)
apply simp_all
apply (subst left_residual [THEN sym])
by (simp add: lemma_2_5_15)
lemma cong_times: "R \<in> congruences \<Longrightarrow> R a b \<Longrightarrow> R c d \<Longrightarrow> R (a * c) (b * d)"
by (simp add: congruences_def)
lemma cong_impl_l: "R \<in> congruences \<Longrightarrow> R a b \<Longrightarrow> R c d \<Longrightarrow> R (a l\<rightarrow> c) (b l\<rightarrow> d)"
by (simp add: congruences_def)
lemma cong_impl_r: "R \<in> congruences \<Longrightarrow> R a b \<Longrightarrow> R c d \<Longrightarrow> R (a r\<rightarrow> c) (b r\<rightarrow> d)"
by (simp add: congruences_def)
lemma cong_refl [simp]: "R \<in> congruences \<Longrightarrow> R a a"
by (simp add: congruences_def equivp_reflp)
lemma cong_trans_a: "R \<in> congruences \<Longrightarrow> R a b \<Longrightarrow> R b c \<Longrightarrow> R a c"
apply (simp add: congruences_def)
apply (rule_tac y = b in equivp_transp)
by simp_all
lemma cong_sym: "R \<in> congruences \<Longrightarrow> R a b \<Longrightarrow> R b a"
by (simp add: congruences_def equivp_symp)
definition
"normalfilter R = {a . R a 1}"
lemma lemma_3_14 [simp]:
"R \<in> congruences \<Longrightarrow> (normalfilter R) \<in> normalfilters"
apply (unfold normalfilters_def)
apply safe
apply (simp add: filters_def)
apply safe
apply (simp add: normalfilter_def)
apply (drule_tac x = 1 in spec)
apply (simp add: congruences_def equivp_reflp)
apply (simp add: normalfilter_def)
apply (drule_tac a = a and c = b and b = 1 and d = 1 and R = R in cong_times)
apply simp_all
apply (simp add: normalfilter_def)
apply (simp add: left_lesseq)
apply (cut_tac R = R and a = a and b = 1 and c = b and d = b in cong_impl_l)
apply simp_all
apply (simp add: cong_sym)
apply (simp_all add: normalfilter_def)
apply (cut_tac R = R and a = "a l\<rightarrow> b" and b = 1 and c = a and d = a in cong_times)
apply simp_all
apply (simp add: inf_l_def [THEN sym])
apply (cut_tac R = R and a = a and b = "a \<sqinter> b" and c = b and d = b in cong_impl_r)
apply simp_all
apply (simp add: cong_sym)
apply (cut_tac R = R and c = "a r\<rightarrow> b" and d = 1 and a = a and b = a in cong_times)
apply simp_all
apply (simp add: inf_r_def [THEN sym])
apply (cut_tac R = R and a = a and b = "a \<sqinter> b" and c = b and d = b in cong_impl_l)
apply simp_all
by (simp add: cong_sym)
lemma lemma_3_15_i:
"H \<in> normalfilters \<Longrightarrow> normalfilter (cong H) = H"
by (simp add: normalfilter_def cong_r cong_r_filter)
lemma lemma_3_15_ii:
"R \<in> congruences \<Longrightarrow> cong (normalfilter R) = R"
apply (simp add: fun_eq_iff cong_r cong_r_filter)
apply (simp add: normalfilter_def)
apply safe
apply (cut_tac R = R and a = "x l\<rightarrow> xa" and b = 1 and c = x and d = x in cong_times)
apply simp_all
apply (cut_tac R = R and a = "xa l\<rightarrow> x" and b = 1 and c = xa and d = xa in cong_times)
apply simp_all
apply (simp add: inf_l_def [THEN sym])
apply (rule_tac b = "x \<sqinter> xa" in cong_trans_a)
apply simp_all
apply (subst cong_sym)
apply simp_all
apply (subst inf.commute)
apply simp_all
apply (cut_tac R = R and a = x and b = xa and c = xa and d = xa in cong_impl_l)
apply simp_all
apply (cut_tac R = R and a = xa and b = xa and c = x and d = xa in cong_impl_l)
by simp_all
lemma lemma_3_15_iii: "H1 \<in> normalfilters \<Longrightarrow> H2 \<in> normalfilters \<Longrightarrow> (H1 \<subseteq> H2) = (cong H1 \<le> cong H2)"
apply safe
apply (simp add: cong_l cong_l_filter)
apply blast
apply (subgoal_tac "cong H2 x 1")
apply (simp add: cong_l cong_l_def)
apply (subgoal_tac "cong H1 x 1")
apply blast
by (simp add: cong_l cong_l_def)
definition
"p x y z = ((x l\<rightarrow> y) r\<rightarrow> z) \<sqinter> ((z l\<rightarrow> y) r\<rightarrow> x)"
lemma lemma_3_16_i: "p x x y = y \<and> p x y y = x"
apply safe
apply (simp_all add: p_def)
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp_all add: lemma_2_10_24)
- apply (rule antisym)
+ apply (rule order.antisym)
by (simp_all add: lemma_2_10_24)
definition "M x y z = ((y l\<rightarrow> x) r\<rightarrow> x) \<sqinter> ((z l\<rightarrow> y) r\<rightarrow> y) \<sqinter> ((x l\<rightarrow> z) r\<rightarrow> z)"
lemma "M x x y = x \<and> M x y x = x \<and> M y x x = x"
apply (simp add: M_def)
apply safe
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp_all add: lemma_2_10_24 lemma_2_5_9_b)
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp_all add: lemma_2_10_24 lemma_2_5_9_b)
- apply (rule antisym)
+ apply (rule order.antisym)
by (simp_all add: lemma_2_10_24 lemma_2_5_9_b)
end
end
diff --git a/thys/PseudoHoops/PseudoHoops.thy b/thys/PseudoHoops/PseudoHoops.thy
--- a/thys/PseudoHoops/PseudoHoops.thy
+++ b/thys/PseudoHoops/PseudoHoops.thy
@@ -1,910 +1,910 @@
section\<open>Pseudo-Hoops\<close>
theory PseudoHoops
imports RightComplementedMonoid
begin
lemma drop_assumption:
"p \<Longrightarrow> True"
by simp
class pseudo_hoop_algebra = left_complemented_monoid_algebra + right_complemented_monoid_nole_algebra +
assumes left_right_impl_times: "(a l\<rightarrow> b) * a = a * (a r\<rightarrow> b)"
begin
definition
"inf_rr a b = a * (a r\<rightarrow> b)"
definition
"lesseq_r a b = (a r\<rightarrow> b = 1)"
definition
"less_r a b = (lesseq_r a b \<and> \<not> lesseq_r b a)"
end
(*
sublocale pseudo_hoop_algebra < right: right_complemented_monoid_algebra lesseq_r less_r 1 "( * )" inf_rr "(r\<rightarrow>)";
apply unfold_locales;
apply simp_all;
apply (simp add: less_r_def);
apply (simp add: inf_rr_def);
apply (rule right_impl_times, rule right_impl_ded);
by (simp add: lesseq_r_def);
*)
context pseudo_hoop_algebra begin
lemma right_complemented_monoid_algebra: "class.right_complemented_monoid_algebra lesseq_r less_r 1 (*) inf_rr (r\<rightarrow>)"
(* by unfold_locales;*)
apply unfold_locales
apply simp_all
apply (simp add: less_r_def)
apply (simp add: inf_rr_def)
apply (rule right_impl_times, rule right_impl_ded)
by (simp add: lesseq_r_def)
lemma inf_rr_inf [simp]: "inf_rr = (\<sqinter>)"
by (unfold fun_eq_iff, simp add: inf_rr_def inf_l_def left_right_impl_times)
lemma lesseq_lesseq_r: "lesseq_r a b = (a \<le> b)"
proof -
interpret A: right_complemented_monoid_algebra lesseq_r less_r 1 "(*)" inf_rr "(r\<rightarrow>)"
by (rule right_complemented_monoid_algebra)
show "lesseq_r a b = (a \<le> b)"
apply (subst le_iff_inf)
apply (subst A.dual_algebra.inf.absorb_iff1 [of a b])
apply (unfold inf_rr_inf [THEN sym])
by simp
qed
lemma [simp]: "lesseq_r = (\<le>)"
apply (unfold fun_eq_iff, simp add: lesseq_lesseq_r) done
lemma [simp]: "less_r = (<)"
by (unfold fun_eq_iff, simp add: less_r_def less_le_not_le)
subclass right_complemented_monoid_algebra
apply (cut_tac right_complemented_monoid_algebra)
by simp
end
sublocale pseudo_hoop_algebra <
pseudo_hoop_dual: pseudo_hoop_algebra "\<lambda> a b . b * a" "(\<sqinter>)" "(r\<rightarrow>)" "(\<le>)" "(<)" 1 "(l\<rightarrow>)"
apply unfold_locales
apply (simp add: inf_l_def)
apply simp
apply (simp add: left_impl_times)
apply (simp add: left_impl_ded)
by (simp add: left_right_impl_times)
context pseudo_hoop_algebra begin
lemma commutative_ps: "(\<forall> a b . a * b = b * a) = ((l\<rightarrow>) = (r\<rightarrow>))"
apply (simp add: fun_eq_iff)
apply safe
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: right_residual [THEN sym])
apply (subgoal_tac "x * (x l\<rightarrow> xa) = (x l\<rightarrow> xa) * x")
apply (drule drop_assumption)
apply simp
apply (simp add: left_residual)
apply simp
apply (simp add: left_residual [THEN sym])
apply (simp add: right_residual)
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: left_residual)
apply (simp add: right_residual [THEN sym])
apply (simp add: left_residual)
by (simp add: right_residual [THEN sym])
lemma lemma_2_4_5: "a l\<rightarrow> b \<le> (c l\<rightarrow> a) l\<rightarrow> (c l\<rightarrow> b)"
apply (simp add: left_residual [THEN sym] mult.assoc)
apply (rule_tac y = "(a l\<rightarrow> b) * a" in order_trans)
apply (rule mult_left_mono)
by (simp_all add: left_residual)
end
context pseudo_hoop_algebra begin
lemma lemma_2_4_6: "a r\<rightarrow> b \<le> (c r\<rightarrow> a) r\<rightarrow> (c r\<rightarrow> b)"
by (rule pseudo_hoop_dual.lemma_2_4_5)
primrec
imp_power_l:: "'a => nat \<Rightarrow> 'a \<Rightarrow> 'a" ("(_) l-(_)\<rightarrow> (_)" [65,0,65] 65) where
"a l-0\<rightarrow> b = b" |
"a l-(Suc n)\<rightarrow> b = (a l\<rightarrow> (a l-n\<rightarrow> b))"
primrec
imp_power_r:: "'a => nat \<Rightarrow> 'a \<Rightarrow> 'a" ("(_) r-(_)\<rightarrow> (_)" [65,0,65] 65) where
"a r-0\<rightarrow> b = b" |
"a r-(Suc n)\<rightarrow> b = (a r\<rightarrow> (a r-n\<rightarrow> b))"
lemma lemma_2_4_7_a: "a l-n\<rightarrow> b = a ^ n l\<rightarrow> b"
apply (induct_tac n)
by (simp_all add: left_impl_ded)
lemma lemma_2_4_7_b: "a r-n\<rightarrow> b = a ^ n r\<rightarrow> b"
apply (induct_tac n)
by (simp_all add: right_impl_ded [THEN sym] power_commutes)
lemma lemma_2_5_8_a [simp]: "a * b \<le> a"
by (rule dual_algebra.H)
lemma lemma_2_5_8_b [simp]: "a * b \<le> b"
by (rule H)
lemma lemma_2_5_9_a: "a \<le> b l\<rightarrow> a"
by (simp add: left_residual [THEN sym] dual_algebra.H)
lemma lemma_2_5_9_b: "a \<le> b r\<rightarrow> a"
by (simp add: right_residual [THEN sym] H)
lemma lemma_2_5_11: "a * b \<le> a \<sqinter> b"
by simp
lemma lemma_2_5_12_a: "a \<le> b \<Longrightarrow> c l\<rightarrow> a \<le> c l\<rightarrow> b"
apply (subst left_residual [THEN sym])
apply (subst left_impl_times)
apply (rule_tac y = "(a l\<rightarrow> c) * b" in order_trans)
apply (simp add: mult_left_mono)
by (rule H)
lemma lemma_2_5_13_a: "a \<le> b \<Longrightarrow> b l\<rightarrow> c \<le> a l\<rightarrow> c"
apply (simp add: left_residual [THEN sym])
apply (rule_tac y = "(b l\<rightarrow> c) * b" in order_trans)
apply (simp add: mult_left_mono)
by (simp add: left_residual)
lemma lemma_2_5_14: "(b l\<rightarrow> c) * (a l\<rightarrow> b) \<le> a l\<rightarrow> c"
apply (simp add: left_residual [THEN sym])
apply (simp add: mult.assoc)
apply (subst left_impl_times)
apply (subst mult.assoc [THEN sym])
apply (subst left_residual)
by (rule dual_algebra.H)
lemma lemma_2_5_16: "(a l\<rightarrow> b) \<le> (b l\<rightarrow> c) r\<rightarrow> (a l\<rightarrow> c)"
apply (simp add: right_residual [THEN sym])
by (rule lemma_2_5_14)
lemma lemma_2_5_18: "(a l\<rightarrow> b) \<le> a * c l\<rightarrow> b * c"
apply (simp add: left_residual [THEN sym])
apply (subst mult.assoc [THEN sym])
apply (rule mult_right_mono)
apply (subst left_impl_times)
by (rule H)
end
context pseudo_hoop_algebra begin
lemma lemma_2_5_12_b: "a \<le> b \<Longrightarrow> c r\<rightarrow> a \<le> c r\<rightarrow> b"
by (rule pseudo_hoop_dual.lemma_2_5_12_a)
lemma lemma_2_5_13_b: "a \<le> b \<Longrightarrow> b r\<rightarrow> c \<le> a r\<rightarrow> c"
by (rule pseudo_hoop_dual.lemma_2_5_13_a)
lemma lemma_2_5_15: "(a r\<rightarrow> b) * (b r\<rightarrow> c) \<le> a r\<rightarrow> c"
by (rule pseudo_hoop_dual.lemma_2_5_14)
lemma lemma_2_5_17: "(a r\<rightarrow> b) \<le> (b r\<rightarrow> c) l\<rightarrow> (a r\<rightarrow> c)"
by (rule pseudo_hoop_dual.lemma_2_5_16)
lemma lemma_2_5_19: "(a r\<rightarrow> b) \<le> c * a r\<rightarrow> c * b"
by (rule pseudo_hoop_dual.lemma_2_5_18)
definition
"lower_bound A = {a . \<forall> x \<in> A . a \<le> x}"
definition
"infimum A = {a \<in> lower_bound A . (\<forall> x \<in> lower_bound A . x \<le> a)}"
lemma infimum_unique: "(infimum A = {x}) = (x \<in> infimum A)"
apply safe
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
by (simp_all add: infimum_def)
lemma lemma_2_6_20:
"a \<in> infimum A \<Longrightarrow> b l\<rightarrow> a \<in> infimum (((l\<rightarrow>) b)`A)"
apply (subst infimum_def)
apply safe
apply (simp add: infimum_def lower_bound_def lemma_2_5_12_a)
by (simp add: infimum_def lower_bound_def left_residual [THEN sym])
end
context pseudo_hoop_algebra begin
lemma lemma_2_6_21:
"a \<in> infimum A \<Longrightarrow> b r\<rightarrow> a \<in> infimum (((r\<rightarrow>) b)`A)"
by (rule pseudo_hoop_dual.lemma_2_6_20)
lemma infimum_pair: "a \<in> infimum {x . x = b \<or> x = c} = (a = b \<sqinter> c)"
apply (simp add: infimum_def lower_bound_def)
apply safe
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
lemma lemma_2_6_20_a:
"a l\<rightarrow> (b \<sqinter> c) = (a l\<rightarrow> b) \<sqinter> (a l\<rightarrow> c)"
apply (subgoal_tac "b \<sqinter> c \<in> infimum {x . x = b \<or> x = c}")
apply (drule_tac b = a in lemma_2_6_20)
apply (case_tac "((l\<rightarrow>) a) ` {x . ((x = b) \<or> (x = c))} = {x . x = a l\<rightarrow> b \<or> x = a l\<rightarrow> c}")
apply (simp_all add: infimum_pair)
by auto
end
context pseudo_hoop_algebra begin
lemma lemma_2_6_21_a:
"a r\<rightarrow> (b \<sqinter> c) = (a r\<rightarrow> b) \<sqinter> (a r\<rightarrow> c)"
by (rule pseudo_hoop_dual.lemma_2_6_20_a)
lemma mult_mono: "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> a * c \<le> b * d"
apply (rule_tac y = "a * d" in order_trans)
by (simp_all add: mult_right_mono mult_left_mono)
lemma lemma_2_7_22: "(a l\<rightarrow> b) * (c l\<rightarrow> d) \<le> (a \<sqinter> c) l\<rightarrow> (b \<sqinter> d)"
apply (rule_tac y = "(a \<sqinter> c l\<rightarrow> b) * (a \<sqinter> c l\<rightarrow> d)" in order_trans)
apply (rule mult_mono)
apply (simp_all add: lemma_2_5_13_a)
apply (rule_tac y = "(a \<sqinter> c l\<rightarrow> b) \<sqinter> (a \<sqinter> c l\<rightarrow> d)" in order_trans)
apply (rule lemma_2_5_11)
by (simp add: lemma_2_6_20_a)
end
context pseudo_hoop_algebra begin
lemma lemma_2_7_23: "(a r\<rightarrow> b) * (c r\<rightarrow> d) \<le> (a \<sqinter> c) r\<rightarrow> (b \<sqinter> d)"
apply (rule_tac y = "(c \<sqinter> a) r\<rightarrow> (d \<sqinter> b)" in order_trans)
apply (rule pseudo_hoop_dual.lemma_2_7_22)
by (simp add: inf_commute)
definition
"upper_bound A = {a . \<forall> x \<in> A . x \<le> a}"
definition
"supremum A = {a \<in> upper_bound A . (\<forall> x \<in> upper_bound A . a \<le> x)}"
lemma supremum_unique:
"a \<in> supremum A \<Longrightarrow> b \<in> supremum A \<Longrightarrow> a = b"
apply (simp add: supremum_def upper_bound_def)
- apply (rule antisym)
+ apply (rule order.antisym)
by auto
lemma lemma_2_8_i:
"a \<in> supremum A \<Longrightarrow> a l\<rightarrow> b \<in> infimum ((\<lambda> x . x l\<rightarrow> b)`A)"
apply (subst infimum_def)
apply safe
apply (simp add: supremum_def upper_bound_def lower_bound_def lemma_2_5_13_a)
apply (simp add: supremum_def upper_bound_def lower_bound_def left_residual [THEN sym])
by (simp add: right_residual)
end
context pseudo_hoop_algebra begin
lemma lemma_2_8_i1:
"a \<in> supremum A \<Longrightarrow> a r\<rightarrow> b \<in> infimum ((\<lambda> x . x r\<rightarrow> b)`A)"
- by (rule pseudo_hoop_dual.lemma_2_8_i, simp)
+ by (fact pseudo_hoop_dual.lemma_2_8_i)
definition
times_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "**" 70) where
"(A ** B) = {a . \<exists> x \<in> A . \<exists> y \<in> B . a = x * y}"
lemma times_set_assoc: "A ** (B ** C) = (A ** B) ** C"
apply (simp add: times_set_def)
apply safe
apply (rule_tac x = "xa * xb" in exI)
apply safe
apply (rule_tac x = xa in bexI)
apply (rule_tac x = xb in bexI)
apply simp_all
apply (subst mult.assoc)
apply (rule_tac x = ya in bexI)
apply simp_all
apply (rule_tac x = xb in bexI)
apply simp_all
apply (rule_tac x = "ya * y" in exI)
apply (subst mult.assoc)
apply simp
by auto
primrec power_set :: "'a set \<Rightarrow> nat \<Rightarrow> 'a set" (infixr "*^" 80) where
power_set_0: "(A *^ 0) = {1}"
| power_set_Suc: "(A *^ (Suc n)) = (A ** (A *^ n))"
lemma infimum_singleton [simp]: "infimum {a} = {a}"
apply (simp add: infimum_def lower_bound_def)
by auto
lemma lemma_2_8_ii:
"a \<in> supremum A \<Longrightarrow> (a ^ n) l\<rightarrow> b \<in> infimum ((\<lambda> x . x l\<rightarrow> b)`(A *^ n))"
apply (induct_tac n)
apply simp
apply (simp add: left_impl_ded)
apply (drule_tac a = "a ^ n l\<rightarrow> b" and b = a in lemma_2_6_20)
apply (simp add: infimum_def lower_bound_def times_set_def)
apply safe
apply (drule_tac b = "y l\<rightarrow> b" in lemma_2_8_i)
apply (simp add: infimum_def lower_bound_def times_set_def left_impl_ded)
apply (rule_tac y = "a l\<rightarrow> y l\<rightarrow> b" in order_trans)
apply simp_all
apply (subgoal_tac "(\<forall>xa \<in> A *^ n. x \<le> a l\<rightarrow> xa l\<rightarrow> b)")
apply simp
apply safe
apply (drule_tac b = "xa l\<rightarrow> b" in lemma_2_8_i)
apply (simp add: infimum_def lower_bound_def)
apply safe
apply (subgoal_tac "(\<forall>xb \<in> A. x \<le> xb l\<rightarrow> xa l\<rightarrow> b)")
apply simp
apply safe
apply (subgoal_tac "x \<le> xb * xa l\<rightarrow> b")
apply (simp add: left_impl_ded)
apply (subgoal_tac "(\<exists>x \<in> A. \<exists>y \<in> A *^ n. xb * xa = x * y)")
by auto
lemma power_set_Suc2:
"A *^ (Suc n) = A *^ n ** A"
apply (induct_tac n)
apply (simp add: times_set_def)
apply simp
apply (subst times_set_assoc)
by simp
lemma power_set_add:
"A *^ (n + m) = (A *^ n) ** (A *^ m)"
apply (induct_tac m)
apply simp
apply (simp add: times_set_def)
apply simp
apply (subst times_set_assoc)
apply (subst times_set_assoc)
apply (subst power_set_Suc2 [THEN sym])
by simp
end
context pseudo_hoop_algebra begin
lemma lemma_2_8_ii1:
"a \<in> supremum A \<Longrightarrow> (a ^ n) r\<rightarrow> b \<in> infimum ((\<lambda> x . x r\<rightarrow> b)`(A *^ n))"
apply (induct_tac n)
apply simp
apply (subst power_Suc2)
apply (subst power_set_Suc2)
apply (simp add: right_impl_ded)
apply (drule_tac a = "a ^ n r\<rightarrow> b" and b = a in lemma_2_6_21)
apply (simp add: infimum_def lower_bound_def times_set_def)
apply safe
apply (drule_tac b = "xa r\<rightarrow> b" in lemma_2_8_i1)
apply (simp add: infimum_def lower_bound_def times_set_def right_impl_ded)
apply (rule_tac y = "a r\<rightarrow> xa r\<rightarrow> b" in order_trans)
apply simp_all
apply (subgoal_tac "(\<forall>xa \<in> A *^ n. x \<le> a r\<rightarrow> xa r\<rightarrow> b)")
apply simp
apply safe
apply (drule_tac b = "xa r\<rightarrow> b" in lemma_2_8_i1)
apply (simp add: infimum_def lower_bound_def)
apply safe
apply (subgoal_tac "(\<forall>xb \<in> A. x \<le> xb r\<rightarrow> xa r\<rightarrow> b)")
apply simp
apply safe
apply (subgoal_tac "x \<le> xa * xb r\<rightarrow> b")
apply (simp add: right_impl_ded)
apply (subgoal_tac "(\<exists>x \<in> A *^ n. \<exists>y \<in> A . xa * xb = x * y)")
by auto
lemma lemma_2_9_i:
"b \<in> supremum A \<Longrightarrow> a * b \<in> supremum ((*) a ` A)"
apply (simp add: supremum_def upper_bound_def)
apply safe
apply (simp add: mult_left_mono)
by (simp add: right_residual)
lemma lemma_2_9_i1:
"b \<in> supremum A \<Longrightarrow> b * a \<in> supremum ((\<lambda> x . x * a) ` A)"
apply (simp add: supremum_def upper_bound_def)
apply safe
apply (simp add: mult_right_mono)
by (simp add: left_residual)
lemma lemma_2_9_ii:
"b \<in> supremum A \<Longrightarrow> a \<sqinter> b \<in> supremum ((\<sqinter>) a ` A)"
apply (subst supremum_def)
apply safe
apply (simp add: supremum_def upper_bound_def)
apply safe
apply (rule_tac y = x in order_trans)
apply simp_all
apply (subst inf_commute)
apply (subst inf_l_def)
apply (subst left_right_impl_times)
apply (frule_tac a = "(b r\<rightarrow> a)" in lemma_2_9_i1)
apply (simp add: right_residual)
apply (simp add: supremum_def upper_bound_def)
apply (simp add: right_residual)
apply safe
apply (subgoal_tac "(\<forall>xa \<in> A. b r\<rightarrow> a \<le> xa r\<rightarrow> x)")
apply simp
apply safe
apply (simp add: inf_l_def)
apply (simp add: left_right_impl_times)
apply (rule_tac y = "xa r\<rightarrow> b * (b r\<rightarrow> a)" in order_trans)
apply simp
apply (rule lemma_2_5_12_b)
apply (subst left_residual)
apply (subgoal_tac "(\<forall>xa\<in>A. xa \<le> (b r\<rightarrow> a) l\<rightarrow> x)")
apply simp
apply safe
apply (subst left_residual [THEN sym])
apply (rule_tac y = "xaa * (xaa r\<rightarrow> a)" in order_trans)
apply (rule mult_left_mono)
apply (rule lemma_2_5_13_b)
apply simp
apply (subst right_impl_times)
by simp
lemma lemma_2_10_24:
"a \<le> (a l\<rightarrow> b) r\<rightarrow> b"
by (simp add: right_residual [THEN sym] inf_l_def [THEN sym])
lemma lemma_2_10_25:
"a \<le> (a l\<rightarrow> b) r\<rightarrow> a"
by (rule lemma_2_5_9_b)
end
context pseudo_hoop_algebra begin
lemma lemma_2_10_26:
"a \<le> (a r\<rightarrow> b) l\<rightarrow> b"
by (rule pseudo_hoop_dual.lemma_2_10_24)
lemma lemma_2_10_27:
"a \<le> (a r\<rightarrow> b) l\<rightarrow> a"
by (rule lemma_2_5_9_a)
lemma lemma_2_10_28:
"b l\<rightarrow> ((a l\<rightarrow> b) r\<rightarrow> a) = b l\<rightarrow> a"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (subst left_residual [THEN sym])
apply (rule_tac y = "((a l\<rightarrow> b) r\<rightarrow> a) \<sqinter> a" in order_trans)
apply (subst inf_l_def)
apply (rule_tac y = "(((a l\<rightarrow> b) r\<rightarrow> a) l\<rightarrow> b) * ((a l\<rightarrow> b) r\<rightarrow> a)" in order_trans)
apply (subst left_impl_times)
apply simp_all
apply (rule mult_right_mono)
apply (rule_tac y = "a l\<rightarrow> b" in order_trans)
apply (rule lemma_2_5_13_a)
apply (fact lemma_2_10_25)
apply (fact lemma_2_10_26)
apply (rule lemma_2_5_12_a)
by (fact lemma_2_10_25)
end
context pseudo_hoop_algebra begin
lemma lemma_2_10_29:
"b r\<rightarrow> ((a r\<rightarrow> b) l\<rightarrow> a) = b r\<rightarrow> a"
by (rule pseudo_hoop_dual.lemma_2_10_28)
lemma lemma_2_10_30:
"((b l\<rightarrow> a) r\<rightarrow> a) l\<rightarrow> a = b l\<rightarrow> a"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp_all add: lemma_2_10_26)
apply (rule lemma_2_5_13_a)
by (rule lemma_2_10_24)
end
context pseudo_hoop_algebra begin
lemma lemma_2_10_31:
"((b r\<rightarrow> a) l\<rightarrow> a) r\<rightarrow> a = b r\<rightarrow> a"
by (rule pseudo_hoop_dual.lemma_2_10_30)
lemma lemma_2_10_32:
"(((b l\<rightarrow> a) r\<rightarrow> a) l\<rightarrow> b) l\<rightarrow> (b l\<rightarrow> a) = b l\<rightarrow> a"
proof -
have "((((b l\<rightarrow> a) r\<rightarrow> a) l\<rightarrow> b) l\<rightarrow> (b l\<rightarrow> a) = (((b l\<rightarrow> a) r\<rightarrow> a) l\<rightarrow> b) l\<rightarrow> (((b l\<rightarrow> a) r\<rightarrow> a) l\<rightarrow> a))"
by (simp add: lemma_2_10_30)
also have "\<dots> = ((((b l\<rightarrow> a) r\<rightarrow> a) l\<rightarrow> b) * ((b l\<rightarrow> a) r\<rightarrow> a) l\<rightarrow> a)"
by (simp add: left_impl_ded)
also have "\<dots> = (((b l\<rightarrow> a) r\<rightarrow> a) \<sqinter> b) l\<rightarrow> a"
by (simp add: inf_l_def)
- also have "\<dots> = b l\<rightarrow> a" apply (subgoal_tac "((b l\<rightarrow> a) r\<rightarrow> a) \<sqinter> b = b", simp, rule antisym)
+ also have "\<dots> = b l\<rightarrow> a" apply (subgoal_tac "((b l\<rightarrow> a) r\<rightarrow> a) \<sqinter> b = b", simp, rule order.antisym)
by (simp_all add: lemma_2_10_24)
finally show ?thesis .
qed
end
context pseudo_hoop_algebra begin
lemma lemma_2_10_33:
"(((b r\<rightarrow> a) l\<rightarrow> a) r\<rightarrow> b) r\<rightarrow> (b r\<rightarrow> a) = b r\<rightarrow> a"
by (rule pseudo_hoop_dual.lemma_2_10_32)
end
class pseudo_hoop_sup_algebra = pseudo_hoop_algebra + sup +
assumes
sup_comute: "a \<squnion> b = b \<squnion> a"
and sup_le [simp]: "a \<le> a \<squnion> b"
and le_sup_equiv: "(a \<le> b) = (a \<squnion> b = b)"
begin
lemma sup_le_2 [simp]:
"b \<le> a \<squnion> b"
by (subst sup_comute, simp)
lemma le_sup_equiv_r:
"(a \<squnion> b = b) = (a \<le> b)"
by (simp add: le_sup_equiv)
lemma sup_idemp [simp]:
"a \<squnion> a = a"
by (simp add: le_sup_equiv_r)
end
class pseudo_hoop_sup1_algebra = pseudo_hoop_algebra + sup +
assumes
sup_def: "a \<squnion> b = ((a l\<rightarrow> b) r\<rightarrow> b) \<sqinter> ((b l\<rightarrow> a) r\<rightarrow> a)"
begin
lemma sup_comute1: "a \<squnion> b = b \<squnion> a"
apply (simp add: sup_def)
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
lemma sup_le1 [simp]: "a \<le> a \<squnion> b"
by (simp add: sup_def lemma_2_10_24 lemma_2_5_9_b)
lemma le_sup_equiv1: "(a \<le> b) = (a \<squnion> b = b)"
apply safe
apply (simp add: left_lesseq)
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: sup_def)
apply (simp add: sup_def)
apply (simp_all add: lemma_2_10_24)
apply (simp add: le_iff_inf)
apply (subgoal_tac "(a \<sqinter> b = a \<sqinter> (a \<squnion> b)) \<and> (a \<sqinter> (a \<squnion> b) = a)")
apply simp
apply safe
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (drule drop_assumption)
by (simp add: sup_comute1)
subclass pseudo_hoop_sup_algebra
apply unfold_locales
apply (simp add: sup_comute1)
apply simp
by (simp add: le_sup_equiv1)
end
class pseudo_hoop_sup2_algebra = pseudo_hoop_algebra + sup +
assumes
sup_2_def: "a \<squnion> b = ((a r\<rightarrow> b) l\<rightarrow> b) \<sqinter> ((b r\<rightarrow> a) l\<rightarrow> a)"
context pseudo_hoop_sup1_algebra begin end
sublocale pseudo_hoop_sup2_algebra < sup1_dual: pseudo_hoop_sup1_algebra "(\<squnion>)" "\<lambda> a b . b * a" "(\<sqinter>)" "(r\<rightarrow>)" "(\<le>)" "(<)" 1 "(l\<rightarrow>)"
apply unfold_locales
by (simp add: sup_2_def)
context pseudo_hoop_sup2_algebra begin
lemma sup_comute_2: "a \<squnion> b = b \<squnion> a"
by (rule sup1_dual.sup_comute)
lemma sup_le2 [simp]: "a \<le> a \<squnion> b"
by (rule sup1_dual.sup_le)
lemma le_sup_equiv2: "(a \<le> b) = (a \<squnion> b = b)"
by (rule sup1_dual.le_sup_equiv)
subclass pseudo_hoop_sup_algebra
apply unfold_locales
apply (simp add: sup_comute_2)
apply simp
by (simp add: le_sup_equiv2)
end
class pseudo_hoop_lattice_a = pseudo_hoop_sup_algebra +
assumes sup_inf_le_distr: "a \<squnion> (b \<sqinter> c) \<le> (a \<squnion> b) \<sqinter> (a \<squnion> c)"
begin
lemma sup_lower_upper_bound [simp]:
"a \<le> c \<Longrightarrow> b \<le> c \<Longrightarrow> a \<squnion> b \<le> c"
apply (subst le_iff_inf)
apply (subgoal_tac "(a \<squnion> b) \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c) \<and> a \<squnion> (b \<sqinter> c) \<le> (a \<squnion> b) \<sqinter> (a \<squnion> c) \<and> a \<squnion> (b \<sqinter> c) = a \<squnion> b")
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply safe
apply auto[1]
apply (simp add: le_sup_equiv)
apply (rule sup_inf_le_distr)
by (simp add: le_iff_inf)
end
sublocale pseudo_hoop_lattice_a < lattice "(\<sqinter>)" "(\<le>)" "(<)" "(\<squnion>)"
by (unfold_locales, simp_all)
class pseudo_hoop_lattice_b = pseudo_hoop_sup_algebra +
assumes le_sup_cong: "a \<le> b \<Longrightarrow> a \<squnion> c \<le> b \<squnion> c"
begin
lemma sup_lower_upper_bound_b [simp]:
"a \<le> c \<Longrightarrow> b \<le> c \<Longrightarrow> a \<squnion> b \<le> c"
proof -
assume A: "a \<le> c"
assume B: "b \<le> c"
have "a \<squnion> b \<le> c \<squnion> b" by (cut_tac A, simp add: le_sup_cong)
also have "\<dots> = b \<squnion> c" by (simp add: sup_comute)
also have "\<dots> \<le> c \<squnion> c" by (cut_tac B, rule le_sup_cong, simp)
also have "\<dots> = c" by simp
finally show ?thesis .
qed
lemma sup_inf_le_distr_b:
"a \<squnion> (b \<sqinter> c) \<le> (a \<squnion> b) \<sqinter> (a \<squnion> c)"
apply (rule sup_lower_upper_bound_b)
apply simp
apply simp
apply safe
apply (subst sup_comute)
apply (rule_tac y = "b" in order_trans)
apply simp_all
apply (rule_tac y = "c" in order_trans)
by simp_all
end
context pseudo_hoop_lattice_a begin end
sublocale pseudo_hoop_lattice_b < pseudo_hoop_lattice_a "(\<squnion>)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
by (unfold_locales, rule sup_inf_le_distr_b)
class pseudo_hoop_lattice = pseudo_hoop_sup_algebra +
assumes sup_assoc_1: "a \<squnion> (b \<squnion> c) = (a \<squnion> b) \<squnion> c"
begin
lemma le_sup_cong_c:
"a \<le> b \<Longrightarrow> a \<squnion> c \<le> b \<squnion> c"
proof -
assume A: "a \<le> b"
have "a \<squnion> c \<squnion> (b \<squnion> c) = a \<squnion> (c \<squnion> (b \<squnion> c))" by (simp add: sup_assoc_1)
also have "\<dots> = a \<squnion> ((b \<squnion> c) \<squnion> c)" by (simp add: sup_comute)
also have "\<dots> = a \<squnion> (b \<squnion> (c \<squnion> c))" by (simp add: sup_assoc_1 [THEN sym])
also have "\<dots> = (a \<squnion> b) \<squnion> c" by (simp add: sup_assoc_1)
also have "\<dots> = b \<squnion> c" by (cut_tac A, simp add: le_sup_equiv)
finally show ?thesis by (simp add: le_sup_equiv)
qed
end
sublocale pseudo_hoop_lattice < pseudo_hoop_lattice_b "(\<squnion>)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
by (unfold_locales, rule le_sup_cong_c)
sublocale pseudo_hoop_lattice < semilattice_sup "(\<squnion>)" "(\<le>)" "(<)"
by (unfold_locales, simp_all)
sublocale pseudo_hoop_lattice < lattice "(\<sqinter>)" "(\<le>)" "(<)" "(\<squnion>)"
by unfold_locales
lemma (in pseudo_hoop_lattice_a) supremum_pair [simp]:
"supremum {a, b} = {a \<squnion> b}"
apply (simp add: supremum_def upper_bound_def)
apply safe
apply simp_all
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
sublocale pseudo_hoop_lattice < distrib_lattice "(\<sqinter>)" "(\<le>)" "(<)" "(\<squnion>)"
apply unfold_locales
apply (rule distrib_imp1)
apply (case_tac "xa \<sqinter> (ya \<squnion> za) \<in> supremum ((\<sqinter>) xa ` {ya, za})")
apply (simp add: supremum_pair)
apply (erule notE)
apply (rule lemma_2_9_ii)
by (simp add: supremum_pair)
class bounded_semilattice_inf_top = semilattice_inf + order_top
begin
lemma inf_eq_top_iff [simp]:
"(inf x y = top) = (x = top \<and> y = top)"
- by (simp add: eq_iff)
+ by (simp add: order.eq_iff)
end
sublocale pseudo_hoop_algebra < bounded_semilattice_inf_top "(\<sqinter>)" "(\<le>)" "(<)" "1"
by unfold_locales simp
definition (in pseudo_hoop_algebra)
sup1::"'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>1" 70) where
"a \<squnion>1 b = ((a l\<rightarrow> b) r\<rightarrow> b) \<sqinter> ((b l\<rightarrow> a) r\<rightarrow> a)"
sublocale pseudo_hoop_algebra < sup1: pseudo_hoop_sup1_algebra "(\<squnion>1)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
by (simp add: sup1_def)
definition (in pseudo_hoop_algebra)
sup2::"'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>2" 70) where
"a \<squnion>2 b = ((a r\<rightarrow> b) l\<rightarrow> b) \<sqinter> ((b r\<rightarrow> a) l\<rightarrow> a)"
sublocale pseudo_hoop_algebra < sup2: pseudo_hoop_sup2_algebra "(\<squnion>2)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
by (simp add: sup2_def)
context pseudo_hoop_algebra
begin
lemma lemma_2_15_i:
"1 \<in> supremum {a, b} \<Longrightarrow> a * b = a \<sqinter> b"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule lemma_2_5_11)
apply (simp add: inf_l_def)
apply (subst left_impl_times)
apply (rule mult_right_mono)
apply (subst right_lesseq)
apply (subgoal_tac "a \<squnion>1 b = 1")
apply (simp add: sup1_def)
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
by (simp add: supremum_def upper_bound_def)
lemma lemma_2_15_ii:
"1 \<in> supremum {a, b} \<Longrightarrow> a \<le> c \<Longrightarrow> b \<le> d \<Longrightarrow> 1 \<in> supremum {c, d}"
apply (simp add: supremum_def upper_bound_def)
apply safe
apply (drule_tac x = x in spec)
apply safe
by simp_all
lemma sup_union:
"a \<in> supremum A \<Longrightarrow> b \<in> supremum B \<Longrightarrow> supremum {a, b} = supremum (A \<union> B)"
apply safe
apply (simp_all add: supremum_def upper_bound_def)
apply safe
apply auto
apply (subgoal_tac "(\<forall>x \<in> A \<union> B. x \<le> xa)")
by auto
lemma sup_singleton [simp]: "a \<in> supremum {a}"
by (simp add: supremum_def upper_bound_def)
lemma sup_union_singleton: "a \<in> supremum X \<Longrightarrow> supremum {a, b} = supremum (X \<union> {b})"
apply (rule_tac B = "{b}" in sup_union)
by simp_all
lemma sup_le_union [simp]: "a \<le> b \<Longrightarrow> supremum (A \<union> {a, b}) = supremum (A \<union> {b})"
apply (simp add: supremum_def upper_bound_def)
by auto
lemma sup_sup_union: "a \<in> supremum A \<Longrightarrow> b \<in> supremum (B \<union> {a}) \<Longrightarrow> b \<in> supremum (A \<union> B)"
apply (simp add: supremum_def upper_bound_def)
by auto
end
(*
context monoid_mult
begin
lemma "a ^ 2 = a * a"
by (simp add: power2_eq_square)
end
*)
lemma [simp]:
"n \<le> 2 ^ n"
apply (induct_tac n)
apply auto
apply (rule_tac y = "1 + 2 ^ n" in order_trans)
by auto
context pseudo_hoop_algebra
begin
lemma sup_le_union_2:
"a \<le> b \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> supremum A = supremum ((A - {a}) \<union> {b})"
apply (case_tac "supremum ((A - {a , b}) \<union> {a, b}) = supremum ((A - {a, b}) \<union> {b})")
apply (case_tac "((A - {a, b}) \<union> {a, b} = A) \<and> ((A - {a, b}) \<union> {b} = (A - {a}) \<union> {b})")
apply safe[1]
apply simp
apply simp
apply (erule notE)
apply safe [1]
apply (erule notE)
apply (rule sup_le_union)
by simp
lemma lemma_2_15_iii_0:
"1 \<in> supremum {a, b} \<Longrightarrow> 1 \<in> supremum {a ^ 2, b ^ 2}"
apply (frule_tac a = a in lemma_2_9_i)
apply simp
apply (frule_tac a = a and b = b in sup_union_singleton)
apply (subgoal_tac "supremum ({a * a, a * b} \<union> {b}) = supremum ({a * a, b})")
apply simp_all
apply (frule_tac a = b in lemma_2_9_i)
apply simp
apply (drule_tac a = b and A = "{b * (a * a), b * b}" and b = 1 and B = "{a * a}" in sup_sup_union)
apply simp
apply (case_tac "{a * a, b} = {b, a * a}")
apply simp
apply auto [1]
apply simp
apply (subgoal_tac "supremum {a * a, b * (a * a), b * b} = supremum {a * a, b * b}")
apply(simp add: power2_eq_square)
apply (case_tac "b * (a * a) = b * b")
apply auto[1]
apply (cut_tac A = "{a * a, b * (a * a), b * b}" and a = "b * (a * a)" and b = "a * a" in sup_le_union_2)
apply simp
apply simp
apply simp
apply (subgoal_tac "({a * a, b * (a * a), b * b} - {b * (a * a)} \<union> {a * a}) = {a * a, b * b}")
apply simp
apply auto[1]
apply (case_tac "a * a = a * b")
apply (subgoal_tac "{b, a * a, a * b} = {a * a, b}")
apply simp
apply auto[1]
apply (cut_tac A = "{b, a * a, a * b}" and a = "a * b" and b = "b" in sup_le_union_2)
apply simp
apply simp
apply simp
apply (subgoal_tac "{b, a * a, a * b} - {a * b} \<union> {b} = {a * a, b}")
apply simp
by auto
lemma [simp]: "m \<le> n \<Longrightarrow> a ^ n \<le> a ^ m"
apply (subgoal_tac "a ^ n = (a ^ m) * (a ^ (n-m))")
apply simp
apply (cut_tac a = a and m = "m" and n = "n - m" in power_add)
by simp
lemma [simp]: "a ^ (2 ^ n) \<le> a ^ n"
by simp
lemma lemma_2_15_iii_1: "1 \<in> supremum {a, b} \<Longrightarrow> 1 \<in> supremum {a ^ (2 ^ n), b ^ (2 ^ n)}"
apply (induct_tac n)
apply auto[1]
apply (drule drop_assumption)
apply (drule lemma_2_15_iii_0)
apply (subgoal_tac "\<forall>a . (a ^ (2::nat) ^ n)\<^sup>2 = a ^ (2::nat) ^ Suc n")
apply simp
apply safe
apply (cut_tac a = aa and m = "2 ^ n" and n = 2 in power_mult)
apply auto
apply (subgoal_tac "((2::nat) ^ n * (2::nat)) = ((2::nat) * (2::nat) ^ n)")
by simp_all
lemma lemma_2_15_iii:
"1 \<in> supremum {a, b} \<Longrightarrow> 1 \<in> supremum {a ^ n, b ^ n}"
apply (drule_tac n = n in lemma_2_15_iii_1)
apply (simp add: supremum_def upper_bound_def)
apply safe
apply (drule_tac x = x in spec)
apply safe
apply (rule_tac y = "a ^ n" in order_trans)
apply simp_all
apply (rule_tac y = "b ^ n" in order_trans)
by simp_all
end
end
diff --git a/thys/PseudoHoops/PseudoWaisbergAlgebra.thy b/thys/PseudoHoops/PseudoWaisbergAlgebra.thy
--- a/thys/PseudoHoops/PseudoWaisbergAlgebra.thy
+++ b/thys/PseudoHoops/PseudoWaisbergAlgebra.thy
@@ -1,463 +1,463 @@
section\<open>Pseudo Waisberg Algebra\<close>
theory PseudoWaisbergAlgebra
imports Operations
begin
class impl_lr_algebra = one + left_imp + right_imp +
assumes W1a [simp]: "1 l\<rightarrow> a = a"
and W1b [simp]: "1 r\<rightarrow> a = a"
and W2a: "(a l\<rightarrow> b) r\<rightarrow> b = (b l\<rightarrow> a) r\<rightarrow> a"
and W2b: "(b l\<rightarrow> a) r\<rightarrow> a = (b r\<rightarrow> a) l\<rightarrow> a"
and W2c: "(b r\<rightarrow> a) l\<rightarrow> a = (a r\<rightarrow> b) l\<rightarrow> b"
and W3a: "(a l\<rightarrow> b) l\<rightarrow> ((b l\<rightarrow> c) r\<rightarrow> (a l\<rightarrow> c)) = 1"
and W3b: "(a r\<rightarrow> b) r\<rightarrow> ((b r\<rightarrow> c) l\<rightarrow> (a r\<rightarrow> c)) = 1"
begin
lemma P1_a [simp]: "x l\<rightarrow> x = 1"
apply (cut_tac a = 1 and b = 1 and c = x in W3b)
by simp
lemma P1_b [simp]: "x r\<rightarrow> x = 1"
apply (cut_tac a = 1 and b = 1 and c = x in W3a)
by simp
lemma P2_a: "x l\<rightarrow> y = 1 \<Longrightarrow> y l\<rightarrow> x = 1 \<Longrightarrow> x = y"
apply (subgoal_tac "(y l\<rightarrow> x) r\<rightarrow> x = y")
apply simp
apply (subgoal_tac "(x l\<rightarrow> y) r\<rightarrow> y = y")
apply (unfold W2a)
by simp_all
lemma P2_b: "x r\<rightarrow> y = 1 \<Longrightarrow> y r\<rightarrow> x = 1 \<Longrightarrow> x = y"
apply (subgoal_tac "(y r\<rightarrow> x) l\<rightarrow> x = y")
apply simp
apply (subgoal_tac "(x r\<rightarrow> y) l\<rightarrow> y = y")
apply (unfold W2c)
by simp_all
lemma P2_c: "x l\<rightarrow> y = 1 \<Longrightarrow> y r\<rightarrow> x = 1 \<Longrightarrow> x = y"
apply (subgoal_tac "(y r\<rightarrow> x) l\<rightarrow> x = y")
apply simp
apply (subgoal_tac "(x l\<rightarrow> y) r\<rightarrow> y = y")
apply (unfold W2b) [1]
apply (unfold W2c) [1]
by simp_all
lemma P3_a: "(x l\<rightarrow> 1) r\<rightarrow> 1 = 1"
apply (unfold W2a)
by simp
lemma P3_b: "(x r\<rightarrow> 1) l\<rightarrow> 1 = 1"
apply (unfold W2c)
by simp
lemma P4_a [simp]: "x l\<rightarrow> 1 = 1"
apply (subgoal_tac "x l\<rightarrow> ((x l\<rightarrow> 1) r\<rightarrow> 1) = 1")
apply (simp add: P3_a)
apply (cut_tac a = 1 and b = x and c = 1 in W3a)
by simp
lemma P4_b [simp]: "x r\<rightarrow> 1 = 1"
apply (subgoal_tac "x r\<rightarrow> ((x r\<rightarrow> 1) l\<rightarrow> 1) = 1")
apply (simp add: P3_b)
apply (cut_tac a = 1 and b = x and c = 1 in W3b)
by simp
lemma P5_a: "x l\<rightarrow> y = 1 \<Longrightarrow> y l\<rightarrow> z = 1 \<Longrightarrow> x l\<rightarrow> z = 1"
apply (cut_tac a = x and b = y and c = z in W3a)
by simp
lemma P5_b: "x r\<rightarrow> y = 1 \<Longrightarrow> y r\<rightarrow> z = 1 \<Longrightarrow> x r\<rightarrow> z = 1"
apply (cut_tac a = x and b = y and c = z in W3b)
by simp
lemma P6_a: "x l\<rightarrow> (y r\<rightarrow> x) = 1"
apply (cut_tac a = y and b = 1 and c = x in W3b)
by simp
lemma P6_b: "x r\<rightarrow> (y l\<rightarrow> x) = 1"
apply (cut_tac a = y and b = 1 and c = x in W3a)
by simp
lemma P7: "(x l\<rightarrow> (y r\<rightarrow> z) = 1) = (y r\<rightarrow> (x l\<rightarrow> z) = 1)"
proof
fix x y z assume A: "x l\<rightarrow> y r\<rightarrow> z = 1" show "y r\<rightarrow> x l\<rightarrow> z = 1"
apply (rule_tac y = "(z r\<rightarrow> y) l\<rightarrow> y" in P5_b)
apply (simp add: P6_b)
apply (unfold W2c)
apply (subgoal_tac "(x l\<rightarrow> (y r\<rightarrow> z)) l\<rightarrow> (((y r\<rightarrow> z) l\<rightarrow> z) r\<rightarrow> x l\<rightarrow> z) = 1")
apply (unfold A) [1]
apply simp
by (simp add: W3a)
next
fix x y z assume A: "y r\<rightarrow> x l\<rightarrow> z = 1" show "x l\<rightarrow> y r\<rightarrow> z = 1"
apply (rule_tac y = "(z l\<rightarrow> x) r\<rightarrow> x" in P5_a)
apply (simp add: P6_a)
apply (unfold W2a)
apply (subgoal_tac "(y r\<rightarrow> x l\<rightarrow> z) r\<rightarrow> (((x l\<rightarrow> z) r\<rightarrow> z) l\<rightarrow> y r\<rightarrow> z) = 1")
apply (unfold A) [1]
apply simp
by (simp add: W3b)
qed
lemma P8_a: "(x l\<rightarrow> y) r\<rightarrow> ((z l\<rightarrow> x) l\<rightarrow> (z l\<rightarrow> y)) = 1"
by (simp add: W3a P7 [THEN sym])
lemma P8_b: "(x r\<rightarrow> y) l\<rightarrow> ((z r\<rightarrow> x) r\<rightarrow> (z r\<rightarrow> y)) = 1"
by (simp add: W3b P7)
lemma P9: "x l\<rightarrow> (y r\<rightarrow> z) = y r\<rightarrow> (x l\<rightarrow> z)"
apply (rule P2_c)
apply (subst P7)
apply (rule_tac y = "(z r\<rightarrow> y) l\<rightarrow> y" in P5_b)
apply (simp add: P6_b)
apply (subst W2c)
apply (rule P8_a)
apply (subst P7 [THEN sym])
apply (rule_tac y = "(z l\<rightarrow> x) r\<rightarrow> x" in P5_a)
apply (simp add: P6_a)
apply (subst W2a)
by (simp add: P8_b)
definition
"lesseq_a a b = (a l\<rightarrow> b = 1)"
definition
"less_a a b = (lesseq_a a b \<and> \<not> lesseq_a b a)"
definition
"lesseq_b a b = (a r\<rightarrow> b = 1)"
definition
"less_b a b = (lesseq_b a b \<and> \<not> lesseq_b b a)"
definition
"sup_a a b = (a l\<rightarrow> b) r\<rightarrow> b"
end
sublocale impl_lr_algebra < order_a:order lesseq_a less_a
apply unfold_locales
apply (simp add: less_a_def)
apply (simp_all add: lesseq_a_def)
apply (rule P5_a)
apply simp_all
apply (rule P2_a)
by simp_all
sublocale impl_lr_algebra < order_b:order lesseq_b less_b
apply unfold_locales
apply (simp add: less_b_def)
apply (simp_all add: lesseq_b_def)
apply (rule P5_b)
apply simp_all
apply (rule P2_b)
by simp_all
sublocale impl_lr_algebra < slattice_a:semilattice_sup sup_a lesseq_a less_a
apply unfold_locales
apply (simp_all add: lesseq_a_def sup_a_def)
apply (simp add: P9)
apply (simp add: P9)
apply (subst W2a)
apply (subgoal_tac "((z l\<rightarrow> y) r\<rightarrow> y) l\<rightarrow> ((y l\<rightarrow> x) r\<rightarrow> x) = 1")
apply simp
apply (subgoal_tac "((z l\<rightarrow> y) r\<rightarrow> y) l\<rightarrow> ((x l\<rightarrow> y) r\<rightarrow> y) = 1")
apply (simp add: W2a)
apply (subgoal_tac "((z l\<rightarrow> y) r\<rightarrow> y) l\<rightarrow> (x l\<rightarrow> y) r\<rightarrow> y = ((x l\<rightarrow> y) r\<rightarrow> (z l\<rightarrow> y)) r\<rightarrow> (((z l\<rightarrow> y) r\<rightarrow> y) l\<rightarrow> (x l\<rightarrow> y) r\<rightarrow> y)")
apply (simp add: W3b)
apply (subgoal_tac "(x l\<rightarrow> y) r\<rightarrow> z l\<rightarrow> y = 1")
apply simp
apply (cut_tac a = z and b = x and c = y in W3a)
by simp
sublocale impl_lr_algebra < slattice_b:semilattice_sup sup_a lesseq_b less_b
apply unfold_locales
apply (simp_all add: lesseq_b_def sup_a_def)
apply (simp_all add: W2b)
apply (simp add: P9 [THEN sym])
apply (simp add: P9 [THEN sym])
apply (subst W2c)
apply (subgoal_tac "((z r\<rightarrow> y) l\<rightarrow> y) r\<rightarrow> ((y r\<rightarrow> x) l\<rightarrow> x) = 1")
apply simp
apply (subgoal_tac "((z r\<rightarrow> y) l\<rightarrow> y) r\<rightarrow> ((x r\<rightarrow> y) l\<rightarrow> y) = 1")
apply (simp add: W2c)
apply (subgoal_tac "((z r\<rightarrow> y) l\<rightarrow> y) r\<rightarrow> (x r\<rightarrow> y) l\<rightarrow> y = ((x r\<rightarrow> y) l\<rightarrow> (z r\<rightarrow> y)) l\<rightarrow> (((z r\<rightarrow> y) l\<rightarrow> y) r\<rightarrow> (x r\<rightarrow> y) l\<rightarrow> y)")
apply (simp add: W3a)
apply (subgoal_tac "(x r\<rightarrow> y) l\<rightarrow> z r\<rightarrow> y = 1")
apply simp
apply (cut_tac a = z and b = x and c = y in W3b)
by simp
context impl_lr_algebra
begin
lemma lesseq_a_b: "lesseq_b = lesseq_a"
apply (simp add: fun_eq_iff)
apply clarify
apply (cut_tac x = x and y = xa in slattice_a.le_iff_sup)
apply (cut_tac x = x and y = xa in slattice_b.le_iff_sup)
by simp
lemma P10: "(a l\<rightarrow> b = 1) = (a r\<rightarrow> b = 1)"
apply (cut_tac lesseq_a_b)
by (simp add: fun_eq_iff lesseq_a_def lesseq_b_def)
end
class one_ord = one + ord
class impl_lr_ord_algebra = impl_lr_algebra + one_ord +
assumes
order: "a \<le> b = (a l\<rightarrow> b = 1)"
and
strict: "a < b = (a \<le> b \<and> \<not> b \<le> a)"
begin
lemma order_l: "(a \<le> b) = (a l\<rightarrow> b = 1)"
by (simp add: order)
lemma order_r: "(a \<le> b) = (a r\<rightarrow> b = 1)"
by (simp add: order P10)
lemma P11_a: "a \<le> b l\<rightarrow> a"
by (simp add: order_r P6_b)
lemma P11_b: "a \<le> b r\<rightarrow> a"
by (simp add: order_l P6_a)
lemma P12: "(a \<le> b l\<rightarrow> c) = (b \<le> a r\<rightarrow> c)"
apply (subst order_r)
apply (subst order_l)
by (simp add: P7)
lemma P13_a: "a \<le> b \<Longrightarrow> b l\<rightarrow> c \<le> a l\<rightarrow> c"
apply (subst order_r)
apply (simp add: order_l)
apply (cut_tac a = a and b = b and c = c in W3a)
by simp
lemma P13_b: "a \<le> b \<Longrightarrow> b r\<rightarrow> c \<le> a r\<rightarrow> c"
apply (subst order_l)
apply (simp add: order_r)
apply (cut_tac a = a and b = b and c = c in W3b)
by simp
lemma P14_a: "a \<le> b \<Longrightarrow> c l\<rightarrow> a \<le> c l\<rightarrow> b"
apply (simp add: order_l)
apply (cut_tac x = a and y = b and z = c in P8_a)
by simp
lemma P14_b: "a \<le> b \<Longrightarrow> c r\<rightarrow> a \<le> c r\<rightarrow> b"
apply (simp add: order_r)
apply (cut_tac x = a and y = b and z = c in P8_b)
by simp
subclass order
apply (subgoal_tac "(\<le>) = lesseq_a \<and> (<) = less_a")
apply simp
apply unfold_locales
apply safe
by (simp_all add: fun_eq_iff lesseq_a_def less_a_def order_l strict)
end
class one_zero_uminus = one + zero + left_uminus + right_uminus
class impl_neg_lr_algebra = impl_lr_ord_algebra + one_zero_uminus +
assumes
W4: "-l 1 = -r 1"
and W5a: "(-l a r\<rightarrow> -l b) l\<rightarrow> (b l\<rightarrow> a) = 1"
and W5b: "(-r a l\<rightarrow> -r b) l\<rightarrow> (b r\<rightarrow> a) = 1"
and zero_def: "0 = -l 1"
begin
lemma zero_r_def: "0 = -r 1"
by (simp add: zero_def W4)
lemma C1_a [simp]: "(-l x r\<rightarrow> 0) l\<rightarrow> x = 1"
apply (unfold zero_def)
apply (cut_tac a = x and b = 1 in W5a)
by simp
lemma C1_b [simp]: "(-r x l\<rightarrow> 0) r\<rightarrow> x = 1"
apply (unfold zero_r_def)
apply (cut_tac a = x and b = 1 in W5b)
by (simp add: P10)
lemma C2_b [simp]: "0 r\<rightarrow> x = 1"
apply (cut_tac x= "-r x l\<rightarrow> 0" and y = x and z = 0 in P8_b)
by (simp add: P6_b)
lemma C2_a [simp]: "0 l\<rightarrow> x = 1"
by (simp add: P10)
lemma C3_a: "x l\<rightarrow> 0 = -l x"
proof -
have A: "-l x l\<rightarrow> (x l\<rightarrow> 0) = 1"
apply (cut_tac x = "-l x" and y = "-l (-r 1)" in P6_a)
apply (cut_tac a = "-r 1" and b = "x" in W5a)
apply (unfold zero_r_def)
apply (rule_tac y = " -l (-r (1::'a)) r\<rightarrow> -l x" in P5_a)
by simp_all
have B: "(x l\<rightarrow> 0) r\<rightarrow> -l x = 1"
apply (cut_tac a = "-l x r\<rightarrow> 0" and b = x and c = 0 in W3a)
apply simp
apply (cut_tac b = "-l x" and a = 0 in W2c)
by simp
show "x l\<rightarrow> 0 = -l x"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: order_r B)
by (simp add: order_l A)
qed
lemma C3_b: "x r\<rightarrow> 0 = -r x"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: order_l)
apply (cut_tac x = x in C1_b)
apply (cut_tac a = "-r x l\<rightarrow> 0" and b = x and c = 0 in W3b)
apply simp
apply (cut_tac b = "-r x" and a = 0 in W2a)
apply simp
apply (cut_tac x = "-r x" and y = "-r (-l 1)" in P6_b)
apply (cut_tac a = "-l 1" and b = "x" in W5b)
apply (unfold zero_def order_r)
apply (rule_tac y = " -r (-l (1::'a)) l\<rightarrow> -r x" in P5_b)
by (simp_all add: P10)
lemma C4_a [simp]: "-r (-l x) = x"
apply (unfold C3_b [THEN sym] C3_a [THEN sym])
apply (subst W2a)
by simp
lemma C4_b [simp]: "-l (-r x) = x"
apply (unfold C3_b [THEN sym] C3_a [THEN sym])
apply (subst W2c)
by simp
lemma C5_a: "-r x l\<rightarrow> -r y = y r\<rightarrow> x"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: order_l W5b)
apply (cut_tac a = "-r y" and b = "-r x" in W5a)
by (simp add: order_l)
lemma C5_b: "-l x r\<rightarrow> -l y = y l\<rightarrow> x"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: order_l W5a)
apply (cut_tac a = "-l y" and b = "-l x" in W5b)
by (simp add: order_l)
lemma C6: "-r x l\<rightarrow> y = -l y r\<rightarrow> x"
apply (cut_tac x = x and y = "-l y" in C5_a)
by simp
lemma C7_a: "(x \<le> y) = (-l y \<le> -l x)"
apply (subst order_l)
apply (subst order_r)
by (simp add: C5_b)
lemma C7_b: "(x \<le> y) = (-r y \<le> -r x)"
apply (subst order_r)
apply (subst order_l)
by (simp add: C5_a)
end
class pseudo_wajsberg_algebra = impl_neg_lr_algebra +
assumes
W6: "-r (a l\<rightarrow> -l b) = -l (b r\<rightarrow> -r a)"
begin
definition
"mult a b = -r (a l\<rightarrow> -l b)"
definition
"inf_a a b = -l (a r\<rightarrow> -r (a l\<rightarrow> b))"
definition
"inf_b a b = -r (b l\<rightarrow> -l (b r\<rightarrow> a))"
end
sublocale pseudo_wajsberg_algebra < slattice_inf_a:semilattice_inf inf_a "(\<le>)" "(<)"
apply unfold_locales
apply (simp_all add: inf_a_def)
apply (subst C7_b)
apply (simp add: order_l P9 C5_a P10 [THEN sym] P6_a)
apply (subst C7_b)
apply (simp add: order_l P9 C5_a P10 [THEN sym] P6_a)
apply (subst W6 [THEN sym])
apply (subst C7_a)
apply simp
proof -
fix x y z
assume A: "x \<le> y"
assume B: "x \<le> z"
have C: "x l\<rightarrow> y = 1" by (simp add: order_l [THEN sym] A)
have E: "((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> -l x = ((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((x l\<rightarrow> y) l\<rightarrow> -l x)"
by (simp add: C)
have F: "((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((x l\<rightarrow> y) l\<rightarrow> -l x) = ((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((-l y r\<rightarrow> -l x) l\<rightarrow> -l x)"
by (simp add: C5_b)
have G: "((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((-l y r\<rightarrow> -l x) l\<rightarrow> -l x) = ((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((-l x r\<rightarrow> -l y) l\<rightarrow> -l y)"
by (simp add: W2c)
have H: "((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((-l x r\<rightarrow> -l y) l\<rightarrow> -l y) = ((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((y l\<rightarrow> x) l\<rightarrow> -l y)"
by (simp add: C5_b)
have I: "((y l\<rightarrow> z) l\<rightarrow> -l y) l\<rightarrow> ((y l\<rightarrow> x) l\<rightarrow> -l y) = 1"
apply (simp add: order_l [THEN sym] P14_a)
apply (rule P13_a)
apply (rule P14_a)
by (simp add: B)
show "(y l\<rightarrow> z) l\<rightarrow> -l y \<le> -l x"
by (simp add: order_l E F G H I)
next
qed
sublocale pseudo_wajsberg_algebra < slattice_inf_b:semilattice_inf inf_b "(\<le>)" "(<)"
apply unfold_locales
apply (simp_all add: inf_b_def)
apply (subst C7_a)
apply (simp add: order_r P9 [THEN sym] C5_b P10 P6_b)
apply (subst C7_a)
apply (simp add: order_r P9 [THEN sym] C5_b P10 P6_b)
apply (subst W6)
apply (subst C7_b)
apply simp
proof -
fix x y z
assume A: "x \<le> y"
assume B: "x \<le> z"
have C: "x r\<rightarrow> z = 1" by (simp add: order_r [THEN sym] B)
have E: "((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> -r x = ((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((x r\<rightarrow> z) r\<rightarrow> -r x)"
by (simp add: C)
have F: "((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((x r\<rightarrow> z) r\<rightarrow> -r x) = ((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((-r z l\<rightarrow> -r x) r\<rightarrow> -r x)"
by (simp add: C5_a)
have G: "((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((-r z l\<rightarrow> -r x) r\<rightarrow> -r x) = ((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((-r x l\<rightarrow> -r z) r\<rightarrow> -r z)"
by (simp add: W2a)
have H: "((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((-r x l\<rightarrow> -r z) r\<rightarrow> -r z) = ((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((z r\<rightarrow> x) r\<rightarrow> -r z)"
by (simp add: C5_a)
have I: "((z r\<rightarrow> y) r\<rightarrow> -r z) r\<rightarrow> ((z r\<rightarrow> x) r\<rightarrow> -r z) = 1"
apply (simp add: order_r [THEN sym])
apply (rule P13_b)
apply (rule P14_b)
by (simp add: A)
show "(z r\<rightarrow> y) r\<rightarrow> -r z \<le> -r x"
by (simp add: order_r E F G H I)
next
qed
context pseudo_wajsberg_algebra
begin
lemma inf_a_b: "inf_a = inf_b"
apply (simp add: fun_eq_iff)
apply clarify
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
end
end
diff --git a/thys/PseudoHoops/SpecialPseudoHoops.thy b/thys/PseudoHoops/SpecialPseudoHoops.thy
--- a/thys/PseudoHoops/SpecialPseudoHoops.thy
+++ b/thys/PseudoHoops/SpecialPseudoHoops.thy
@@ -1,892 +1,885 @@
section\<open>Some Classes of Pseudo-Hoops\<close>
theory SpecialPseudoHoops
imports PseudoHoopFilters PseudoWaisbergAlgebra
begin
class cancel_pseudo_hoop_algebra = pseudo_hoop_algebra +
assumes mult_cancel_left: "a * b = a * c \<Longrightarrow> b = c"
and mult_cancel_right: "b * a = c * a \<Longrightarrow> b = c"
begin
lemma cancel_left_a: "b l\<rightarrow> (a * b) = a"
apply (rule_tac a = b in mult_cancel_right)
apply (subst inf_l_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
lemma cancel_right_a: "b r\<rightarrow> (b * a) = a"
apply (rule_tac a = b in mult_cancel_left)
apply (subst inf_r_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
end
class cancel_pseudo_hoop_algebra_2 = pseudo_hoop_algebra +
assumes cancel_left: "b l\<rightarrow> (a * b) = a"
and cancel_right: "b r\<rightarrow> (b * a) = a"
begin
subclass cancel_pseudo_hoop_algebra
apply unfold_locales
apply (subgoal_tac "b = a r\<rightarrow> (a * b) \<and> a r\<rightarrow> (a * b) = a r\<rightarrow> (a * c) \<and> a r\<rightarrow> (a * c) = c")
apply simp
apply (rule conjI)
apply (subst cancel_right)
apply simp
apply (rule conjI)
apply simp
apply (subst cancel_right)
apply simp
apply (subgoal_tac "b = a l\<rightarrow> (b * a) \<and> a l\<rightarrow> (b * a) = a l\<rightarrow> (c * a) \<and> a l\<rightarrow> (c * a) = c")
apply simp
apply (rule conjI)
apply (subst cancel_left)
apply simp
apply (rule conjI)
apply simp
apply (subst cancel_left)
by simp
end
context cancel_pseudo_hoop_algebra
begin
lemma lemma_4_2_i: "a l\<rightarrow> b = (a * c) l\<rightarrow> (b * c)"
apply (subgoal_tac "a l\<rightarrow> b = a l\<rightarrow> (c l\<rightarrow> (b * c)) \<and> a l\<rightarrow> (c l\<rightarrow> (b * c)) = (a * c) l\<rightarrow> (b * c)")
apply simp
apply (rule conjI)
apply (simp add: cancel_left_a)
by (simp add: left_impl_ded)
lemma lemma_4_2_ii: "a r\<rightarrow> b = (c * a) r\<rightarrow> (c * b)"
apply (subgoal_tac "a r\<rightarrow> b = a r\<rightarrow> (c r\<rightarrow> (c * b)) \<and> a r\<rightarrow> (c r\<rightarrow> (c * b)) = (c * a) r\<rightarrow> (c * b)")
apply simp
apply (rule conjI)
apply (simp add: cancel_right_a)
by (simp add: right_impl_ded)
lemma lemma_4_2_iii: "(a * c \<le> b * c) = (a \<le> b)"
by (simp add: left_lesseq lemma_4_2_i [THEN sym])
lemma lemma_4_2_iv: "(c * a \<le> c * b) = (a \<le> b)"
by (simp add: right_lesseq lemma_4_2_ii [THEN sym])
end
class wajsberg_pseudo_hoop_algebra = pseudo_hoop_algebra +
assumes wajsberg1: "(a l\<rightarrow> b) r\<rightarrow> b = (b l\<rightarrow> a) r\<rightarrow> a"
and wajsberg2: "(a r\<rightarrow> b) l\<rightarrow> b = (b r\<rightarrow> a) l\<rightarrow> a"
context wajsberg_pseudo_hoop_algebra
begin
lemma lemma_4_3_i_a: "a \<squnion>1 b = (a l\<rightarrow> b) r\<rightarrow> b"
by (simp add: sup1_def wajsberg1)
lemma lemma_4_3_i_b: "a \<squnion>1 b = (b l\<rightarrow> a) r\<rightarrow> a"
by (simp add: sup1_def wajsberg1)
lemma lemma_4_3_ii_a: "a \<squnion>2 b = (a r\<rightarrow> b) l\<rightarrow> b"
by (simp add: sup2_def wajsberg2)
lemma lemma_4_3_ii_b: "a \<squnion>2 b = (b r\<rightarrow> a) l\<rightarrow> a"
by (simp add: sup2_def wajsberg2)
end
sublocale wajsberg_pseudo_hoop_algebra < lattice1:pseudo_hoop_lattice_b "(\<squnion>1)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
apply (simp add: lemma_4_3_i_a)
by (simp add: lemma_2_5_13_b lemma_2_5_13_a)
class zero_one = zero + one
class bounded_wajsberg_pseudo_hoop_algebra = zero_one + wajsberg_pseudo_hoop_algebra +
assumes zero_smallest [simp]: "0 \<le> a"
begin
end
sublocale wajsberg_pseudo_hoop_algebra < lattice2:pseudo_hoop_lattice_b "(\<squnion>2)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
apply (simp add: lemma_4_3_ii_a)
by (simp add: lemma_2_5_13_b lemma_2_5_13_a)
lemma (in wajsberg_pseudo_hoop_algebra) sup1_eq_sup2: "(\<squnion>1) = (\<squnion>2)"
apply (simp add: fun_eq_iff)
apply safe
apply (cut_tac a = x and b = xa in lattice1.supremum_pair)
apply (cut_tac a = x and b = xa in lattice2.supremum_pair)
by blast
context bounded_wajsberg_pseudo_hoop_algebra
begin
definition
"negl a = a l\<rightarrow> 0"
definition
"negr a = a r\<rightarrow> 0"
lemma [simp]: "0 l\<rightarrow> a = 1"
by (simp add: order [THEN sym])
lemma [simp]: "0 r\<rightarrow> a = 1"
by (simp add: order_r [THEN sym])
end
sublocale bounded_wajsberg_pseudo_hoop_algebra < wajsberg: pseudo_wajsberg_algebra "1" "(l\<rightarrow>)" "(r\<rightarrow>)" "(\<le>)" "(<)" "0" "negl" "negr"
apply unfold_locales
apply simp_all
apply (simp add: lemma_4_3_i_a [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (simp add: lemma_4_3_i_a [THEN sym] lemma_4_3_ii_a [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (simp add: lemma_4_3_i_a [THEN sym] lemma_4_3_ii_a [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (subst left_lesseq [THEN sym])
apply (simp add: lemma_2_5_16)
apply (subst right_lesseq [THEN sym])
apply (simp add: lemma_2_5_17)
apply (simp add: left_lesseq)
apply (simp add: less_def)
apply (simp_all add: negl_def negr_def)
apply (subst left_lesseq [THEN sym])
apply (subgoal_tac "b l\<rightarrow> a = ((b l\<rightarrow> 0) r\<rightarrow> 0) l\<rightarrow> ((a l\<rightarrow> 0) r\<rightarrow> 0)")
apply (simp add: lemma_2_5_17)
apply (subst wajsberg1)
apply simp
apply (subst wajsberg1)
apply simp
apply (subst left_lesseq [THEN sym])
apply (subgoal_tac "b r\<rightarrow> a = ((b r\<rightarrow> 0) l\<rightarrow> 0) r\<rightarrow> ((a r\<rightarrow> 0) l\<rightarrow> 0)")
apply (simp add: lemma_2_5_16)
apply (subst wajsberg2)
apply simp
apply (subst wajsberg2)
apply simp
apply (simp add: left_impl_ded [THEN sym])
apply (simp add: right_impl_ded [THEN sym])
apply (simp add: lemma_4_3_i_a [THEN sym] lemma_4_3_ii_a [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
context pseudo_wajsberg_algebra
begin
lemma "class.bounded_wajsberg_pseudo_hoop_algebra mult inf_a (l\<rightarrow>) (\<le>) (<) 1 (r\<rightarrow>) (0::'a)"
apply unfold_locales
apply (simp add: inf_a_def mult_def W6)
apply (simp add: strict)
apply (simp_all add: mult_def order_l strict)
apply (simp add: zero_def [THEN sym] C3_a)
apply (simp add: W6 inf_a_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (simp add: C6 P9 [THEN sym] C5_b)
apply (simp add: inf_b_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (simp add: inf_b_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (simp add: W6)
apply (simp add: C6 [THEN sym])
apply (simp add: P9 C5_a)
apply (simp add: inf_b_def [THEN sym])
apply (simp add: W6 inf_a_def [THEN sym])
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
apply (simp add: W2a)
by (simp add: W2c)
end
class basic_pseudo_hoop_algebra = pseudo_hoop_algebra +
assumes B1: "(a l\<rightarrow> b) l\<rightarrow> c \<le> ((b l\<rightarrow> a) l\<rightarrow> c) l\<rightarrow> c"
and B2: "(a r\<rightarrow> b) r\<rightarrow> c \<le> ((b r\<rightarrow> a) r\<rightarrow> c) r\<rightarrow> c"
begin
lemma lemma_4_5_i: "(a l\<rightarrow> b) \<squnion>1 (b l\<rightarrow> a) = 1"
apply (cut_tac a = a and b = b and c = "(a l\<rightarrow> b) \<squnion>1 (b l\<rightarrow> a)" in B1)
apply (subgoal_tac "(a l\<rightarrow> b) l\<rightarrow> (a l\<rightarrow> b) \<squnion>1 (b l\<rightarrow> a) = 1 \<and> ((b l\<rightarrow> a) l\<rightarrow> (a l\<rightarrow> b) \<squnion>1 (b l\<rightarrow> a)) = 1")
apply (erule conjE)
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply simp
apply safe
apply (subst left_lesseq [THEN sym])
apply simp
apply (subst left_lesseq [THEN sym])
by simp
lemma lemma_4_5_ii: "(a r\<rightarrow> b) \<squnion>2 (b r\<rightarrow> a) = 1"
apply (cut_tac a = a and b = b and c = "(a r\<rightarrow> b) \<squnion>2 (b r\<rightarrow> a)" in B2)
apply (subgoal_tac "(a r\<rightarrow> b) r\<rightarrow> (a r\<rightarrow> b) \<squnion>2 (b r\<rightarrow> a) = 1 \<and> ((b r\<rightarrow> a) r\<rightarrow> (a r\<rightarrow> b) \<squnion>2 (b r\<rightarrow> a)) = 1")
apply (erule conjE)
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply simp
apply safe
apply (subst right_lesseq [THEN sym])
apply simp
apply (subst right_lesseq [THEN sym])
by simp
lemma lemma_4_5_iii: "a l\<rightarrow> b = (a \<squnion>1 b) l\<rightarrow> b"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule_tac y = "((a l\<rightarrow> b) r\<rightarrow> b) l\<rightarrow> b" in order_trans)
apply (rule lemma_2_10_26)
apply (rule lemma_2_5_13_a)
apply (simp add: sup1_def)
apply (rule lemma_2_5_13_a)
by simp
lemma lemma_4_5_iv: "a r\<rightarrow> b = (a \<squnion>2 b) r\<rightarrow> b"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (rule_tac y = "((a r\<rightarrow> b) l\<rightarrow> b) r\<rightarrow> b" in order_trans)
apply (rule lemma_2_10_24)
apply (rule lemma_2_5_13_b)
apply (simp add: sup2_def)
apply (rule lemma_2_5_13_b)
by simp
lemma lemma_4_5_v: "(a \<squnion>1 b) l\<rightarrow> c = (a l\<rightarrow> c) \<sqinter> (b l\<rightarrow> c)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply safe
apply (rule lemma_2_5_13_a)
apply simp
apply (rule lemma_2_5_13_a)
apply simp
apply (subst right_lesseq)
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (rule_tac y = "(a l\<rightarrow> b) l\<rightarrow> ((a l\<rightarrow> c) \<sqinter> (b l\<rightarrow> c) r\<rightarrow> a \<squnion>1 b l\<rightarrow> c)" in order_trans)
apply (subst left_residual [THEN sym])
apply simp
apply (subst lemma_4_5_iii)
apply (subst right_residual [THEN sym])
apply (subst left_residual [THEN sym])
apply (rule_tac y = "b \<sqinter> c" in order_trans)
apply (subst (2) inf_l_def)
apply (rule_tac y = "((a l\<rightarrow> c) \<sqinter> (b l\<rightarrow> c)) * ((a \<squnion>1 b) \<sqinter> b)" in order_trans)
apply (subst (3) inf_l_def)
apply (simp add: mult.assoc)
apply (subgoal_tac "(a \<squnion>1 b \<sqinter> b) = b")
apply simp
- apply (rule antisym, simp)
+ apply (rule order.antisym, simp)
apply simp
apply simp
apply (rule_tac y = "((b l\<rightarrow> a) l\<rightarrow> ((a l\<rightarrow> c) \<sqinter> (b l\<rightarrow> c) r\<rightarrow> a \<squnion>1 b l\<rightarrow> c)) l\<rightarrow> ((a l\<rightarrow> c) \<sqinter> (b l\<rightarrow> c) r\<rightarrow> a \<squnion>1 b l\<rightarrow> c)" in order_trans)
apply (rule B1)
apply (subgoal_tac "(b l\<rightarrow> a) l\<rightarrow> ((a l\<rightarrow> c) \<sqinter> (b l\<rightarrow> c) r\<rightarrow> a \<squnion>1 b l\<rightarrow> c) = 1")
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (subst left_residual [THEN sym])
apply simp
apply (subst lemma_4_5_iii)
apply (subst right_residual [THEN sym])
apply (subst left_residual [THEN sym])
apply (rule_tac y = "a \<sqinter> c" in order_trans)
apply (subst (2) inf_l_def)
apply (rule_tac y = "((a l\<rightarrow> c) \<sqinter> (b l\<rightarrow> c)) * ((a \<squnion>1 b) \<sqinter> a)" in order_trans)
apply (subst (3) inf_l_def)
apply (subst sup1.sup_comute1)
apply (simp add: mult.assoc)
apply (subgoal_tac "(a \<squnion>1 b \<sqinter> a) = a")
apply simp
- apply (rule antisym, simp)
+ apply (rule order.antisym, simp)
apply simp
by simp
lemma lemma_4_5_vi: "(a \<squnion>2 b) r\<rightarrow> c = (a r\<rightarrow> c) \<sqinter> (b r\<rightarrow> c)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply safe
apply (rule lemma_2_5_13_b)
apply simp
apply (rule lemma_2_5_13_b)
apply simp
apply (subst left_lesseq)
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (rule_tac y = "(a r\<rightarrow> b) r\<rightarrow> ((a r\<rightarrow> c) \<sqinter> (b r\<rightarrow> c) l\<rightarrow> a \<squnion>2 b r\<rightarrow> c)" in order_trans)
apply (subst right_residual [THEN sym])
apply simp
apply (subst lemma_4_5_iv)
apply (subst left_residual [THEN sym])
apply (subst right_residual [THEN sym])
apply (rule_tac y = "b \<sqinter> c" in order_trans)
apply (subst (2) inf_r_def)
apply (rule_tac y = "((a \<squnion>2 b) \<sqinter> b) * ((a r\<rightarrow> c) \<sqinter> (b r\<rightarrow> c))" in order_trans)
apply (subst (2) inf_r_def)
apply (simp add: mult.assoc)
apply (subgoal_tac "(a \<squnion>2 b \<sqinter> b) = b")
apply simp
- apply (rule antisym, simp)
+ apply (rule order.antisym, simp)
apply simp
apply simp
apply (rule_tac y = "((b r\<rightarrow> a) r\<rightarrow> ((a r\<rightarrow> c) \<sqinter> (b r\<rightarrow> c) l\<rightarrow> a \<squnion>2 b r\<rightarrow> c)) r\<rightarrow> ((a r\<rightarrow> c) \<sqinter> (b r\<rightarrow> c) l\<rightarrow> a \<squnion>2 b r\<rightarrow> c)" in order_trans)
apply (rule B2)
apply (subgoal_tac "(b r\<rightarrow> a) r\<rightarrow> ((a r\<rightarrow> c) \<sqinter> (b r\<rightarrow> c) l\<rightarrow> a \<squnion>2 b r\<rightarrow> c) = 1")
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (subst right_residual [THEN sym])
apply simp
apply (subst lemma_4_5_iv)
apply (subst left_residual [THEN sym])
apply (subst right_residual [THEN sym])
apply (rule_tac y = "a \<sqinter> c" in order_trans)
apply (subst (2) inf_r_def)
apply (rule_tac y = "((a \<squnion>2 b) \<sqinter> a) * ((a r\<rightarrow> c) \<sqinter> (b r\<rightarrow> c))" in order_trans)
apply (subst (2) inf_r_def)
apply (subst (2) sup2.sup_comute)
apply (simp add: mult.assoc)
apply (subgoal_tac "(a \<squnion>2 b \<sqinter> a) = a")
apply simp
- apply (rule antisym, simp)
+ apply (rule order.antisym, simp)
apply simp
by simp
lemma lemma_4_5_a: "a \<le> c \<Longrightarrow> b \<le> c \<Longrightarrow> a \<squnion>1 b \<le> c"
apply (subst left_lesseq)
apply (subst lemma_4_5_v)
by simp
lemma lemma_4_5_b: "a \<le> c \<Longrightarrow> b \<le> c \<Longrightarrow> a \<squnion>2 b \<le> c"
apply (subst right_lesseq)
apply (subst lemma_4_5_vi)
by simp
lemma lemma_4_5: "a \<squnion>1 b = a \<squnion>2 b"
- apply (rule antisym)
+ apply (rule order.antisym)
by (simp_all add: lemma_4_5_a lemma_4_5_b)
end
sublocale basic_pseudo_hoop_algebra < basic_lattice:lattice "(\<sqinter>)" "(\<le>)" "(<)" "(\<squnion>1)"
apply unfold_locales
by (simp_all add: lemma_4_5_a)
context pseudo_hoop_lattice begin end
sublocale basic_pseudo_hoop_algebra < pseudo_hoop_lattice "(\<squnion>1)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
by (simp_all add: basic_lattice.sup_assoc)
class sup_assoc_pseudo_hoop_algebra = pseudo_hoop_algebra +
assumes sup1_assoc: "a \<squnion>1 (b \<squnion>1 c) = (a \<squnion>1 b) \<squnion>1 c"
and sup2_assoc: "a \<squnion>2 (b \<squnion>2 c) = (a \<squnion>2 b) \<squnion>2 c"
sublocale sup_assoc_pseudo_hoop_algebra < sup1_lattice: pseudo_hoop_lattice "(\<squnion>1)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
by (simp add: sup1_assoc)
sublocale sup_assoc_pseudo_hoop_algebra < sup2_lattice: pseudo_hoop_lattice "(\<squnion>2)" "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
by (simp add: sup2_assoc)
class sup_assoc_pseudo_hoop_algebra_1 = sup_assoc_pseudo_hoop_algebra +
assumes union_impl: "(a l\<rightarrow> b) \<squnion>1 (b l\<rightarrow> a) = 1"
and union_impr: "(a r\<rightarrow> b) \<squnion>1 (b r\<rightarrow> a) = 1"
lemma (in pseudo_hoop_algebra) [simp]: "infimum {a, b} = {a \<sqinter> b}"
apply (simp add: infimum_def lower_bound_def)
apply safe
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
lemma (in pseudo_hoop_lattice) sup_impl_inf:
"(a \<squnion> b) l\<rightarrow> c = (a l\<rightarrow> c) \<sqinter> (b l\<rightarrow>c)"
apply (cut_tac A = "{a, b}" and a = "a \<squnion> b" and b = c in lemma_2_8_i)
by simp_all
lemma (in pseudo_hoop_lattice) sup_impr_inf:
"(a \<squnion> b) r\<rightarrow> c = (a r\<rightarrow> c) \<sqinter> (b r\<rightarrow>c)"
apply (cut_tac A = "{a, b}" and a = "a \<squnion> b" and b = c in lemma_2_8_i1)
by simp_all
lemma (in pseudo_hoop_lattice) sup_times:
"a * (b \<squnion> c) = (a * b) \<squnion> (a * c)"
apply (cut_tac A = "{b, c}" and b = "b \<squnion> c" and a = a in lemma_2_9_i)
by simp_all
lemma (in pseudo_hoop_lattice) sup_times_right:
"(b \<squnion> c) * a = (b * a) \<squnion> (c * a)"
apply (cut_tac A = "{b, c}" and b = "b \<squnion> c" and a = a in lemma_2_9_i1)
by simp_all
context basic_pseudo_hoop_algebra begin end
sublocale sup_assoc_pseudo_hoop_algebra_1 < basic_1: basic_pseudo_hoop_algebra "(*)" "(\<sqinter>)" "(l\<rightarrow>)" "(\<le>)" "(<)" 1 "(r\<rightarrow>)"
apply unfold_locales
apply (subst left_residual [THEN sym])
apply (rule_tac y = "(a l\<rightarrow> b) \<squnion>1 (b l\<rightarrow> a) l\<rightarrow> c" in order_trans)
apply (subst sup1_lattice.sup_impl_inf)
apply (simp add: lemma_2_5_11)
apply (simp add: union_impl)
apply (subst right_residual [THEN sym])
apply (rule_tac y = "(b r\<rightarrow> a) \<squnion>1 (a r\<rightarrow> b) r\<rightarrow> c" in order_trans)
apply (subst sup1_lattice.sup_impr_inf)
apply (simp add: lemma_2_5_11)
by (simp add: union_impr)
context basic_pseudo_hoop_algebra
begin
lemma lemma_4_8_i: "a * (b \<sqinter> c) = (a * b) \<sqinter> (a * c)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (subgoal_tac "a * (b \<sqinter> c) = (a * (b * (b r\<rightarrow> c))) \<squnion>1 (a * (c * (c r\<rightarrow> b)))")
apply simp
apply (drule drop_assumption)
apply (rule_tac y = "(((a * b) \<sqinter> (a * c)) * (b r\<rightarrow> c)) \<squnion>1 (((a * b) \<sqinter> (a * c)) * (c r\<rightarrow> b))" in order_trans)
apply (subst sup_times [THEN sym])
apply (simp add: lemma_4_5 lemma_4_5_ii)
apply (simp add: mult.assoc [THEN sym])
apply safe
apply (rule_tac y = "a * b * (b r\<rightarrow> c)" in order_trans)
apply simp
apply simp
apply (rule_tac y = "a * c * (c r\<rightarrow> b)" in order_trans)
apply simp
apply simp
apply (simp add: inf_r_def [THEN sym])
apply (subgoal_tac "b \<sqinter> c = c \<sqinter> b")
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
lemma lemma_4_8_ii: "(b \<sqinter> c) * a = (b * a) \<sqinter> (c * a)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp
apply (subgoal_tac "(b \<sqinter> c) * a = (((b l\<rightarrow> c) * b) * a) \<squnion>1 (((c l\<rightarrow> b) * c) * a)")
apply simp
apply (drule drop_assumption)
apply (rule_tac y = "((b l\<rightarrow> c) * ((b * a) \<sqinter> (c * a))) \<squnion>1 ((c l\<rightarrow> b) * ((b * a) \<sqinter> (c * a)))" in order_trans)
apply (subst sup_times_right [THEN sym])
apply (simp add: lemma_4_5_i)
apply (simp add: mult.assoc)
apply safe
apply (rule_tac y = "(b l\<rightarrow> c) * (b * a)" in order_trans)
apply simp_all
apply (rule_tac y = "(c l\<rightarrow> b) * (c * a)" in order_trans)
apply simp_all
apply (simp add: inf_l_def [THEN sym])
apply (subgoal_tac "b \<sqinter> c = c \<sqinter> b")
apply simp
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
lemma lemma_4_8_iii: "(a l\<rightarrow> b) l\<rightarrow> (b l\<rightarrow> a) = b l\<rightarrow> a"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (cut_tac a = a and b = b in lemma_4_5_i)
apply (unfold sup1_def right_lesseq, simp)
by (simp add: lemma_2_5_9_a)
lemma lemma_4_8_iv: "(a r\<rightarrow> b) r\<rightarrow> (b r\<rightarrow> a) = b r\<rightarrow> a"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (cut_tac a = a and b = b in lemma_4_5_ii)
apply (unfold sup2_def left_lesseq, simp)
by (simp add: lemma_2_5_9_b)
end
context wajsberg_pseudo_hoop_algebra
begin
subclass sup_assoc_pseudo_hoop_algebra_1
apply unfold_locales
apply (simp add: lattice1.sup_assoc)
apply (simp add: lattice2.sup_assoc)
apply (simp add: lemma_4_3_i_a)
apply (subgoal_tac "(a l\<rightarrow> b) l\<rightarrow> (b l\<rightarrow> a) = b l\<rightarrow> a")
apply simp
apply (subst lemma_2_10_30 [THEN sym])
apply (subst wajsberg1)
apply (simp add: lemma_2_10_32)
apply (subst sup1_eq_sup2)
apply (simp add: lemma_4_3_ii_a)
apply (subgoal_tac "(a r\<rightarrow> b) r\<rightarrow> (b r\<rightarrow> a) = b r\<rightarrow> a")
apply simp
apply (subst lemma_2_10_31 [THEN sym])
apply (subst wajsberg2)
by (simp add: lemma_2_10_33)
end
class bounded_basic_pseudo_hoop_algebra = zero_one + basic_pseudo_hoop_algebra +
assumes zero_smallest [simp]: "0 \<le> a"
-begin
-end
class inf_a =
fixes inf_a :: "'a => 'a => 'a" (infixl "\<sqinter>1" 65)
-
class pseudo_bl_algebra = zero + sup + inf + monoid_mult + ord + left_imp + right_imp +
assumes bounded_lattice: "class.bounded_lattice inf (\<le>) (<) sup 0 1"
and left_residual_bl: "(x * a \<le> b) = (x \<le> a l\<rightarrow> b)"
and right_residual_bl: "(a * x \<le> b) = (x \<le> a r\<rightarrow> b)"
and inf_l_bl_def: "a \<sqinter> b = (a l\<rightarrow> b) * a"
and inf_r_bl_def: "a \<sqinter> b = a * (a r\<rightarrow> b)"
and impl_sup_bl: "(a l\<rightarrow> b) \<squnion> (b l\<rightarrow> a) = 1"
and impr_sup_bl: "(a r\<rightarrow> b) \<squnion> (b r\<rightarrow> a) = 1"
-begin
-end
-context pseudo_bl_algebra begin end
-
-sublocale bounded_basic_pseudo_hoop_algebra < basic:pseudo_bl_algebra 1 "(*)" "0" "(\<sqinter>)" "(\<squnion>1)" "(l\<rightarrow>)" "(r\<rightarrow>)" "(\<le>)" "(<)"
+sublocale bounded_basic_pseudo_hoop_algebra < basic: pseudo_bl_algebra 1 "(*)" "0" "(\<sqinter>)" "(\<squnion>1)" "(l\<rightarrow>)" "(r\<rightarrow>)" "(\<le>)" "(<)"
apply unfold_locales
apply (rule zero_smallest)
apply (rule left_residual)
apply (rule right_residual)
apply (rule inf_l_def)
apply (simp add: inf_r_def [THEN sym])
apply (rule lemma_4_5_i)
apply (simp add: lemma_4_5)
by (rule lemma_4_5_ii)
sublocale pseudo_bl_algebra < bounded_lattice: bounded_lattice "inf" "(\<le>)" "(<)" "sup" "0" "1"
by (rule bounded_lattice)
context pseudo_bl_algebra
begin
lemma impl_one_bl [simp]: "a l\<rightarrow> a = 1"
- apply (rule bounded_lattice.antisym)
+ apply (rule bounded_lattice.order.antisym)
apply simp_all
apply (subst left_residual_bl [THEN sym])
by simp
lemma impr_one_bl [simp]: "a r\<rightarrow> a = 1"
- apply (rule bounded_lattice.antisym)
+ apply (rule bounded_lattice.order.antisym)
apply simp_all
apply (subst right_residual_bl [THEN sym])
by simp
lemma impl_ded_bl: "((a * b) l\<rightarrow> c) = (a l\<rightarrow> (b l\<rightarrow> c))"
- apply (rule bounded_lattice.antisym)
+ apply (rule bounded_lattice.order.antisym)
apply (case_tac "(a * b l\<rightarrow> c \<le> a l\<rightarrow> b l\<rightarrow> c) = ((a * b l\<rightarrow> c) * a \<le> b l\<rightarrow> c)
\<and> ((a * b l\<rightarrow> c) * a \<le> b l\<rightarrow> c) = (((a * b l\<rightarrow> c) * a) * b \<le> c)
\<and> (((a * b l\<rightarrow> c) * a) * b \<le> c) = ((a * b l\<rightarrow> c) * (a * b) \<le> c)
\<and> ((a * b l\<rightarrow> c) * (a * b) \<le> c) = ((a * b l\<rightarrow> c) \<le> (a * b l\<rightarrow> c))")
apply simp
apply (erule notE)
apply (rule conjI)
apply (simp add: left_residual_bl)
apply (rule conjI)
apply (simp add: left_residual_bl)
apply (rule conjI)
apply (simp add: mult.assoc)
apply (simp add: left_residual_bl)
apply (simp add: left_residual_bl [THEN sym])
apply (rule_tac y="(b l\<rightarrow> c) * b" in bounded_lattice.order_trans)
apply (simp add: mult.assoc [THEN sym])
apply (subst inf_l_bl_def [THEN sym])
apply (subst bounded_lattice.inf_commute)
apply (subst inf_l_bl_def)
apply (subst mult.assoc)
apply (subst left_residual_bl)
apply simp
apply (subst left_residual_bl)
by simp
lemma impr_ded_bl: "(b * a r\<rightarrow> c) = (a r\<rightarrow> (b r\<rightarrow> c))"
- apply (rule bounded_lattice.antisym)
+ apply (rule bounded_lattice.order.antisym)
apply (case_tac "(b * a r\<rightarrow> c \<le> a r\<rightarrow> b r\<rightarrow> c) = (a * (b * a r\<rightarrow> c) \<le> b r\<rightarrow> c)
\<and> (a * (b * a r\<rightarrow> c) \<le> b r\<rightarrow> c) = (b * (a * (b * a r\<rightarrow> c)) \<le> c)
\<and> (b * ( a* (b * a r\<rightarrow> c)) \<le> c) = ((b * a) * (b * a r\<rightarrow> c) \<le> c)
\<and> ((b * a) * (b * a r\<rightarrow> c) \<le> c) = ((b * a r\<rightarrow> c) \<le> (b * a r\<rightarrow> c))")
apply simp
apply (erule notE)
apply (rule conjI)
apply (simp add: right_residual_bl)
apply (rule conjI)
apply (simp add: right_residual_bl)
apply (rule conjI)
apply (simp add: mult.assoc)
apply (simp add: right_residual_bl)
apply (simp add: right_residual_bl [THEN sym])
apply (rule_tac y="b * (b r\<rightarrow> c)" in bounded_lattice.order_trans)
apply (simp add: mult.assoc)
apply (subst inf_r_bl_def [THEN sym])
apply (subst bounded_lattice.inf_commute)
apply (subst inf_r_bl_def)
apply (subst mult.assoc [THEN sym])
apply (subst right_residual_bl)
apply simp
apply (subst right_residual_bl)
by simp
lemma lesseq_impl_bl: "(a \<le> b) = (a l\<rightarrow> b = 1)"
apply (rule iffI)
- apply (rule bounded_lattice.antisym)
+ apply (rule bounded_lattice.order.antisym)
apply simp
apply (simp add: left_residual_bl [THEN sym])
apply (subgoal_tac "1 \<le> a l\<rightarrow> b")
apply (subst (asm) left_residual_bl [THEN sym])
by simp_all
end
context pseudo_bl_algebra
begin
subclass pseudo_hoop_lattice
apply unfold_locales
apply (rule inf_l_bl_def)
apply (simp add: bounded_lattice.less_le_not_le)
apply (simp add: mult_1_left)
apply (simp add: mult_1_right)
apply (simp add: impl_one_bl)
apply (simp add: inf_l_bl_def [THEN sym])
apply (rule bounded_lattice.inf_commute)
apply (rule impl_ded_bl)
apply (rule lesseq_impl_bl)
apply (rule inf_r_bl_def)
apply (simp add: impr_one_bl)
apply (simp add: inf_r_bl_def [THEN sym])
apply (rule bounded_lattice.inf_commute)
apply (rule impr_ded_bl)
apply (simp add: inf_r_bl_def [THEN sym] inf_l_bl_def [THEN sym])
apply (rule bounded_lattice.sup_commute)
apply simp
apply safe
- apply (rule bounded_lattice.antisym)
+ apply (rule bounded_lattice.order.antisym)
apply simp_all
apply (subgoal_tac "a \<le> a \<squnion> b")
apply simp
apply (drule drop_assumption)
apply simp
by (simp add: bounded_lattice.sup_assoc)
subclass bounded_basic_pseudo_hoop_algebra
apply unfold_locales
apply simp_all
apply (simp add: left_residual [THEN sym])
apply (rule_tac y = "((a l\<rightarrow> b) \<squnion> (b l\<rightarrow> a)) l\<rightarrow> c" in bounded_lattice.order_trans)
apply (simp add: sup_impl_inf)
apply (simp add: impl_sup_bl)
apply (simp add: right_residual [THEN sym])
apply (rule_tac y = "((a r\<rightarrow> b) \<squnion> (b r\<rightarrow> a)) r\<rightarrow> c" in bounded_lattice.order_trans)
apply (simp add: sup_impr_inf)
by (simp add: impr_sup_bl)
end
class product_pseudo_hoop_algebra = basic_pseudo_hoop_algebra +
assumes P1: "b l\<rightarrow> b * b \<le> (a \<sqinter> (a l\<rightarrow> b)) l\<rightarrow> b"
and P2: "b r\<rightarrow> b * b \<le> (a \<sqinter> (a r\<rightarrow> b)) r\<rightarrow> b"
and P3: "((a l\<rightarrow> b) l\<rightarrow> b) * (c * a l\<rightarrow> d * a) * (c * b l\<rightarrow> d * b) \<le> c l\<rightarrow> d"
and P4: "((a r\<rightarrow> b) r\<rightarrow> b) * (a * c r\<rightarrow> a * d) * (b * c r\<rightarrow> b * d) \<le> c r\<rightarrow> d"
class cancel_basic_pseudo_hoop_algebra = basic_pseudo_hoop_algebra + cancel_pseudo_hoop_algebra
begin
subclass product_pseudo_hoop_algebra
apply unfold_locales
apply (rule_tac y = "1 l\<rightarrow> b" in order_trans)
apply (cut_tac a = 1 and b = b and c = b in lemma_4_2_i)
apply simp
apply (simp add: lemma_2_5_9_a)
apply (rule_tac y = "1 r\<rightarrow> b" in order_trans)
apply (cut_tac a = 1 and b = b and c = b in lemma_4_2_ii)
apply simp
apply (simp add: lemma_2_5_9_b)
apply (simp add: lemma_4_2_i [THEN sym])
by (simp add: lemma_4_2_ii [THEN sym])
end
class simple_pseudo_hoop_algebra = pseudo_hoop_algebra +
assumes simple: "normalfilters \<inter> properfilters = {{1}}"
class proper = one +
assumes proper: "\<exists> a . a \<noteq> 1"
class strong_simple_pseudo_hoop_algebra = pseudo_hoop_algebra +
assumes strong_simple: "properfilters = {{1}}"
begin
subclass proper
apply unfold_locales
apply (cut_tac strong_simple)
apply (simp add: properfilters_def)
apply (case_tac "{1} = UNIV")
apply blast
by blast
lemma lemma_4_12_i_ii: "a \<noteq> 1 \<Longrightarrow> filterof({a}) = UNIV"
apply (cut_tac strong_simple)
apply (simp add: properfilters_def)
apply (subgoal_tac "filterof {a} \<notin> {F \<in> filters. F \<noteq> UNIV}")
apply (drule drop_assumption)
apply (drule drop_assumption)
apply simp
apply simp
apply safe
apply (subgoal_tac "a \<in> filterof {a}")
apply simp
apply (subst filterof_def)
by simp
lemma lemma_4_12_i_iii: "a \<noteq> 1 \<Longrightarrow> (\<exists> n . a ^ n \<le> b)"
apply (drule lemma_4_12_i_ii)
apply (simp add: prop_3_2_i)
by blast
lemma lemma_4_12_i_iv: "a \<noteq> 1 \<Longrightarrow> (\<exists> n . a l-n\<rightarrow> b = 1)"
apply (subst lemma_2_4_7_a)
apply (subst left_lesseq [THEN sym])
by (simp add: lemma_4_12_i_iii)
lemma lemma_4_12_i_v: "a \<noteq> 1 \<Longrightarrow> (\<exists> n . a r-n\<rightarrow> b = 1)"
apply (subst lemma_2_4_7_b)
apply (subst right_lesseq [THEN sym])
by (simp add: lemma_4_12_i_iii)
end
lemma (in pseudo_hoop_algebra) [simp]: "{1} \<in> filters"
apply (simp add: filters_def)
apply safe
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
class strong_simple_pseudo_hoop_algebra_a = pseudo_hoop_algebra + proper +
assumes strong_simple_a: "a \<noteq> 1 \<Longrightarrow> filterof({a}) = UNIV"
begin
subclass strong_simple_pseudo_hoop_algebra
apply unfold_locales
apply (simp add: properfilters_def)
apply safe
apply simp_all
apply (case_tac "xb = 1")
apply simp
apply (cut_tac a = xb in strong_simple_a)
apply simp
apply (simp add: filterof_def)
apply blast
apply (cut_tac proper)
by blast
end
sublocale strong_simple_pseudo_hoop_algebra < strong_simple_pseudo_hoop_algebra_a
apply unfold_locales
by (simp add: lemma_4_12_i_ii)
lemma (in pseudo_hoop_algebra) power_impl: "b l\<rightarrow> a = a \<Longrightarrow> b ^ n l\<rightarrow> a = a"
apply (induct_tac n)
apply simp
by (simp add: left_impl_ded)
lemma (in pseudo_hoop_algebra) power_impr: "b r\<rightarrow> a = a \<Longrightarrow> b ^ n r\<rightarrow> a = a"
apply (induct_tac n)
apply simp
by (simp add: right_impl_ded)
context strong_simple_pseudo_hoop_algebra
begin
lemma lemma_4_13_i: "b l\<rightarrow> a = a \<Longrightarrow> a = 1 \<or> b = 1"
apply safe
apply (drule_tac a = b and b = a in lemma_4_12_i_iii)
apply safe
apply (subst (asm) left_lesseq)
apply (drule_tac n = n in power_impl)
by simp
lemma lemma_4_13_ii: "b r\<rightarrow> a = a \<Longrightarrow> a = 1 \<or> b = 1"
apply safe
apply (drule_tac a = b and b = a in lemma_4_12_i_iii)
apply safe
apply (subst (asm) right_lesseq)
apply (drule_tac n = n in power_impr)
by simp
end
class basic_pseudo_hoop_algebra_A = basic_pseudo_hoop_algebra +
assumes left_impl_one: "b l\<rightarrow> a = a \<Longrightarrow> a = 1 \<or> b = 1"
and right_impl_one: "b r\<rightarrow> a = a \<Longrightarrow> a = 1 \<or> b = 1"
begin
subclass linorder
apply unfold_locales
apply (cut_tac a = "x" and b = "y" in lemma_4_8_iii)
apply (drule left_impl_one)
apply (simp add: left_lesseq)
by blast
lemma [simp]: "(a l\<rightarrow> b) r\<rightarrow> b \<le> (b l\<rightarrow> a) r\<rightarrow> a"
apply (case_tac "a = b")
apply simp
apply (cut_tac x = a and y =b in linear)
apply safe
apply (subst (asm) left_lesseq)
apply (simp add: lemma_2_10_24)
apply (subst (asm) left_lesseq)
apply simp
apply (subst left_lesseq)
apply (cut_tac b = "((a l\<rightarrow> b) r\<rightarrow> b) l\<rightarrow> a" and a = "a l\<rightarrow> b" in left_impl_one)
apply (simp add: lemma_2_10_32)
apply (simp add: left_lesseq [THEN sym])
apply safe
apply (erule notE)
by simp
end
context basic_pseudo_hoop_algebra_A begin
lemma [simp]: "(a r\<rightarrow> b) l\<rightarrow> b \<le> (b r\<rightarrow> a) l\<rightarrow> a"
apply (case_tac "a = b")
apply simp
apply (cut_tac x = a and y =b in linear)
apply safe
apply (subst (asm) right_lesseq)
apply simp
apply (simp add: lemma_2_10_26)
apply (unfold right_lesseq)
apply (cut_tac b = "((a r\<rightarrow> b) l\<rightarrow> b) r\<rightarrow> a" and a = "a r\<rightarrow> b" in right_impl_one)
apply (simp add: lemma_2_10_33)
apply (simp add: right_lesseq [THEN sym])
apply safe
apply (erule notE)
by simp
subclass wajsberg_pseudo_hoop_algebra
apply unfold_locales
- apply (rule antisym)
+ apply (rule order.antisym)
apply simp_all
- apply (rule antisym)
+ apply (rule order.antisym)
by simp_all
end
class strong_simple_basic_pseudo_hoop_algebra = strong_simple_pseudo_hoop_algebra + basic_pseudo_hoop_algebra
begin
subclass basic_pseudo_hoop_algebra_A
apply unfold_locales
apply (simp add: lemma_4_13_i)
by (simp add: lemma_4_13_ii)
subclass wajsberg_pseudo_hoop_algebra
by unfold_locales
end
end
diff --git a/thys/Quantales/Quantale_Left_Sided.thy b/thys/Quantales/Quantale_Left_Sided.thy
--- a/thys/Quantales/Quantale_Left_Sided.thy
+++ b/thys/Quantales/Quantale_Left_Sided.thy
@@ -1,568 +1,568 @@
(* Title: Left-Sided Elements in Quantales
Author: Georg Struth
Maintainer: Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Left-Sided Elements in Quantales\<close>
theory Quantale_Left_Sided
imports Quantic_Nuclei_Conuclei
begin
context quantale
begin
subsection \<open>Basic Definitions\<close>
text \<open>Left-sided elements are well investigated in quantale theory. They could be defined in weaker settings, for instance,
ord with a top element.\<close>
definition "lsd x = (\<top> \<cdot> x \<le> x)"
definition "rsd x = (x \<cdot> \<top> \<le> x)"
definition "twosd x = (lsd x \<and> rsd x)"
definition "slsd x = (\<top> \<cdot> x = x)"
definition "srsd x = (x \<cdot> \<top> = x)"
definition "stwsd x = (slsd x \<and> srsd x)"
lemma lsided_bres: "(lsd x) = (x \<le> \<top> \<rightarrow> x)"
by (simp add: bres_galois lsd_def)
lemma lsided_fres: "(lsd x) = (\<top> \<le> x \<leftarrow> x)"
by (simp add: fres_galois lsd_def)
lemma lsided_fres_eq: "(lsd x) = (x \<leftarrow> x = \<top>)"
using fres_galois top_le lsd_def by force
lemma lsided_slsided: "lsd x \<Longrightarrow> x \<cdot> x = x \<Longrightarrow> slsd x"
using fres_sol lsided_fres_eq slsd_def by force
lemma lsided_prop: "lsd x \<Longrightarrow> y \<cdot> x \<le> x"
by (simp add: fres_galois lsided_fres_eq)
lemma rsided_fres: "(rsd x) = (x \<le> x \<leftarrow> \<top>)"
by (simp add: fres_galois rsd_def)
lemma rsided_bres: "(rsd x) = (\<top> \<le> x \<rightarrow> x)"
by (simp add: bres_galois rsd_def)
lemma rsided_bres_eq: "(rsd x) = (x \<rightarrow> x = \<top>)"
using top_le rsided_bres by blast
lemma rsided_srsided: "rsd x \<Longrightarrow> x \<cdot> x = x \<Longrightarrow> srsd x"
using bres_sol rsided_bres_eq srsd_def by auto
lemma rsided_prop: "rsd x \<Longrightarrow> x \<cdot> y \<le> x"
by (simp add: bres_galois rsided_bres_eq)
lemma lsided_top: "lsd \<top>"
by (simp add: lsd_def)
lemma lsided_bot: "lsd \<bottom>"
by (simp add: lsd_def)
lemma rsided_top: "rsd \<top>"
by (simp add: rsd_def)
lemma rsided_bot: "rsd \<bottom>"
by (simp add: rsd_def)
end
text \<open>Right-sided elements are henceforth not considered. Their properties arise by opposition
duality, which is not formalised.\<close>
text \<open>The following functions have left-sided elements as fixpoints.\<close>
definition lsl:: "'a::quantale \<Rightarrow> 'a" ("\<nu>") where
"\<nu> x = \<top> \<cdot> x"
definition lsu :: "'a::quantale \<Rightarrow> 'a" ("\<nu>\<^sup>\<natural>") where
"\<nu>\<^sup>\<natural> x = \<top> \<rightarrow> x"
text \<open>These functions are adjoints.\<close>
lemma ls_galois: "\<nu> \<stileturn> \<nu>\<^sup>\<natural>"
by (simp add: adj_def bres_galois lsl_def lsu_def)
text \<open>Due to this, the following properties hold.\<close>
lemma lsl_iso: "mono \<nu>"
using adj_iso1 ls_galois by blast
lemma lsl_iso_var: "x \<le> y \<Longrightarrow> \<nu> x \<le> \<nu> y"
by (simp add: lsl_iso monoD)
lemma lsu_iso: "mono \<nu>\<^sup>\<natural>"
using adj_iso2 ls_galois by blast
lemma lsu_iso_var: "x \<le> y \<Longrightarrow> \<nu>\<^sup>\<natural> x \<le> \<nu>\<^sup>\<natural> y"
by (simp add: lsu_iso monoD)
lemma lsl_bot [simp]: "\<nu> \<bottom> = \<bottom>"
by (simp add: lsl_def)
lemma lsu_top [simp]: "\<nu>\<^sup>\<natural> \<top> = \<top>"
by (simp add: lsu_def bres_galois_imp top_le)
lemma lsu_Inf_pres: "Inf_pres \<nu>\<^sup>\<natural>"
unfolding fun_eq_iff by (metis ls_galois radj_Inf_pres)
lemma lsl_Sup_pres: "Sup_pres (\<nu>::'a::quantale \<Rightarrow> 'a)"
by (simp add: fun_eq_iff, metis SUP_cong Sup_distl lsl_def)
lemma lsu_Inf_closed: "Inf_closed_set (range \<nu>\<^sup>\<natural>)"
by (simp add: Inf_pres_Inf_closed lsu_Inf_pres)
lemma lsl_Sup_closed: "Sup_closed_set (range (\<nu>::'a::quantale \<Rightarrow> 'a))"
by (simp add: Sup_pres_Sup_closed lsl_Sup_pres)
lemma lsl_lsu_canc1: "\<nu> \<circ> \<nu>\<^sup>\<natural> \<le> id"
by (simp add: adj_cancel1 ls_galois)
lemma lsl_lsu_canc2: "id \<le> \<nu>\<^sup>\<natural> \<circ> \<nu>"
by (simp add: adj_cancel2 ls_galois)
lemma clop_lsu_lsl: "clop (\<nu>\<^sup>\<natural> \<circ> \<nu>)"
by (simp add: clop_adj ls_galois)
lemma coclop_lsl_lsu: "coclop (\<nu> \<circ> \<nu>\<^sup>\<natural>)"
by (simp add: coclop_adj ls_galois)
lemma dang1: "\<nu> (\<nu> x \<sqinter> y) \<le> \<nu> x"
unfolding lsl_def by (metis bres_galois bres_interchange bres_top_top inf.coboundedI1)
lemma lsl_trans: "\<nu> \<circ> \<nu> \<le> \<nu>"
unfolding lsl_def fun_eq_iff comp_def
by (metis (no_types, lifting) bres_galois bres_interchange bres_top_top le_funI)
lemma lsl_lsu_prop: "\<nu> \<circ> \<nu>\<^sup>\<natural> \<le> \<nu>\<^sup>\<natural>"
unfolding le_fun_def comp_def
by (metis adj_def dang1 dual_order.trans le_iff_inf ls_galois order_refl top_greatest)
lemma lsu_lsl_prop: "\<nu> \<le> \<nu>\<^sup>\<natural> \<circ> \<nu>"
unfolding le_fun_def comp_def by (metis adj_def comp_def le_fun_def ls_galois lsl_trans)
context unital_quantale
begin
text \<open>Left-sidedness and strict left-sidedness now coincide.\<close>
lemma lsided_eq: "lsd = slsd"
- unfolding fun_eq_iff by (simp add: eq_iff le_top lsd_def slsd_def)
+ unfolding fun_eq_iff by (simp add: order.eq_iff le_top lsd_def slsd_def)
lemma lsided_eq_var1: "(x \<le> \<top> \<rightarrow> x) = (x = \<top> \<rightarrow> x)"
- using bres_galois dual_order.trans eq_iff le_top by blast
+ using bres_galois dual_order.trans order.eq_iff le_top by blast
lemma lsided_eq_var2: "lsd x = (x = \<top> \<rightarrow> x)"
using bres_galois lsided_eq lsided_eq_var1 lsd_def by simp
end
lemma lsided_def3: "(\<nu> (x::'a::unital_quantale) = x) = (\<nu>\<^sup>\<natural> x = x)"
unfolding lsl_def lsu_def by (metis lsided_eq lsided_eq_var2 slsd_def)
lemma Fix_lsl_lsu: "Fix (\<nu>::'a::unital_quantale \<Rightarrow> 'a) = Fix \<nu>\<^sup>\<natural>"
unfolding Fix_def by (simp add: lsided_def3)
lemma Fix_lsl_left_slsided: "Fix \<nu> = {(x::'a::unital_quantale). lsd x}"
- unfolding Fix_def lsl_def lsd_def using eq_iff le_top by blast
+ unfolding Fix_def lsl_def lsd_def using order.eq_iff le_top by blast
lemma Fix_lsl_iff [simp]: "(x \<in> Fix \<nu>) = (\<nu> x = x)"
by (simp add: Fix_def)
lemma Fix_lsu_iff [simp]: "(x \<in> Fix \<nu>\<^sup>\<natural>) = (\<nu>\<^sup>\<natural> x = x)"
by (simp add: Fix_def)
lemma lsl_lsu_prop_eq [simp]: "(\<nu>::'a::unital_quantale \<Rightarrow> 'a) \<circ> \<nu>\<^sup>\<natural> = \<nu>\<^sup>\<natural>"
by (smt antisym clop_extensive_var clop_lsu_lsl comp_apply le_fun_def lsided_eq_var1 lsl_lsu_prop lsu_def lsu_lsl_prop)
lemma lsu_lsl_prop_eq [simp]: "\<nu>\<^sup>\<natural> \<circ> \<nu> = (\<nu>::'a::unital_quantale \<Rightarrow> 'a)"
by (metis adj_cancel_eq1 ls_galois lsl_lsu_prop_eq)
subsection \<open>Connection with Closure and Coclosure Operators, Nuclei and Conuclei\<close>
text \<open>lsl is therefore a closure operator, lsu a cocolosure operator.\<close>
lemma lsl_clop: "clop (\<nu>::'a::unital_quantale \<Rightarrow> 'a)"
by (metis clop_lsu_lsl lsu_lsl_prop_eq)
lemma lsu_coclop: "coclop (\<nu>\<^sup>\<natural>::'a::unital_quantale \<Rightarrow> 'a)"
by (metis coclop_lsl_lsu lsl_lsu_prop_eq)
lemma lsl_range_fix: "range (\<nu>::'a::unital_quantale \<Rightarrow> 'a) = Fix \<nu>"
by (simp add: clop_closure_set lsl_clop)
lemma lsu_range_fix: "range (\<nu>\<^sup>\<natural>::'a::unital_quantale \<Rightarrow> 'a) = Fix \<nu>\<^sup>\<natural>"
by (simp add: coclop_coclosure_set lsu_coclop)
lemma range_lsl_iff [simp]: "((x::'a::unital_quantale) \<in> range \<nu>) = (\<nu> x = x)"
by (simp add: lsl_range_fix)
lemma range_lsu_iff [simp]: "((x::'a::unital_quantale) \<in> range \<nu>\<^sup>\<natural>) = (\<nu>\<^sup>\<natural> x = x)"
by (simp add: lsu_range_fix)
text \<open>lsl and lsu are therefore both Sup and Inf closed.\<close>
lemma lsu_Sup_closed: "Sup_closed_set (Fix (\<nu>\<^sup>\<natural>::'a::unital_quantale \<Rightarrow> 'a))"
by (metis Fix_lsl_lsu lsl_Sup_closed lsl_range_fix)
lemma lsl_Inf_closed: "Inf_closed_set (Fix (\<nu>::'a::unital_quantale \<Rightarrow> 'a))"
by (metis Fix_lsl_lsu lsu_Inf_closed lsu_range_fix)
text \<open>lsl is even a quantic conucleus.\<close>
lemma lsu_lax: "\<nu>\<^sup>\<natural> (x::'a::unital_quantale) \<cdot> \<nu>\<^sup>\<natural> y \<le> \<nu>\<^sup>\<natural> (x \<cdot> y)"
unfolding lsu_def by (meson bres_canc1 bres_galois bres_interchange le_top order_trans)
lemma lsu_one: "\<nu>\<^sup>\<natural> 1 \<le> (1::'a::unital_quantale)"
unfolding lsu_def using bres_canc1 dual_order.trans le_top by blast
lemma lsl_one: "1 \<le> \<nu> (1::'a::unital_quantale)"
unfolding lsl_def by simp
lemma lsu_conuc: "conucleus (\<nu>\<^sup>\<natural>::'a::unital_quantale \<Rightarrow> 'a)"
by (simp add: lsu_coclop conucleus_def lsu_lax)
text \<open>It is therefore closed under composition.\<close>
lemma lsu_comp_closed_var [simp]: "\<nu>\<^sup>\<natural> (\<nu>\<^sup>\<natural> x \<cdot> \<nu>\<^sup>\<natural> y) = \<nu>\<^sup>\<natural> x \<cdot> \<nu>\<^sup>\<natural> (y::'a::unital_quantale)"
by (simp add: conuc_comp_closed lsu_conuc)
lemma lsu_comp_closed: "comp_closed_set (Fix (\<nu>\<^sup>\<natural>::'a::unital_quantale \<Rightarrow> 'a))"
by (simp add: conuc_Sup_qclosed lsu_conuc)
text \<open>lsl is not a quantic nucleus unless composition is commutative.\<close>
lemma "\<nu> x \<cdot> \<nu> y \<le> \<nu> (x \<cdot> (y::'a::unital_quantale))" (*nitpick*)
oops
text \<open>Yet it is closed under composition (because the set of fixpoints are the same).\<close>
lemma lsl_comp_closed: "comp_closed_set (Fix (\<nu>::'a::unital_quantale \<Rightarrow> 'a))"
by (simp add: Fix_lsl_lsu lsu_comp_closed)
lemma lsl_comp_closed_var [simp]: "\<nu> (\<nu> x \<cdot> \<nu> (y::'a::unital_quantale)) = \<nu> x \<cdot> \<nu> y"
by (metis Fix_lsl_iff lsl_def lsl_range_fix mult.assoc rangeI)
text \<open>The following simple properties go beyond nuclei and conuclei.\<close>
lemma lsl_lsu_ran: "range \<nu>\<^sup>\<natural> = range (\<nu>::'a::unital_quantale \<Rightarrow> 'a)"
by (simp add: Fix_lsl_lsu lsl_range_fix lsu_range_fix)
lemma lsl_top [simp]: "\<nu> \<top> = (\<top>::'a::unital_quantale)"
by (simp add: clop_top lsl_clop)
lemma lsu_bot [simp]: "\<nu>\<^sup>\<natural> \<bottom> = (\<bottom>::'a::unital_quantale)"
using lsided_def3 lsl_bot by blast
lemma lsu_inj_on: "inj_on \<nu>\<^sup>\<natural> (Fix (\<nu>\<^sup>\<natural>::'a::unital_quantale \<Rightarrow> 'a))"
by (metis coclop_coclosure inj_onI lsu_coclop lsu_range_fix)
lemma lsl_inj_on: "inj_on \<nu> (Fix (\<nu>::'a::unital_quantale \<Rightarrow> 'a))"
by (metis clop_closure inj_onI lsl_clop lsl_range_fix)
text \<open>Additional preservation properties of lsl and lsu are useful in proofs.\<close>
lemma lsl_Inf_closed_var [simp]: "\<nu> (\<Sqinter>x \<in> X. \<nu> x) = (\<Sqinter>x \<in> X. \<nu> (x::'a::unital_quantale))"
by (metis (no_types, hide_lams) INF_image lsided_def3 lsu_Inf_pres lsu_lsl_prop_eq o_apply)
lemma lsl_Sup_closed_var [simp]: "\<nu> (\<Squnion>x \<in> X. \<nu> x) = (\<Squnion>x \<in> X. \<nu> (x::'a::unital_quantale))"
unfolding lsl_def by (metis Sup_distl mult.assoc top_times_top)
lemma lsu_Inf_closed_var [simp]: "\<nu>\<^sup>\<natural> (\<Sqinter>x \<in> X. \<nu>\<^sup>\<natural> x) = (\<Sqinter>x \<in> X. \<nu>\<^sup>\<natural> (x::'a::unital_quantale))"
by (metis INF_image lsided_def3 lsl_Inf_closed_var lsl_lsu_prop_eq)
lemma lsu_Sup_closed_var [simp]: "\<nu>\<^sup>\<natural> (\<Squnion>x \<in> X. \<nu>\<^sup>\<natural> x) = (\<Squnion>x \<in> X. \<nu>\<^sup>\<natural> (x::'a::unital_quantale))"
by (metis SUP_image lsided_def3 lsl_Sup_closed_var lsl_lsu_prop_eq)
lemma lsl_inf_closed_var [simp]: "\<nu> (\<nu> x \<sqinter> \<nu> y) = \<nu> (x::'a::unital_quantale) \<sqinter> \<nu> y"
by (smt antisym clop_extensive_var dang1 inf_sup_aci(1) le_inf_iff lsl_clop)
lemma lsl_sup_closed_var [simp]: "\<nu> (\<nu> x \<squnion> \<nu> y) = \<nu> (x::'a::unital_quantale) \<squnion> \<nu> y"
by (meson Sup_sup_closed lsl_Sup_closed range_eqI range_lsl_iff sup_closed_set_def)
lemma lsu_inf_closed_var [simp]: "\<nu>\<^sup>\<natural> (\<nu>\<^sup>\<natural> x \<sqinter> \<nu>\<^sup>\<natural> y) = \<nu>\<^sup>\<natural> (x::'a::unital_quantale) \<sqinter> \<nu>\<^sup>\<natural> y"
by (metis (no_types, lifting) lsided_def3 lsl_inf_closed_var lsl_lsu_prop_eq o_apply)
lemma lsu_sup_closed_var [simp]: "\<nu>\<^sup>\<natural> (\<nu>\<^sup>\<natural> x \<squnion> \<nu>\<^sup>\<natural> y) = \<nu>\<^sup>\<natural> (x::'a::unital_quantale) \<squnion> \<nu>\<^sup>\<natural> y"
by (metis (no_types, lifting) lsided_def3 lsl_lsu_prop_eq lsl_sup_closed_var o_def)
lemma lsu_fres_closed [simp]: "\<nu>\<^sup>\<natural> (\<nu>\<^sup>\<natural> x \<leftarrow> \<nu>\<^sup>\<natural> y) = \<nu>\<^sup>\<natural> x \<leftarrow> \<nu>\<^sup>\<natural> (y::'a::unital_quantale)"
unfolding lsu_def by (simp add: bres_curry fres_bres)
lemma lsl_fres_closed [simp]: "\<nu> (\<nu> x \<leftarrow> \<nu> y) = \<nu> x \<leftarrow> \<nu> (y::'a::unital_quantale)"
unfolding lsl_def by (metis fres_interchange lsd_def lsided_eq mult.assoc slsd_def top_times_top)
lemma lsl_almost_lax: "x \<cdot> \<nu> y \<le> \<nu> (y::'a::unital_quantale)"
by (metis inf_top.right_neutral lsided_eq lsided_prop lsl_def lsl_inf_closed_var mult.right_neutral slsd_def)
text \<open>Finally, here are two counterexamples.\<close>
lemma "\<nu>\<^sup>\<natural> (\<nu>\<^sup>\<natural> x \<rightarrow> \<nu>\<^sup>\<natural> y) = \<nu>\<^sup>\<natural> x \<rightarrow> \<nu>\<^sup>\<natural> (y::'a::unital_quantale)" (*nitpick*)
oops
lemma "\<nu> (\<nu> x \<rightarrow> \<nu> y) = \<nu> x \<rightarrow> \<nu> (y::'a::unital_quantale)" (*nitpick*)
oops
context ab_quantale
begin
lemma lsided_times_top: "lsd (\<top> \<cdot> x)"
by (metis lsd_def mult_isor top_greatest mult_assoc)
lemma lsided_le2: "lsd x \<Longrightarrow> x \<cdot> y \<le> x \<sqinter> (y \<cdot> \<top>)"
using lsided_prop psrpq.mult_isol mult_commute by auto
lemma lsided_le3: "lsd x \<Longrightarrow> (x \<sqinter> y) \<cdot> z \<le> x"
by (metis inf_le1 lsided_prop mult_isol order_trans mult_commute)
lemma lsided_le4: "lsd x \<Longrightarrow> (x \<sqinter> y) \<cdot> z \<le> x \<sqinter> (y \<cdot> z)"
by (simp add: mult_isor lsided_le3)
lemma lsided_times_le_inf: "lsd x \<Longrightarrow> lsd y \<Longrightarrow> x \<cdot> y \<le> x \<sqinter> y"
using lsided_prop lsided_le2 by force
end
text \<open>Now lsl is a quantic nucleus.\<close>
lemma lsl_lax: "\<nu> x \<cdot> \<nu> y \<le> \<nu> (x \<cdot> (y::'aa::ab_unital_quantale))"
by (metis (no_types, hide_lams) lsl_almost_lax lsl_def mult.commute mult.left_commute)
lemma lsl_nuc: "nucleus (\<nu>::'a::ab_unital_quantale \<Rightarrow> 'a)"
by (simp add: lsl_clop nucleus_def lsl_lax)
text \<open>The following properties go beyond nuclei.\<close>
lemma lsl_comp_pres: "\<nu> (x \<cdot> y) = \<nu> x \<cdot> \<nu> (y::'a::ab_unital_quantale)"
by (metis (no_types, lifting) lsl_comp_closed_var lsl_nuc nuc_prop1 nuc_prop2)
lemma lsl_bres_closed [simp]: "\<nu> (\<nu> x \<rightarrow> \<nu> y) = \<nu> x \<rightarrow> \<nu> (y::'a::ab_unital_quantale)"
by (simp add: lsl_nuc nuc_bres_closed)
lemma lsu_bres_pres [simp]: "\<nu>\<^sup>\<natural> (\<nu>\<^sup>\<natural> x \<rightarrow> \<nu>\<^sup>\<natural> y) = \<nu>\<^sup>\<natural> x \<rightarrow> \<nu>\<^sup>\<natural> (y::'a::ab_unital_quantale)"
by (simp add: bres_fres_eq)
lemma lsl_prelocalic_var: "\<nu> x \<cdot> y \<le> \<nu> x \<sqinter> \<nu> (y::'a::ab_unital_quantale)"
by (metis inf_top.right_neutral lsided_le4 lsided_times_top lsl_def)
lemma dang3: "(\<nu> x \<sqinter> y) \<cdot> z \<le> \<nu> x \<sqinter> (y \<cdot> (z::'a::ab_unital_quantale))"
by (metis lsided_le4 lsided_times_top lsl_def)
lemma lsl_prelocalic: "\<nu> x \<cdot> \<nu> y \<le> \<nu> x \<sqinter> \<nu> (y::'a::ab_unital_quantale)"
by (metis lsided_times_le_inf lsided_times_top lsl_def)
lemma lsl_local: "x \<cdot> \<nu> y \<le> \<nu> (x \<cdot> (y::'a::ab_unital_quantale))"
by (simp add: lsl_def mult.left_commute)
lemma lsl_local_eq: "x \<cdot> \<nu> y = \<nu> (x \<cdot> (y::'a::ab_unital_quantale))"
by (simp add: lsl_def mult.left_commute)
text \<open>Relative pseudocomplementation can be defined on the subquantale induced by lsu.\<close>
definition "rpc x y = \<nu>\<^sup>\<natural> (-x \<squnion> (y::'a::bool_unital_quantale))"
lemma lsl_rpc [simp]: "\<nu> (rpc x y) = rpc x y"
by (metis lsl_lsu_prop_eq o_apply rpc_def)
lemma lsu_rpc [simp]: "\<nu>\<^sup>\<natural> (rpc x y) = rpc x y"
using lsided_def3 lsl_rpc by blast
lemma lsl_rpc_galois: "(x \<sqinter> \<nu> z \<le> y) = (z \<le> rpc x (y::'a::bool_unital_quantale))"
unfolding rpc_def by (metis adj_def inf_commute ls_galois shunt1)
lemma lsl_rpc_galois_var: "x \<sqinter> \<nu> z \<le> y \<longleftrightarrow> \<nu> z \<le> rpc x y"
by (metis adj_def ls_galois lsl_rpc_galois lsu_rpc)
lemma rpc_expl_aux: "\<Squnion>{z. x \<sqinter> \<nu> z \<le> y} = \<Squnion>{z. x \<sqinter> z \<le> y \<and> \<nu> z = (z::'a::bool_unital_quantale)}"
proof (rule antisym)
show "\<Squnion>{z. x \<sqinter> \<nu> z \<le> y} \<le> \<Squnion>{z. x \<sqinter> z \<le> y \<and> \<nu> z = z}"
apply (rule Sup_mono, safe)
by (smt lsided_eq lsided_prop lsl_def lsl_lsu_prop_eq lsl_rpc_galois mem_Collect_eq o_apply rpc_def slsd_def)
show " \<Squnion>{z. x \<sqinter> z \<le> y \<and> \<nu> z = z} \<le> \<Squnion>{z. x \<sqinter> \<nu> z \<le> y}"
by (simp add: Collect_mono_iff Sup_subset_mono)
qed
lemma rpc_expl: "rpc x y = \<Squnion>{z. x \<sqinter> z \<le> y \<and> \<nu> z = (z::'a::bool_unital_quantale)}"
proof-
have "rpc x y = \<Squnion>{z. z \<le> rpc x y}"
using Sup_atMost atMost_def by metis
also have "... = \<Squnion>{z. x \<sqinter> \<nu> z \<le> y}"
using lsl_rpc_galois by metis
finally show ?thesis
by (simp add: rpc_expl_aux)
qed
subsection \<open>Non-Preservation and Lack of Closure\<close>
context bool_ab_unital_quantale
begin
text \<open>A few counterexamples deal in particular with lack of closure with respect to boolean complementation.\<close>
lemma "\<nu>\<^sup>\<natural> (x \<cdot> y) = \<nu>\<^sup>\<natural> x \<cdot> \<nu>\<^sup>\<natural> y" (*nitpick *)
oops
lemma "\<nu> 1 = 1" (*nitpick*)
oops
lemma "\<nu>\<^sup>\<natural> x = \<nu> x" (*nitpick*)
oops
lemma "\<nu>\<^sup>\<natural> (\<Squnion> P) = \<Squnion>{\<nu>\<^sup>\<natural> p |p. p \<in> P}" (*nitpick*)
oops
lemma "rpc (\<nu>\<^sup>\<natural> x) (\<nu>\<^sup>\<natural> y) = -(\<nu>\<^sup>\<natural> x) \<squnion> (\<nu>\<^sup>\<natural> y)" (*nitpick*)
oops
lemma "rpc (\<nu> x) (\<nu> y) = -(\<nu> x) \<squnion> (\<nu> y)" (*nitpick*)
oops
lemma "\<nu> (-(\<nu>\<^sup>\<natural> x) \<squnion> (\<nu>\<^sup>\<natural> y)) = -(\<nu>\<^sup>\<natural> x) \<squnion> (\<nu>\<^sup>\<natural> y)" (*nitpick*)
oops
lemma "\<nu> (-(\<nu> x) \<squnion> (\<nu> y)) = -(\<nu> x) \<squnion> (\<nu> y)" (*nitpick*)
oops
lemma "\<nu> x \<cdot> \<nu> y = \<nu> x \<sqinter> \<nu> y" (*nitpick this rules out localic nuclei.*)
oops
lemma "\<nu> (- (\<nu> x)) = - (\<nu> x)" (*nitpick*)
oops
lemma "\<nu>\<^sup>\<natural> (- (\<nu>\<^sup>\<natural> x)) = - (\<nu>\<^sup>\<natural> x)" (*nitpick*)
oops
lemma "\<nu> (- (\<nu> x) \<squnion> (\<nu> y)) = - (\<nu> x) \<squnion> (\<nu> y)" (*nitpick*)
oops
lemma "\<nu>\<^sup>\<natural> (- (\<nu>\<^sup>\<natural> x) \<squnion> (\<nu>\<^sup>\<natural> y)) = - (\<nu>\<^sup>\<natural> x) \<squnion> (\<nu>\<^sup>\<natural> y)" (*nitpick*)
oops
end
subsection \<open>Properties of Quotient Algebras and Subalgebras\<close>
text \<open>Finally I replay Rosenthal's quotient and subalgebra proofs for nuclei and conuclei in the concrete setting of left-sided elements. Ideally
it should be sufficient to show that lsl is a (quantale with) nucleus and then pick up the quotient algebra proof from the nucleus section. But there is no
way of achieving this, apart from creating a type class for quantales with the lsl operation. This seems bizarre, since lsl is just a definition.
On the other hand, an approach in which interpretations are used instead of instantiations might do the job.\<close>
text \<open>The first proof shows that lsu, as a conucleus, yields a subalgebra.\<close>
typedef (overloaded) 'a lsu_set = "Fix (\<nu>\<^sup>\<natural>::'a::unital_quantale \<Rightarrow> 'a)"
using Fix_lsu_iff lsu_bot by blast
setup_lifting type_definition_lsu_set
instantiation lsu_set :: (unital_quantale) quantale
begin
lift_definition times_lsu_set :: "'a lsu_set \<Rightarrow> 'a lsu_set \<Rightarrow> 'a lsu_set" is times
using comp_closed_set_def lsu_comp_closed by blast
lift_definition Inf_lsu_set :: "'a lsu_set set \<Rightarrow> 'a lsu_set" is Inf
by (metis Fix_lsl_lsu Inf_closed_set_def lsl_Inf_closed subset_eq)
lift_definition Sup_lsu_set :: "'a lsu_set set \<Rightarrow> 'a lsu_set" is Sup
using Sup_closed_set_def lsu_Sup_closed subsetI by blast
lift_definition bot_lsu_set :: "'a lsu_set" is "\<bottom>"
by simp
lift_definition sup_lsu_set :: "'a lsu_set \<Rightarrow> 'a lsu_set \<Rightarrow> 'a lsu_set" is sup
by (metis Fix_lsu_iff lsu_sup_closed_var)
lift_definition top_lsu_set :: "'a lsu_set" is "\<top>"
by simp
lift_definition inf_lsu_set :: "'a lsu_set \<Rightarrow> 'a lsu_set \<Rightarrow> 'a lsu_set" is inf
by (metis Fix_lsu_iff lsu_inf_closed_var)
lift_definition less_eq_lsu_set :: "'a lsu_set \<Rightarrow> 'a lsu_set \<Rightarrow> bool" is less_eq .
lift_definition less_lsu_set :: "'a lsu_set \<Rightarrow> 'a lsu_set \<Rightarrow> bool" is less .
instance
by (intro_classes; transfer, simp_all add: mult.assoc Inf_lower Inf_greatest Sup_upper Sup_least less_le_not_le Sup_distr Sup_distl)
end
text \<open>This proof checks simply closure conditions, as one would expect for establishing a subalgebra.\<close>
instance lsu_set :: (bool_ab_unital_quantale) distrib_ab_quantale
apply (intro_classes; transfer)
apply (simp add: mult.commute)
using sup_inf_distrib1 by blast
typedef (overloaded) 'a lsl_set = "range (\<nu>::'a:: unital_quantale \<Rightarrow> 'a)"
by blast
setup_lifting type_definition_lsl_set
text \<open>The final statement shows that lsu, as a nucleus, yields a quotient algebra.\<close>
instantiation lsl_set :: (ab_unital_quantale) ab_unital_quantale
begin
lift_definition one_lsl_set :: "'a::ab_unital_quantale lsl_set" is "\<nu> 1"
by blast
lift_definition Inf_lsl_set :: "'a lsl_set set \<Rightarrow> 'a lsl_set" is "\<lambda>X. \<nu> (\<Sqinter>X)"
by blast
lift_definition Sup_lsl_set :: "'a lsl_set set \<Rightarrow> 'a lsl_set" is "\<lambda>X. \<nu> (\<Squnion>X)"
by blast
lift_definition bot_lsl_set :: "'a lsl_set" is "\<nu> \<bottom>"
by blast
lift_definition sup_lsl_set :: "'a lsl_set \<Rightarrow> 'a lsl_set \<Rightarrow> 'a lsl_set" is "\<lambda>x y. \<nu> x \<squnion> \<nu> y"
by auto (metis lsl_sup_closed_var)
lift_definition top_lsl_set :: "'a lsl_set" is "\<nu> \<top>"
by blast
lift_definition inf_lsl_set :: "'a lsl_set \<Rightarrow> 'a lsl_set \<Rightarrow> 'a lsl_set" is "\<lambda>x y. \<nu> x \<sqinter> \<nu> y"
by (meson lsl_inf_closed_var range_lsl_iff)
lift_definition less_eq_lsl_set :: "'a lsl_set \<Rightarrow> 'a lsl_set \<Rightarrow> bool" is "\<lambda>x y . \<nu> x \<le> \<nu> y" .
lift_definition less_lsl_set :: "'a lsl_set \<Rightarrow> 'a lsl_set \<Rightarrow> bool" is "\<lambda>x y. \<nu> x \<le> \<nu> y \<and> x \<noteq> y" .
lift_definition times_lsl_set :: "'a lsl_set \<Rightarrow> 'a lsl_set \<Rightarrow> 'a lsl_set" is "\<lambda> x y. \<nu> (x \<cdot> y)"
by blast
instance
apply (standard; transfer)
apply (simp add: lsl_local_eq mult.commute mult.left_commute)
apply (simp add: mult.commute)
apply auto[1]
apply fastforce+
apply auto[1]
apply fastforce+
apply (metis Inf_lower lsl_iso_var range_lsl_iff)
apply (metis Inf_greatest lsl_iso_var range_lsl_iff)
apply (metis Sup_upper lsl_iso_var range_lsl_iff)
apply (metis Sup_least lsl_iso_var range_lsl_iff)
apply fastforce+
apply (smt Sup_distl image_cong lsl_local_eq mult.commute)
apply (smt Sup_distl image_cong lsl_local_eq mult.commute)
apply (simp add: lsl_def cong del: image_cong_simp)
apply (simp add: lsl_local_eq)
done
end
text \<open>This proof checks preservation properties instead of closure ones.\<close>
end
diff --git a/thys/Quantales/Quantale_Star.thy b/thys/Quantales/Quantale_Star.thy
--- a/thys/Quantales/Quantale_Star.thy
+++ b/thys/Quantales/Quantale_Star.thy
@@ -1,158 +1,158 @@
(* Title: Star and Fixpoints in Quantales
Author: Georg Struth
Maintainer: Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Star and Fixpoints in Quantales\<close>
theory Quantale_Star
imports Quantales
Kleene_Algebra.Kleene_Algebra
begin
text \<open>This component formalises properties of the star in quantales. For pre-quantales these are modelled as
fixpoints. For weak quantales they are given by iteration.\<close>
subsection \<open>Star and Fixpoints in Pre-Quantales\<close>
context unital_near_quantale
begin
definition "qiter_fun x y = (\<squnion>) x \<circ> (\<cdot>) y"
definition "qiterl x y = lfp (qiter_fun x y)"
definition "qiterg x y = gfp (qiter_fun x y)"
abbreviation "qiterl_id \<equiv> qiterl 1"
abbreviation "qiterq_id \<equiv> qiterg 1"
definition "qstar x = (\<Squnion>i. x ^ i)"
lemma qiter_fun_exp: "qiter_fun x y z = x \<squnion> y \<cdot> z"
unfolding qiter_fun_def by simp
end
text \<open>Many properties of fixpoints have been developed for Isabelle's monotone functions. These carry
two type parameters, and must therefore be used outside of contexts.\<close>
lemma iso_qiter_fun: "mono (\<lambda>z::'a::unital_pre_quantale. qiter_fun x y z)"
by (metis monoI proto_pre_quantale_class.mult_isol qiter_fun_exp sup.idem sup.mono sup.orderI)
text \<open>I derive the left unfold and induction laws of Kleene algebra. I link with the Kleene algebra components
at the end of this section, to bring properties into scope.\<close>
lemma qiterl_unfoldl [simp]: "x \<squnion> y \<cdot> qiterl x y = qiterl (x::'a::unital_pre_quantale) y"
by (metis iso_qiter_fun lfp_unfold qiter_fun_exp qiterl_def)
lemma qiterg_unfoldl [simp]: "x \<squnion> y \<cdot> qiterg x y = qiterg (x::'a::unital_pre_quantale) y"
by (metis gfp_fixpoint iso_qiter_fun qiter_fun_exp qiterg_def)
lemma qiterl_inductl: "x \<squnion> y \<cdot> z \<le> z \<Longrightarrow> qiterl (x::'a::unital_near_quantale) y \<le> z"
by (simp add: lfp_lowerbound qiterl_def qiter_fun_def)
lemma qiterg_coinductl: "z \<le> x \<squnion> y \<cdot> z \<Longrightarrow> z \<le> qiterg (x::'a::unital_near_quantale) y"
by (simp add: gfp_upperbound qiterg_def qiter_fun_def)
context unital_near_quantale
begin
lemma powers_distr: "qstar x \<cdot> y = (\<Squnion>i. x ^ i \<cdot> y)"
by (simp add: qstar_def Sup_distr image_comp)
lemma Sup_iter_unfold: "x ^ 0 \<squnion> (\<Squnion>n. x ^ (Suc n)) = (\<Squnion>n. x ^ n)"
using fSup_unfold by presburger
lemma Sup_iter_unfold_var: "1 \<squnion> (\<Squnion>n. x \<cdot> x ^ n) = (\<Squnion>n. x ^ n)"
by (simp only: Sup_iter_unfold[symmetric]) auto
lemma power_inductr: "z \<squnion> y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x ^ i \<le> y"
- by (induct i, simp_all, metis power_commutes eq_iff mult_isor order.trans mult_assoc power.power.power_Suc)
+ by (induct i, simp_all, metis power_commutes order.eq_iff mult_isor order.trans mult_assoc power.power.power_Suc)
end
context unital_pre_quantale
begin
lemma powers_subdistl: "(\<Squnion>i. x \<cdot> y ^ i) \<le> x \<cdot> qstar y"
by (simp add: SUP_le_iff SUP_upper mult_isol qstar_def)
lemma qstar_subcomm: "qstar x \<cdot> x \<le> x \<cdot> qstar x"
by (simp add: powers_distr powers_subdistl power_commutes)
lemma power_inductl: "z \<squnion> x \<cdot> y \<le> y \<Longrightarrow> x ^ i \<cdot> z \<le> y"
by (induct i, simp_all, metis dual_order.trans mult_isol mult_assoc)
end
subsection \<open>Star and Iteration in Weak Quantales\<close>
context unital_weak_quantale
begin
text \<open>In unital weak quantales one can derive the star axioms of Kleene algebra for iteration. This
generalises the language and relation models from our Kleene algebra components.\<close>
lemma powers_distl: "x \<cdot> qstar y = (\<Squnion>i. x \<cdot> y ^ i)"
by (simp add: qstar_def weak_Sup_distl image_comp)
lemma qstar_unfoldl: "1 \<squnion> x \<cdot> qstar x \<le> qstar x"
by (simp only: powers_distl Sup_iter_unfold_var, simp add: qstar_def)
lemma qstar_comm: "x \<cdot> qstar x = qstar x \<cdot> x"
using power_commutes powers_distr powers_distl by auto
lemma qstar_unfoldr [simp]: "1 \<squnion> qstar x \<cdot> x = qstar x"
using qstar_comm qstar_unfoldl Sup_iter_unfold_var qstar_def powers_distl by auto
lemma qstar_inductl: "z \<squnion> x \<cdot> y \<le> y \<Longrightarrow> qstar x \<cdot> z \<le> y"
by (subst powers_distr, auto intro!: Sup_least power_inductl)
lemma qstar_inductr: "z \<squnion> y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> qstar x \<le> y"
by (subst powers_distl, auto intro!: Sup_least power_inductr)
text \<open>Hence in this setting one also obtains the right Kleene algebra axioms.\<close>
end
sublocale unital_weak_quantale \<subseteq> uwqlka: left_kleene_algebra_zerol "(\<squnion>)" "(\<cdot>)" 1 \<bottom> "(\<le>)" "(<)" qstar
apply unfold_locales
using qstar_unfoldl apply blast
by (simp add: qstar_inductl)
sublocale unital_quantale \<subseteq> uqka: kleene_algebra "(\<squnion>)" "(\<cdot>)" 1 \<bottom> "(\<le>)" "(<)" qstar
by unfold_locales (simp add: qstar_inductr)
text \<open>The star is indeed the least fixpoint.\<close>
lemma qstar_qiterl: "qstar (x::'a::unital_weak_quantale) = qiterl_id x"
by (simp add: antisym qiterl_inductl uwqlka.star_inductl_one)
context unital_weak_quantale
begin
subsection \<open>Deriving the Star Axioms of Action Algebras\<close>
text \<open>Finally the star axioms of action algebras are derived.\<close>
lemma act_star1: "1 \<squnion> x \<squnion> (qstar x) \<cdot> (qstar x) \<le> (qstar x)"
by simp
lemma (in unital_quantale) act_star3: "qstar (x \<rightarrow> x) \<le> x \<rightarrow> x"
by (simp add: bres_canc1 bres_galois_imp uqka.star_inductr_var)
lemma act_star3_var: "qstar (x \<leftarrow> x) \<le> x \<leftarrow> x"
using fres_galois qstar_inductl by auto
end
text \<open>An integration of action algebras requires first resolving some notational issues within the components
where these algebras are located.\<close>
end
diff --git a/thys/Quantales/Quantales.thy b/thys/Quantales/Quantales.thy
--- a/thys/Quantales/Quantales.thy
+++ b/thys/Quantales/Quantales.thy
@@ -1,696 +1,696 @@
(* Title: Quantales
Author: Georg Struth
Maintainer: Georg Struth <g.struth@sheffield.ac.uk>
Contributions by Brijesh Dongol, Victor Gomes, Ian Hayes
*)
section \<open>Quantales\<close>
theory Quantales
imports
"Order_Lattice_Props.Closure_Operators"
Kleene_Algebra.Dioid
begin
subsection \<open>Families of Proto-Quantales\<close>
text \<open>Proto-Quanales are complete lattices equipped with an operation of composition or multiplication
that need not be associative. The notation in this component differs from Rosenthal's \cite{Rosenthal90}, but is consistent with the one we use
for semirings and Kleene algebras.\<close>
class proto_near_quantale = complete_lattice + times +
assumes Sup_distr: "\<Squnion>X \<cdot> y = (\<Squnion>x \<in> X. x \<cdot> y)"
lemma Sup_pres_multr: "Sup_pres (\<lambda>(z::'a::proto_near_quantale). z \<cdot> y)"
unfolding fun_eq_iff comp_def Sup_distr by simp
lemma sup_pres_multr: "sup_pres (\<lambda>(z::'a::proto_near_quantale). z \<cdot> y)"
using Sup_pres_multr Sup_sup_pres by fastforce
lemma bot_pres_multr: "bot_pres (\<lambda>(z::'a::proto_near_quantale). z \<cdot> y)"
by (metis SUP_empty Sup_distr Sup_empty)
context proto_near_quantale
begin
lemma mult_botl [simp]: "\<bottom> \<cdot> x = \<bottom>"
proof -
have "\<bottom> \<cdot> x = (\<Squnion>a\<in>{}. a \<cdot> x)"
using Sup_distr Sup_empty by blast
thus ?thesis
by simp
qed
lemma sup_distr: "(x \<squnion> y) \<cdot> z = (x \<cdot> z) \<squnion> (y \<cdot> z)"
by (smt SUP_empty SUP_insert Sup_distr sup_Sup sup_bot.right_neutral)
lemma mult_isor: "x \<le> y \<Longrightarrow> x \<cdot> z \<le> y \<cdot> z"
by (metis sup.absorb_iff1 sup_distr)
text \<open>Left and right residuals can be defined in every proto-nearquantale.\<close>
definition bres :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<rightarrow>" 60) where
"x \<rightarrow> z = \<Squnion>{y. x \<cdot> y \<le> z}"
definition fres :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<leftarrow>" 60) where
"z \<leftarrow> y = \<Squnion>{x. x \<cdot> y \<le> z}"
text \<open>The left one is a right adjoint to composition. For the right one, additional assumptions are needed\<close>
lemma bres_galois_imp: "x \<cdot> y \<le> z \<Longrightarrow> y \<le> x \<rightarrow> z"
by (simp add: Sup_upper bres_def)
lemma fres_galois: "(x \<cdot> y \<le> z) = (x \<le> z \<leftarrow> y)"
proof
show "x \<cdot> y \<le> z \<Longrightarrow> x \<le> z \<leftarrow> y"
by (simp add: Sup_upper fres_def)
next
assume "x \<le> z \<leftarrow> y"
hence "x \<cdot> y \<le> \<Squnion>{x. x \<cdot> y \<le> z} \<cdot> y"
by (simp add: fres_def mult_isor)
also have "... = \<Squnion>{x \<cdot> y |x. x \<cdot> y \<le> z}"
by (simp add: Sup_distr setcompr_eq_image)
also have "... \<le> z"
by (rule Sup_least, auto)
finally show "x \<cdot> y \<le> z" .
qed
end
lemma fres_adj: "(\<lambda>(x::'a::proto_near_quantale). x \<cdot> y) \<stileturn> (\<lambda>x. x \<leftarrow> y)"
by (simp add: adj_def fres_galois)
context proto_near_quantale
begin
lemma fres_canc1: "(y \<leftarrow> x) \<cdot> x \<le> y"
by (simp add: fres_galois)
lemma fres_canc2: "y \<le> (y \<cdot> x) \<leftarrow> x"
using fres_galois by force
lemma inf_fres: "y \<cdot> x = \<Sqinter>{z. y \<le> z \<leftarrow> x}"
by (metis (mono_tags, lifting) fres_canc2 Inf_eqI fres_galois mem_Collect_eq)
lemma bres_iso: "x \<le> y \<Longrightarrow> z \<rightarrow> x \<le> z \<rightarrow> y"
using Sup_le_iff bres_def bres_galois_imp by force
lemma bres_anti: "x \<le> y \<Longrightarrow> y \<rightarrow> z \<le> x \<rightarrow> z"
by (smt Sup_le_iff bres_def bres_galois_imp fres_galois order_trans mem_Collect_eq)
lemma fres_iso: "x \<le> y \<Longrightarrow> x \<leftarrow> z \<le> y \<leftarrow> z"
using fres_galois dual_order.trans by blast
lemma bres_top_top [simp]: "\<top> \<rightarrow> \<top> = \<top>"
by (simp add: bres_galois_imp dual_order.antisym)
lemma fres_top_top [simp]: "\<top> \<leftarrow> \<top> = \<top>"
using fres_galois top_greatest top_le by blast
lemma bres_bot_bot [simp]: "\<bottom> \<rightarrow> \<bottom> = \<top>"
by (simp add: bres_galois_imp top_le)
lemma left_sided_localp: "\<top> \<cdot> y = y \<Longrightarrow> x \<cdot> y \<le> y"
by (metis mult_isor top_greatest)
lemma fres_sol: "((y \<leftarrow> x) \<cdot> x = y) = (\<exists>z. z \<cdot> x = y)"
using dual_order.antisym fres_canc1 fres_canc2 mult_isor by fastforce
lemma sol_fres: "((y \<cdot> x) \<leftarrow> x = y) = (\<exists>z. y = z \<leftarrow> x)"
- by (metis fres_canc1 fres_canc2 fres_sol eq_iff fres_galois)
+ by (metis fres_canc1 fres_canc2 fres_sol order.eq_iff fres_galois)
end
class proto_pre_quantale = proto_near_quantale +
assumes Sup_subdistl: "(\<Squnion>y \<in> Y. x \<cdot> y) \<le> x \<cdot> \<Squnion>Y"
begin
lemma sup_subdistl: "(x \<cdot> y) \<squnion> (x \<cdot> z) \<le> x \<cdot> (y \<squnion> z)"
by (smt SUP_empty SUP_insert Sup_subdistl sup_Sup sup_bot_right)
lemma mult_isol: "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
by (metis le_iff_sup le_sup_iff sup_subdistl)
lemma fres_anti: "x \<le> y \<Longrightarrow> z \<leftarrow> y \<le> z \<leftarrow> x"
using dual_order.trans fres_galois mult_isol by blast
end
class weak_proto_quantale = proto_near_quantale +
assumes weak_Sup_distl: "Y \<noteq> {} \<Longrightarrow> x \<cdot> \<Squnion>Y = (\<Squnion>y \<in> Y. x \<cdot> y)"
begin
subclass proto_pre_quantale
proof unfold_locales
have a: "\<And>x Y. Y = {} \<Longrightarrow> (\<Squnion>y \<in> Y. x \<cdot> y) \<le> x \<cdot> \<Squnion>Y"
by simp
have b: "\<And>x Y. Y \<noteq> {} \<Longrightarrow> (\<Squnion>y \<in> Y. x \<cdot> y) \<le> x \<cdot> \<Squnion>Y"
by (simp add: weak_Sup_distl)
show "\<And>x Y. (\<Squnion>y \<in> Y. x \<cdot> y) \<le> x \<cdot> \<Squnion>Y"
using a b by blast
qed
lemma sup_distl: "x \<cdot> (y \<squnion> z) = (x \<cdot> y) \<squnion> (x \<cdot> z)"
using weak_Sup_distl[where Y="{y, z}"] by (fastforce intro!: Sup_eqI)
lemma "y \<le> x \<rightarrow> z \<longrightarrow> x \<cdot> y \<le> z" (* nitpick [expect = genuine] *)
oops
end
class proto_quantale = proto_near_quantale +
assumes Sup_distl: "x \<cdot> \<Squnion>Y = (\<Squnion>y \<in> Y. x \<cdot> y)"
lemma Sup_pres_multl: "Sup_pres (\<lambda>(z::'a::proto_quantale). x \<cdot> z)"
unfolding fun_eq_iff comp_def Sup_distl by simp
lemma sup_pres_multl: "sup_pres (\<lambda>(z::'a::proto_quantale). x \<cdot> z)"
by (metis (no_types, lifting) SUP_insert Sup_distl Sup_empty Sup_insert sup_bot_right)
lemma bot_pres_multl: "bot_pres (\<lambda>(z::'a::proto_quantale). x \<cdot> z)"
by (metis SUP_empty Sup_distl Sup_empty)
context proto_quantale
begin
subclass weak_proto_quantale
by standard (simp add: Sup_distl)
lemma mult_botr [simp]: "x \<cdot> \<bottom> = \<bottom>"
by (smt image_empty Sup_distl Sup_empty)
text \<open>Now there is also an adjunction for the other residual.\<close>
lemma bres_galois: "x \<cdot> y \<le> z \<longleftrightarrow> y \<le> x \<rightarrow> z"
proof
show "x \<cdot> y \<le> z \<Longrightarrow> y \<le> x \<rightarrow> z"
by (simp add: Sup_upper bres_def)
next
assume "y \<le> x \<rightarrow> z"
hence "x \<cdot> y \<le> x \<cdot> \<Squnion>{y. x \<cdot> y \<le> z}"
by (simp add: bres_def mult_isol)
also have "... = \<Squnion>{x \<cdot> y |y. x \<cdot> y \<le> z}"
by (simp add: Sup_distl setcompr_eq_image)
also have "... \<le> z"
by (rule Sup_least, safe)
finally show "x \<cdot> y \<le> z" .
qed
end
lemma bres_adj: "(\<lambda>(y::'a::proto_quantale). x \<cdot> y) \<stileturn> (\<lambda>y. x \<rightarrow> y)"
by (simp add: adj_def bres_galois)
context proto_quantale
begin
lemma bres_canc1: "x \<cdot> (x \<rightarrow> y) \<le> y"
by (simp add: bres_galois)
lemma bres_canc2: "y \<le> x \<rightarrow> (x \<cdot> y)"
by (simp add: bres_galois_imp)
lemma inf_bres: "x \<cdot> y = \<Sqinter>{z. y \<le> x \<rightarrow> z}"
using bres_galois fres_galois inf_fres by force
lemma bres_sol: "(x \<cdot> (x \<rightarrow> y) = y) = (\<exists>z. x \<cdot> z = y)"
- using bres_galois antisym mult_isol by force
+ using bres_galois order.antisym mult_isol by force
lemma sol_bres: "(x \<rightarrow> (x \<cdot> y) = y) = (\<exists>z. y = x \<rightarrow> z)"
- by (metis bres_canc1 bres_canc2 bres_iso eq_iff)
+ by (metis bres_canc1 bres_canc2 bres_iso order.eq_iff)
end
lemma bres_fres_clop: "clop (\<lambda>x::'a::proto_quantale. y \<leftarrow> (x \<rightarrow> y))"
unfolding clop_def comp_def mono_def le_fun_def
by (metis bres_anti bres_canc1 bres_galois_imp fres_anti fres_galois id_apply)
lemma fres_bres_clop: "clop (\<lambda>x::'a::proto_quantale. (y \<leftarrow> x) \<rightarrow> y)"
unfolding clop_def comp_def mono_def le_fun_def
by (metis bres_anti bres_canc1 bres_galois_imp fres_anti fres_canc1 fres_galois id_apply)
subsection \<open>Families of Quantales\<close>
class near_quantale = proto_near_quantale + semigroup_mult
sublocale near_quantale \<subseteq> nsrnq: near_dioid "(\<squnion>)" "(\<cdot>)" "(\<le>)" "(<)"
apply unfold_locales
apply (simp add: sup_assoc)
apply (simp add: sup_commute)
apply (simp_all add: sup_distr)
apply (simp add: le_iff_sup)
by auto
context near_quantale
begin
lemma fres_curry: "(z \<leftarrow> y) \<leftarrow> x = z \<leftarrow> (x \<cdot> y)"
- by (metis eq_iff fres_canc1 fres_galois mult_assoc)
+ by (metis order.eq_iff fres_canc1 fres_galois mult_assoc)
end
class unital_near_quantale = near_quantale + monoid_mult
sublocale unital_near_quantale \<subseteq> nsrnqo: near_dioid_one "(\<squnion>)" "(\<cdot>)" "1""(\<le>)" "(<)"
by (unfold_locales, simp_all)
context unital_near_quantale
begin
definition iter :: "'a \<Rightarrow> 'a" where
"iter x \<equiv> \<Sqinter>i. x ^ i"
lemma iter_ref [simp]: "iter x \<le> 1"
by (metis iter_def Inf_lower power.power_0 rangeI)
lemma le_top: "x \<le> \<top> \<cdot> x"
by (metis mult.left_neutral mult_isor top_greatest)
lemma top_times_top [simp]: "\<top> \<cdot> \<top> = \<top>"
by (simp add: le_top top_le)
lemma bres_one: "1 \<le> x \<rightarrow> x"
by (simp add: bres_galois_imp)
lemma fres_one: "1 \<le> x \<leftarrow> x"
using fres_galois by fastforce
end
class pre_quantale = proto_pre_quantale + semigroup_mult
begin
subclass near_quantale ..
lemma fres_interchange: "z \<cdot> (x \<leftarrow> y) \<le> (z \<cdot> x) \<leftarrow> y"
using Sup_upper fres_canc1 fres_def mult_isol mult_assoc by fastforce
end
sublocale pre_quantale \<subseteq> psrpq: pre_dioid "(\<squnion>)" "(\<cdot>)" "(\<le>)" "(<)"
by (unfold_locales, simp add: mult_isol)
class unital_pre_quantale = pre_quantale + monoid_mult
begin
subclass unital_near_quantale ..
text \<open>Abstract rules of Hoare logic without the star can be derived.\<close>
lemma h_w1: "x \<le> x' \<Longrightarrow> x' \<cdot> y \<le> z \<Longrightarrow> x \<cdot> y \<le> z"
by (simp add: fres_galois)
lemma h_w2: "x \<cdot> y \<le> z' \<Longrightarrow> z' \<le> z \<Longrightarrow> x \<cdot> y \<le> z"
using order_trans by blast
lemma h_seq: "x \<cdot> v \<le> z \<Longrightarrow> y \<cdot> w \<le> v \<Longrightarrow> x \<cdot> y \<cdot> w \<le> z"
using dual_order.trans mult_isol mult_assoc by presburger
lemma h_sup: "x \<cdot> w \<le> z \<Longrightarrow> y \<cdot> w \<le> z \<Longrightarrow> (x \<squnion> y) \<cdot> w \<le> z"
by (simp add: fres_galois)
lemma h_Sup: "\<forall>x \<in> X. x \<cdot> w \<le> z \<Longrightarrow> \<Squnion>X \<cdot> w \<le> z"
by (simp add: Sup_least fres_galois)
end
sublocale unital_pre_quantale \<subseteq> psrpqo: pre_dioid_one "(\<squnion>)" "(\<cdot>)" "1" "(\<le>)" "(<)"..
class weak_quantale = weak_proto_quantale + semigroup_mult
begin
subclass pre_quantale ..
text \<open>The following counterexample shows an important consequence of weakness:
the absence of right annihilation.\<close>
lemma "x \<cdot> \<bottom> = \<bottom>" (*nitpick[expect=genuine]*)
oops
end
class unital_weak_quantale = weak_quantale + monoid_mult
lemma (in unital_weak_quantale) "x \<cdot> \<bottom> = \<bottom>" (*nitpick[expect=genuine]*)
oops
subclass (in unital_weak_quantale) unital_pre_quantale ..
sublocale unital_weak_quantale \<subseteq> wswq: dioid_one_zerol "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" "(\<le>)" "(<)"
by (unfold_locales, simp_all add: sup_distl)
class quantale = proto_quantale + semigroup_mult
begin
subclass weak_quantale ..
lemma Inf_subdistl: "x \<cdot> \<Sqinter>Y \<le> (\<Sqinter>y \<in> Y. x \<cdot> y)"
by (auto intro!: Inf_greatest Inf_lower mult_isol)
lemma Inf_subdistr: "\<Sqinter> X \<cdot> y \<le> (\<Sqinter>x \<in> X. x \<cdot> y)"
by (auto intro!: Inf_greatest Inf_lower mult_isor)
lemma fres_bot_bot [simp]: "\<bottom> \<leftarrow> \<bottom> = \<top>"
by (simp add: fres_def)
lemma bres_interchange: "(x \<rightarrow> y) \<cdot> z \<le> x \<rightarrow> (y \<cdot> z)"
by (metis bres_canc1 bres_galois mult_isor mult_assoc)
lemma bres_curry: "x \<rightarrow> (y \<rightarrow> z) = (y \<cdot> x) \<rightarrow> z"
by (metis bres_canc1 bres_galois dual_order.antisym mult_assoc)
lemma fres_bres: "x \<rightarrow> (y \<leftarrow> z) = (x \<rightarrow> y) \<leftarrow> z"
proof-
{fix w
have "(w \<le> x \<rightarrow> (y \<leftarrow> z)) = (x \<cdot> w \<le> y \<leftarrow> z)"
by (simp add: bres_galois)
also have "... = (x \<cdot> w \<cdot> z \<le> y)"
by (simp add: fres_galois)
also have "... = (w \<cdot> z \<le> x \<rightarrow> y)"
by (simp add: bres_galois mult_assoc)
also have "... = (w \<le> (x \<rightarrow> y) \<leftarrow> z)"
by (simp add: fres_galois)
finally have "(w \<le> x \<rightarrow> (y \<leftarrow> z)) = (w \<le> (x \<rightarrow> y) \<leftarrow> z)".}
thus ?thesis
- using eq_iff by blast
+ using order.eq_iff by blast
qed
end
class quantale_with_dual = quantale + complete_lattice_with_dual
class unital_quantale = quantale + monoid_mult
class unital_quantale_with_dual = unital_quantale + quantale_with_dual
subclass (in unital_quantale) unital_weak_quantale ..
sublocale unital_quantale \<subseteq> wswq: dioid_one_zero "(\<squnion>)" "(\<cdot>)" "1" "\<bottom>" "(\<le>)" "(<)"
by (unfold_locales, simp)
class ab_quantale = quantale + ab_semigroup_mult
begin
lemma bres_fres_eq: "x \<rightarrow> y = y \<leftarrow> x"
by (simp add: fres_def bres_def mult_commute)
end
class ab_unital_quantale = ab_quantale + unital_quantale
sublocale complete_heyting_algebra \<subseteq> chaq: ab_unital_quantale "(\<sqinter>)" _ _ _ _ _ _ _ _ \<top>
by (unfold_locales, simp add: inf.assoc, simp_all add: inf.assoc ch_dist inf.commute)
class distrib_quantale = quantale + distrib_lattice
class bool_quantale = quantale + complete_boolean_algebra_alt
class distrib_unital_quantale = unital_quantale + distrib_lattice
class bool_unital_quantale = unital_quantale + complete_boolean_algebra_alt
class distrib_ab_quantale = distrib_quantale + ab_quantale
class bool_ab_quantale = bool_quantale + ab_quantale
class distrib_ab_unital_quantale = distrib_quantale + unital_quantale
class bool_ab_unital_quantale = bool_ab_quantale + unital_quantale
sublocale complete_boolean_algebra \<subseteq> cba_quantale: bool_ab_unital_quantale inf _ _ _ _ _ _ _ _ _ _ \<top>
by (unfold_locales, simp add: inf.assoc, simp_all add: inf.commute Setcompr_eq_image inf_Sup Sup_inf)
context complete_boolean_algebra
begin
text \<open>In this setting, residuation is classical implication.\<close>
lemma cba_bres1: "x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> cba_quantale.bres y z"
using cba_quantale.bres_galois inf.commute by fastforce
lemma cba_bres2: "x \<le> -y \<squnion> z \<longleftrightarrow> x \<le> cba_quantale.bres y z"
using cba_bres1 shunt1 by auto
lemma cba_bres_prop: "cba_quantale.bres x y = -x \<squnion> y"
- using cba_bres2 eq_iff by blast
+ using cba_bres2 order.eq_iff by blast
end
subsection \<open>Quantales Based on Sup-Lattices and Inf-Lattices\<close>
text \<open>These classes are defined for convenience in instantiation and interpretation proofs, or likewise.
They are useful, e.g., in the context of predicate transformers, where only one of Sup or Inf may be well behaved.\<close>
class Sup_quantale = Sup_lattice + semigroup_mult +
assumes Supq_distr: "\<Squnion>X \<cdot> y = (\<Squnion>x \<in> X. x \<cdot> y)"
and Supq_distl: "x \<cdot> \<Squnion>Y = (\<Squnion>y \<in> Y. x \<cdot> y)"
class unital_Sup_quantale = Sup_quantale + monoid_mult
class Inf_quantale = Inf_lattice + monoid_mult +
assumes Supq_distr: "\<Sqinter>X \<cdot> y = (\<Sqinter>x \<in> X. x \<cdot> y)"
and Supq_distl: "x \<cdot> \<Sqinter>Y = (\<Sqinter>y \<in> Y. x \<cdot> y)"
class unital_Inf_quantale = Inf_quantale + monoid_mult
sublocale Inf_quantale \<subseteq> qdual: Sup_quantale _ Inf "(\<ge>)"
by (unfold_locales, simp_all add: Supq_distr Supq_distl)
sublocale unital_Inf_quantale \<subseteq> uqdual: unital_Sup_quantale _ _ Inf "(\<ge>)"..
sublocale Sup_quantale \<subseteq> supq: quantale _ Infs Sup_class.Sup infs "(\<le>)" le sups bots tops
by (unfold_locales, simp_all add: Supq_distr Supq_distl)
sublocale unital_Sup_quantale \<subseteq> usupq: unital_quantale _ _ Infs Sup_class.Sup infs "(\<le>)" le sups bots tops..
subsection \<open>Products of Quantales\<close>
definition "Inf_prod X = ((\<Sqinter>x \<in> X. fst x), (\<Sqinter>x \<in> X. snd x))"
definition "inf_prod x y = (fst x \<sqinter> fst y, snd x \<sqinter> snd y)"
definition "bot_prod = (bot,bot)"
definition "Sup_prod X = ((\<Squnion>x \<in> X. fst x), (\<Squnion>x \<in> X. snd x))"
definition "sup_prod x y = (fst x \<squnion> fst y, snd x \<squnion> snd y)"
definition "top_prod = (top,top)"
definition "less_eq_prod x y \<equiv> less_eq (fst x) (fst y) \<and> less_eq (snd x) (snd y)"
definition "less_prod x y \<equiv> less_eq (fst x) (fst y) \<and> less_eq (snd x) (snd y) \<and> x \<noteq> y"
definition "times_prod' x y = (fst x \<cdot> fst y, snd x \<cdot> snd y)"
definition "one_prod = (1,1)"
definition "dual_prod x = (\<partial> (fst x),\<partial> (snd x))"
interpretation prod: complete_lattice Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::complete_lattice \<times> 'b::complete_lattice)"
by standard (auto simp add: Inf_prod_def Sup_prod_def inf_prod_def sup_prod_def bot_prod_def top_prod_def less_eq_prod_def less_prod_def Sup_distl Sup_distr intro: Inf_lower Inf_greatest Sup_upper Sup_least)
interpretation prod: complete_lattice_with_dual Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::complete_lattice_with_dual \<times> 'b::complete_lattice_with_dual)" dual_prod
by standard (simp_all add: dual_prod_def fun_eq_iff inj_def Sup_prod_def Inf_prod_def inj_dual_iff Sup_dual_def_var image_comp)
interpretation prod: proto_near_quantale Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::proto_near_quantale \<times> 'b::proto_near_quantale)" times_prod'
by standard (simp add: times_prod'_def Sup_prod_def Sup_distr image_comp)
interpretation prod: proto_quantale Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::proto_quantale \<times> 'b::proto_quantale)" times_prod'
by standard (simp add: times_prod'_def Sup_prod_def less_eq_prod_def Sup_distl image_comp)
interpretation prod: unital_quantale one_prod times_prod' Inf_prod Sup_prod inf_prod less_eq_prod less_prod sup_prod bot_prod "top_prod :: ('a::unital_quantale \<times> 'b::unital_quantale)"
by standard (simp_all add: one_prod_def times_prod'_def ac_simps image_comp)
subsection \<open>Quantale Morphisms\<close>
text \<open>There are various ways of defining quantale morphisms, depending on the application. Following Rosenthal,
I present the most important one.\<close>
abbreviation comp_pres :: "('a::times \<Rightarrow> 'b::times) \<Rightarrow> bool" where
"comp_pres f \<equiv> (\<forall>x y. f (x \<cdot> y) = f x \<cdot> f y)"
abbreviation un_pres :: "('a::one \<Rightarrow> 'b::one) \<Rightarrow> bool" where
"un_pres f \<equiv> (f 1 = 1)"
definition "comp_closed_set X = (\<forall>x \<in> X. \<forall>y \<in> X. x \<cdot> y \<in> X)"
definition "un_closed_set X = (1 \<in> X)"
definition quantale_homset :: "('a::quantale \<Rightarrow> 'b::quantale) set" where
"quantale_homset = {f. comp_pres f \<and> Sup_pres f}"
lemma quantale_homset_iff: "f \<in> quantale_homset = (comp_pres f \<and> Sup_pres f)"
unfolding quantale_homset_def by clarsimp
definition unital_quantale_homset :: "('a::unital_quantale \<Rightarrow> 'b::unital_quantale) set" where
"unital_quantale_homset = {f. comp_pres f \<and> Sup_pres f \<and> un_pres f}"
lemma unital_quantale_homset_iff: "f \<in> unital_quantale_homset = (comp_pres f \<and> Sup_pres f \<and> un_pres f)"
unfolding unital_quantale_homset_def by clarsimp
text \<open>Though Infs can be defined from Sups in any quantale, quantale morphisms do not generally preserve Infs.
A different kind of morphism is needed if this is to be guaranteed.\<close>
lemma "f \<in> quantale_homset \<Longrightarrow> Inf_pres f" (*nitpick*)
oops
text \<open>The images of quantale morphisms are closed under compositions and Sups, hence they form quantales.\<close>
lemma quantale_hom_q_pres: "f \<in> quantale_homset \<Longrightarrow> Sup_closed_set (range f) \<and> comp_closed_set (range f)"
apply safe
apply (simp add: Sup_pres_Sup_closed quantale_homset_iff)
unfolding quantale_homset_iff comp_closed_set_def by (metis (no_types, lifting) imageE range_eqI)
text \<open>Yet the image need not be Inf-closed.\<close>
lemma "f \<in> quantale_homset \<Longrightarrow> Inf_closed_set (range f)" (*nitpick*)
oops
text \<open>Of course Sups are preserved by quantale-morphisms, hence they are the same in subsets as in the original set.
Infs in the subset, however, exist, since they subset forms a quantale in which Infs can be defined, but these are generally
different from the Infs in the superstructure.
This fact is hidden in Isabelle's definition of complete lattices, where Infs are axiomatised. There is no easy way in general to
show that images of quantale morphisms form quantales, though the statement for Sup-quantales is straightforward. I show this for quantic nuclei
and left-sided elements.\<close>
typedef (overloaded) ('a,'b) quantale_homset = "quantale_homset::('a::quantale \<Rightarrow> 'b::quantale) set"
proof-
have a: "comp_pres (\<lambda>x::'a::quantale. bot::'b)"
by simp
have b: "Sup_pres (\<lambda>x::'a::quantale. bot::'b)"
unfolding fun_eq_iff comp_def by simp
hence "(\<lambda>x::'a::quantale. bot::'b) \<in> quantale_homset"
by (simp add: quantale_homset_iff)
thus ?thesis
by auto
qed
setup_lifting type_definition_quantale_homset
text \<open>Interestingly, the following type is not (gobally) inhabited.\<close>
typedef (overloaded) ('a,'b) unital_quantale_homset = "unital_quantale_homset::('a::unital_quantale \<Rightarrow> 'b::unital_quantale) set" (*nitpick*)
oops
lemma quantale_hom_radj:
fixes f :: "'a::quantale_with_dual \<Rightarrow> 'b::quantale_with_dual"
shows "f \<in> quantale_homset \<Longrightarrow> f \<stileturn> radj f"
unfolding quantale_homset_iff by (simp add: Sup_pres_ladj_aux)
lemma quantale_hom_prop1:
fixes f :: "'a::quantale_with_dual \<Rightarrow> 'b::quantale_with_dual"
shows "f \<in> quantale_homset \<Longrightarrow> radj f (f x \<rightarrow> y) = x \<rightarrow> radj f y"
proof-
assume h: "f \<in> quantale_homset"
have "f x \<cdot> f (radj f (f x \<rightarrow> y)) \<le> y"
by (meson h adj_def bres_galois order_refl quantale_hom_radj)
hence "f (x \<cdot> radj f (f x \<rightarrow> y)) \<le> y"
by (metis h quantale_homset_iff)
hence "x \<cdot> radj f (f x \<rightarrow> y) \<le> radj f y"
using adj_def h quantale_hom_radj by blast
hence le: "radj f (f x \<rightarrow> y) \<le> x \<rightarrow> radj f y"
by (simp add: bres_galois)
have "x \<cdot> (x \<rightarrow> radj f y) \<le> radj f y"
by (simp add: bres_canc1)
hence "f (x \<cdot> (x \<rightarrow> radj f y)) \<le> y"
using adj_def h quantale_hom_radj by blast
hence "f x \<cdot> f (x \<rightarrow> radj f y) \<le> y"
by (metis h quantale_homset_iff)
hence "f (x \<rightarrow> radj f y) \<le> f x \<rightarrow> y"
by (simp add: bres_galois)
hence "x \<rightarrow> radj f y \<le> radj f (f x \<rightarrow> y)"
using adj_def h quantale_hom_radj by blast
thus ?thesis
by (simp add: dual_order.antisym le)
qed
lemma quantale_hom_prop2:
fixes f :: "'a::quantale_with_dual \<Rightarrow> 'b::quantale_with_dual"
shows "f \<in> quantale_homset \<Longrightarrow> radj f (y \<leftarrow> f x) = radj f y \<leftarrow> x"
proof-
assume h: "f \<in> quantale_homset"
have "f (radj f (y \<leftarrow> f x)) \<cdot> f x \<le> y"
by (meson adj_def fres_galois h order_refl quantale_hom_radj)
hence "f (radj f (y \<leftarrow> f x) \<cdot> x) \<le> y"
by (metis h quantale_homset_iff)
hence "radj f (y \<leftarrow> f x) \<cdot> x\<le> radj f y"
using adj_def h quantale_hom_radj by blast
hence le: "radj f (y \<leftarrow> f x) \<le> radj f y \<leftarrow> x"
by (simp add: fres_galois)
have "(radj f y \<leftarrow> x) \<cdot> x \<le> radj f y"
by (simp add: fres_canc1)
hence "f ((radj f y \<leftarrow> x) \<cdot> x) \<le> y"
using adj_def h quantale_hom_radj by blast
hence "f (radj f y \<leftarrow> x) \<cdot> f x\<le> y"
by (metis h quantale_homset_iff)
hence "f (radj f y \<leftarrow> x) \<le> y \<leftarrow> f x"
by (simp add: fres_galois)
hence "radj f y \<leftarrow> x\<le> radj f (y \<leftarrow> f x)"
using adj_def h quantale_hom_radj by blast
thus ?thesis
by (simp add: dual_order.antisym le)
qed
definition quantale_closed_maps :: "('a::quantale \<Rightarrow> 'b::quantale) set" where
"quantale_closed_maps = {f. (\<forall>x y. f x \<cdot> f y \<le> f (x \<cdot> y))}"
lemma quantale_closed_maps_iff: "f \<in> quantale_closed_maps = (\<forall> x y. f x \<cdot> f y \<le> f (x \<cdot> y))"
unfolding quantale_closed_maps_def by clarsimp
definition quantale_closed_Sup_maps :: "('a::quantale \<Rightarrow> 'b::quantale) set" where
"quantale_closed_Sup_maps = {f. (\<forall> x y. f x \<cdot> f y \<le> f (x \<cdot> y)) \<and> Sup_pres f}"
lemma quantale_closed_Sup_maps_iff: "f \<in> quantale_closed_Sup_maps = (\<forall> x y. f x \<cdot> f y \<le> f (x \<cdot> y) \<and> Sup_pres f)"
unfolding quantale_closed_Sup_maps_def by clarsimp
definition quantale_closed_unital_maps :: "('a::unital_quantale \<Rightarrow> 'b::unital_quantale) set" where
"quantale_closed_unital_maps = {f. (\<forall> x y. f x \<cdot> f y \<le> f (x \<cdot> y)) \<and> 1 \<le> f 1}"
lemma quantale_closed_unital_maps_iff: "f \<in> quantale_closed_unital_maps = (\<forall> x y. f x \<cdot> f y \<le> f (x \<cdot> y) \<and> 1 \<le> f 1)"
unfolding quantale_closed_unital_maps_def by clarsimp
definition quantale_closed_unital_Sup_maps :: "('a::unital_quantale \<Rightarrow> 'b::unital_quantale) set" where
"quantale_closed_unital_Sup_maps = {f. (\<forall> x y. f x \<cdot> f y \<le> f (x \<cdot> y)) \<and> Sup_pres f \<and> 1 \<le> f 1}"
lemma quantale_closed_unital_Sup_maps_iff: "f \<in> quantale_closed_unital_Sup_maps = (\<forall> x y. f x \<cdot> f y \<le> f (x \<cdot> y) \<and> Sup_pres f \<and> 1 \<le> f 1)"
unfolding quantale_closed_unital_Sup_maps_def by clarsimp
text \<open>Closed maps are the right adjoints of quantale morphisms.\<close>
lemma quantale_hom_closed_map:
fixes f :: "'a::quantale_with_dual \<Rightarrow> 'b::quantale_with_dual"
shows "(f \<in> quantale_homset) \<Longrightarrow> (radj f \<in> quantale_closed_maps)"
proof-
assume h: "f \<in> quantale_homset"
have "\<forall>x y. f (radj f x) \<cdot> f (radj f y) \<le> x \<cdot> y"
by (metis adj_def h order_refl psrpq.mult_isol_var quantale_hom_radj)
hence "\<forall>x y. f (radj f x \<cdot> radj f y) \<le> x \<cdot> y"
by (metis h quantale_homset_iff)
hence "\<forall>x y. radj f x \<cdot> radj f y \<le> radj f (x \<cdot> y)"
using adj_def h quantale_hom_radj by blast
thus ?thesis
by (simp add: quantale_closed_maps_iff)
qed
lemma unital_quantale_hom_closed_unital_map:
fixes f :: "'a::unital_quantale_with_dual \<Rightarrow> 'b::unital_quantale_with_dual"
shows "(f \<in> unital_quantale_homset) \<Longrightarrow> (radj f \<in> quantale_closed_unital_maps)"
by (metis (no_types, hide_lams) adj_def order_refl quantale_closed_maps_iff quantale_closed_unital_maps_iff quantale_hom_closed_map quantale_hom_radj quantale_homset_iff unital_quantale_homset_iff)
end
diff --git a/thys/Quick_Sort_Cost/Randomised_Quick_Sort.thy b/thys/Quick_Sort_Cost/Randomised_Quick_Sort.thy
--- a/thys/Quick_Sort_Cost/Randomised_Quick_Sort.thy
+++ b/thys/Quick_Sort_Cost/Randomised_Quick_Sort.thy
@@ -1,717 +1,717 @@
(*
File: Randomised_Quick_Sort.thy
Author: Manuel Eberl <eberlm@in.tum.de>
Definition and cost analysis of randomised QuickSort (i.e. pivot chosen uniformly at random).
*)
section \<open>Randomised QuickSort\<close>
theory Randomised_Quick_Sort
imports
"HOL-Probability.Probability"
"Landau_Symbols.Landau_More"
Comparison_Sort_Lower_Bound.Linorder_Relations
begin
subsection \<open>Deletion by index\<close>
text \<open>
The following function deletes the $n$-th element of a list.
\<close>
fun delete_index :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"delete_index _ [] = []"
| "delete_index 0 (x # xs) = xs"
| "delete_index (Suc n) (x # xs) = x # delete_index n xs"
lemma delete_index_altdef: "delete_index n xs = take n xs @ drop (Suc n) xs"
by (induction n xs rule: delete_index.induct) simp_all
lemma delete_index_ge_length: "n \<ge> length xs \<Longrightarrow> delete_index n xs = xs"
by (simp add: delete_index_altdef)
lemma length_delete_index [simp]: "n < length xs \<Longrightarrow> length (delete_index n xs) = length xs - 1"
by (simp add: delete_index_altdef)
lemma delete_index_Cons:
"delete_index n (x # xs) = (if n = 0 then xs else x # delete_index (n - 1) xs)"
by (cases n) simp_all
lemma insert_set_delete_index:
"n < length xs \<Longrightarrow> insert (xs ! n) (set (delete_index n xs)) = set xs"
by (induction n xs rule: delete_index.induct) auto
lemma add_mset_delete_index:
"i < length xs \<Longrightarrow> add_mset (xs ! i) (mset (delete_index i xs)) = mset xs"
by (induction i xs rule: delete_index.induct) simp_all
lemma nth_delete_index:
"i < length xs \<Longrightarrow> n < length xs \<Longrightarrow>
delete_index n xs ! i = (if i < n then xs ! i else xs ! Suc i)"
by (auto simp: delete_index_altdef nth_append min_def)
lemma set_delete_index_distinct:
assumes "distinct xs" "n < length xs"
shows "set (delete_index n xs) = set xs - {xs ! n}"
using assms by (induction n xs rule: delete_index.induct) fastforce+
lemma distinct_delete_index [simp, intro]:
assumes "distinct xs"
shows "distinct (delete_index n xs)"
proof (cases "n < length xs")
case True
with assms show ?thesis
by (induction n xs rule: delete_index.induct) (auto simp: set_delete_index_distinct)
qed (simp_all add: delete_index_ge_length assms)
lemma mset_delete_index [simp]:
"i < length xs \<Longrightarrow> mset (delete_index i xs) = mset xs - {# xs!i #}"
by (induction i xs rule: delete_index.induct) simp_all
subsection \<open>Definition\<close>
text \<open>
The following is a functional randomised version of QuickSort that also records the number
of comparisons that were made. The randomisation is in the selection of the pivot element:
In each step, the next pivot is chosen uniformly at random from all remaining list elements.
The function takes the ordering relation to use as a first argument in the form of a set of
pairs.
\<close>
function rquicksort :: "('a \<times> 'a) set \<Rightarrow> 'a list \<Rightarrow> ('a list \<times> nat) pmf" where
"rquicksort R xs =
(if xs = [] then
return_pmf ([], 0)
else
do {
i \<leftarrow> pmf_of_set {..<length xs};
let x = xs ! i;
case partition (\<lambda>y. (y,x) \<in> R) (delete_index i xs) of
(ls, rs) \<Rightarrow> do {
(ls, n1) \<leftarrow> rquicksort R ls;
(rs, n2) \<leftarrow> rquicksort R rs;
return_pmf (ls @ [x] @ rs, length xs - 1 + n1 + n2)
}
})"
by auto
termination proof (relation "Wellfounded.measure (length \<circ> snd)", goal_cases)
show "wf (Wellfounded.measure (length \<circ> snd))" by simp
qed (subst (asm) set_pmf_of_set; force intro!: le_less_trans[OF length_filter_le])+
declare rquicksort.simps [simp del]
lemma rquicksort_Nil [simp]: "rquicksort R [] = return_pmf ([], 0)"
by (simp add: rquicksort.simps)
subsection \<open>Correctness proof\<close>
lemma set_pmf_of_set_lessThan_length [simp]:
"xs \<noteq> [] \<Longrightarrow> set_pmf (pmf_of_set {..<length xs}) = {..<length xs}"
by (subst set_pmf_of_set) auto
text \<open>
We can now prove that any list that can be returned by QuickSort is sorted w.\,r.\,t.\ the
given relation. (as long as that relation is reflexive, transitive, and total)
\<close>
theorem rquicksort_correct:
assumes "trans R" and "total_on (set xs) R" and "\<forall>x\<in>set xs. (x,x) \<in> R"
assumes "(ys, n) \<in> set_pmf (rquicksort R xs)"
shows "sorted_wrt R ys \<and> mset ys = mset xs"
using assms(2-)
proof (induction xs arbitrary: ys n rule: length_induct)
case (1 xs)
have IH: "sorted_wrt R zs" "mset zs = mset ys"
if "(zs, n) \<in> set_pmf (rquicksort R ys)" "length ys < length xs" "set ys \<subseteq> set xs" for zs ys n
using that "1.IH" total_on_subset[OF "1.prems"(1) that(3)] "1.prems"(2) by blast+
show ?case
proof (cases "xs = []")
case False
with "1.prems" obtain ls rs n1 n2 i where *:
"i < length xs" "(ls, n1) \<in> set_pmf (rquicksort R [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R])"
"(rs, n2) \<in> set_pmf (rquicksort R [y\<leftarrow>delete_index i xs. (y, xs ! i) \<notin> R])"
"ys = ls @ [xs ! i] @ rs"
by (subst (asm) rquicksort.simps[of _ xs]) (auto simp: Let_def o_def)
note ys = \<open>ys = ls @ [xs ! i] @ rs\<close>
define ls' where "ls' = [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R]"
define rs' where "rs' = [y\<leftarrow>delete_index i xs. (y, xs ! i) \<notin> R]"
from \<open>i < length xs\<close> have less: "length ls' < length xs" "length rs' < length xs"
unfolding ls'_def rs'_def by (intro le_less_trans[OF length_filter_le]; force)+
have ls: "(ls, n1) \<in> set_pmf (rquicksort R ls')" and rs: "(rs, n2) \<in> set_pmf (rquicksort R rs')"
using * unfolding ls'_def rs'_def by blast+
have subset: "set ls' \<subseteq> set xs" "set rs' \<subseteq> set xs"
using insert_set_delete_index[of i xs] \<open>i < length xs\<close>
by (auto simp: ls'_def rs'_def)
have sorted: "sorted_wrt R ls" "sorted_wrt R rs"
and mset: "mset ls = mset ls'" "mset rs = mset rs'"
by (rule IH[of ls n1 ls'] IH[of rs n2 rs'] less ls rs subset)+
have ls_le: "(x, xs ! i) \<in> R" if "x \<in> set ls" for x
proof -
from that have "x \<in># mset ls" by simp
also note mset(1)
finally show ?thesis by (simp add: ls'_def)
qed
have rs_ge: "(x, xs ! i) \<notin> R" "(xs ! i, x) \<in> R" if "x \<in> set rs" for x
proof -
from that have "x \<in># mset rs" by simp
also note mset(2)
finally have x: "x \<in> set rs'" by simp
thus "(x, xs ! i) \<notin> R" by (simp_all add: rs'_def)
from x and subset and \<open>i < length xs\<close> have "x \<in> set xs" "xs ! i \<in> set xs" by auto
with "1.prems" and \<open>(x, xs ! i) \<notin> R\<close> show "(xs ! i, x) \<in> R"
unfolding total_on_def by (cases "xs ! i = x") auto
qed
have "sorted_wrt R ys" unfolding ys
by (intro sorted_wrt_append \<open>trans R\<close> sorted_wrt_singleton sorted)
(auto intro: rs_ge ls_le transD[OF \<open>trans R\<close>, of _ "xs!i"])
moreover have "mset ys = mset xs" unfolding ys using \<open>i < length xs\<close>
by (simp add: mset ls'_def rs'_def add_mset_delete_index)
ultimately show ?thesis ..
qed (insert "1.prems", simp_all)
qed
subsection \<open>Cost analysis\<close>
text \<open>
The following distribution describes the number of comparisons made by randomised
QuickSort in terms of the list length. (This is only valid if all list elements are distinct)
A succinct explanation of this cost analysis is given by Jacek Cicho\'{n}~\cite{cichon}.
\<close>
fun rqs_cost :: "nat \<Rightarrow> nat pmf" where
"rqs_cost 0 = return_pmf 0"
| "rqs_cost (Suc n) =
do {i \<leftarrow> pmf_of_set {..n}; a \<leftarrow> rqs_cost i; b \<leftarrow> rqs_cost (n - i); return_pmf (n + a + b)}"
lemma finite_set_pmf_rqs_cost [intro!]: "finite (set_pmf (rqs_cost n))"
by (induction n rule: rqs_cost.induct) simp_all
text \<open>
We connect the @{const rqs_cost} function to the @{const rquicksort} function by showing that
projecting out the number of comparisons from a run of @{const rquicksort} on a list with distinct
elements yields the same distribution as @{const rqs_cost} for the length of that list.
\<close>
theorem snd_rquicksort:
assumes "linorder_on A R" and "set xs \<subseteq> A" and "distinct xs"
shows "map_pmf snd (rquicksort R xs) = rqs_cost (length xs)"
using assms(2-)
proof (induction xs rule: length_induct)
case (1 xs)
have IH: "map_pmf snd (rquicksort R ys) = rqs_cost (length ys)"
if "length ys < length xs" "mset ys \<subseteq># mset xs" for ys
proof -
from set_mset_mono[OF that(2)] have "set ys \<subseteq> set xs" by simp
also note \<open>set xs \<subseteq> A\<close>
finally have "set ys \<subseteq> A" .
moreover from \<open>distinct xs\<close> and that(2) have "distinct ys"
by (rule distinct_mset_mono)
ultimately show ?thesis using that and "1.IH" by blast
qed
define n where "n = length xs"
define cnt where "cnt = (\<lambda>i. length [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R])"
have cnt_altdef: "cnt i = linorder_rank R (set xs) (xs ! i)" if i: "i < n" for i
proof -
have "cnt i = length [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R]" by (simp add: cnt_def)
also have "\<dots> = card (set [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R])"
by (intro distinct_card [symmetric] distinct_filter distinct_delete_index "1.prems")
also have "set [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R] =
{x \<in> set xs-{xs!i}. (x, xs ! i) \<in> R}"
using "1.prems" and i by (simp add: set_delete_index_distinct n_def)
also have "card \<dots> = linorder_rank R (set xs) (xs ! i)" by (simp add: linorder_rank_def)
finally show ?thesis .
qed
from "1.prems" have "bij_betw ((!) xs) {..<n} (set xs)"
by (intro bij_betw_byWitness[where f' = "index xs"]) (auto simp: n_def index_nth_id)
moreover have "bij_betw (linorder_rank R (set xs)) (set xs) {..<card (set xs)}"
using assms(1) by (rule bij_betw_linorder_rank) (insert "1.prems", auto)
ultimately have "bij_betw (linorder_rank R (set xs) \<circ> (\<lambda>i. xs ! i)) {..<n} {..<card (set xs)}"
by (rule bij_betw_trans)
hence bij: "bij_betw (\<lambda>i. linorder_rank R (set xs) (xs ! i)) {..<n} {..<n}"
using "1.prems" by (simp add: n_def o_def distinct_card)
show ?case
proof (cases "xs = []")
case False
hence "n > 0" by (simp add: n_def)
hence [simp]: "n \<noteq> 0" by (intro notI) auto
from False have "map_pmf snd (rquicksort R xs) =
pmf_of_set {..<length xs} \<bind>
(\<lambda>i. map_pmf (\<lambda>z. length xs - 1 + fst z + snd z)
(pair_pmf (map_pmf snd (rquicksort R [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R]))
(map_pmf snd (rquicksort R [y\<leftarrow>delete_index i xs. (y, xs ! i) \<notin> R]))))"
by (subst rquicksort.simps)
(simp add: map_bind_pmf bind_map_pmf Let_def case_prod_unfold o_def pair_pmf_def)
also have "\<dots> = pmf_of_set {..<length xs} \<bind>
(\<lambda>i. map_pmf (\<lambda>z. n - 1 + fst z + snd z)
(pair_pmf (rqs_cost (cnt i)) (rqs_cost (n - 1 - cnt i))))"
proof (intro bind_pmf_cong refl, goal_cases)
case (1 i)
with \<open>xs \<noteq> []\<close> have i: "i < length xs" by auto
from i have "map_pmf snd (rquicksort R [y\<leftarrow>delete_index i xs. (y, xs ! i) \<notin> R]) =
rqs_cost (length [y\<leftarrow>delete_index i xs. (y, xs ! i) \<notin> R])"
by (intro IH)
- (auto intro!: le_less_trans[OF length_filter_le] simp: mset_filter
- intro: subset_mset.order.trans multiset_filter_subset diff_subset_eq_self)
+ (auto intro!: le_less_trans[OF length_filter_le]
+ intro: subset_mset.trans multiset_filter_subset diff_subset_eq_self)
also have "length [y\<leftarrow>delete_index i xs. (y, xs ! i) \<notin> R] = n - 1 - cnt i"
unfolding n_def cnt_def
using sum_length_filter_compl[of "\<lambda>y. (y, xs ! i) \<in> R" "delete_index i xs"] i by simp
finally have "map_pmf snd (rquicksort R [y\<leftarrow>delete_index i xs. (y, xs ! i) \<notin> R]) =
rqs_cost (n - 1 - cnt i)" .
moreover have "map_pmf snd (rquicksort R [y\<leftarrow>delete_index i xs. (y, xs ! i) \<in> R]) =
rqs_cost (cnt i)" unfolding cnt_def using i
by (intro IH)
- (auto intro!: le_less_trans[OF length_filter_le] simp: mset_filter
- intro: subset_mset.order.trans multiset_filter_subset diff_subset_eq_self)
+ (auto intro!: le_less_trans[OF length_filter_le]
+ intro: subset_mset.trans multiset_filter_subset diff_subset_eq_self)
ultimately show ?case by (simp only: n_def)
qed
also have "\<dots> = map_pmf cnt (pmf_of_set {..<n}) \<bind>
(\<lambda>i. map_pmf (\<lambda>z. n - 1 + fst z + snd z) (pair_pmf (rqs_cost i) (rqs_cost (n - 1 - i))))"
(is "_ = bind_pmf _ ?f") by (simp add: bind_map_pmf n_def)
also have "map_pmf cnt (pmf_of_set {..<n}) =
map_pmf (\<lambda>i. linorder_rank R (set xs) (xs ! i)) (pmf_of_set {..<n})"
using \<open>n > 0\<close> by (intro map_pmf_cong refl, subst (asm) set_pmf_of_set) (auto simp: cnt_altdef)
also from \<open>n > 0\<close> have "\<dots> = pmf_of_set {..<n}" by (intro map_pmf_of_set_bij_betw bij) auto
also have "pmf_of_set {..<n} \<bind> ?f = rqs_cost n"
by (cases n) (simp_all add: lessThan_Suc_atMost bind_map_pmf map_bind_pmf pair_pmf_def)
finally show ?thesis by (simp add: n_def)
qed simp_all
qed
subsection \<open>Expected cost\<close>
text \<open>
It is relatively straightforward to see that the following recursive function
(sometimes called the `QuickSort equation') describes the expectation of @{const rqs_cost},
i.e. the expected number of comparisons of QuickSort when run on a list with distinct elements.
\<close>
fun rqs_cost_exp :: "nat \<Rightarrow> real" where
"rqs_cost_exp 0 = 0"
| "rqs_cost_exp (Suc n) = real n + (\<Sum>i\<le>n. rqs_cost_exp i + rqs_cost_exp (n - i)) / real (Suc n)"
lemmas rqs_cost_exp_0 = rqs_cost_exp.simps(1)
lemmas rqs_cost_exp_Suc [simp del] = rqs_cost_exp.simps(2)
lemma rqs_cost_exp_Suc_0 [simp]: "rqs_cost_exp (Suc 0) = 0" by (simp add: rqs_cost_exp_Suc)
text \<open>
The following theorem shows that @{const rqs_cost_exp} is indeed the expectation of
@{const rqs_cost}.
\<close>
theorem expectation_rqs_cost: "measure_pmf.expectation (rqs_cost n) real = rqs_cost_exp n"
proof (induction n rule: rqs_cost.induct)
case (2 n)
note IH = "2.IH"
have "measure_pmf.expectation (rqs_cost (Suc n)) real =
(\<Sum>a\<le>n. inverse (real (Suc n)) *
measure_pmf.expectation (rqs_cost a \<bind> (\<lambda>aa. rqs_cost (n - a) \<bind>
(\<lambda>b. return_pmf (n + aa + b)))) real)"
unfolding rqs_cost.simps by (subst pmf_expectation_bind_pmf_of_set) auto
also have "\<dots> = (\<Sum>i\<le>n. inverse (real (Suc n)) * (real n + rqs_cost_exp i + rqs_cost_exp (n - i)))"
proof (intro sum.cong refl, goal_cases)
case (1 i)
have "rqs_cost i \<bind> (\<lambda>a. rqs_cost (n - i) \<bind> (\<lambda>b. return_pmf (n + a + b))) =
map_pmf (\<lambda>(a,b). n + a + b) (pair_pmf (rqs_cost i) (rqs_cost (n - i)))"
by (simp add: pair_pmf_def map_bind_pmf)
also have "measure_pmf.expectation \<dots> real =
measure_pmf.expectation (pair_pmf (rqs_cost i) (rqs_cost (n - i)))
(\<lambda>z. real n + (real (fst z) + real (snd z)))"
by (subst integral_map_pmf) (simp add: case_prod_unfold add_ac)
also have "\<dots> = real n + measure_pmf.expectation (pair_pmf (rqs_cost i) (rqs_cost (n - i)))
(\<lambda>z. real (fst z) + real (snd z))" (is "_ = _ + ?A")
by (subst Bochner_Integration.integral_add) (auto intro!: integrable_measure_pmf_finite)
also have "?A = measure_pmf.expectation (map_pmf fst (pair_pmf (rqs_cost i) (rqs_cost (n - i)))) real +
measure_pmf.expectation (map_pmf snd (pair_pmf (rqs_cost i) (rqs_cost (n - i)))) real"
unfolding integral_map_pmf
by (subst Bochner_Integration.integral_add) (auto intro!: integrable_measure_pmf_finite)
also have "\<dots> = measure_pmf.expectation (rqs_cost i) real +
measure_pmf.expectation (rqs_cost (n - i)) real"
unfolding map_fst_pair_pmf map_snd_pair_pmf ..
also from 1 have "\<dots> = rqs_cost_exp i + rqs_cost_exp (n - i)" by (simp_all add: IH)
finally show ?case by simp
qed
also have "\<dots> = (\<Sum>i\<le>n. inverse (real (Suc n)) * real n) +
(\<Sum>i\<le>n. rqs_cost_exp i + rqs_cost_exp (n - i)) / real (Suc n)"
by (simp add: sum.distrib field_simps sum_distrib_left sum_distrib_right
sum_divide_distrib [symmetric] del: of_nat_Suc)
also have "(\<Sum>i\<le>n. inverse (real (Suc n)) * real n) = real n" by simp
also have "\<dots> + (\<Sum>i\<le>n. rqs_cost_exp i + rqs_cost_exp (n - i)) / real (Suc n) = rqs_cost_exp (Suc n)"
by (simp add: rqs_cost_exp_Suc)
finally show ?case .
qed simp_all
text \<open>
We will now obtain a closed-form solution for @{const rqs_cost_exp}. First of all, we can
reindex the right-most sum in the recursion step and obtain:
\<close>
lemma rqs_cost_exp_Suc':
"rqs_cost_exp (Suc n) = real n + 2 / real (Suc n) * (\<Sum>i\<le>n. rqs_cost_exp i)"
proof -
have "rqs_cost_exp (Suc n) = real n + (\<Sum>i\<le>n. rqs_cost_exp i + rqs_cost_exp (n - i)) / real (Suc n)"
by (rule rqs_cost_exp_Suc)
also have "(\<Sum>i\<le>n. rqs_cost_exp i + rqs_cost_exp (n - i)) = (\<Sum>i\<le>n. rqs_cost_exp i) + (\<Sum>i\<le>n. rqs_cost_exp (n - i))"
by (simp add: sum.distrib)
also have "(\<Sum>i\<le>n. rqs_cost_exp (n - i)) = (\<Sum>i\<le>n. rqs_cost_exp i)"
by (intro sum.reindex_bij_witness[of _ "\<lambda>i. n - i" "\<lambda>i. n - i"]) auto
also have "\<dots> + \<dots> = 2 * \<dots>" by simp
also have "\<dots> / real (Suc n) = 2 / real (Suc n) * (\<Sum>i\<le>n. rqs_cost_exp i)" by simp
finally show ?thesis .
qed
text \<open>
Next, we can apply some standard techniques to transform this equation into a simple
linear recurrence, which we can then solve easily in terms of harmonic numbers:
\<close>
theorem rqs_cost_exp_eq [code]: "rqs_cost_exp n = 2 * real (n + 1) * harm n - 4 * real n"
proof -
define F where "F = (\<lambda>n. rqs_cost_exp n / (real n + 1))"
have [simp]: "F 0 = 0" "F (Suc 0) = 0" by (simp_all add: F_def)
have F_Suc: "F (Suc m) = F m + real (2*m) / (real ((m+1)*(m+2)))" if "m > 0" for m
proof (cases m)
case (Suc n)
have A: "rqs_cost_exp (Suc (Suc n)) * real (Suc (Suc n)) =
real ((n+1)*(n+2)) + 2 * (\<Sum>i\<le>n. rqs_cost_exp i) + 2 * rqs_cost_exp (Suc n)"
by (subst rqs_cost_exp_Suc') (simp_all add: field_simps)
have B: "rqs_cost_exp (Suc n) * real (Suc n) = real (n*(n+1)) + 2 * (\<Sum>i\<le>n. rqs_cost_exp i)"
by (subst rqs_cost_exp_Suc') (simp_all add: field_simps)
have "rqs_cost_exp (Suc (Suc n)) * real (Suc (Suc n)) - rqs_cost_exp (Suc n) * real (Suc n) =
real ((n+1)*(n+2)) - real (n*(n+1)) + 2 * rqs_cost_exp (Suc n)"
by (subst A, subst B) simp_all
also have "real ((n+1)*(n+2)) - real (n*(n+1)) = real (2*(n+1))" by simp
finally have "rqs_cost_exp (Suc (Suc n)) * real (n+2) = rqs_cost_exp (Suc n) * real (n+3) + real (2*(n+1))"
by (simp add: algebra_simps)
hence "rqs_cost_exp (Suc (Suc n)) / real (n+3) =
rqs_cost_exp (Suc n) / real (n+2) + real (2*(n+1)) / (real (n+2)*real (n+3))"
by (simp add: divide_simps del: of_nat_Suc of_nat_add)
thus ?thesis by (simp add: F_def algebra_simps Suc)
qed simp_all
have F_eq: "F n = 2 * (\<Sum>k=1..n. real (k - 1) / real (k * (k + 1)))" for n
proof (cases "n \<ge> 1")
case True
thus ?thesis by (induction n rule: dec_induct) (simp_all add: F_Suc algebra_simps)
qed (simp_all add: not_le)
have "F n = 2 * (\<Sum>k=1..n. real (k - 1) / real (k * (k + 1)))" (is "_ = 2 * ?S") by (fact F_eq)
also have "?S = (\<Sum>k=1..n. 2 / real (Suc k) - 1 / real k)"
by (intro sum.cong) (simp_all add: field_simps of_nat_diff)
also have "\<dots> = 2 * (\<Sum>k=1..n. inverse (real (Suc k))) - harm n"
by (subst sum_subtractf) (simp add: harm_def sum.distrib sum_distrib_left divide_simps)
also have "(\<Sum>k=1..n. inverse (real (Suc k))) = (\<Sum>k=Suc 1..Suc n. inverse (real k))"
by (intro sum.reindex_bij_witness[of _ "\<lambda>x. x - 1" Suc]) auto
also have "\<dots> = harm (Suc n) - 1" unfolding harm_def by (subst (2) sum.atLeast_Suc_atMost) simp_all
finally have "F n = 2 * harm n + 4 * (1 / (n + 1) - 1)" by (simp add: harm_Suc field_simps)
also have "\<dots> * real (n + 1) = 2 * real (n + 1) * harm n - 4 * real n"
by (simp add: field_simps)
also have "F n * real (n + 1) = rqs_cost_exp n" by (simp add: F_def add_ac)
finally show ?thesis .
qed
(* TODO Move *)
lemma asymp_equiv_harm [asymp_equiv_intros]: "harm \<sim>[at_top] (\<lambda>n. ln (real n))"
proof -
have "(\<lambda>n. harm n - ln (real n)) \<in> O(\<lambda>_. 1)" using euler_mascheroni_LIMSEQ
by (intro bigoI_tendsto[where c = euler_mascheroni]) simp_all
also have "(\<lambda>_. 1) \<in> o(\<lambda>n. ln (real n))" by auto
finally have "(\<lambda>n. ln (real n) + (harm n - ln (real n))) \<sim>[at_top] (\<lambda>n. ln (real n))"
by (subst asymp_equiv_add_right) simp_all
thus ?thesis by simp
qed
corollary rqs_cost_exp_asymp_equiv: "rqs_cost_exp \<sim>[at_top] (\<lambda>n. 2 * n * ln n)"
proof -
have "rqs_cost_exp = (\<lambda>n. 2 * real (n + 1) * harm n - 4 * real n)" using rqs_cost_exp_eq ..
also have "\<dots> = (\<lambda>n. 2 * real n * harm n + (2 * harm n - 4 * real n))"
by (simp add: algebra_simps)
finally have "rqs_cost_exp \<sim>[at_top] \<dots>" by simp
also have "\<dots> \<sim>[at_top] (\<lambda>n. 2 * real n * harm n)"
proof (subst asymp_equiv_add_right)
have "(\<lambda>x. 1 * harm x) \<in> o(\<lambda>x. real x * harm x)"
by (intro landau_o.small_big_mult smallo_real_nat_transfer) simp_all
moreover have "harm \<in> \<omega>(\<lambda>_. 1 :: real)"
by (intro smallomegaI_filterlim_at_top_norm) (auto simp: harm_at_top)
hence "(\<lambda>x. real x * 1) \<in> o(\<lambda>x. real x * harm x)"
by (intro landau_o.big_small_mult) (simp_all add: smallomega_iff_smallo)
ultimately show "(\<lambda>n. 2 * harm n - 4 * real n) \<in> o(\<lambda>n. 2 * real n * harm n)"
by (intro sum_in_smallo) simp_all
qed simp_all
also have "\<dots> \<sim>[at_top] (\<lambda>n. 2 * real n * ln (real n))" by (intro asymp_equiv_intros)
finally show ?thesis .
qed
lemma harm_mono: "m \<le> n \<Longrightarrow> harm m \<le> (harm n :: real)"
unfolding harm_def by (intro sum_mono2) auto
lemma harm_Suc_0 [simp]: "harm (Suc 0) = 1"
by (simp add: harm_def)
lemma harm_ge_1: "n > 0 \<Longrightarrow> harm n \<ge> (1::real)"
using harm_mono[of 1 n] by simp
lemma mono_rqs_cost_exp: "mono rqs_cost_exp"
proof (rule incseq_SucI)
fix n show "rqs_cost_exp n \<le> rqs_cost_exp (Suc n)"
proof (cases "n = 0")
case False
have "0 < (1 * 2 * (real n + 1) - 2 * real n) / (real n + 1)" by simp
also have "\<dots> \<le> (harm n * 2 * (real n + 1) - 2 * real n) / (real n + 1)" using False
by (intro divide_right_mono diff_right_mono mult_right_mono) (auto simp: harm_ge_1)
also have "\<dots> = rqs_cost_exp (Suc n) - rqs_cost_exp n"
by (simp add: rqs_cost_exp_eq harm_Suc field_simps)
finally show ?thesis by simp
qed auto
qed
lemma rqs_cost_exp_leI: "m \<le> n \<Longrightarrow> rqs_cost_exp m \<le> rqs_cost_exp n"
using mono_rqs_cost_exp by (simp add: mono_def)
subsection \<open>Version for lists with repeated elements\<close>
definition threeway_partition where
"threeway_partition x R xs =
(filter (\<lambda>y. (y,x) \<in> R \<and> (x,y) \<notin> R) xs,
filter (\<lambda>y. (x,y) \<in> R \<and> (y,x) \<in> R) xs,
filter (\<lambda>y. (x,y) \<in> R \<and> (y,x) \<notin> R) xs)"
text \<open>
The following version of randomised Quicksort uses a three-way partitioning function
in order to also achieve expected logarithmic running time on lists with repeated elements.
\<close>
function rquicksort' :: "('a \<times> 'a) set \<Rightarrow> 'a list \<Rightarrow> ('a list \<times> nat) pmf" where
"rquicksort' R xs =
(if xs = [] then
return_pmf ([], 0)
else
do {
i \<leftarrow> pmf_of_set {..<length xs};
let x = xs ! i;
case threeway_partition x R (delete_index i xs) of
(ls, es, rs) \<Rightarrow> do {
(ls, n1) \<leftarrow> rquicksort' R ls;
(rs, n2) \<leftarrow> rquicksort' R rs;
return_pmf (ls @ x # es @ rs, length xs - 1 + n1 + n2)
}
})"
by auto
termination proof (relation "Wellfounded.measure (length \<circ> snd)", goal_cases)
show "wf (Wellfounded.measure (length \<circ> snd))" by simp
qed (subst (asm) set_pmf_of_set;
force intro!: le_less_trans[OF length_filter_le] simp: threeway_partition_def)+
declare rquicksort'.simps [simp del]
lemma rquicksort'_Nil [simp]: "rquicksort' R [] = return_pmf ([], 0)"
by (simp add: rquicksort'.simps)
context
begin
qualified definition lesss :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"lesss R x xs = filter (\<lambda>y. (y, x) \<in> R \<and> (x, y) \<notin> R) xs"
qualified definition greaters :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"greaters R x xs = filter (\<lambda>y. (x, y) \<in> R \<and> (y, x) \<notin> R) xs"
qualified lemma lesss_Cons:
"lesss R x (y # ys) =
(if (y, x) \<in> R \<and> (x, y) \<notin> R then y # lesss R x ys else lesss R x ys)"
by (simp add: lesss_def)
qualified lemma length_lesss_le [intro]: "length (lesss R x xs) \<le> length xs"
by (simp add: lesss_def)
qualified lemma length_lesss_less [intro]:
assumes "x \<in> set xs"
shows "length (lesss R x xs) < length xs"
using assms by (induction xs) (auto simp: lesss_Cons intro: le_less_trans)
qualified lemma greaters_Cons:
"greaters R x (y # ys) =
(if (x, y) \<in> R \<and> (y, x) \<notin> R then y # greaters R x ys else greaters R x ys)"
by (simp add: greaters_def)
qualified lemma length_greaters_le [intro]: "length (greaters R x xs) \<le> length xs"
by (simp add: greaters_def)
qualified lemma length_greaters_less [intro]:
assumes "x \<in> set xs"
shows "length (greaters R x xs) < length xs"
using assms by (induction xs) (auto simp: greaters_Cons intro: le_less_trans)
text \<open>
The following function counts the comparisons made by the modified randomised Quicksort.
\<close>
function rqs'_cost :: "('a \<times> 'a) set \<Rightarrow> 'a list \<Rightarrow> nat pmf" where
"rqs'_cost R xs =
(if xs = [] then
return_pmf 0
else
do {
i \<leftarrow> pmf_of_set {..<length xs};
let x = xs ! i;
map_pmf (\<lambda>(n1,n2). length xs - 1 + n1 + n2)
(pair_pmf (rqs'_cost R (lesss R x xs)) (rqs'_cost R (greaters R x xs)))
})"
by auto
termination by (relation "Wellfounded.measure (length \<circ> snd)") auto
declare rqs'_cost.simps [simp del]
lemma rqs'_cost_nonempty:
"xs \<noteq> [] \<Longrightarrow> rqs'_cost R xs =
do {
i \<leftarrow> pmf_of_set {..<length xs};
let x = xs ! i;
n1 \<leftarrow> rqs'_cost R (lesss R x xs);
n2 \<leftarrow> rqs'_cost R (greaters R x xs);
return_pmf (length xs - 1 + n1 + n2)
}"
by (subst rqs'_cost.simps) (auto simp: pair_pmf_def Let_def map_bind_pmf)
lemma finite_set_pmf_rqs'_cost [simp, intro]:
"finite (set_pmf (rqs'_cost R xs))"
by (induction R xs rule: rqs'_cost.induct) (auto simp: rqs'_cost.simps Let_def)
(* TODO: Move? *)
lemma expectation_pair_pmf_fst [simp]:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (fst x)) = measure_pmf.expectation p f"
proof -
have "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (fst x)) =
measure_pmf.expectation (map_pmf fst (pair_pmf p q)) f" by simp
also have "map_pmf fst (pair_pmf p q) = p"
by (simp add: map_fst_pair_pmf)
finally show ?thesis .
qed
lemma expectation_pair_pmf_snd [simp]:
fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
shows "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (snd x)) = measure_pmf.expectation q f"
proof -
have "measure_pmf.expectation (pair_pmf p q) (\<lambda>x. f (snd x)) =
measure_pmf.expectation (map_pmf snd (pair_pmf p q)) f" by simp
also have "map_pmf snd (pair_pmf p q) = q"
by (simp add: map_snd_pair_pmf)
finally show ?thesis .
qed
qualified lemma length_lesss_le_sorted:
assumes "sorted_wrt R xs" "i < length xs"
shows "length (lesss R (xs ! i) xs) \<le> i"
using assms by (induction arbitrary: i rule: sorted_wrt.induct)
(force simp: lesss_def nth_Cons le_Suc_eq split: nat.splits)+
qualified lemma length_greaters_le_sorted:
assumes "sorted_wrt R xs" "i < length xs"
shows "length (greaters R (xs ! i) xs) \<le> length xs - i - 1"
using assms
by (induction arbitrary: i rule: sorted_wrt.induct)
(force simp: greaters_def nth_Cons le_Suc_eq split: nat.splits)+
qualified lemma length_lesss_le':
assumes "i < length xs" "linorder_on A R" "set xs \<subseteq> A"
shows "length (lesss R (insort_wrt R xs ! i) xs) \<le> i"
proof -
define x where "x = insort_wrt R xs ! i"
define less where "less = (\<lambda>x y. (x,y) \<in> R \<and> (y,x) \<notin> R)"
have "length (lesss R x xs) = size {# y \<in># mset xs. less y x #}"
by (simp add: lesss_def size_mset [symmetric] less_def mset_filter del: size_mset)
also have "mset xs = mset (insort_wrt R xs)" by simp
also have "size {#y \<in># mset (insort_wrt R xs). less y x#} =
length (lesss R x (insort_wrt R xs))"
by (simp only: mset_filter [symmetric] size_mset lesss_def less_def)
also have "\<dots> \<le> i" unfolding x_def by (rule length_lesss_le_sorted) (use assms in auto)
finally show ?thesis unfolding x_def .
qed
qualified lemma length_greaters_le':
assumes "i < length xs" "linorder_on A R" "set xs \<subseteq> A"
shows "length (greaters R (insort_wrt R xs ! i) xs) \<le> length xs - i - 1"
proof -
define x where "x = insort_wrt R xs ! i"
define less where "less = (\<lambda>x y. (x,y) \<in> R \<and> (y,x) \<notin> R)"
have "length (greaters R x xs) = size {# y \<in># mset xs. less x y #}"
by (simp add: greaters_def size_mset [symmetric] less_def mset_filter del: size_mset)
also have "mset xs = mset (insort_wrt R xs)" by simp
also have "size {#y \<in># mset (insort_wrt R xs). less x y#} =
length (greaters R x (insort_wrt R xs))"
by (simp only: mset_filter [symmetric] size_mset greaters_def less_def)
also have "\<dots> \<le> length (insort_wrt R xs) - i - 1" unfolding x_def
by (rule length_greaters_le_sorted) (use assms in auto)
finally show ?thesis unfolding x_def by simp
qed
text \<open>
We can show quite easily that the expected number of comparisons in this modified
QuickSort is bounded above by the expected number of comparisons on a list of the same length
with no repeated elements.
\<close>
theorem rqs'_cost_expectation_le:
assumes "linorder_on A R" "set xs \<subseteq> A"
shows "measure_pmf.expectation (rqs'_cost R xs) real \<le> rqs_cost_exp (length xs)"
using assms
proof (induction R xs rule: rqs'_cost.induct)
case (1 R xs)
show ?case
proof (cases "xs = []")
case False
define n where "n = length xs - 1"
have length_eq: "length xs = Suc n" using False by (simp add: n_def)
define E where "E = (\<lambda>xs. measure_pmf.expectation (rqs'_cost R xs) real)"
define f where "f = (\<lambda>x. rqs_cost_exp (length (lesss R x xs)) +
rqs_cost_exp (length (greaters R x xs)))"
have "rqs'_cost R xs =
do {
i \<leftarrow> pmf_of_set {..<length xs};
map_pmf (\<lambda>(n1, y). length xs - Suc 0 + n1 + y)
(pair_pmf (rqs'_cost R (lesss R (xs ! i) xs))
(rqs'_cost R (greaters R (xs ! i) xs)))
}"
using False by (subst rqs'_cost.simps) (simp_all add: Let_def)
also have "measure_pmf.expectation \<dots> real = real n +
(\<Sum>k<length xs. E (lesss R (xs ! k) xs) + E (greaters R (xs ! k) xs)) / real (length xs)"
using False
by (subst pmf_expectation_bind_pmf_of_set)
(auto intro!: finite_imageI finite_cartesian_product simp: case_prod_unfold
integrable_measure_pmf_finite sum_divide_distrib [symmetric] field_simps
length_eq sum.distrib E_def)
also have "\<dots> \<le> real n + (\<Sum>k<length xs. f (xs ! k)) / real (length xs)"
unfolding E_def f_def using False "1.prems"
by (intro add_mono order.refl divide_right_mono sum_mono "1.IH"[OF _ _ refl] False)
(auto simp: lesss_def greaters_def)
also have "(\<Sum>k<length xs. f (xs ! k)) = (\<Sum>x\<in>#mset xs. f x)"
by (simp only: mset_map [symmetric] sum_mset_sum_list sum_list_sum_nth)
(simp_all add: atLeast0LessThan)
also have "mset xs = mset (insort_wrt R xs)"
by simp
also have "(\<Sum>x\<in>#\<dots>. f x) = (\<Sum>i<length xs. f (insort_wrt R xs ! i))"
by (simp only: mset_map [symmetric] sum_mset_sum_list sum_list_sum_nth)
(simp_all add: atLeast0LessThan)
also have "\<dots> \<le> (\<Sum>i<length xs. rqs_cost_exp i + rqs_cost_exp (length xs - i - 1))"
unfolding f_def
proof (intro sum_mono add_mono rqs_cost_exp_leI)
fix i assume i: "i \<in> {..<length xs}"
show "length (lesss R (insort_wrt R xs ! i) xs) \<le> i"
using i "1.prems" by (intro length_lesss_le'[where A = A]) auto
show "length (greaters R (insort_wrt R xs ! i) xs) \<le> length xs - i - 1"
using i "1.prems" by (intro length_greaters_le'[where A = A]) auto
qed
also have "\<dots> = (\<Sum>i\<le>n. rqs_cost_exp i + rqs_cost_exp (n - i))"
by (intro sum.cong) (auto simp: length_eq)
also have "real n + \<dots> / real (length xs) = rqs_cost_exp (length xs)"
by (simp add: length_eq rqs_cost_exp.simps(2))
finally show ?thesis by (simp add: divide_right_mono)
qed (auto simp: rqs'_cost.simps)
qed
end
end
diff --git a/thys/Refine_Monadic/Generic/RefineG_Domain.thy b/thys/Refine_Monadic/Generic/RefineG_Domain.thy
--- a/thys/Refine_Monadic/Generic/RefineG_Domain.thy
+++ b/thys/Refine_Monadic/Generic/RefineG_Domain.thy
@@ -1,343 +1,343 @@
section \<open>General Domain Theory\<close>
theory RefineG_Domain
imports "../Refine_Misc"
begin
subsection \<open>General Order Theory Tools\<close>
lemma chain_f_apply: "Complete_Partial_Order.chain (fun_ord le) F
\<Longrightarrow> Complete_Partial_Order.chain le {y . \<exists>f\<in>F. y = f x}"
unfolding Complete_Partial_Order.chain_def
by (auto simp: fun_ord_def)
lemma ccpo_lift:
assumes "class.ccpo lub le lt"
shows "class.ccpo (fun_lub lub) (fun_ord le) (mk_less (fun_ord le))"
proof -
interpret ccpo: ccpo lub le lt by fact
show ?thesis
apply unfold_locales
apply (simp_all only: mk_less_def fun_ord_def fun_lub_def)
apply simp
using ccpo.order_trans apply blast
- using ccpo.antisym apply blast
+ using ccpo.order.antisym apply blast
using ccpo.ccpo_Sup_upper apply (blast intro: chain_f_apply)
using ccpo.ccpo_Sup_least apply (blast intro: chain_f_apply)
done
qed
(* TODO: Move *)
lemma fun_lub_simps[simp]:
"fun_lub lub {} = (\<lambda>x. lub {})"
"fun_lub lub {f} = (\<lambda>x. lub {f x})"
unfolding fun_lub_def by auto
subsection \<open>Flat Ordering\<close>
lemma flat_ord_chain_cases:
assumes A: "Complete_Partial_Order.chain (flat_ord b) C"
obtains "C={}"
| "C={b}"
| x where "x\<noteq>b" and "C={x}"
| x where "x\<noteq>b" and "C={b,x}"
proof -
have "\<exists>m1 m2. C \<subseteq> {m1,m2} \<and> (m1 = b \<or> m2 = b)"
apply (rule ccontr)
proof clarsimp
assume "\<forall>m1 m2. C \<subseteq> {m1, m2} \<longrightarrow> m1\<noteq>b \<and> m2\<noteq>b"
then obtain m1 m2 where "m1\<in>C" "m2\<in>C"
"m1\<noteq>m2" "m1\<noteq>b" "m2\<noteq>b"
by blast
with A show False
unfolding Complete_Partial_Order.chain_def flat_ord_def
by auto
qed
then obtain m where "C \<subseteq> {m,b}" by blast
thus ?thesis
apply (cases "m=b")
using that
apply blast+
done
qed
lemma flat_lub_simps[simp]:
"flat_lub b {} = b"
"flat_lub b {x} = x"
"flat_lub b (insert b X) = flat_lub b X"
unfolding flat_lub_def
by auto
lemma flat_ord_simps[simp]:
"flat_ord b b x"
by (simp add: flat_ord_def)
interpretation flat_ord: ccpo "flat_lub b" "flat_ord b" "mk_less (flat_ord b)"
apply unfold_locales
apply (simp_all only: mk_less_def flat_ord_def) apply auto [4]
apply (erule flat_ord_chain_cases, auto) []
apply (erule flat_ord_chain_cases, auto) []
done
interpretation flat_le_mono_setup: mono_setup_loc "flat_ord b"
by standard auto
subsubsection \<open>Flat function Ordering\<close>
abbreviation "flatf_ord b == fun_ord (flat_ord b)"
abbreviation "flatf_lub b == fun_lub (flat_lub b)"
interpretation flatf_ord: ccpo "flatf_lub b" "flatf_ord b" "mk_less (flatf_ord b)"
apply (rule ccpo_lift)
apply unfold_locales
done
subsubsection \<open>Fixed Points in Flat Ordering\<close>
text \<open>
Fixed points in a flat ordering are used to express recursion.
The bottom element is interpreted as non-termination.
\<close>
abbreviation "flat_mono b == monotone (flat_ord b) (flat_ord b)"
abbreviation "flatf_mono b == monotone (flatf_ord b) (flatf_ord b)"
abbreviation "flatf_fp b \<equiv> flatf_ord.fixp b"
lemma flatf_fp_mono[refine_mono]:
\<comment> \<open>The fixed point combinator is monotonic\<close>
assumes "flatf_mono b f"
and "flatf_mono b g"
and "\<And>Z x. flat_ord b (f Z x) (g Z x)"
shows "flat_ord b (flatf_fp b f x) (flatf_fp b g x)"
proof -
have "flatf_ord b (flatf_fp b f) (flatf_fp b g)"
apply (rule flatf_ord.fixp_mono[OF assms(1,2)])
using assms(3) by (simp add: fun_ord_def)
thus ?thesis unfolding fun_ord_def by blast
qed
lemma flatf_admissible_pointwise:
"(\<And>x. P x b) \<Longrightarrow>
ccpo.admissible (flatf_lub b) (flatf_ord b) (\<lambda>g. \<forall>x. P x (g x))"
apply (intro ccpo.admissibleI allI impI)
apply (drule_tac x=x in chain_f_apply)
apply (erule flat_ord_chain_cases)
apply (auto simp add: fun_lub_def) [2]
apply (force simp add: fun_lub_def) []
apply (auto simp add: fun_lub_def) []
done
text \<open>
If a property is defined pointwise, and holds for the bottom element,
we can use fixed-point induction for it.
In the induction step, we can assume that the function is less or equal
to the fixed-point.
This rule covers refinement and transfer properties,
such as: Refinement of fixed-point combinators and transfer of
fixed-point combinators to different domains.
\<close>
lemma flatf_fp_induct_pointwise:
\<comment> \<open>Fixed-point induction for pointwise properties\<close>
fixes a :: 'a
assumes cond_bot: "\<And>a x. pre a x \<Longrightarrow> post a x b"
assumes MONO: "flatf_mono b B"
assumes PRE0: "pre a x"
assumes STEP: "\<And>f a x.
\<lbrakk>\<And>a' x'. pre a' x' \<Longrightarrow> post a' x' (f x'); pre a x;
flatf_ord b f (flatf_fp b B) \<rbrakk>
\<Longrightarrow> post a x (B f x)"
shows "post a x (flatf_fp b B x)"
proof -
define ub where "ub = flatf_fp b B"
have "(\<forall>x a. pre a x \<longrightarrow> post a x (flatf_fp b B x))
\<and> flatf_ord b (flatf_fp b B) ub"
apply (rule flatf_ord.fixp_induct[OF _ MONO])
apply (rule admissible_conj)
apply (rule flatf_admissible_pointwise)
apply (blast intro: cond_bot)
apply (rule ccpo.admissibleI)
apply (blast intro: flatf_ord.ccpo_Sup_least)
apply (auto intro: cond_bot simp: fun_ord_def flat_ord_def) []
apply (rule conjI)
unfolding ub_def
apply (blast intro: STEP)
apply (subst flatf_ord.fixp_unfold[OF MONO])
apply (blast intro: monotoneD[OF MONO])
done
with PRE0 show ?thesis by blast
qed
text \<open>
The next rule covers transfer between fixed points.
It allows to lift a pointwise transfer condition
\<open>P x y \<longrightarrow> tr (f x) (f y)\<close> to fixed points.
Note that one of the fixed points may be an arbitrary fixed point.
\<close>
lemma flatf_fixp_transfer:
\<comment> \<open>Transfer rule for fixed points\<close>
assumes TR_BOT[simp]: "\<And>x'. tr b x'"
assumes MONO: "flatf_mono b B"
assumes FP': "fp' = B' fp'"
assumes R0: "P x x'"
assumes RS: "\<And>f f' x x'.
\<lbrakk>\<And>x x'. P x x' \<Longrightarrow> tr (f x) (f' x'); P x x'; fp' = f'\<rbrakk>
\<Longrightarrow> tr (B f x) (B' f' x')"
shows "tr (flatf_fp b B x) (fp' x')"
apply (rule flatf_fp_induct_pointwise[where pre="\<lambda>x y. P y x", OF _ MONO])
apply simp
apply (rule R0)
apply (subst FP')
apply (blast intro: RS)
done
subsubsection \<open>Relation of Flat Ordering to Complete Lattices\<close>
text \<open>
In this section, we establish the relation between flat orderings
and complete lattices. This relation is exploited to show properties
of fixed points wrt. a refinement ordering.
\<close>
abbreviation "flat_le \<equiv> flat_ord bot"
abbreviation "flat_ge \<equiv> flat_ord top"
abbreviation "flatf_le \<equiv> flatf_ord bot"
abbreviation "flatf_ge \<equiv> flatf_ord top"
text \<open>The flat ordering implies the lattice ordering\<close>
lemma flat_ord_compat:
fixes x y :: "'a :: complete_lattice"
shows
"flat_le x y \<Longrightarrow> x \<le> y"
"flat_ge x y \<Longrightarrow> x \<ge> y"
unfolding flat_ord_def by auto
lemma flatf_ord_compat:
fixes x y :: "'a \<Rightarrow> ('b :: complete_lattice)"
shows
"flatf_le x y \<Longrightarrow> x \<le> y"
"flatf_ge x y \<Longrightarrow> x \<ge> y"
by (auto simp: flat_ord_compat le_fun_def fun_ord_def)
abbreviation "flat_mono_le \<equiv> flat_mono bot"
abbreviation "flat_mono_ge \<equiv> flat_mono top"
abbreviation "flatf_mono_le \<equiv> flatf_mono bot"
abbreviation "flatf_mono_ge \<equiv> flatf_mono top"
abbreviation "flatf_gfp \<equiv> flatf_ord.fixp top"
abbreviation "flatf_lfp \<equiv> flatf_ord.fixp bot"
text \<open>If a functor is monotonic wrt. both the flat and the
lattice ordering, the fixed points wrt. these orderings coincide.
\<close>
lemma lfp_eq_flatf_lfp:
assumes FM: "flatf_mono_le B" and M: "mono B"
shows "lfp B = flatf_lfp B"
proof -
from lfp_unfold[OF M, symmetric] have "B (lfp B) = lfp B" .
hence "flatf_le (B (lfp B)) (lfp B)" by simp
with flatf_ord.fixp_lowerbound[OF FM] have "flatf_le (flatf_lfp B) (lfp B)" .
with flatf_ord_compat have "flatf_lfp B \<le> lfp B" by blast
also
from flatf_ord.fixp_unfold[OF FM, symmetric] have "B (flatf_lfp B) = flatf_lfp B" .
hence "B (flatf_lfp B) \<le> flatf_lfp B" by simp
with lfp_lowerbound[where A="flatf_lfp B"] have "lfp B \<le> flatf_lfp B" .
finally show "lfp B = flatf_lfp B" ..
qed
lemma gfp_eq_flatf_gfp:
assumes FM: "flatf_mono_ge B" and M: "mono B"
shows "gfp B = flatf_gfp B"
proof -
from gfp_unfold[OF M, symmetric] have "B (gfp B) = gfp B" .
hence "flatf_ge (B (gfp B)) (gfp B)" by simp
with flatf_ord.fixp_lowerbound[OF FM] have "flatf_ge (flatf_gfp B) (gfp B)" .
with flatf_ord_compat have "gfp B \<le> flatf_gfp B" by blast
also
from flatf_ord.fixp_unfold[OF FM, symmetric] have "B (flatf_gfp B) = flatf_gfp B" .
hence "flatf_gfp B \<le> B (flatf_gfp B)" by simp
with gfp_upperbound[where X="flatf_gfp B"] have "flatf_gfp B \<le> gfp B" .
finally show "gfp B = flatf_gfp B" .
qed
(* TODO: This belongs to "General Recursion"*)
text \<open>
The following lemma provides a well-founded induction scheme for arbitrary
fixed point combinators.
\<close>
lemma wf_fixp_induct:
\<comment> \<open>Well-Founded induction for arbitrary fixed points\<close>
fixes a :: 'a
assumes fixp_unfold: "fp B = B (fp B)"
assumes WF: "wf V"
assumes P0: "pre a x"
assumes STEP: "\<And>f a x. \<lbrakk>
\<And>a' x'. \<lbrakk> pre a' x'; (x',x)\<in>V \<rbrakk> \<Longrightarrow> post a' x' (f x'); fp B = f; pre a x
\<rbrakk> \<Longrightarrow> post a x (B f x)"
shows "post a x (fp B x)"
proof -
have "\<forall>a. pre a x \<longrightarrow> post a x (fp B x)"
using WF
apply (induct x rule: wf_induct_rule)
apply (intro allI impI)
apply (subst fixp_unfold)
apply (rule STEP)
apply (simp)
apply (simp)
apply (simp)
done
with P0 show ?thesis by blast
qed
lemma flatf_lfp_transfer:
\<comment> \<open>Transfer rule for least fixed points\<close>
fixes B::"(_ \<Rightarrow> 'a::order_bot) \<Rightarrow> _"
assumes TR_BOT[simp]: "\<And>x. tr bot x"
assumes MONO: "flatf_mono_le B"
assumes MONO': "flatf_mono_le B'"
assumes R0: "P x x'"
assumes RS: "\<And>f f' x x'.
\<lbrakk>\<And>x x'. P x x' \<Longrightarrow> tr (f x) (f' x'); P x x'; flatf_lfp B' = f'\<rbrakk>
\<Longrightarrow> tr (B f x) (B' f' x')"
shows "tr (flatf_lfp B x) (flatf_lfp B' x')"
apply (rule flatf_fixp_transfer[where tr=tr and fp'="flatf_lfp B'" and P=P])
apply (fact|rule flatf_ord.fixp_unfold[OF MONO'])+
done
lemma flatf_gfp_transfer:
\<comment> \<open>Transfer rule for greatest fixed points\<close>
fixes B::"(_ \<Rightarrow> 'a::order_top) \<Rightarrow> _"
assumes TR_TOP[simp]: "\<And>x. tr x top"
assumes MONO: "flatf_mono_ge B"
assumes MONO': "flatf_mono_ge B'"
assumes R0: "P x x'"
assumes RS: "\<And>f f' x x'.
\<lbrakk>\<And>x x'. P x x' \<Longrightarrow> tr (f x) (f' x'); P x x'; flatf_gfp B = f\<rbrakk>
\<Longrightarrow> tr (B f x) (B' f' x')"
shows "tr (flatf_gfp B x) (flatf_gfp B' x')"
apply (rule flatf_fixp_transfer[where
tr="\<lambda>x y. tr y x" and fp'="flatf_gfp B" and P="\<lambda>x y. P y x"])
apply (fact|assumption|rule flatf_ord.fixp_unfold[OF MONO] RS)+
done
lemma meta_le_everything_if_top: "(m=top) \<Longrightarrow> (\<And>x. x \<le> (m::'a::order_top))"
by auto
lemmas flatf_lfp_refine = flatf_lfp_transfer[
where tr = "\<lambda>a b. a \<le> cf b" for cf, OF bot_least]
lemmas flatf_gfp_refine = flatf_gfp_transfer[
where tr = "\<lambda>a b. a \<le> cf b" for cf, OF meta_le_everything_if_top]
lemma flat_ge_sup_mono[refine_mono]: "\<And>a a'::'a::complete_lattice.
flat_ge a a' \<Longrightarrow> flat_ge b b' \<Longrightarrow> flat_ge (sup a b) (sup a' b')"
by (auto simp: flat_ord_def)
declare sup_mono[refine_mono]
end
diff --git a/thys/Refine_Monadic/Refine_Misc.thy b/thys/Refine_Monadic/Refine_Misc.thy
--- a/thys/Refine_Monadic/Refine_Misc.thy
+++ b/thys/Refine_Monadic/Refine_Misc.thy
@@ -1,717 +1,717 @@
section \<open>Miscellanneous Lemmas and Tools\<close>
theory Refine_Misc
imports
Automatic_Refinement.Automatic_Refinement
Refine_Mono_Prover
begin
text \<open>Basic configuration for monotonicity prover:\<close>
lemmas [refine_mono] = monoI monotoneI[of "(\<le>)" "(\<le>)"]
lemmas [refine_mono] = TrueI le_funI order_refl
lemma case_prod_mono[refine_mono]:
"\<lbrakk>\<And>a b. p=(a,b) \<Longrightarrow> f a b \<le> f' a b\<rbrakk> \<Longrightarrow> case_prod f p \<le> case_prod f' p"
by (auto split: prod.split)
lemma case_option_mono[refine_mono]:
assumes "fn \<le> fn'"
assumes "\<And>v. x=Some v \<Longrightarrow> fs v \<le> fs' v"
shows "case_option fn fs x \<le> case_option fn' fs' x"
using assms by (auto split: option.split)
lemma case_list_mono[refine_mono]:
assumes "fn \<le> fn'"
assumes "\<And>x xs. l=x#xs \<Longrightarrow> fc x xs \<le> fc' x xs"
shows "case_list fn fc l \<le> case_list fn' fc' l"
using assms by (auto split: list.split)
lemma if_mono[refine_mono]:
assumes "b \<Longrightarrow> m1 \<le> m1'"
assumes "\<not>b \<Longrightarrow> m2 \<le> m2'"
shows "(if b then m1 else m2) \<le> (if b then m1' else m2')"
using assms by auto
lemma let_mono[refine_mono]:
"f x \<le> f' x' \<Longrightarrow> Let x f \<le> Let x' f'" by auto
subsection \<open>Uncategorized Lemmas\<close>
lemma all_nat_split_at: "\<forall>i::'a::linorder<k. P i \<Longrightarrow> P k \<Longrightarrow> \<forall>i>k. P i
\<Longrightarrow> \<forall>i. P i"
by (metis linorder_neq_iff)
subsection \<open>Well-Foundedness\<close>
lemma wf_no_infinite_down_chainI:
assumes "\<And>f. \<lbrakk>\<And>i. (f (Suc i), f i)\<in>r\<rbrakk> \<Longrightarrow> False"
shows "wf r"
by (metis assms wf_iff_no_infinite_down_chain)
text \<open>This lemma transfers well-foundedness over a simulation relation.\<close>
lemma sim_wf:
assumes WF: "wf (S'\<inverse>)"
assumes STARTR: "(x0,x0')\<in>R"
assumes SIM: "\<And>s s' t. \<lbrakk> (s,s')\<in>R; (s,t)\<in>S; (x0',s')\<in>S'\<^sup>* \<rbrakk>
\<Longrightarrow> \<exists>t'. (s',t')\<in>S' \<and> (t,t')\<in>R"
assumes CLOSED: "Domain S \<subseteq> S\<^sup>*``{x0}"
shows "wf (S\<inverse>)"
proof (rule wf_no_infinite_down_chainI, simp)
txt \<open>
Informal proof:
Assume there is an infinite chain in \<open>S\<close>.
Due to the closedness property of \<open>S\<close>, it can be extended to
start at \<open>x0\<close>.
Now, we inductively construct an infinite chain in \<open>S'\<close>, such that
each element of the new chain is in relation with the corresponding
element of the original chain:
The first element is \<open>x0'\<close>.
For any element \<open>i+1\<close>, the simulation property yields the next
element.
This chain contradicts well-foundedness of \<open>S'\<close>.
\<close>
fix f
assume CHAIN: "\<And>i. (f i, f (Suc i))\<in>S"
txt \<open>Extend to start with \<open>x0\<close>\<close>
obtain f' where CHAIN': "\<And>i. (f' i, f' (Suc i))\<in>S" and [simp]: "f' 0 = x0"
proof -
{
fix x
assume S: "x = f 0"
from CHAIN have "f 0 \<in> Domain S" by auto
with CLOSED have "(x0,x)\<in>S\<^sup>*" by (auto simp: S)
then obtain g k where G0: "g 0 = x0" and X: "x = g k"
and CH: "(\<forall>i<k. (g i, g (Suc i))\<in>S)"
proof induct
case base thus ?case by force
next
case (step y z) show ?case proof (rule step.hyps(3))
fix g k
assume "g 0 = x0" and "y = g k"
and "\<forall>i<k. (g i, g (Suc i))\<in>S"
thus ?case using \<open>(y,z)\<in>S\<close>
by (rule_tac step.prems[where g="g(Suc k := z)" and k="Suc k"])
auto
qed
qed
define f' where "f' i = (if i<k then g i else f (i-k))" for i
have "\<exists>f'. f' 0 = x0 \<and> (\<forall>i. (f' i, f' (Suc i))\<in>S)"
apply (rule_tac x=f' in exI)
apply (unfold f'_def)
apply (rule conjI)
using S X G0 apply (auto) []
apply (rule_tac k=k in all_nat_split_at)
apply (auto)
apply (simp add: CH)
apply (subgoal_tac "k = Suc i")
apply (simp add: S[symmetric] CH X)
apply simp
apply (simp add: CHAIN)
apply (subgoal_tac "Suc i - k = Suc (i-k)")
using CHAIN apply simp
apply simp
done
}
then obtain f' where "\<forall>i. (f' i,f' (Suc i))\<in>S" and "f' 0 = x0" by blast
thus ?thesis by (blast intro!: that)
qed
txt \<open>Construct chain in \<open>S'\<close>\<close>
define g' where "g' = rec_nat x0' (\<lambda>i x. SOME x'.
(x,x')\<in>S' \<and> (f' (Suc i),x')\<in>R \<and> (x0', x')\<in>S'\<^sup>* )"
{
fix i
note [simp] = g'_def
have "(g' i, g' (Suc i))\<in>S' \<and> (f' (Suc i),g' (Suc i))\<in>R
\<and> (x0',g' (Suc i))\<in>S'\<^sup>*"
proof (induct i)
case 0
from SIM[OF STARTR] CHAIN'[of 0] obtain t' where
"(x0',t')\<in>S'" and "(f' (Suc 0),t')\<in>R" by auto
moreover hence "(x0',t')\<in>S'\<^sup>*" by auto
ultimately show ?case
by (auto intro: someI2 simp: STARTR)
next
case (Suc i)
with SIM[OF _ CHAIN'[of "Suc i"]]
obtain t' where LS: "(g' (Suc i),t')\<in>S'" and "(f' (Suc (Suc i)),t')\<in>R"
by auto
moreover from LS Suc have "(x0', t')\<in>S'\<^sup>*" by auto
ultimately show ?case
apply simp
apply (rule_tac a="t'" in someI2)
apply auto
done
qed
} hence S'CHAIN: "\<forall>i. (g' i, g'(Suc i))\<in>S'" by simp
txt \<open>This contradicts well-foundedness\<close>
with WF show False
by (erule_tac wf_no_infinite_down_chainE[where f=g']) simp
qed
text \<open>Well-founded relation that approximates a finite set from below.\<close>
definition "finite_psupset S \<equiv> { (Q',Q). Q\<subset>Q' \<and> Q' \<subseteq> S }"
lemma finite_psupset_wf[simp, intro]: "finite S \<Longrightarrow> wf (finite_psupset S)"
unfolding finite_psupset_def by (blast intro: wf_bounded_supset)
definition "less_than_bool \<equiv> {(a,b). a<(b::bool)}"
lemma wf_less_than_bool[simp, intro!]: "wf (less_than_bool)"
unfolding less_than_bool_def
by (auto simp: wf_def)
lemma less_than_bool_iff[simp]:
"(x,y)\<in>less_than_bool \<longleftrightarrow> x=False \<and> y=True"
by (auto simp: less_than_bool_def)
definition "greater_bounded N \<equiv> inv_image less_than (\<lambda>x. N-x)"
lemma wf_greater_bounded[simp, intro!]: "wf (greater_bounded N)" by (auto simp: greater_bounded_def)
lemma greater_bounded_Suc_iff[simp]: "(Suc x,x)\<in>greater_bounded N \<longleftrightarrow> Suc x \<le> N"
by (auto simp: greater_bounded_def)
subsection \<open>Monotonicity and Orderings\<close>
lemma mono_const[simp, intro!]: "mono (\<lambda>_. c)" by (auto intro: monoI)
lemma mono_if: "\<lbrakk>mono S1; mono S2\<rbrakk> \<Longrightarrow>
mono (\<lambda>F s. if b s then S1 F s else S2 F s)"
apply rule
apply (rule le_funI)
apply (auto dest: monoD[THEN le_funD])
done
lemma mono_infI: "mono f \<Longrightarrow> mono g \<Longrightarrow> mono (inf f g)"
unfolding inf_fun_def
apply (rule monoI)
apply (metis inf_mono monoD)
done
lemma mono_infI':
"mono f \<Longrightarrow> mono g \<Longrightarrow> mono (\<lambda>x. inf (f x) (g x) :: 'b::lattice)"
by (rule mono_infI[unfolded inf_fun_def])
lemma mono_infArg:
fixes f :: "'a::lattice \<Rightarrow> 'b::order"
shows "mono f \<Longrightarrow> mono (\<lambda>x. f (inf x X))"
apply (rule monoI)
apply (erule monoD)
apply (metis inf_mono order_refl)
done
lemma mono_Sup:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> Sup (f`S) \<le> f (Sup S)"
apply (rule Sup_least)
apply (erule imageE)
apply simp
apply (erule monoD)
apply (erule Sup_upper)
done
lemma mono_SupI:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
assumes "mono f"
assumes "S'\<subseteq>f`S"
shows "Sup S' \<le> f (Sup S)"
by (metis Sup_subset_mono assms mono_Sup order_trans)
lemma mono_Inf:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "mono f \<Longrightarrow> f (Inf S) \<le> Inf (f`S)"
apply (rule Inf_greatest)
apply (erule imageE)
apply simp
apply (erule monoD)
apply (erule Inf_lower)
done
lemma mono_funpow: "mono (f::'a::order \<Rightarrow> 'a) \<Longrightarrow> mono (f^^i)"
apply (induct i)
apply (auto intro!: monoI)
apply (auto dest: monoD)
done
lemma mono_id[simp, intro!]:
"mono id"
"mono (\<lambda>x. x)"
by (auto intro: monoI)
declare SUP_insert[simp]
lemma (in semilattice_inf) le_infD1:
"a \<le> inf b c \<Longrightarrow> a \<le> b" by (rule le_infE)
lemma (in semilattice_inf) le_infD2:
"a \<le> inf b c \<Longrightarrow> a \<le> c" by (rule le_infE)
lemma (in semilattice_inf) inf_leI:
"\<lbrakk> \<And>x. \<lbrakk> x\<le>a; x\<le>b \<rbrakk> \<Longrightarrow> x\<le>c \<rbrakk> \<Longrightarrow> inf a b \<le> c"
by (metis inf_le1 inf_le2)
lemma top_Sup: "(top::'a::complete_lattice)\<in>A \<Longrightarrow> Sup A = top"
by (metis Sup_upper top_le)
lemma bot_Inf: "(bot::'a::complete_lattice)\<in>A \<Longrightarrow> Inf A = bot"
by (metis Inf_lower le_bot)
lemma mono_compD: "mono f \<Longrightarrow> x\<le>y \<Longrightarrow> f o x \<le> f o y"
apply (rule le_funI)
apply (auto dest: monoD le_funD)
done
subsubsection \<open>Galois Connections\<close>
locale galois_connection =
fixes \<alpha>::"'a::complete_lattice \<Rightarrow> 'b::complete_lattice" and \<gamma>
assumes galois: "c \<le> \<gamma>(a) \<longleftrightarrow> \<alpha>(c) \<le> a"
begin
lemma \<alpha>\<gamma>_defl: "\<alpha>(\<gamma>(x)) \<le> x"
proof -
have "\<gamma> x \<le> \<gamma> x" by simp
with galois show "\<alpha>(\<gamma>(x)) \<le> x" by blast
qed
lemma \<gamma>\<alpha>_infl: "x \<le> \<gamma>(\<alpha>(x))"
proof -
have "\<alpha> x \<le> \<alpha> x" by simp
with galois show "x \<le> \<gamma>(\<alpha>(x))" by blast
qed
lemma \<alpha>_mono: "mono \<alpha>"
proof
fix x::'a and y::'a
assume "x\<le>y"
also note \<gamma>\<alpha>_infl[of y]
finally show "\<alpha> x \<le> \<alpha> y" by (simp add: galois)
qed
lemma \<gamma>_mono: "mono \<gamma>"
by rule (metis \<alpha>\<gamma>_defl galois inf_absorb1 le_infE)
lemma dist_\<gamma>[simp]:
"\<gamma> (inf a b) = inf (\<gamma> a) (\<gamma> b)"
apply (rule order_antisym)
apply (rule mono_inf[OF \<gamma>_mono])
apply (simp add: galois)
apply (simp add: galois[symmetric])
done
lemma dist_\<alpha>[simp]:
"\<alpha> (sup a b) = sup (\<alpha> a) (\<alpha> b)"
by (metis (no_types) \<alpha>_mono galois mono_sup order_antisym
sup_ge1 sup_ge2 sup_least)
end
subsubsection \<open>Fixed Points\<close>
lemma mono_lfp_eqI:
assumes MONO: "mono f"
assumes FIXP: "f a \<le> a"
assumes LEAST: "\<And>x. f x = x \<Longrightarrow> a\<le>x"
shows "lfp f = a"
apply (rule antisym)
apply (rule lfp_lowerbound)
apply (rule FIXP)
by (metis LEAST MONO lfp_unfold)
lemma mono_gfp_eqI:
assumes MONO: "mono f"
assumes FIXP: "a \<le> f a"
assumes GREATEST: "\<And>x. f x = x \<Longrightarrow> x\<le>a"
shows "gfp f = a"
apply (rule antisym)
apply (metis GREATEST MONO gfp_unfold)
apply (rule gfp_upperbound)
apply (rule FIXP)
done
lemma lfp_le_gfp': "mono f \<Longrightarrow> lfp f x \<le> gfp f x"
by (rule le_funD[OF lfp_le_gfp])
(* Just a reformulation of lfp_induct *)
lemma lfp_induct':
assumes M: "mono f"
assumes IS: "\<And>m. \<lbrakk> m \<le> lfp f; m \<le> P \<rbrakk> \<Longrightarrow> f m \<le> P"
shows "lfp f \<le> P"
apply (rule lfp_induct[OF M])
apply (rule IS)
by simp_all
lemma lfp_gen_induct:
\<comment> \<open>Induction lemma for generalized lfps\<close>
assumes M: "mono f"
notes MONO'[refine_mono] = monoD[OF M]
assumes I0: "m0 \<le> P"
assumes IS: "\<And>m. \<lbrakk>
m \<le> lfp (\<lambda>s. sup m0 (f s)); \<comment> \<open>Assume already established invariants\<close>
m \<le> P; \<comment> \<open>Assume invariant\<close>
f m \<le> lfp (\<lambda>s. sup m0 (f s)) \<comment> \<open>Assume that step preserved est. invars\<close>
\<rbrakk> \<Longrightarrow> f m \<le> P" \<comment> \<open>Show that step preserves invariant\<close>
shows "lfp (\<lambda>s. sup m0 (f s)) \<le> P"
apply (rule lfp_induct')
apply (meson MONO' monoI order_mono_setup.refl sup_mono)
apply (rule sup_least)
apply (rule I0)
apply (rule IS, assumption+)
apply (subst lfp_unfold)
apply (meson MONO' monoI order_mono_setup.refl sup_mono)
apply (rule le_supI2)
apply (rule monoD[OF M])
.
subsubsection \<open>Connecting Complete Lattices and
Chain-Complete Partial Orders\<close>
(* Note: Also connected by subclass now. However, we need both directions
of embedding*)
lemma (in complete_lattice) is_ccpo: "class.ccpo Sup (\<le>) (<)"
apply unfold_locales
apply (erule Sup_upper)
apply (erule Sup_least)
done
lemma (in complete_lattice) is_dual_ccpo: "class.ccpo Inf (\<ge>) (>)"
apply unfold_locales
apply (rule less_le_not_le)
apply (rule order_refl)
apply (erule (1) order_trans)
- apply (erule (1) antisym)
+ apply (erule (1) order.antisym)
apply (erule Inf_lower)
apply (erule Inf_greatest)
done
lemma ccpo_mono_simp: "monotone (\<le>) (\<le>) f \<longleftrightarrow> mono f"
unfolding monotone_def mono_def by simp
lemma ccpo_monoI: "mono f \<Longrightarrow> monotone (\<le>) (\<le>) f"
by (simp add: ccpo_mono_simp)
lemma ccpo_monoD: "monotone (\<le>) (\<le>) f \<Longrightarrow> mono f"
by (simp add: ccpo_mono_simp)
lemma dual_ccpo_mono_simp: "monotone (\<ge>) (\<ge>) f \<longleftrightarrow> mono f"
unfolding monotone_def mono_def by auto
lemma dual_ccpo_monoI: "mono f \<Longrightarrow> monotone (\<ge>) (\<ge>) f"
by (simp add: dual_ccpo_mono_simp)
lemma dual_ccpo_monoD: "monotone (\<ge>) (\<ge>) f \<Longrightarrow> mono f"
by (simp add: dual_ccpo_mono_simp)
lemma ccpo_lfp_simp: "\<And>f. mono f \<Longrightarrow> ccpo.fixp Sup (\<le>) f = lfp f"
apply (rule antisym)
defer
apply (rule lfp_lowerbound)
apply (drule ccpo.fixp_unfold[OF is_ccpo ccpo_monoI, symmetric])
apply simp
apply (rule ccpo.fixp_lowerbound[OF is_ccpo ccpo_monoI], assumption)
apply (simp add: lfp_unfold[symmetric])
done
lemma ccpo_gfp_simp: "\<And>f. mono f \<Longrightarrow> ccpo.fixp Inf (\<ge>) f = gfp f"
apply (rule antisym)
apply (rule gfp_upperbound)
apply (drule ccpo.fixp_unfold[OF is_dual_ccpo dual_ccpo_monoI, symmetric])
apply simp
apply (rule ccpo.fixp_lowerbound[OF is_dual_ccpo dual_ccpo_monoI], assumption)
apply (simp add: gfp_unfold[symmetric])
done
abbreviation "chain_admissible P \<equiv> ccpo.admissible Sup (\<le>) P"
abbreviation "is_chain \<equiv> Complete_Partial_Order.chain (\<le>)"
lemmas chain_admissibleI[intro?] = ccpo.admissibleI[where lub=Sup and ord="(\<le>)"]
abbreviation "dual_chain_admissible P \<equiv> ccpo.admissible Inf (\<lambda>x y. y\<le>x) P"
abbreviation "is_dual_chain \<equiv> Complete_Partial_Order.chain (\<lambda>x y. y\<le>x)"
lemmas dual_chain_admissibleI[intro?] =
ccpo.admissibleI[where lub=Inf and ord="(\<lambda>x y. y\<le>x)"]
lemma dual_chain_iff[simp]: "is_dual_chain C = is_chain C"
unfolding chain_def
apply auto
done
lemmas chain_dualI = iffD1[OF dual_chain_iff]
lemmas dual_chainI = iffD2[OF dual_chain_iff]
lemma is_chain_empty[simp, intro!]: "is_chain {}"
by (rule chainI) auto
lemma is_dual_chain_empty[simp, intro!]: "is_dual_chain {}"
by (rule dual_chainI) auto
lemma point_chainI: "is_chain M \<Longrightarrow> is_chain ((\<lambda>f. f x)`M)"
by (auto intro: chainI le_funI dest: chainD le_funD)
text \<open>We transfer the admissible induction lemmas to complete
lattices.\<close>
lemma lfp_cadm_induct:
"\<lbrakk>chain_admissible P; P (Sup {}); mono f; \<And>x. P x \<Longrightarrow> P (f x)\<rbrakk> \<Longrightarrow> P (lfp f)"
by (simp only: ccpo_mono_simp[symmetric] ccpo_lfp_simp[symmetric])
(rule ccpo.fixp_induct[OF is_ccpo])
lemma gfp_cadm_induct:
"\<lbrakk>dual_chain_admissible P; P (Inf {}); mono f; \<And>x. P x \<Longrightarrow> P (f x)\<rbrakk> \<Longrightarrow> P (gfp f)"
by (simp only: dual_ccpo_mono_simp[symmetric] ccpo_gfp_simp[symmetric])
(rule ccpo.fixp_induct[OF is_dual_ccpo])
subsubsection \<open>Continuity and Kleene Fixed Point Theorem\<close>
definition "cont f \<equiv> \<forall>C. C\<noteq>{} \<longrightarrow> f (Sup C) = Sup (f`C)"
definition "strict f \<equiv> f bot = bot"
definition "inf_distrib f \<equiv> strict f \<and> cont f"
lemma contI[intro?]: "\<lbrakk>\<And>C. C\<noteq>{} \<Longrightarrow> f (Sup C) = Sup (f`C)\<rbrakk> \<Longrightarrow> cont f"
unfolding cont_def by auto
lemma contD: "cont f \<Longrightarrow> C\<noteq>{} \<Longrightarrow> f (Sup C) = Sup (f ` C)"
unfolding cont_def by auto
lemma contD': "cont f \<Longrightarrow> C\<noteq>{} \<Longrightarrow> f (Sup C) = Sup (f ` C)"
by (fact contD)
lemma strictD[dest]: "strict f \<Longrightarrow> f bot = bot"
unfolding strict_def by auto
\<comment> \<open>We only add this lemma to the simpset for functions on the same type.
Otherwise, the simplifier tries it much too often.\<close>
lemma strictD_simp[simp]: "strict f \<Longrightarrow> f (bot::'a::bot) = (bot::'a)"
unfolding strict_def by auto
lemma strictI[intro?]: "f bot = bot \<Longrightarrow> strict f"
unfolding strict_def by auto
lemma inf_distribD[simp]:
"inf_distrib f \<Longrightarrow> strict f"
"inf_distrib f \<Longrightarrow> cont f"
unfolding inf_distrib_def by auto
lemma inf_distribI[intro?]: "\<lbrakk>strict f; cont f\<rbrakk> \<Longrightarrow> inf_distrib f"
unfolding inf_distrib_def by auto
lemma inf_distribD'[simp]:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "inf_distrib f \<Longrightarrow> f (Sup C) = Sup (f`C)"
apply (cases "C={}")
apply (auto dest: inf_distribD contD')
done
lemma inf_distribI':
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
assumes B: "\<And>C. f (Sup C) = Sup (f`C)"
shows "inf_distrib f"
apply (rule)
apply (rule)
apply (rule B[of "{}", simplified])
apply (rule)
apply (rule B)
done
lemma cont_is_mono[simp]:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "cont f \<Longrightarrow> mono f"
apply (rule monoI)
apply (drule_tac C="{x,y}" in contD)
apply (auto simp: le_iff_sup)
done
lemma inf_distrib_is_mono[simp]:
fixes f :: "'a::complete_lattice \<Rightarrow> 'b::complete_lattice"
shows "inf_distrib f \<Longrightarrow> mono f"
by simp
text \<open>Only proven for complete lattices here. Also holds for CCPOs.\<close>
theorem gen_kleene_lfp:
fixes f:: "'a::complete_lattice \<Rightarrow> 'a"
assumes CONT: "cont f"
shows "lfp (\<lambda>x. sup m (f x)) = (SUP i. (f^^i) m)"
proof (rule antisym)
let ?f = "(\<lambda>x. sup m (f x))"
let ?K="{ (f^^i) m | i . True}"
note MONO=cont_is_mono[OF CONT]
note MONO'[refine_mono] = monoD[OF MONO]
{
fix i
have "(f^^i) m \<le> lfp ?f"
apply (induct i)
apply simp
apply (subst lfp_unfold)
apply (meson MONO' monoI order_mono_setup.refl sup_mono)
apply simp
apply (subst lfp_unfold)
apply (meson MONO' monoI order_mono_setup.refl sup_mono)
apply simp
by (metis MONO' le_supI2)
} thus "(SUP i. (f^^i) m) \<le> lfp ?f" by (blast intro: SUP_least)
next
let ?f = "(\<lambda>x. sup m (f x))"
show "lfp ?f \<le> (SUP i. (f^^i) m)"
apply (rule lfp_lowerbound)
apply (rule sup_least)
apply (rule order_trans[OF _ SUP_upper[where i=0]], simp_all) []
apply (simp add: contD [OF CONT])
apply (rule Sup_subset_mono)
apply (auto)
apply (rule_tac x="Suc i" in range_eqI)
apply simp
done
qed
theorem kleene_lfp:
fixes f:: "'a::complete_lattice \<Rightarrow> 'a"
assumes CONT: "cont f"
shows "lfp f = (SUP i. (f^^i) bot)"
using gen_kleene_lfp[OF CONT,where m=bot] by simp
theorem (* Detailed proof *)
fixes f:: "'a::complete_lattice \<Rightarrow> 'a"
assumes CONT: "cont f"
shows "lfp f = (SUP i. (f^^i) bot)"
proof (rule antisym)
let ?K="{ (f^^i) bot | i . True}"
note MONO=cont_is_mono[OF CONT]
{
fix i
have "(f^^i) bot \<le> lfp f"
apply (induct i)
apply simp
apply simp
by (metis MONO lfp_unfold monoD)
} thus "(SUP i. (f^^i) bot) \<le> lfp f" by (blast intro: SUP_least)
next
show "lfp f \<le> (SUP i. (f^^i) bot)"
apply (rule lfp_lowerbound)
apply (simp add: contD [OF CONT])
apply (rule Sup_subset_mono)
apply auto
apply (rule_tac x="Suc i" in range_eqI)
apply auto
done
qed
(* Alternative proof of gen_kleene_lfp that re-uses standard Kleene, but is more tedious *)
lemma SUP_funpow_contracting:
fixes f :: "'a \<Rightarrow> ('a::complete_lattice)"
assumes C: "cont f"
shows "f (SUP i. (f^^i) m) \<le> (SUP i. (f^^i) m)"
proof -
have 1: "\<And>i x. f ((f^^i) x) = (f^^(Suc i)) x"
by simp
have "f (SUP i. (f^^i) m) = (SUP i. f ((f ^^ i) m))"
by (subst contD[OF C]) (simp_all add: image_comp)
also have "\<dots> \<le> (SUP i. (f^^i) m)"
apply (rule SUP_least)
apply (simp, subst 1)
apply (rule SUP_upper)
..
finally show ?thesis .
qed
lemma gen_kleene_chain_conv:
fixes f :: "'a::complete_lattice \<Rightarrow> 'a"
assumes C: "cont f"
shows "(SUP i. (f^^i) m) = (SUP i. ((\<lambda>x. sup m (f x))^^i) bot)"
proof -
let ?f' = "\<lambda>x. sup m (f x)"
show ?thesis
proof (intro antisym SUP_least)
from C have C': "cont ?f'"
unfolding cont_def
by (simp add: SUP_sup_distrib[symmetric])
fix i
show "(f ^^ i) m \<le> (SUP i. (?f' ^^ i) bot)"
proof (induction i)
case 0 show ?case
by (rule order_trans[OF _ SUP_upper[where i=1]]) auto
next
case (Suc i)
from cont_is_mono[OF C, THEN monoD, OF Suc]
have "(f ^^ (Suc i)) m \<le> f (SUP i. ((\<lambda>x. sup m (f x)) ^^ i) bot)"
by simp
also have "\<dots> \<le> sup m \<dots>" by simp
also note SUP_funpow_contracting[OF C']
finally show ?case .
qed
next
fix i
show "(?f'^^i) bot \<le> (SUP i. (f^^i) m)"
proof (induction i)
case 0 thus ?case by simp
next
case (Suc i)
from monoD[OF cont_is_mono[OF C] Suc]
have "(?f'^^Suc i) bot \<le> sup m (f (SUP i. (f ^^ i) m))"
by (simp add: le_supI2)
also have "\<dots> \<le> (SUP i. (f ^^ i) m)"
apply (rule sup_least)
apply (rule order_trans[OF _ SUP_upper[where i=0]], simp_all) []
apply (rule SUP_funpow_contracting[OF C])
done
finally show ?case .
qed
qed
qed
theorem
assumes C: "cont f"
shows "lfp (\<lambda>x. sup m (f x)) = (SUP i. (f^^i) m)"
apply (subst gen_kleene_chain_conv[OF C])
apply (rule kleene_lfp)
using C
unfolding cont_def
apply (simp add: SUP_sup_distrib[symmetric])
done
lemma (in galois_connection) dual_inf_dist_\<gamma>: "\<gamma> (Inf C) = Inf (\<gamma>`C)"
apply (rule antisym)
apply (rule Inf_greatest)
apply clarify
apply (rule monoD[OF \<gamma>_mono])
apply (rule Inf_lower)
apply simp
apply (subst galois)
apply (rule Inf_greatest)
apply (subst galois[symmetric])
apply (rule Inf_lower)
apply simp
done
lemma (in galois_connection) inf_dist_\<alpha>: "inf_distrib \<alpha>"
apply (rule inf_distribI')
apply (rule antisym)
apply (subst galois[symmetric])
apply (rule Sup_least)
apply (subst galois)
apply (rule Sup_upper)
apply simp
apply (rule Sup_least)
apply clarify
apply (rule monoD[OF \<alpha>_mono])
apply (rule Sup_upper)
apply simp
done
subsection \<open>Maps\<close>
subsubsection \<open>Key-Value Set\<close>
lemma map_to_set_simps[simp]:
"map_to_set Map.empty = {}"
"map_to_set [a\<mapsto>b] = {(a,b)}"
"map_to_set (m|`K) = map_to_set m \<inter> K\<times>UNIV"
"map_to_set (m(x:=None)) = map_to_set m - {x}\<times>UNIV"
"map_to_set (m(x\<mapsto>v)) = map_to_set m - {x}\<times>UNIV \<union> {(x,v)}"
"map_to_set m \<inter> dom m\<times>UNIV = map_to_set m"
"m k = Some v \<Longrightarrow> (k,v)\<in>map_to_set m"
"single_valued (map_to_set m)"
apply (simp_all)
by (auto simp: map_to_set_def restrict_map_def split: if_split_asm
intro: single_valuedI)
lemma map_to_set_inj:
"(k,v)\<in>map_to_set m \<Longrightarrow> (k,v')\<in>map_to_set m \<Longrightarrow> v = v'"
by (auto simp: map_to_set_def)
end
diff --git a/thys/RefinementReactive/Temporal.thy b/thys/RefinementReactive/Temporal.thy
--- a/thys/RefinementReactive/Temporal.thy
+++ b/thys/RefinementReactive/Temporal.thy
@@ -1,260 +1,260 @@
theory Temporal imports Main
begin
section\<open>Linear Temporal Logic\<close>
text\<open>
In this section we introduce an algebraic axiomatization of Linear Temporal Logic (LTL).
We model LTL formulas semantically as predicates on traces. For example the LTL formula
$\alpha = \Box\; \diamondsuit\; (x = 1)$ is modeled as a predicate
$\alpha : (nat \Rightarrow nat) \Rightarrow bool$, where
$\alpha \;x = True$ if $x\;i=1$ for infinitely many $i:nat$. In this formula $\Box$
and $\diamondsuit$ denote the always and eventually operators, respectively.
Formulas with multiple variables are modeled similarly. For example a formula $\alpha$ in two
variables is modeled as $\alpha : (nat \Rightarrow \tv a) \Rightarrow (nat \Rightarrow \tv b) \Rightarrow bool$,
and for example $(\Box\; \alpha) \; x\; y$ is defined as $(\forall i . \alpha \; x[i..]\; y[i..])$,
where $x[i..]\;j = x\;(i+j)$. We would like to construct an algebraic structure (Isabelle class)
which has the temporal operators as operations, and which has instatiations to
$(nat \Rightarrow \tv a) \Rightarrow bool$, $(nat \Rightarrow \tv a) \Rightarrow (nat \Rightarrow \tv b) \Rightarrow bool$,
and so on. Ideally our structure should be such that if we have this structure on a type $\tv a::temporal$,
then we could extend it to $(nat \Rightarrow \tv b) \Rightarrow \tv a$ in a way similar to the
way Boolean algebras are extended from a type $\tv a::boolean\_algebra$ to $\tv b\Rightarrow \tv a$.
Unfortunately, if we use for example $\Box$ as primitive operation on our temporal structure,
then we cannot extend $\Box$ from $\tv a::temporal$ to $(nat \Rightarrow \tv b) \Rightarrow \tv a$. A
possible extension of $\Box$ could be
$$(\Box\; \alpha)\;x = \bigwedge_{i:nat} \Box (\alpha\; x[i..]) \mbox{ and } \Box \; b = b$$
where $\alpha: (nat \Rightarrow \tv b) \Rightarrow \tv a$ and $b:bool$. However, if we apply this
definition to $\alpha : (nat \Rightarrow \tv a) \Rightarrow (nat \Rightarrow \tv b) \Rightarrow bool$,
then we get
$$(\Box\; \alpha) \; x\; y = (\forall i\;j. \alpha \; x[i..]\; y[j..])$$
which is not correct.
To evercome this problem we introduce as a primitive operation $!!:\tv a \Rightarrow nat \Rightarrow \tv a$,
where $\tv a$ is the type of temporal formulas, and $\alpha !! i$ is the formula $\alpha$ at time point $i$.
If $\alpha$ is a formula in two variables as before, then
$$(\alpha !! i)\; x\;y = \alpha\; x[i..]\;y[i..].$$
and we define for example the the operator always by
$$\Box \alpha = \bigwedge_{i:nat} \alpha !! i$$
\<close>
notation
bot ("\<bottom>") and
top ("\<top>") and
inf (infixl "\<sqinter>" 70)
and sup (infixl "\<squnion>" 65)
class temporal = complete_boolean_algebra +
fixes at :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 150)
assumes [simp]: "a !! i !! j = a !! (i + j)"
assumes [simp]: "a !! 0 = a"
assumes [simp]: "\<top> !! i = \<top>"
assumes [simp]: "-(a !! i) = (-a) !! i"
assumes [simp]: "(a \<sqinter> b) !! i = (a !! i) \<sqinter> (b !! i)"
begin
definition always :: "'a \<Rightarrow> 'a" ("\<box> (_)" [900] 900) where
"\<box> p = (INF i . p !! i)"
definition eventually :: "'a \<Rightarrow> 'a" ("\<diamond> (_)" [900] 900) where
"\<diamond> p = (SUP i . p !! i)"
definition "next" :: "'a \<Rightarrow> 'a" ("\<circle> (_)" [900] 900) where
"\<circle> p = p !! (Suc 0)"
definition until :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "until" 65) where
"(p until q) = (SUP n . (Inf (at p ` {i . i < n})) \<sqinter> (q !! n))"
end
text\<open>
Next lemma, in the context of complete boolean algebras, will be used
to prove $-(p\ until\ -p) = \Box\; p$.
\<close>
context complete_boolean_algebra
begin
lemma until_always: "(INF n. (SUP i \<in> {i. i < n} . - p i) \<squnion> ((p :: nat \<Rightarrow> 'a) n)) \<le> p n"
proof -
have "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> (INF i\<in>{i. i \<le> n}. p i)"
proof (induction n)
have "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> (SUP i\<in>{i. i < 0}. - p i) \<squnion> p 0"
by (rule INF_lower, simp)
also have "... \<le> (INF i\<in>{i. i \<le> 0}. p i)"
by simp
finally show "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> (INF i\<in>{i. i \<le> 0}. p i)"
by simp
next
fix n::nat assume "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> (INF i \<in> {i. i \<le> n}. p i)"
also have "\<And> i . i \<le> n \<Longrightarrow> ... \<le> p i" by (rule INF_lower, simp)
finally have [simp]: "\<And> i . i \<le> n \<Longrightarrow> (INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> p i"
by simp
show "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> (INF i \<in> {i. i \<le> Suc n}. p i)"
proof (rule INF_greatest, safe, cases)
fix i::nat
assume "i \<le> n" from this show "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> p i" by simp
next
fix i::nat
have A: "{i. i \<le> n} = {i . i < Suc n}" by auto
have B: "(SUP i\<in>{i. i \<le> n}. - p i) \<le> - (INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n)"
by (metis (lifting, mono_tags) \<open>(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> (INF i\<in>{i. i \<le> n}. p i)\<close> compl_mono uminus_INF)
assume "i \<le> Suc n" and "\<not> i \<le> n"
from this have [simp]: "i = Suc n" by simp
have "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> (INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<sqinter> ((SUP i\<in>{i. i \<le> n}. - p i) \<squnion> p (Suc n))"
by (simp add: A, rule INF_lower, simp)
also have "... \<le> ((INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<sqinter> ((- (INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n)) \<squnion> p (Suc n)))"
by (rule inf_mono, simp_all, rule_tac y = "- (INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n)" in order_trans, simp_all add: B)
also have "... \<le> p i"
by (simp add: inf_sup_distrib1 inf_compl_bot)
finally show "(INF n. (SUP i\<in>{i. i < n}. - p i) \<squnion> p n) \<le> p i" by simp
qed
qed
also have "(INF i\<in>{i. i \<le> n}. p i) \<le> p n" by (rule INF_lower, auto)
finally show "(INF n. (SUP i \<in> {i. i < n} . - p i) \<squnion> ((p :: nat \<Rightarrow> 'a) n)) \<le> p n" by simp
qed
end
text\<open>
We prove now a number of results of the temporal class.
\<close>
context temporal
begin
lemma [simp]: "(a \<squnion> b) !! i = (a !! i) \<squnion> (b !! i)"
by (subst compl_eq_compl_iff [THEN sym], simp)
lemma always_less [simp]: "\<box> p \<le> p"
proof -
have "\<box> p \<le> p !! 0"
by (unfold always_def, rule INF_lower, simp)
also have "p !! 0 = p" by simp
finally show "\<box> p \<le> p" by simp
qed
lemma always_and: "\<box> (p \<sqinter> q) = (\<box> p) \<sqinter> (\<box> q)"
by (simp add: always_def INF_inf_distrib)
lemma eventually_or: "\<diamond> (p \<squnion> q) = (\<diamond> p) \<squnion> (\<diamond> q)"
by (simp add: eventually_def SUP_sup_distrib)
lemma neg_until_always: "-(p until -p) = \<box> p"
- proof (rule antisym)
+ proof (rule order.antisym)
show "- (p until - p) \<le> \<box> p"
by (simp add: until_def always_def uminus_SUP uminus_INF, rule INF_greatest, cut_tac p = "\<lambda> n . p !! n" in until_always, simp)
next
have "\<And> n . \<box> p \<le> p !! n"
by (simp add: always_def INF_lower)
also have "\<And> n . p !! n \<le> (SUP x\<in>{i. i < n}. (- p) !! x) \<squnion> p !! n"
by simp
finally show "\<box> p \<le> -(p until -p)"
apply (simp add: until_def uminus_SUP uminus_INF)
by (rule INF_greatest, simp)
qed
lemma neg_always_eventually: "\<box> p = - \<diamond> (- p)"
by (simp add: fun_eq_iff always_def eventually_def until_def uminus_SUP)
lemma neg_true_until_always: "-(\<top> until -p) = \<box> p"
by (simp add: fun_eq_iff always_def until_def uminus_SUP uminus_INF)
lemma true_until_eventually: "(\<top> until p) = \<diamond> p"
by (cut_tac p = "-p" in neg_always_eventually, cut_tac p = "-p" in neg_true_until_always, simp)
end
text\<open>
Boolean algebras with $b!!i = b$ form a temporal class.
\<close>
instantiation bool :: temporal
begin
definition at_bool_def [simp]: "(p::bool) !! i = p"
instance proof
qed auto
end
type_synonym 'a trace = "nat \<Rightarrow> 'a"
text\<open>
Asumming that $\tv a::temporal$ is a type of class $temporal$, and $\tv b$ is an arbitrary type,
we would like to create the instantiation of $\tv b\ trace \Rightarrow \tv a$ as a temporal
class. However Isabelle allows only instatiations of functions from a class to another
class. To solve this problem we introduce a new class called trace with an operation
$\mathit{suffix}::\tv a \Rightarrow nat \Rightarrow \tv a$ where
$\mathit{suffix}\;a\;i\;j = (a[i..])\; j = a\;(i+j)$ when
$a$ is a trace with elements of some type $\tv b$ ($\tv a = nat \Rightarrow \tv b$).
\<close>
class trace =
fixes suffix :: "'a \<Rightarrow> nat \<Rightarrow> 'a" ("_[_ ..]" [80, 15] 80)
assumes [simp]: "a[i..][j..] = a[i + j..]"
assumes [simp]: "a[0..] = a"
begin
definition "next_trace" :: "'a \<Rightarrow> 'a" ("\<odot> (_)" [900] 900) where
"\<odot> p = p[Suc 0..]"
end
instantiation "fun" :: (trace, temporal) temporal
begin
definition at_fun_def: "(P:: 'a \<Rightarrow> 'b) !! i = (\<lambda> x . (P (x[i..])) !! i)"
instance proof qed (simp_all add: at_fun_def add.commute fun_eq_iff le_fun_def)
end
text\<open>
In the last part of our formalization, we need to instantiate the functions
from $nat$ to some arbitrary type $\tv a$ as a trace class. However, this again is not
possible using the instatiation mechanism of Isabelle. We solve this problem
by creating another class called $nat$, and then we instatiate the functions
from $\tv a::nat$ to $\tv b$ as traces. The class $nat$ is defined such that if we
have a type $\tv a::nat$, then $\tv a$ is isomorphic to the type $nat$.
\<close>
class nat = zero + plus + minus +
fixes RepNat :: "'a \<Rightarrow> nat"
fixes AbsNat :: "nat \<Rightarrow> 'a"
assumes [simp]: "RepNat (AbsNat n) = n"
and [simp]: "AbsNat (RepNat x) = x"
and zero_Nat_def: "0 = AbsNat 0"
and plus_Nat_def: "a + b = AbsNat (RepNat a + RepNat b)"
and minus_Nat_def: "a - b = AbsNat (RepNat a - RepNat b)"
begin
lemma AbsNat_plus: "AbsNat (i + j) = AbsNat i + AbsNat j"
by (simp add: plus_Nat_def)
lemma AbsNat_zero [simp]: "AbsNat 0 + i = i"
by (simp add: plus_Nat_def)
subclass comm_monoid_diff
apply (unfold_locales)
apply (simp_all add: plus_Nat_def zero_Nat_def minus_Nat_def add.assoc)
by (simp add: add.commute)
end
text\<open>
The type natural numbers is an instantiation of the class $nat$.
\<close>
instantiation nat :: nat
begin
definition RepNat_nat_def [simp]: "(RepNat:: nat \<Rightarrow> nat) = id"
definition AbsNat_nat_def [simp]: "(AbsNat:: nat \<Rightarrow> nat) = id"
instance proof
qed auto
end
text\<open>
Finally, functions from $\tv a::nat$ to some arbitrary type $\tv b$ are instatiated
as a trace class.
\<close>
instantiation "fun" :: (nat, type) trace
begin
definition at_trace_def [simp]: "((t :: 'a \<Rightarrow> 'b)[i..]) j = (t (AbsNat i + j))"
instance proof
qed (simp_all add: fun_eq_iff AbsNat_plus add.assoc)
end
text\<open>
By putting together all class definitions and instatiations introduced so far, we obtain the
temporal class structure for predicates on traces with arbitrary number of parameters.
For example in the next lemma $r$ and $r'$ are predicate relations, and the operator
always is available for them as a consequence of the above construction.
\<close>
lemma "(\<box> r) OO (\<box> r') \<le> (\<box> (r OO r'))"
by (simp add: le_fun_def always_def at_fun_def, auto)
end
diff --git a/thys/Regular_Algebras/Regular_Algebra_Variants.thy b/thys/Regular_Algebras/Regular_Algebra_Variants.thy
--- a/thys/Regular_Algebras/Regular_Algebra_Variants.thy
+++ b/thys/Regular_Algebras/Regular_Algebra_Variants.thy
@@ -1,149 +1,149 @@
(* Title: Regular Algebras
Author: Simon Foster, Georg Struth
Maintainer: Simon Foster <s.foster at york.ac.uk>
Georg Struth <g.struth at sheffield.ac.uk>
*)
section \<open>Variants of Regular Algebra\<close>
theory Regular_Algebra_Variants
imports Regular_Algebras Pratts_Counterexamples
begin
text \<open>Replacing Kozen's induction axioms by Boffa's leads to incompleteness.\<close>
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x. x \<cdot> x = x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
shows "\<And>x y. (x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
(* nitpick [expect=genuine] --"5-element counterexample"*)
oops
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x. x \<cdot> x = x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
shows "\<And>x y. x \<le> y \<longrightarrow> x\<^sup>\<star> \<le> y\<^sup>\<star>"
(* nitpick [expect=genuine] --"5-element counterexample"*)
oops
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x y. 1 + x \<le> y \<and> y \<cdot> y \<le> y \<longrightarrow> x\<^sup>\<star> \<le> y"
shows "\<And>x. 1 + x \<le> x\<^sup>\<star>"
by (metis add_iso_r assms(1) mult_oner subdistl)
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x y. 1 + x \<le> y \<and> y \<cdot> y \<le> y \<longrightarrow> x\<^sup>\<star> \<le> y"
shows "\<And>x. x\<^sup>\<star> = (1 + x)\<^sup>\<star>"
oops \<comment> \<open>need to reconsider this\<close>
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x y. 1 + x + y \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y"
shows "\<And>x. x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
oops
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x y z. x \<cdot> y = y \<cdot> z \<Longrightarrow> x\<^sup>\<star> \<cdot> y = y \<cdot> z\<^sup>\<star>"
shows "1\<^sup>\<star> = 1"
(* nitpick -- "3-element counterexample"*)
oops
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x y z. x \<cdot> y \<le> y \<cdot> z \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y \<cdot> z\<^sup>\<star>"
and "\<And>x y z. y \<cdot> x \<le> z \<cdot> y \<Longrightarrow> y \<cdot> x\<^sup>\<star> \<le> z\<^sup>\<star> \<cdot> y"
shows "1\<^sup>\<star> = 1"
(* nitpick [expect=genuine] -- "3-element counterexample"*)
oops
lemma (in star_dioid)
assumes "\<And>x. 1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and "\<And>x. 1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and "\<And>x y. x = y \<cdot> x \<Longrightarrow> x = y\<^sup>\<star> \<cdot> x"
and "\<And>x y. x = x \<cdot> y \<Longrightarrow> x = x \<cdot> y\<^sup>\<star>"
shows "\<And>x y. y \<cdot> x \<le> y \<Longrightarrow> y \<cdot> x\<^sup>\<star> \<le> y"
oops
class C3l_var = conway_dioid +
assumes C3l_var: "z + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> y"
class C3r_var = conway_dioid +
assumes C3r_var: "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"
class C3_var = C3l_var + C3r_var
sublocale C3l_var \<subseteq> C3l_algebra
apply unfold_locales
by (simp add: local.C3l_var)
sublocale C3l_algebra \<subseteq> C3l_var
by (unfold_locales, metis star_inductl_var)
sublocale C3_var \<subseteq> C3_algebra
apply unfold_locales
by (simp add: local.C3r_var)
sublocale C3_algebra \<subseteq> C3_var
by (unfold_locales, metis star_inductr_var)
class Brtc_algebra = star_dioid +
assumes rtc1: "1 + x\<^sup>\<star> \<cdot> x\<^sup>\<star> + x \<le> x\<^sup>\<star>"
and rtc2: "1 + x + y \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y"
sublocale B2_algebra \<subseteq> Brtc_algebra
proof
fix x y
show "1 + x\<^sup>\<star> \<cdot> x\<^sup>\<star> + x \<le> x\<^sup>\<star>"
by (simp add: local.star_ext)
show "1 + x + y \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y"
- by (metis B23 local.join.le_sup_iff eq_iff mult_1_right subdistl)
+ by (metis B23 local.join.le_sup_iff order.eq_iff mult_1_right subdistl)
qed
sublocale Brtc_algebra \<subseteq> B2_algebra
proof
fix x y
show "1 + x \<le> x\<^sup>\<star>"
by (metis rtc1 join.le_sup_iff)
show "x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
- proof (rule antisym)
+ proof (rule order.antisym)
show "x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis rtc1 join.le_sup_iff)
show "x\<^sup>\<star> \<le> x\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (metis rtc1 add.commute join.le_sup_iff less_eq_def mult_isor mult_onel)
qed
show "\<lbrakk> 1 + x \<le> y; y \<cdot> y = y \<rbrakk> \<Longrightarrow> x\<^sup>\<star> \<le> y"
by (metis rtc2 eq_refl less_eq_def)
qed
class wB1_algebra = conway_dioid +
assumes wR: "x \<cdot> x \<le> x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
sublocale wB1_algebra \<subseteq> B1_algebra
by (unfold_locales, metis order_refl wR)
lemma (in B1_algebra) one_plus_star: "x\<^sup>\<star> = (1 + x)\<^sup>\<star>"
by (metis C11_var R add_idem' mult_onel mult_oner)
sublocale B1_algebra \<subseteq> wB1_algebra
proof unfold_locales
fix x :: 'a
assume "x \<cdot> x \<le> x"
hence "x \<cdot> (1 + x) = x"
by (simp add: local.distrib_left local.join.sup_absorb1)
hence "(1 + x) \<cdot> (1 + x) = 1 + x"
using add.left_commute distrib_right' by simp
thus "x\<^sup>\<star> = 1 + x"
by (metis R add_assoc' add_idem' one_plus_star)
qed
end
diff --git a/thys/Regular_Algebras/Regular_Algebras.thy b/thys/Regular_Algebras/Regular_Algebras.thy
--- a/thys/Regular_Algebras/Regular_Algebras.thy
+++ b/thys/Regular_Algebras/Regular_Algebras.thy
@@ -1,1288 +1,1288 @@
(* Title: Regular Algebras
Author: Simon Foster, Georg Struth
Maintainer: Simon Foster <s.foster at york.ac.uk>
Georg Struth <g.struth at sheffield.ac.uk>
*)
section \<open>Regular Algebras\<close>
theory Regular_Algebras
imports Dioid_Power_Sum Kleene_Algebra.Finite_Suprema Kleene_Algebra.Kleene_Algebra
begin
subsection \<open>Conway's Classical Axioms\<close>
text \<open>Conway's classical axiomatisation of Regular Algebra from~\cite{Conway}.\<close>
class star_dioid = dioid_one_zero + star_op + plus_ord
class conway_dioid = star_dioid +
assumes C11: "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
and C12: "(x \<cdot> y)\<^sup>\<star> = 1 + x \<cdot>(y \<cdot> x)\<^sup>\<star> \<cdot> y"
class strong_conway_dioid = conway_dioid +
assumes C13: "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
class C_algebra = strong_conway_dioid +
assumes C14: "x\<^sup>\<star> = (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
text \<open>We tried to dualise using sublocales, but this causes an infinite loop on dual.dual.dual....\<close>
lemma (in conway_dioid) C11_var: "(x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
proof -
have "x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star> + x\<^sup>\<star> \<cdot> y \<cdot> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (metis C12 distrib_left mult.assoc mult_oner)
also have "... = (1 + x\<^sup>\<star> \<cdot> y \<cdot> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star>) \<cdot> x\<^sup>\<star>"
by (metis distrib_right mult.assoc mult_onel)
finally show ?thesis
by (metis C11 C12 mult_onel mult_oner)
qed
lemma (in conway_dioid) dual_conway_dioid:
"class.conway_dioid (+) (\<odot>) 1 0 (\<le>) (<) star"
proof
fix x y z :: 'a
show "(x \<odot> y) \<odot> z = x \<odot>(y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir times.opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis opp_mult_def distrib_right')
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<odot> y)\<^sup>\<star> \<odot> x\<^sup>\<star>"
by (metis C11_var opp_mult_def)
show "(x \<odot> y)\<^sup>\<star> = 1 + x \<odot> (y \<odot> x)\<^sup>\<star> \<odot> y"
by (metis C12 mult.assoc opp_mult_def)
qed
lemma (in strong_conway_dioid) dual_strong_conway_dioid: "class.strong_conway_dioid ((+) ) ((\<odot>) ) 1 0 (\<le>) (<) star"
proof
fix x y z :: 'a
show "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis opp_mult_def distrib_right')
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<odot> y)\<^sup>\<star> \<odot> x\<^sup>\<star>"
by (metis C11_var opp_mult_def)
show "(x \<odot> y)\<^sup>\<star> = 1 + x \<odot> (y \<odot> x)\<^sup>\<star> \<odot> y"
by (metis C12 mult.assoc opp_mult_def)
show "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
by (metis C13)
qed
text\<open>Nitpick finds counterexamples to the following claims.\<close>
lemma (in conway_dioid) "1\<^sup>\<star> = 1"
nitpick [expect=genuine] \<comment> \<open>3-element counterexample\<close>
oops
lemma (in conway_dioid) "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
nitpick [expect=genuine] \<comment> \<open>3-element counterexample\<close>
oops
context C_algebra
begin
lemma C_unfoldl [simp]: "1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
by (metis C12 mult_onel mult_oner)
lemma C_slide: "(x \<cdot> y)\<^sup>\<star> \<cdot> x = x \<cdot> (y \<cdot> x)\<^sup>\<star>"
proof-
have "(x \<cdot> y)\<^sup>\<star> \<cdot> x = x + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y \<cdot> x"
by (metis C12 mult_onel distrib_right')
also have "... = x \<cdot> (1 + (y \<cdot> x)\<^sup>\<star> \<cdot> y \<cdot> x)"
by (metis distrib_left mult.assoc mult_oner)
finally show ?thesis
by (metis C12 mult.assoc mult_onel mult_oner)
qed
lemma powsum_ub: "i \<le> n \<Longrightarrow> x\<^bsup>i\<^esup> \<le> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
proof (induct n)
case 0 show ?case
- by (metis (hide_lams, mono_tags) "0.prems" eq_iff le_0_eq power_0 powsum_00)
+ by (metis (hide_lams, mono_tags) "0.prems" order.eq_iff le_0_eq power_0 powsum_00)
next
case (Suc n) show ?case
proof -
{ assume aa1: "Suc n \<noteq> i"
have ff1: "x\<^bsub>0\<^esub>\<^bsup>Suc n\<^esup> \<le> x\<^bsub>0\<^esub>\<^bsup>Suc n\<^esup> \<and> Suc n \<noteq> i"
using aa1 by fastforce
have ff2: "\<exists>x\<^sub>1. x\<^bsub>0\<^esub>\<^bsup>n\<^esup> + x\<^sub>1 \<le> x\<^bsub>0\<^esub>\<^bsup>Suc n\<^esup> \<and> Suc n \<noteq> i"
using ff1 powsum2 by auto
have "x\<^bsup>i\<^esup> \<le> x\<^bsub>0\<^esub>\<^bsup>Suc n\<^esup>"
by (metis Suc.hyps Suc.prems ff2 le_Suc_eq local.dual_order.trans local.join.le_supE)
}
thus "x\<^bsup>i\<^esup> \<le> x\<^bsub>0\<^esub>\<^bsup>Suc n\<^esup>"
using local.less_eq_def local.powsum_split_var2 by blast
qed
qed
lemma C14_aux: "m \<le> n \<Longrightarrow> x\<^bsup>m\<^esup> \<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star> = (x\<^bsup>n\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>m\<^esup>"
proof -
assume assm: "m \<le> n"
hence "x\<^bsup>m\<^esup> \<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star> = x\<^bsup>m\<^esup> \<cdot> (x\<^bsup>n-m\<^esup> \<cdot> x\<^bsup>m\<^esup>)\<^sup>\<star>"
by (metis (full_types) le_add_diff_inverse2 power_add)
also have "... = (x\<^bsup>n-m\<^esup> \<cdot> x\<^bsup>m\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>m\<^esup>"
by (metis (hide_lams, mono_tags) C_slide ab_semigroup_add_class.add.commute power_add)
finally show ?thesis
by (metis (full_types) assm le_add_diff_inverse ab_semigroup_add_class.add.commute power_add)
qed
end
context dioid_one_zero
begin
lemma opp_power_def:
"power.power 1 (\<odot>) x n = x\<^bsup>n\<^esup>"
proof (induction n)
case 0 thus ?case
by (metis power.power.power_0)
next
case (Suc n) thus ?case
by (metis power.power.power_Suc power_Suc2 times.opp_mult_def)
qed
lemma opp_powsum_def:
"dioid_one_zero.powsum (+) (\<odot>) 1 0 x m n = x\<^bsub>m\<^esub>\<^bsup>n\<^esup>"
proof -
have "sum (power.power 1 (\<odot>) x) {m..n + m} = sum ((^) x) {m..n + m}"
by (induction n, simp_all add:opp_power_def)
thus ?thesis
by (simp add: dioid_one_zero.powsum_def[of _ _ _ _ "(\<le>)" "(<)"] dual_dioid_one_zero powsum_def)
qed
end
lemma C14_dual:
fixes x::"'a::C_algebra"
shows "x\<^sup>\<star> = x\<^bsub>0\<^esub>\<^bsup>n\<^esup> \<cdot> (x\<^bsup>n+1\<^esup>)\<^sup>\<star>"
proof -
have "x\<^sup>\<star> = (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
by (rule C14)
also have "... = (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> (\<Sum>i=0..n. x^i)"
by (subst powsum_def, auto)
also have "... = (\<Sum>i=0..n. (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x^i)"
by (metis le0 sum_interval_distl)
also have "... = (\<Sum>i=0..n. x^i \<cdot> (x\<^bsup>n+1\<^esup>)\<^sup>\<star>)"
by (auto intro: sum_interval_cong simp only:C14_aux)
also have "... = x\<^bsub>0\<^esub>\<^bsup>n\<^esup> \<cdot> (x\<^bsup>n+1\<^esup>)\<^sup>\<star>"
by (simp only: sum_interval_distr[THEN sym] powsum_def Nat.add_0_right)
finally show ?thesis .
qed
lemma C_algebra: "class.C_algebra (+) (\<odot>) (1::'a::C_algebra) 0 (\<le>) (<) star"
proof
fix x y :: 'a and n :: nat
show "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<odot> y)\<^sup>\<star> \<odot> x\<^sup>\<star>"
by (metis C11_var opp_mult_def)
show "(x \<odot> y)\<^sup>\<star> = 1 + x \<odot> (y \<odot> x)\<^sup>\<star> \<odot> y"
by (metis C12 mult.assoc opp_mult_def)
show "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
by (metis C13)
show "x\<^sup>\<star> = power.power 1 (\<odot>) x (n + 1)\<^sup>\<star> \<odot> dioid_one_zero.powsum (+) (\<odot>) 1 0 x 0 n"
by (metis C14_dual opp_mult_def opp_power_def opp_powsum_def)
qed (simp_all add: opp_mult_def mult.assoc distrib_left)
subsection \<open>Boffa's Axioms\<close>
text \<open>Boffa's two axiomatisations of Regular Algebra from~\cite{Boffa1,Boffa2}.\<close>
class B1_algebra = conway_dioid +
assumes R: "x \<cdot> x = x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
class B2_algebra = star_dioid +
assumes B21: "1 + x \<le> x\<^sup>\<star>"
and B22 [simp]: "x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and B23: "\<lbrakk> 1 + x \<le> y; y \<cdot> y = y \<rbrakk> \<Longrightarrow> x\<^sup>\<star> \<le> y"
lemma (in B1_algebra) B1_algebra:
"class.B1_algebra (+) (\<odot>) 1 0 (\<le>) (<) star"
proof
fix x y z :: 'a
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<odot> y)\<^sup>\<star> \<odot> x\<^sup>\<star>"
by (metis C11_var opp_mult_def)
show "(x \<odot> y)\<^sup>\<star> = 1 + x \<odot> (y \<odot> x)\<^sup>\<star> \<odot> y"
by (metis C12 mult.assoc opp_mult_def)
show "x \<odot> x = x \<Longrightarrow> x\<^sup>\<star> = 1 + x"
by (metis R opp_mult_def)
qed
lemma (in B2_algebra) B2_algebra:
"class.B2_algebra (+) (\<odot>) 1 0 (\<le>) (<) star"
proof
fix x y z :: 'a
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left times.opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "1 + x \<le> x\<^sup>\<star>"
by (metis B21)
show "x\<^sup>\<star> \<odot> x\<^sup>\<star> = x\<^sup>\<star>"
by (metis B22 opp_mult_def)
show "\<lbrakk> 1 + x \<le> y; y \<odot> y = y \<rbrakk> \<Longrightarrow> x\<^sup>\<star> \<le> y"
by (metis B23 opp_mult_def)
qed
instance B1_algebra \<subseteq> B2_algebra
proof
fix x y :: 'a
show "1 + x \<le> x\<^sup>\<star>"
by (metis C12 add_iso_r distrib_right join.sup.cobounded1 mult_onel)
show two: "x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
by (metis (no_types, lifting) C11_var C12 R add_idem' mult_onel mult_oner)
show "\<lbrakk> 1 + x \<le> y; y \<cdot> y = y \<rbrakk> \<Longrightarrow> x\<^sup>\<star> \<le> y"
by (metis (no_types, lifting) C11_var R two distrib_left join.sup.bounded_iff less_eq_def mult.assoc mult.right_neutral)
qed
context B2_algebra
begin
lemma star_ref: "1 \<le> x\<^sup>\<star>"
using local.B21 by auto
lemma star_plus_one [simp]: "1 + x\<^sup>\<star> = x\<^sup>\<star>"
by (metis less_eq_def star_ref)
lemma star_trans: "x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis B22 order_refl)
lemma star_trans_eq [simp]: "x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
by (metis B22)
lemma star_invol [simp]: "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
- by (metis B21 B22 B23 antisym star_plus_one)
+ by (metis B21 B22 B23 order.antisym star_plus_one)
lemma star_1l: "x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis local.B21 local.join.sup.boundedE local.mult_isor local.star_trans_eq)
lemma star_one [simp]: "1\<^sup>\<star> = 1"
- by (metis B23 add_idem antisym mult_oner order_refl star_ref)
+ by (metis B23 add_idem order.antisym mult_oner order_refl star_ref)
lemma star_subdist: "x\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (meson local.B21 local.B23 local.join.sup.bounded_iff local.star_trans_eq)
lemma star_iso: "x \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y\<^sup>\<star>"
by (metis less_eq_def star_subdist)
lemma star2: "(1 + x)\<^sup>\<star> = x\<^sup>\<star>"
by (metis B21 add.commute less_eq_def star_invol star_subdist)
lemma star_unfoldl: "1 + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis local.join.sup.bounded_iff star_1l star_ref)
lemma star_unfoldr: "1 + x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
by (metis (full_types) B21 local.join.sup.bounded_iff mult_isol star_trans_eq)
lemma star_ext: "x \<le> x\<^sup>\<star>"
by (metis B21 local.join.sup.bounded_iff)
lemma star_1r: "x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
by (metis mult_isol star_ext star_trans_eq)
lemma star_unfoldl_eq [simp]: "1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
proof -
have "(1 + x \<cdot> x\<^sup>\<star>) \<cdot> (1 + x \<cdot> x\<^sup>\<star>) = 1 \<cdot> (1 + x \<cdot> x\<^sup>\<star>) + x \<cdot> x\<^sup>\<star> \<cdot> (1 + x \<cdot> x\<^sup>\<star>)"
by (metis distrib_right)
also have "... = 1 + x \<cdot> x\<^sup>\<star> + (x \<cdot> x\<^sup>\<star> \<cdot> x \<cdot> x\<^sup>\<star>)"
by (metis add_assoc' add_idem' distrib_left mult.assoc mult_onel mult_oner)
also have "... = 1 + x \<cdot> x\<^sup>\<star>"
by (metis add.assoc add.commute distrib_left less_eq_def mult.assoc star_1l star_trans_eq)
finally show ?thesis
- by (metis B23 local.join.sup.mono local.join.sup.cobounded1 distrib_left eq_iff mult_1_right star_plus_one star_unfoldl)
+ by (metis B23 local.join.sup.mono local.join.sup.cobounded1 distrib_left order.eq_iff mult_1_right star_plus_one star_unfoldl)
qed
lemma star_unfoldr_eq [simp]: "1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
proof -
have "(1 + x\<^sup>\<star> \<cdot> x) \<cdot> (1 + x\<^sup>\<star> \<cdot> x) = 1 \<cdot> (1 + x\<^sup>\<star> \<cdot> x) + x\<^sup>\<star> \<cdot> x \<cdot> (1 + x\<^sup>\<star> \<cdot> x)"
by (metis distrib_right)
also have "... = 1 + x\<^sup>\<star> \<cdot> x + (x\<^sup>\<star> \<cdot> x \<cdot> x\<^sup>\<star> \<cdot> x)"
by (metis add.assoc add_idem' distrib_left mult_1_left mult_1_right mult.assoc)
also have "... = 1 + x\<^sup>\<star> \<cdot>x"
by (metis add_assoc' distrib_left mult.assoc mult_oner distrib_right' star_trans_eq star_unfoldl_eq)
finally show ?thesis
- by (metis B21 B23 add.commute local.join.sup.mono local.join.sup.cobounded1 eq_iff eq_refl mult_1_left distrib_right' star_unfoldl_eq star_unfoldr)
+ by (metis B21 B23 add.commute local.join.sup.mono local.join.sup.cobounded1 order.eq_iff eq_refl mult_1_left distrib_right' star_unfoldl_eq star_unfoldr)
qed
lemma star_prod_unfold_le: "(x \<cdot> y)\<^sup>\<star> \<le> 1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y"
proof -
have "(1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y) \<cdot> (1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y) =
1 \<cdot> (1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y) + (x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y) \<cdot> (1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y)"
by (metis distrib_right')
also have "... = 1 + x \<cdot>(y \<cdot> x)\<^sup>\<star> \<cdot> y + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y \<cdot> x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y"
by (metis add.assoc local.join.sup.cobounded1 distrib_left less_eq_def mult_1_right mult.assoc mult_onel)
finally have "(1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y) \<cdot> (1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y) = 1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y"
by (metis add.assoc distrib_left distrib_right mult.assoc mult_oner star_trans_eq star_unfoldr_eq)
moreover have "(x \<cdot> y) \<le> 1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y"
by (metis local.join.sup.cobounded2 mult_1_left mult.assoc mult_double_iso order_trans star_ref)
ultimately show ?thesis
by (simp add: local.B23)
qed
lemma star_prod_unfold [simp]: " 1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y = (x \<cdot> y)\<^sup>\<star>"
proof -
have "1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y \<le> 1 + x \<cdot> (1 + y \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> x) \<cdot> y"
by (metis local.join.sup.mono mult_double_iso order_refl star_prod_unfold_le)
also have "... = 1 + x \<cdot> y + x \<cdot> y \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> x \<cdot> y"
by (metis add.assoc distrib_left mult_1_left mult.assoc distrib_right')
finally have "1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y \<le> (x \<cdot> y)\<^sup>\<star>"
by (metis add.assoc distrib_left mult_1_right mult.assoc star_unfoldl_eq star_unfoldr_eq)
thus ?thesis
- by (metis antisym star_prod_unfold_le)
+ by (metis order.antisym star_prod_unfold_le)
qed
lemma star_slide1: "(x \<cdot> y)\<^sup>\<star> \<cdot> x \<le> x \<cdot> (y \<cdot> x)\<^sup>\<star>"
proof -
have "(x \<cdot> y)\<^sup>\<star> \<cdot> x = (1 + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y) \<cdot> x"
by (metis star_prod_unfold)
also have "... = (x + x \<cdot> (y \<cdot> x)\<^sup>\<star> \<cdot> y \<cdot> x)"
by (metis mult_onel distrib_right')
also have "... = x \<cdot> (1 + (y \<cdot> x)\<^sup>\<star> \<cdot> y \<cdot> x)"
by (metis distrib_left mult.assoc mult_oner)
finally show ?thesis
by (metis eq_refl mult.assoc star_unfoldr_eq)
qed
lemma star_slide_var1: "x\<^sup>\<star> \<cdot> x \<le> x \<cdot> x\<^sup>\<star>"
by (metis mult_onel mult_oner star_slide1)
lemma star_slide: "(x \<cdot> y)\<^sup>\<star> \<cdot> x = x \<cdot> (y \<cdot> x)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "(x \<cdot> y)\<^sup>\<star> \<cdot> x \<le> x \<cdot> (y \<cdot> x)\<^sup>\<star>"
by (metis star_slide1)
have "x \<cdot> (y \<cdot> x)\<^sup>\<star> = x \<cdot> (1 + y \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> x)"
by (metis star_prod_unfold)
also have "... = x + x \<cdot> y \<cdot> (x \<cdot> y)\<^sup>\<star> \<cdot> x"
by (metis distrib_left mult.assoc mult_oner)
also have "... = (1 + x \<cdot> y \<cdot> (x \<cdot> y)\<^sup>\<star>) \<cdot> x"
by (metis mult_onel distrib_right')
finally show "x \<cdot> (y \<cdot> x)\<^sup>\<star> \<le> (x \<cdot> y)\<^sup>\<star> \<cdot> x"
by (metis mult_isor star_unfoldl)
qed
lemma star_rtc1: "1 + x + x\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
using local.B21 local.join.sup_least local.star_trans by blast
lemma star_rtc1_eq: "1 + x + x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
by (metis B21 B22 less_eq_def)
lemma star_subdist_var_1: "x \<le> (x + y)\<^sup>\<star>"
using local.join.le_supE local.star_ext by blast
lemma star_subdist_var_2: "x \<cdot> y \<le> (x + y)\<^sup>\<star>"
by (metis (full_types) local.join.le_supE mult_isol_var star_ext star_trans_eq)
lemma star_subdist_var_3: "x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (metis add.commute mult_isol_var star_subdist star_trans_eq)
lemma R_lemma:
assumes "x \<cdot> x = x"
shows "x\<^sup>\<star> = 1 + x"
-proof (rule antisym)
+proof (rule order.antisym)
show "1 + x \<le> x\<^sup>\<star>"
by (metis B21)
have "(1 + x) \<cdot> (1 + x) = 1 + x"
by (metis add.commute add_idem' add.left_commute assms distrib_left mult_onel mult_oner distrib_right')
thus "x\<^sup>\<star> \<le> 1 + x"
by (metis B23 order_refl)
qed
lemma star_denest_var_0: "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have one_below: "1 \<le> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (metis mult_isol_var star_one star_ref star_trans_eq)
have x_below: "x \<le> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (metis mult_isol mult_oner order_trans star_ext star_ref star_slide)
have y_below: "y \<le> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (metis mult_isol_var mult_onel mult_oner order_trans star_ext star_slide star_unfoldl_eq subdistl)
from one_below x_below y_below have "1 + x + y \<le> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by simp
moreover have "(x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (metis star_trans_eq star_slide mult.assoc)
ultimately show "(x + y)\<^sup>\<star> \<le> (x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (metis B23 add_assoc' mult.assoc)
show "(x\<^sup>\<star> \<cdot> y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<le> (x + y)\<^sup>\<star>"
by (metis (full_types) add.commute mult_isol_var star_invol star_iso star_subdist_var_1 star_trans_eq)
qed
lemma star_denest: "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (metis R_lemma add.commute star_denest_var_0 star_plus_one star_prod_unfold star_slide star_trans_eq)
lemma star_sum_var: "(x + y)\<^sup>\<star> = (x\<^sup>\<star> + y\<^sup>\<star>)\<^sup>\<star>"
by (metis star_denest star_invol)
lemma star_denest_var: "(x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (metis star_denest_var_0 star_slide)
lemma star_denest_var_2: "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (metis star_denest star_denest_var)
lemma star_denest_var_3: "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star> \<cdot> (y\<^sup>\<star> \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (metis B22 add_comm mult.assoc star_denest star_denest_var_2)
lemma star_denest_var_4: "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> = (y\<^sup>\<star> \<cdot> x\<^sup>\<star>)\<^sup>\<star>"
by (metis add_comm star_denest)
lemma star_denest_var_5: "x\<^sup>\<star> \<cdot> (y \<cdot> x\<^sup>\<star>)\<^sup>\<star> = y\<^sup>\<star> \<cdot> (x \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (metis add.commute star_denest_var)
lemma star_denest_var_6: "(x + y)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> (x + y)\<^sup>\<star>"
by (metis mult.assoc star_denest star_denest_var_3)
lemma star_denest_var_7: "(x + y)\<^sup>\<star> = (x + y)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis star_denest star_denest_var star_denest_var_3 star_denest_var_5 star_slide)
lemma star_denest_var_8: "(x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star> \<cdot> y\<^sup>\<star> \<cdot> (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star>"
by (metis star_denest star_denest_var_6)
lemma star_denest_var_9: " (x\<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> = (x \<^sup>\<star> \<cdot> y\<^sup>\<star>)\<^sup>\<star> \<cdot> x\<^sup>\<star> \<cdot> y\<^sup>\<star>"
by (metis star_denest star_denest_var_7)
lemma star_slide_var: "x\<^sup>\<star> \<cdot> x = x \<cdot> x\<^sup>\<star>"
by (metis mult_1_left mult_oner star_slide)
lemma star_sum_unfold: "(x + y)\<^sup>\<star> = x\<^sup>\<star> + x\<^sup>\<star> \<cdot> y \<cdot> (x + y)\<^sup>\<star>"
by (metis distrib_left mult_1_right mult.assoc star_denest_var star_unfoldl_eq)
lemma troeger: "x\<^sup>\<star> \<cdot> (y \<cdot> ((x + y)\<^sup>\<star> \<cdot> z) + z) = (x + y)\<^sup>\<star> \<cdot> z"
proof -
have "x\<^sup>\<star> \<cdot> (y \<cdot> ((x + y)\<^sup>\<star> \<cdot> z) + z) = (x\<^sup>\<star> + x\<^sup>\<star> \<cdot> y \<cdot> (x + y)\<^sup>\<star>) \<cdot> z"
by (metis add_comm distrib_left mult.assoc distrib_right')
thus ?thesis
by (metis star_sum_unfold)
qed
lemma meyer_1: "x\<^sup>\<star> = (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "(1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> \<cdot> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> = ((x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star>) \<cdot> ((x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star>)"
by (metis mult.assoc mult_onel distrib_right')
also have "... = ((x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star>) \<cdot> (x \<cdot> x)\<^sup>\<star> + ((x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star>) \<cdot> x \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis distrib_left mult.assoc)
also have "... = (x \<cdot> x) \<^sup>\<star> \<cdot> (x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star> \<cdot> (x \<cdot> x)\<^sup>\<star> + (x \<cdot> x)\<^sup>\<star> \<cdot> x \<cdot> (x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star> \<cdot> x \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis combine_common_factor distrib_right)
also have "... = (x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star> + x \<cdot> x \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis add.assoc add_idem' mult.assoc star_slide star_trans_eq)
also have "... = 1 + x \<cdot> x \<cdot> (x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star> + x \<cdot> x \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis star_unfoldl_eq)
also have "... = (x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis add_comm add_idem' add.left_commute star_unfoldl_eq)
finally have "(1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> \<cdot> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> = (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis mult_1_left distrib_right')
moreover have "1 + x \<le> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis mult_1_right star_unfoldl_eq subdistl)
ultimately show "x\<^sup>\<star> \<le> (1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis B23 mult.assoc)
next
have "(1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> = (x \<cdot> x)\<^sup>\<star> + x \<cdot> (x \<cdot> x)\<^sup>\<star>"
by (metis mult_1_left distrib_right')
thus "(1 + x) \<cdot> (x \<cdot> x)\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis local.add_zeror local.join.sup_least local.mult_isol_var local.mult_oner local.star_ext local.star_invol local.star_iso local.star_subdist_var_2 local.star_trans_eq local.subdistl_eq)
qed
lemma star_zero [simp]: "0\<^sup>\<star> = 1"
by (metis add_zeror star2 star_one)
lemma star_subsum [simp]: "x\<^sup>\<star> + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
by (metis add.assoc add_idem star_slide_var star_unfoldl_eq)
lemma prod_star_closure: "x \<le> z\<^sup>\<star> \<Longrightarrow> y \<le> z\<^sup>\<star> \<Longrightarrow> x \<cdot> y \<le> z\<^sup>\<star>"
by (metis mult_isol_var star_trans_eq)
end
sublocale B2_algebra \<subseteq> B1_algebra
by unfold_locales (metis star_denest_var_0, metis star_prod_unfold, metis R_lemma)
context B2_algebra
begin
lemma power_le_star: "x\<^bsup>n\<^esup> \<le> x\<^sup>\<star>"
by (induct n, simp_all add: star_ref prod_star_closure star_ext)
lemma star_power_slide:
assumes "k \<le> n"
shows "x\<^bsup>k \<^esup>\<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star> = (x\<^bsup>n\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>k\<^esup>"
proof -
from assms have "x\<^bsup>k \<^esup>\<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star> = (x\<^bsup>k \<^esup>\<cdot> x\<^bsup>n-k\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>k\<^esup>"
by (metis (full_types) le_add_diff_inverse2 power_add star_slide)
with assms show ?thesis
by (metis (full_types) le_add_diff_inverse2 ab_semigroup_add_class.add.commute power_add)
qed
lemma powsum_le_star: "x\<^bsub>m\<^esub>\<^bsup>n\<^esup> \<le> x\<^sup>\<star>"
by (induct n, simp_all add: powsum2, metis power_le_star, metis power_Suc power_le_star)
lemma star_sum_power_slide:
assumes "m \<le> n"
shows "x\<^bsub>0\<^esub>\<^bsup>m \<^esup>\<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star> = (x\<^bsup>n\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m\<^esup>"
using assms
proof (induct m)
case 0 thus ?case
by (metis mult_onel mult_oner powsum_00)
next
case (Suc m) note hyp = this
have "x\<^bsub>0\<^esub>\<^bsup>Suc m\<^esup> \<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star> = (x\<^bsub>0\<^esub>\<^bsup>m\<^esup> + x\<^bsup>Suc m\<^esup>) \<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star>"
by (simp add:powsum2)
also have "... = x\<^bsub>0\<^esub>\<^bsup>m \<^esup>\<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star> + x\<^bsup>Suc m \<^esup>\<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star>"
by (metis distrib_right')
also have "... = (x\<^bsup>n\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m\<^esup> + x\<^bsup>Suc m \<^esup>\<cdot> (x\<^bsup>n\<^esup>)\<^sup>\<star>"
by (metis Suc.hyps Suc.prems Suc_leD)
also have "... = (x\<^bsup>n\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m\<^esup> + (x\<^bsup>n\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>Suc m\<^esup>"
by (metis Suc.prems star_power_slide)
also have "... = (x\<^bsup>n\<^esup>)\<^sup>\<star> \<cdot> (x\<^bsub>0\<^esub>\<^bsup>m\<^esup> + x\<^bsup>Suc m\<^esup>)"
by (metis distrib_left)
finally show ?case
by (simp add:powsum2)
qed
lemma aarden_aux:
assumes "y \<le> y \<cdot> x + z"
shows "y \<le> y \<cdot> x\<^bsup>(Suc n) \<^esup>+ z \<cdot> x\<^sup>\<star>"
proof (induct n)
case 0
have "y \<cdot> x + z \<le> y \<cdot> x\<^bsup>(Suc 0)\<^esup>+ z \<cdot> x\<^sup>\<star>"
by (metis (mono_tags) One_nat_def add.commute add_iso mult_1_right power_one_right star_plus_one subdistl)
thus ?case
by (metis assms order_trans)
next
case (Suc n)
have "y \<cdot> x + z \<le> (y \<cdot> x\<^bsup>(Suc n) \<^esup>+ z \<cdot> x\<^sup>\<star>) \<cdot> x + z"
by (metis Suc add_iso mult_isor)
also have "... = y \<cdot> x\<^bsup>(Suc n) \<^esup>\<cdot> x + z \<cdot> (x\<^sup>\<star> \<cdot> x + 1)"
by (subst distrib_right, metis add_assoc' distrib_left mult.assoc mult_oner)
finally show ?case
by (metis add.commute assms mult.assoc order_trans power_Suc2 star_unfoldr_eq)
qed
lemma conway_powerstar1: "(x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n \<^esup>\<cdot> (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup> = (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
proof (cases n)
case 0 thus ?thesis
by simp
next
case (Suc m) thus ?thesis
proof -
assume assm: "n = Suc m"
have "(x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1 \<^esup>\<cdot> (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1\<^esup> = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1 \<^esup>\<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1\<^esup>"
by (subgoal_tac "m + 1 \<le> m + 2", metis mult.assoc star_sum_power_slide, simp)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1 \<^esup>\<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1\<^esup>"
by (metis star_trans_eq)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>(Suc m)+(Suc m)\<^esup>"
by (simp add: mult.assoc powsum_prod)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> (x\<^bsub>0\<^esub>\<^bsup>Suc m \<^esup>+ x\<^bsub>m + 2\<^esub>\<^bsup>m\<^esup>)"
by (metis monoid_add_class.add.left_neutral powsum_split_var3 add_2_eq_Suc')
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>Suc m \<^esup>+ (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>(m + 2)+ 0\<^esub>\<^bsup>m\<^esup>"
by (simp add: local.distrib_left)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>Suc m \<^esup>+ (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>m+2 \<^esup>\<cdot> x\<^bsub>0\<^esub>\<^bsup>m\<^esup>"
by (subst powsum_shift[THEN sym], metis mult.assoc)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> (x\<^bsub>0\<^esub>\<^bsup>m \<^esup>+ x\<^bsup>m+1\<^esup>) + (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>m+2 \<^esup>\<cdot> x\<^bsub>0\<^esub>\<^bsup>m\<^esup>"
by (simp add:powsum2)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m \<^esup>+ (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>m+2 \<^esup>\<cdot> x\<^bsub>0\<^esub>\<^bsup>m \<^esup>+ (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>m+1\<^esup>"
by (metis add.assoc add.commute add.left_commute distrib_left mult.assoc)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m \<^esup>+ (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m \<^esup>+ (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsup>m+1\<^esup>"
by (metis add_idem' distrib_right' star_subsum)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> (x\<^bsub>0\<^esub>\<^bsup>m \<^esup>+ x\<^bsup>m+1\<^esup>)"
by (metis add_idem' distrib_left)
also have "... = (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1\<^esup>"
by (simp add:powsum2)
finally show ?thesis
by (simp add: assm)
qed
qed
lemma conway_powerstar2: "1 + x \<le> (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
proof (cases n)
case 0 show ?thesis
using "0" local.B21 by auto
next
case (Suc m) show ?thesis
proof -
have one: "x \<le> (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1\<^esup>"
by (metis Suc_eq_plus1 powsum_ext mult_isor mult_onel order_trans star_ref)
have two: "1 \<le> (x\<^bsup>m+2\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>m+1\<^esup>"
by (metis Suc_eq_plus1 local.join.le_supE mult_isor mult_onel powsum_split_var1 star_ref)
from one two show ?thesis
by (metis Suc Suc_eq_plus1 add_2_eq_Suc' local.join.sup_least)
qed
qed
theorem powerstar: "x\<^sup>\<star> = (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
-proof (rule antisym)
+proof (rule order.antisym)
show "x\<^sup>\<star> \<le> (x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
by (metis conway_powerstar1 conway_powerstar2 mult.assoc B23)
have "(x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup> \<le> (x\<^sup>\<star>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
by (metis mult_isor power_le_star star_iso)
also have "... = x\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n\<^esup>"
by (metis star_invol)
also have "... \<le> x\<^sup>\<star> \<cdot> x\<^sup>\<star>"
by (simp add: local.prod_star_closure powsum_le_star)
finally show "(x\<^bsup>n+1\<^esup>)\<^sup>\<star> \<cdot> x\<^bsub>0\<^esub>\<^bsup>n \<^esup>\<le> x\<^sup>\<star>"
by (metis star_trans_eq)
qed
end
sublocale B2_algebra \<subseteq> strong_conway_dioid
by unfold_locales (metis star_invol)
sublocale B2_algebra \<subseteq> C_algebra
by unfold_locales (metis powerstar)
text \<open>The following fact could neither be verified nor falsified in Isabelle. It does not hold for other reasons.\<close>
lemma (in C_algebra) "x\<cdot>x = x \<longrightarrow> x\<^sup>\<star> = 1+x"
oops
subsection \<open>Boffa Monoid Identities\<close>
typedef ('a , 'b) boffa_mon = "{f :: 'a::{finite,monoid_mult} \<Rightarrow> 'b::B1_algebra. True}"
by auto
notation
Rep_boffa_mon ("_\<^bsub>_\<^esub>")
lemma "finite (range (Rep_boffa_mon M))"
by (metis finite_code finite_imageI)
abbreviation boffa_pair :: "('a, 'b) boffa_mon \<Rightarrow> 'a::{finite,monoid_mult} \<Rightarrow> 'a \<Rightarrow> 'b::B1_algebra" where
"boffa_pair x i j \<equiv> \<Sum> { x\<^bsub>k\<^esub> | k. i\<cdot>k = j}"
notation
boffa_pair ("_\<^bsub>_,_\<^esub>")
abbreviation conway_assms where
"conway_assms x \<equiv> (\<forall> i j. (x\<^bsub>i \<^esub>\<cdot> x\<^bsub>j\<^esub> \<le> x\<^bsub>i\<cdot>j\<^esub>) \<and> (x\<^bsub>i,i\<^esub>)\<^sup>\<star> = x\<^bsub>i,i\<^esub>)"
lemma pair_one: "x\<^bsub>1,1\<^esub> = x\<^bsub>1\<^esub>"
by (simp)
definition conway_assm1 where "conway_assm1 x = (\<forall> i j. x\<^bsub>i \<^esub>\<cdot> x\<^bsub>j\<^esub> \<le> x\<^bsub>i\<cdot>j\<^esub>)"
definition conway_assm2 where "conway_assm2 x = (\<forall>i. x\<^bsub>i,i\<^esub>\<^sup>\<star> = x\<^bsub>i,i\<^esub>)"
lemma pair_star:
assumes "conway_assm2 x"
shows "x\<^bsub>1\<^esub>\<^sup>\<star> = x\<^bsub>1\<^esub>"
proof -
have "x\<^bsub>1\<^esub>\<^sup>\<star> = x\<^bsub>1,1\<^esub>\<^sup>\<star>"
by simp
also from assms have "... = x\<^bsub>1,1\<^esub>"
by (metis (mono_tags) conway_assm2_def)
finally show ?thesis
by simp
qed
lemma conway_monoid_one:
assumes "conway_assm2 x"
shows "x\<^bsub>1\<^esub> = 1 + x\<^bsub>1\<^esub>"
proof -
from assms have "x\<^bsub>1\<^esub> = x\<^bsub>1\<^esub>\<^sup>\<star>"
by (metis pair_star)
thus ?thesis
by (metis star_plus_one)
qed
lemma conway_monoid_split:
assumes "conway_assm2 x"
shows "\<Sum> {x\<^bsub>i \<^esub>| i . i \<in> UNIV} = 1 + \<Sum> {x\<^bsub>i \<^esub>| i . i \<in> UNIV}"
proof -
have "\<Sum> {x\<^bsub>i \<^esub>| i . i \<in> UNIV} = \<Sum> {x\<^bsub>i \<^esub>| i . i \<in> (insert 1 (UNIV - {1}))}"
by (metis UNIV_I insert_Diff_single insert_absorb)
also have "... = \<Sum> (Rep_boffa_mon x ` (insert 1 (UNIV - {1})))"
by (metis fset_to_im)
also have "... = x\<^bsub>1\<^esub> + \<Sum> (Rep_boffa_mon x ` (UNIV - {1}))"
by (subst sum_fun_insert, auto)
also have "... = x\<^bsub>1\<^esub> + \<Sum> { x\<^bsub>i\<^esub> | i. i\<in>(UNIV - {1})}"
by (metis fset_to_im)
also from assms have unfld:"... = 1 + x\<^bsub>1\<^esub> + \<Sum> { x\<^bsub>i\<^esub> | i. i\<in>(UNIV - {1})}"
by (metis (lifting, no_types) conway_monoid_one)
finally show ?thesis
by (metis (lifting, no_types) ac_simps unfld)
qed
lemma boffa_mon_aux1: "{x\<^bsub>i\<cdot>j \<^esub>| i j. i \<in> UNIV \<and> j \<in> UNIV} = {x\<^bsub>i\<^esub> | i. i \<in> UNIV}"
by (auto, metis monoid_mult_class.mult.left_neutral)
lemma sum_intro' [intro]:
"\<lbrakk>finite (A :: 'a::join_semilattice_zero set); finite B; \<forall>a\<in>A. \<exists>b\<in>B. a \<le> b \<rbrakk> \<Longrightarrow> \<Sum>A \<le> \<Sum>B"
by (metis sum_intro)
lemma boffa_aux2:
"conway_assm1 x \<Longrightarrow>
\<Sum>{x\<^bsub>i\<^esub>\<cdot>x\<^bsub>j \<^esub>| i j. i \<in> UNIV \<and> j \<in> UNIV} \<le> \<Sum>{x\<^bsub>i\<cdot>j\<^esub> | i j. i \<in> UNIV \<and> j \<in> UNIV}"
unfolding conway_assm1_def
using [[simproc add: finite_Collect]]
by force
lemma boffa_aux3:
assumes "conway_assm1 x"
shows "(\<Sum> {x\<^bsub>i\<^esub> | i. i\<in>UNIV}) + (\<Sum> {x\<^bsub>i \<^esub>\<cdot> x\<^bsub>j\<^esub> | i j . i\<in>UNIV \<and> j\<in>UNIV}) = (\<Sum> {x\<^bsub>i\<^esub> | i. i\<in>UNIV})"
proof -
from assms
have "(\<Sum> {x\<^bsub>i\<^esub> | i. i\<in>UNIV}) + (\<Sum> {x\<^bsub>i \<^esub>\<cdot> x\<^bsub>j\<^esub> | i j . i\<in>UNIV \<and> j\<in>UNIV}) \<le> (\<Sum> {x\<^bsub>i\<^esub> | i. i\<in>UNIV})+(\<Sum> {x\<^bsub>i\<cdot>j\<^esub> | i j . i\<in>UNIV \<and> j\<in>UNIV})"
apply (subst add_iso_r)
apply (subst boffa_aux2)
by simp_all
also have "... = (\<Sum> {x\<^bsub>i\<^esub> | i. i\<in>UNIV})"
by (metis (mono_tags) add_idem boffa_mon_aux1)
ultimately show ?thesis
by (simp add: dual_order.antisym)
qed
lemma conway_monoid_identity:
assumes "conway_assm1 x" "conway_assm2 x"
shows "(\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV})\<^sup>\<star> = (\<Sum> {x\<^bsub>i\<^esub>| i. i\<in>UNIV})"
proof -
have one:"(\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV}) \<cdot> (\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV}) = (1 + (\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV})) \<cdot> (1 + (\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV}))"
by (metis (mono_tags) assms(2) conway_monoid_split)
also have "... = 1 + (\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV}) + ((\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV}) \<cdot> (\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV}))"
by (metis (lifting, no_types) calculation less_eq_def mult_isol mult_isol_equiv_subdistl mult_oner)
also have "... = 1 + (\<Sum> {x\<^bsub>i\<^esub>|i. i\<in>UNIV}) + (\<Sum> {x\<^bsub>i \<^esub>\<cdot> x\<^bsub>j\<^esub> | i j. i\<in>UNIV \<and> j\<in>UNIV})"
by (simp only: dioid_sum_prod finite_UNIV)
finally have "\<Sum> {x\<^bsub>i\<^esub> |i. i \<in> UNIV} \<cdot> \<Sum> {x\<^bsub>i\<^esub> |i. i \<in> UNIV} = \<Sum> {x\<^bsub>i\<^esub> |i. i \<in> UNIV}"
apply (simp only:)
proof -
assume a1: "\<Sum>{x\<^bsub>i\<^esub> |i. i \<in> UNIV} \<cdot> \<Sum>{x\<^bsub>i\<^esub> |i. i \<in> UNIV} = 1 + \<Sum>{x\<^bsub>i\<^esub> |i. i \<in> UNIV} + \<Sum>{x\<^bsub>i\<^esub> \<cdot> x\<^bsub>j\<^esub> |i j. i \<in> UNIV \<and> j \<in> UNIV}"
hence "\<Sum>{x\<^bsub>R\<^esub> |R. R \<in> UNIV} \<cdot> \<Sum>{x\<^bsub>R\<^esub> |R. R \<in> UNIV} = \<Sum>{x\<^bsub>R\<^esub> |R. R \<in> UNIV}"
using assms(1) assms(2) boffa_aux3 conway_monoid_split by fastforce
thus "1 + \<Sum>{x\<^bsub>i\<^esub> |i. i \<in> UNIV} + \<Sum>{x\<^bsub>i\<^esub> \<cdot> x\<^bsub>j\<^esub> |i j. i \<in> UNIV \<and> j \<in> UNIV} = \<Sum>{x\<^bsub>i\<^esub> |i. i \<in> UNIV}"
using a1 by simp
qed
thus ?thesis
by (metis (mono_tags) one B1_algebra_class.R star_trans_eq)
qed
subsection \<open>Conway's Conjectures\<close>
class C0_algebra = strong_conway_dioid +
assumes C0: "x \<cdot> y = y \<cdot> z \<Longrightarrow> x\<^sup>\<star> \<cdot> y = y \<cdot> z\<^sup>\<star>"
class C1l_algebra = strong_conway_dioid +
assumes C1l: "x \<cdot> y \<le> y \<cdot> z \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y \<cdot> z\<^sup>\<star>"
class C1r_algebra = strong_conway_dioid +
assumes C1r: "y \<cdot> x \<le> z \<cdot> y \<Longrightarrow> y \<cdot> x\<^sup>\<star> \<le> z\<^sup>\<star> \<cdot> y"
class C2l_algebra = conway_dioid +
assumes C2l: "x = y \<cdot> x \<Longrightarrow> x = y\<^sup>\<star> \<cdot> x"
class C2r_algebra = conway_dioid +
assumes C2r: "x = x \<cdot> y \<Longrightarrow> x = x \<cdot> y\<^sup>\<star>"
class C3l_algebra = conway_dioid +
assumes C3l: "x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y"
class C3r_algebra = conway_dioid +
assumes C3r: "y \<cdot> x \<le> y \<Longrightarrow> y \<cdot> x\<^sup>\<star> \<le> y"
sublocale C1r_algebra \<subseteq> dual: C1l_algebra
"(+)" "(\<odot>)" "1" "0" "(\<le>)" "(<)" "star"
proof
fix x y z :: 'a
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<odot> y)\<^sup>\<star> \<odot> x\<^sup>\<star>"
by (metis C11_var opp_mult_def)
show "(x \<odot> y)\<^sup>\<star> = 1 + x \<odot> (y \<odot> x)\<^sup>\<star> \<odot> y"
by (metis C12 mult.assoc opp_mult_def)
show "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
by (metis C13)
show "x \<odot> y \<le> y \<odot> z \<Longrightarrow> x\<^sup>\<star> \<odot> y \<le> y \<odot> z\<^sup>\<star>"
by (metis C1r opp_mult_def)
qed
sublocale C2r_algebra \<subseteq> dual: C2l_algebra
"(+)" "(\<odot>)" "1" "0" "(\<le>)" "(<)" "star"
proof
fix x y z :: 'a
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<odot> y)\<^sup>\<star> \<odot> x\<^sup>\<star>"
by (metis C11_var opp_mult_def)
show "(x \<odot> y)\<^sup>\<star> = 1 + x \<odot> (y \<odot> x)\<^sup>\<star> \<odot> y"
by (metis C12 mult.assoc opp_mult_def)
show "x = y \<odot> x \<Longrightarrow> x = y\<^sup>\<star> \<odot> x"
by (metis C2r opp_mult_def)
qed
sublocale C3r_algebra \<subseteq> dual: C3l_algebra
"(+)" "(\<odot>)" "1" "0" "(\<le>)" "(<)" "star"
proof
fix x y z :: 'a
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "(x + y)\<^sup>\<star> = (x\<^sup>\<star> \<odot> y)\<^sup>\<star> \<odot> x\<^sup>\<star>"
by (metis C11_var opp_mult_def)
show "(x \<odot> y)\<^sup>\<star> = 1 + x \<odot> (y \<odot> x)\<^sup>\<star> \<odot> y"
by (metis C12 mult.assoc opp_mult_def)
show " x \<odot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<odot> y \<le> y"
by (metis C3r opp_mult_def)
qed
lemma (in C3l_algebra) k2_var: "z + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> y"
by (metis local.C3l local.join.le_supE local.join.sup.absorb2 local.subdistl)
instance C2l_algebra \<subseteq> B1_algebra
by (intro_classes, metis C2l monoid_mult_class.mult.left_neutral mult_oner conway_dioid_class.C12)
instance C2r_algebra \<subseteq> B1_algebra
by (intro_classes, metis C2r conway_dioid_class.C12)
text \<open>The following claims are refuted by Nitpick\<close>
lemma (in conway_dioid)
assumes "x \<cdot> y = y \<cdot> z \<Longrightarrow> x\<^sup>\<star> \<cdot> y = y \<cdot> z\<^sup>\<star>"
shows "1\<^sup>\<star> = 1"
(* nitpick [expect=genuine] -- "3-element counterexample"*)
oops
lemma (in conway_dioid)
assumes "x \<cdot> y \<le> y \<cdot> z \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y \<cdot> z\<^sup>\<star>"
shows "1\<^sup>\<star> = 1"
(* nitpick [expect=genuine] -- "3-element counterexample"*)
oops
text \<open>The following fact could not be refuted by Nitpick or Quickcheck; but an infinite counterexample exists.\<close>
lemma (in B1_algebra) "x = x\<cdot>y\<longrightarrow> x = x\<cdot>y\<^sup>\<star>"
oops
instance C3l_algebra \<subseteq> C2l_algebra
by (intro_classes, metis C3l conway_dioid_class.C12 dual_order.antisym join.sup.cobounded1 mult_isol_var mult_onel order_refl)
sublocale C2l_algebra \<subseteq> C3l_algebra
proof
fix x y
show "x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y"
proof -
assume "x \<cdot> y \<le> y"
hence "(x + 1) \<cdot> y = y"
by (metis less_eq_def mult_onel distrib_right')
hence "(x + 1)\<^sup>\<star> \<cdot> y = y"
by (metis C2l)
hence "x\<^sup>\<star> \<cdot> y = y"
by (metis C11 C2l add_comm mult_1_left mult_1_right)
thus "x\<^sup>\<star> \<cdot> y \<le> y"
by (metis eq_refl)
qed
qed
sublocale C1l_algebra \<subseteq> C3l_algebra
by unfold_locales (metis mult_oner C1l C12 C13 add_zeror annir)
sublocale C3l_algebra \<subseteq> C1l_algebra
proof
fix x y z
show "(x\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star>"
- by (metis local.C11_var local.C12 local.C3l local.eq_iff local.eq_refl local.join.sup.absorb2 local.join.sup_ge1 local.mult_onel local.mult_oner)
+ by (metis local.C11_var local.C12 local.C3l order.eq_iff local.eq_refl local.join.sup.absorb2 local.join.sup_ge1 local.mult_onel local.mult_oner)
show "x \<cdot> y \<le> y \<cdot> z \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y \<cdot> z\<^sup>\<star>"
proof -
assume assm: "x \<cdot> y \<le> y \<cdot> z"
have r1:"y \<le> y \<cdot> z\<^sup>\<star>"
by (metis C12 mult_isol mult_oner order_prop)
from assm have "x \<cdot> y \<cdot> z\<^sup>\<star> \<le> y \<cdot> z \<cdot>z\<^sup>\<star>"
by (metis mult_isor)
also have "... \<le> y \<cdot> z\<^sup>\<star>"
by (metis local.C12 local.join.sup_commute local.mult_onel local.mult_oner local.subdistl mult_assoc)
finally have "y + x \<cdot> y \<cdot> z\<^sup>\<star> \<le> y \<cdot> z\<^sup>\<star>"
by (simp add: r1)
thus "x\<^sup>\<star> \<cdot> y \<le> y \<cdot> z\<^sup>\<star>"
by (metis k2_var mult.assoc)
qed
qed
sublocale C1l_algebra \<subseteq> C2l_algebra
by (unfold_locales, metis C12 C3l add.commute local.join.sup.cobounded1 distrib_right less_eq_def mult_1_left order_refl)
sublocale C3r_algebra \<subseteq> C2r_algebra
by (unfold_locales, metis C12 C3r add.commute local.join.sup.cobounded1 distrib_left less_eq_def mult_1_right order_refl)
sublocale C2r_algebra \<subseteq> C3r_algebra
by unfold_locales (metis dual.C3l opp_mult_def)
sublocale C1r_algebra \<subseteq> C3r_algebra
by unfold_locales (metis dual.C3l opp_mult_def)
sublocale C3r_algebra \<subseteq> C1r_algebra
by (unfold_locales, metis dual.C13, metis dual.C1l opp_mult_def)
class C1_algebra = C1l_algebra + C1r_algebra
class C2_algebra = C2l_algebra + C2r_algebra
class C3_algebra = C3l_algebra + C3r_algebra
sublocale C0_algebra \<subseteq> C2_algebra
by unfold_locales (metis C12 C13 add_zeror annil mult_oner mult_onel C0)+
sublocale C2_algebra \<subseteq> C0_algebra
- by unfold_locales (metis C1l C1r eq_iff)
+ by unfold_locales (metis C1l C1r order.eq_iff)
sublocale C2_algebra \<subseteq> C1_algebra ..
sublocale C1_algebra \<subseteq> C2_algebra ..
sublocale C2_algebra \<subseteq> C3_algebra ..
sublocale C3_algebra \<subseteq> C2_algebra ..
subsection \<open>Kozen's Kleene Algebras\<close>
text \<open>Kozen's Kleene Algebras~\cite{Kozen,Kozensemi}.\<close>
class Kl_base = star_dioid +
assumes Kl: "1 + x \<cdot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
class Kr_base = star_dioid +
assumes Kr: "1 + x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
class K1l_algebra = Kl_base +
assumes star_inductl: "x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> y \<le> y"
class K1r_algebra = Kr_base +
assumes star_inductr: "y \<cdot> x \<le> y \<Longrightarrow> y \<cdot> x\<^sup>\<star> \<le> y"
class K2l_algebra = Kl_base +
assumes star_inductl_var: "z + x \<cdot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<cdot> z \<le> y"
class K2r_algebra = Kr_base +
assumes star_inductr_var: "z + y \<cdot> x \<le> y \<Longrightarrow> z \<cdot> x\<^sup>\<star> \<le> y"
class K1_algebra = K1l_algebra + K1r_algebra
class K2_algebra = K2l_algebra + K2r_algebra
sublocale K1r_algebra \<subseteq> dual: K1l_algebra
"(+)" "(\<odot>)" "1" "0" "(\<le>)" "(<)" "star"
proof
fix x y z :: 'a
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "1 + x \<odot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis Kr opp_mult_def)
show "x \<odot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<odot> y \<le> y"
by (metis star_inductr opp_mult_def)
qed
sublocale K1l_algebra \<subseteq> B2_algebra
proof
fix x y :: 'a
show "1 + x \<le> x\<^sup>\<star>"
by (metis add_iso_r local.join.sup.cobounded1 mult_isol mult_oner order_trans Kl)
show "x\<^sup>\<star> \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
- using local.Kl local.eq_iff local.phl_cons1 local.star_inductl by fastforce
+ using local.Kl order.eq_iff local.phl_cons1 local.star_inductl by fastforce
show "\<lbrakk> 1 + x \<le> y; y \<cdot> y = y \<rbrakk> \<Longrightarrow> x\<^sup>\<star> \<le> y"
by (metis local.distrib_right' local.join.le_sup_iff local.join.sup.order_iff local.mult_isol local.mult_oner local.star_inductl)
qed
sublocale K1r_algebra \<subseteq> B2_algebra
by unfold_locales (metis dual.B21 dual.B22 dual.B23 opp_mult_def)+
sublocale K1l_algebra \<subseteq> C2l_algebra
by (unfold_locales, metis less_eq_def mult_1_left mult.assoc star_inductl star_invol star_one star_plus_one star_trans_eq troeger)
sublocale C2l_algebra \<subseteq> K1l_algebra
- by (unfold_locales, metis C12 eq_iff mult_1_left mult_1_right, metis C3l)
+ by (unfold_locales, metis C12 order.eq_iff mult_1_left mult_1_right, metis C3l)
sublocale K1r_algebra \<subseteq> C2r_algebra
by unfold_locales (metis dual.C2l opp_mult_def)
sublocale C2r_algebra \<subseteq> K1r_algebra
by (unfold_locales, metis dual.star_unfoldl opp_mult_def, metis C3r)
sublocale K1_algebra \<subseteq> C0_algebra
- by unfold_locales (metis C1l C1r eq_iff)
+ by unfold_locales (metis C1l C1r order.eq_iff)
sublocale C0_algebra \<subseteq> K1l_algebra ..
sublocale K2r_algebra \<subseteq> dual: K2l_algebra
"(+)" "(\<odot>)" "1" "0" "(\<le>)" "(<)" "star"
proof
fix x y z :: 'a
show "x \<odot> y \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis distrib_right opp_mult_def)
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "1 + x \<odot> x\<^sup>\<star> \<le> x\<^sup>\<star>"
by (metis opp_mult_def Kr)
show "z + x \<odot> y \<le> y \<Longrightarrow> x\<^sup>\<star> \<odot> z \<le> y"
by (metis opp_mult_def star_inductr_var)
qed
sublocale K1l_algebra \<subseteq> K2l_algebra
by unfold_locales (metis local.join.le_supE star_inductl order_prop subdistl)
sublocale K2l_algebra \<subseteq> K1l_algebra
by (unfold_locales, simp add: local.star_inductl_var)
sublocale K1r_algebra \<subseteq> K2r_algebra
by unfold_locales (metis dual.star_inductl_var opp_mult_def)
sublocale K2r_algebra \<subseteq> K1r_algebra
by unfold_locales (metis dual.star_inductl opp_mult_def)
sublocale kleene_algebra \<subseteq> K1_algebra
by (unfold_locales, metis star_unfoldl, metis star_inductl_var, metis star_unfoldr, metis star_inductr_var)
sublocale K1_algebra \<subseteq> K2_algebra ..
sublocale K2_algebra \<subseteq> koz: kleene_algebra
by (unfold_locales, metis Kl, metis star_inductl_var, metis star_inductr_var)
subsection \<open>Salomaa's Axioms\<close>
text \<open>Salomaa's axiomatisations of Regular Algebra~\cite{Salomaa}.\<close>
class salomaa_base = star_dioid +
fixes ewp :: "'a \<Rightarrow> bool"
assumes S11: "(1 + x)\<^sup>\<star> = x\<^sup>\<star>"
and EWP : "ewp x \<longleftrightarrow> (\<exists>y. x = 1 + y \<and> \<not> ewp y)"
class Sr_algebra = salomaa_base +
assumes S12r: "1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
and Ar : "\<lbrakk> \<not> ewp y; x = x \<cdot> y + z \<rbrakk> \<Longrightarrow> x = z \<cdot> y\<^sup>\<star>"
text \<open>The following claim is ruled out by Nitpick. The unfold law cannot be weakened as in Kleene algebra.\<close>
lemma (in salomaa_base)
assumes S12r': "1 + x\<^sup>\<star> \<cdot> x \<le> x\<^sup>\<star>"
and Ar' : "\<lbrakk> \<not> ewp y; x = x \<cdot> y + z \<rbrakk> \<Longrightarrow> x = z \<cdot> y\<^sup>\<star>"
shows "x\<^sup>\<star> \<le> 1 + x\<^sup>\<star> \<cdot> x"
(*nitpick [expect=genuine] -- "4-element counterexample"*)
oops
class Sl_algebra = salomaa_base +
assumes S12l: "1 + x \<cdot> x\<^sup>\<star> = x\<^sup>\<star>"
and Al : "\<lbrakk> \<not> ewp y; x = y\<cdot>x+z \<rbrakk> \<Longrightarrow> x = y\<^sup>\<star>\<cdot>z"
class S_algebra = Sl_algebra + Sr_algebra
sublocale Sl_algebra \<subseteq> dual: Sr_algebra
"(+)" "(\<odot>)" "1" "0" "(\<le>)" "(<)" "star" "ewp"
proof
fix x y z :: 'a
show "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)"
by (metis mult.assoc opp_mult_def)
show "(x + y) \<odot> z = x \<odot> z + y \<odot> z"
by (metis distrib_left opp_mult_def)
show "x + x = x"
by (metis add_idem)
show "1 \<odot> x = x"
by (metis mult_oner opp_mult_def)
show "x \<odot> 1 = x"
by (metis mult_onel opp_mult_def)
show "0 + x = x"
by (metis add_zerol)
show "0 \<odot> x = 0"
by (metis annir times.opp_mult_def)
show "x \<odot> (y + z) = x \<odot> y + x \<odot> z"
by (metis opp_mult_def distrib_right')
show "x \<odot> 0 = 0"
by (metis annil opp_mult_def)
show "(1 + x)\<^sup>\<star> = x\<^sup>\<star>"
by (metis S11)
show "ewp x = (\<exists>y. x = 1 + y \<and> \<not> ewp y)"
by (metis EWP)
show "1 + x\<^sup>\<star> \<odot> x = x\<^sup>\<star>"
by (metis S12l opp_mult_def)
show "\<lbrakk> \<not> ewp y; x = x \<odot> y + z \<rbrakk> \<Longrightarrow> x = z \<odot> y\<^sup>\<star>"
by (metis opp_mult_def Al)
qed
context Sr_algebra
begin
lemma kozen_induct_r:
assumes "y \<cdot> x + z \<le> y"
shows "z \<cdot> x\<^sup>\<star> \<le> y"
proof (cases "ewp x")
case False thus ?thesis
by (metis add_commute assms local.Ar local.join.le_supE local.join.sup.orderE local.mult_isor)
next
case True thus ?thesis
proof -
assume "ewp x"
then obtain x' where assm1: "x = 1 + x'" and assm2: "\<not> ewp x'"
by (metis EWP)
have "y = (z + y) \<cdot> x\<^sup>\<star>"
- by (metis S11 local.join.le_supE assm1 assm2 assms eq_iff less_eq_def Ar subdistl)
+ by (metis S11 local.join.le_supE assm1 assm2 assms order.eq_iff less_eq_def Ar subdistl)
thus "z \<cdot> x\<^sup>\<star> \<le> y"
by (metis local.join.sup.cobounded1 local.mult_isor)
qed
qed
end
context Sl_algebra
begin
lemma kozen_induct_l:
assumes "x \<cdot> y + z \<le> y"
shows "x\<^sup>\<star>\<cdot>z \<le> y"
by (metis dual.kozen_induct_r times.opp_mult_def assms)
end
sublocale Sr_algebra \<subseteq> K2r_algebra
by unfold_locales (metis S12r order_refl, metis add_comm kozen_induct_r)
sublocale Sr_algebra \<subseteq> K1r_algebra ..
sublocale Sl_algebra \<subseteq> K2l_algebra
by unfold_locales (metis S12l order_refl, metis add_comm kozen_induct_l)
sublocale Sl_algebra \<subseteq> K1l_algebra ..
sublocale S_algebra \<subseteq> K1_algebra ..
sublocale S_algebra \<subseteq> K2_algebra ..
text \<open>The following claim could be refuted.\<close>
lemma (in K2_algebra) "(\<not> 1 \<le> x) \<longrightarrow> x = x \<cdot> y + z \<longrightarrow> x = z \<cdot> y\<^sup>\<star>"
oops
class salomaa_conj_r = salomaa_base +
assumes salomaa_small_unfold: "1 + x\<^sup>\<star> \<cdot> x = x\<^sup>\<star>"
assumes salomaa_small_r: "\<lbrakk> \<not> ewp y ; x = x \<cdot> y + 1 \<rbrakk> \<Longrightarrow> x = y\<^sup>\<star>"
sublocale Sr_algebra \<subseteq> salomaa_conj_r
by (unfold_locales, metis S12r, metis mult_onel Ar)
lemma (in salomaa_conj_r) "(\<not> ewp y) \<and> (x = x \<cdot> y + z) \<longrightarrow> x = z \<cdot> y\<^sup>\<star>"
(*nitpick [expect=genuine] -- "3-element counterexample"*)
oops
end
diff --git a/thys/Relation_Algebra/More_Boolean_Algebra.thy b/thys/Relation_Algebra/More_Boolean_Algebra.thy
--- a/thys/Relation_Algebra/More_Boolean_Algebra.thy
+++ b/thys/Relation_Algebra/More_Boolean_Algebra.thy
@@ -1,293 +1,293 @@
(* Title: Relation Algebra
Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>(More) Boolean Algebra\<close>
theory More_Boolean_Algebra
imports Main
begin
subsection \<open>Laws of Boolean Algebra\<close>
text \<open>The following laws of Boolean algebra support relational proofs. We
might add laws for the binary minus since that would make certain theorems look
more nicely. These are currently not so well supported.\<close>
context boolean_algebra
begin
no_notation
times (infixl "\<cdot>" 70)
and plus (infixl "+" 65)
and Groups.zero_class.zero ("0")
and Groups.one_class.one ("1")
notation
inf (infixl "\<cdot>" 70)
and sup (infixl "+" 65)
and bot ("0")
and top ("1")
lemma meet_assoc: "x \<cdot> (y \<cdot> z) = (x \<cdot> y) \<cdot> z"
by (metis inf_assoc)
lemma aux4 [simp]: "x \<cdot> y + x \<cdot> -y = x"
by (metis inf_sup_distrib1 inf_top_right sup_compl_top)
lemma aux4_comm [simp]: "x \<cdot> -y + x \<cdot> y = x"
by (metis aux4 sup.commute)
lemma aux6 [simp]: "(x + y) \<cdot> -x = y \<cdot> -x"
by (metis inf_compl_bot inf_sup_distrib2 sup_bot_left)
lemma aux6_var [simp]: "(-x + y) \<cdot> x = x \<cdot> y"
by (metis compl_inf_bot inf_commute inf_sup_distrib2 sup_bot_left)
lemma aux9 [simp]: "x + -x \<cdot> y = x + y"
by (metis aux4 aux6 inf.commute inf_sup_absorb)
lemma join_iso: "x \<le> y \<Longrightarrow> x + z \<le> y + z"
by (metis eq_refl sup_mono)
lemma join_isol: "x \<le> y \<Longrightarrow> z + x \<le> z + y"
by (metis join_iso sup.commute)
lemma join_double_iso: "x \<le> y \<Longrightarrow> w + x + z \<le> w + y + z"
by (metis le_iff_inf sup_inf_distrib1 sup_inf_distrib2)
lemma comp_anti: "x \<le> y \<longleftrightarrow> -y \<le> -x"
by (metis compl_le_swap2 double_compl)
lemma meet_iso: "x \<le> y \<Longrightarrow> x \<cdot> z \<le> y \<cdot> z"
by (metis eq_refl inf_mono)
lemma meet_isor: "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
by (metis inf.commute meet_iso)
lemma meet_double_iso: "x \<le> y \<Longrightarrow> w \<cdot> x \<cdot> z \<le> w \<cdot> y \<cdot> z"
by (metis meet_iso meet_isor)
lemma de_morgan_3 [simp]: "-(-x \<cdot> -y) = x + y"
by (metis compl_sup double_compl)
lemma subdist_2_var: "x + y \<cdot> z \<le> x + y"
by (metis eq_refl inf_le1 sup_mono)
lemma dist_alt: "\<lbrakk>x + z = y + z; x \<cdot> z = y \<cdot> z\<rbrakk> \<Longrightarrow> x = y"
by (metis aux4 aux6 sup.commute)
text \<open>Finally we prove the Galois connections for complementation.\<close>
lemma galois_aux: "x \<cdot> y = 0 \<longleftrightarrow> x \<le> -y"
by (metis aux6 compl_sup double_compl inf.commute le_iff_inf sup_bot_right sup_compl_top)
lemma galois_aux2: "x \<cdot> -y = 0 \<longleftrightarrow> x \<le> y"
by (metis double_compl galois_aux)
lemma galois_1: "x \<cdot> -y \<le> z \<longleftrightarrow> x \<le> y + z"
apply (rule iffI)
apply (metis inf_le2 join_iso le_iff_sup le_supE join_isol aux4)
apply (metis meet_iso aux6 le_infE)
done
lemma galois_2: "x \<le> y + -z \<longleftrightarrow> x \<cdot> z \<le> y"
apply (rule iffI)
apply (metis compl_sup double_compl galois_1 inf.commute)
apply (metis inf.commute order_trans subdist_2_var aux4 join_iso)
done
lemma galois_aux3: "x + y = 1 \<longleftrightarrow> -x \<le> y"
by (metis galois_1 inf_top_left top_unique)
lemma galois_aux4: "-x + y = 1 \<longleftrightarrow> x \<le> y"
by (metis double_compl galois_aux3)
subsection \<open>Boolean Algebras with Operators\<close>
text \<open>We follow J\'onsson and Tarski to define pairs of conjugate functions
on Boolean algebras. We also consider material from Maddux's article. This
gives rise to a Galois connection and the notion of Boolean algebras with
operators.
We do not explicitly define families of functions over Boolean algebras as a
type class.
This development should certainly be expanded do deal with complete Boolean
algebras one the one hand and other lattices on the other hand.
Boolean algebras with operators and their variants can be applied in various
ways. The prime example are relation algebras. The modular laws, for instance,
can be derived by instantiation. Other applications are antidomain semirings
where modal operators satisfy conjugations and Galois connections, and algebras
of predicate transformers.\<close>
text\<open>We define conjugation as a predicate which holds if a pair of functions
are conjugates.\<close>
definition is_conjugation :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
where "is_conjugation f g \<equiv> (\<forall>x y . f x \<cdot> y = 0 \<longleftrightarrow> x \<cdot> g y = 0)"
text \<open>We now prove the standard lemmas. First we show that conjugation is
symmetric and that conjugates are uniqely defined.\<close>
lemma is_conjugation_sym: "is_conjugation f g \<longleftrightarrow> is_conjugation g f"
by (metis inf.commute is_conjugation_def)
lemma is_conjugation_unique: "\<lbrakk>is_conjugation f g; is_conjugation f h\<rbrakk> \<Longrightarrow> g = h"
-by (metis galois_aux inf.commute double_compl eq_iff ext is_conjugation_def)
+by (metis galois_aux inf.commute double_compl order.eq_iff ext is_conjugation_def)
text \<open>Next we show that conjugates give rise to adjoints in a Galois
connection.\<close>
lemma conj_galois_1:
assumes "is_conjugation f g"
shows "f x \<le> y \<longleftrightarrow> x \<le> -g (-y)"
by (metis assms is_conjugation_def double_compl galois_aux)
lemma conj_galois_2:
assumes "is_conjugation f g"
shows "g x \<le> y \<longleftrightarrow> x \<le> -f (-y)"
by (metis assms is_conjugation_sym conj_galois_1)
text \<open>Now we prove some of the standard properties of adjoints and
conjugates. In fact, conjugate functions even distribute over all existing
suprema. We display the next proof in detail because it is elegant.\<close>
lemma f_pre_additive:
assumes "is_conjugation f g"
shows "f (x + y) \<le> z \<longleftrightarrow> f x + f y \<le> z"
proof -
have "f (x + y) \<le> z \<longleftrightarrow> x + y \<le> -g (-z)"
by (metis assms conj_galois_1)
also have "... \<longleftrightarrow> x \<le> -g (-z) \<and> y \<le> -g (-z)"
by (metis le_sup_iff)
also have "... \<longleftrightarrow> f x \<le> z \<and> f y \<le> z"
by (metis assms conj_galois_1)
thus ?thesis
by (metis le_sup_iff calculation)
qed
lemma f_additive:
assumes "is_conjugation f g"
shows "f (sup x y) = sup (f x) (f y)"
-by (metis assms eq_iff f_pre_additive)
+by (metis assms order.eq_iff f_pre_additive)
lemma g_pre_additive:
assumes "is_conjugation f g"
shows "g (sup x y) \<le> z \<longleftrightarrow> sup (g x) (g y) \<le> z"
by (metis assms is_conjugation_sym f_pre_additive)
lemma g_additive:
assumes "is_conjugation f g"
shows "g (sup x y) = sup (g x) (g y)"
by (metis assms is_conjugation_sym f_additive)
text \<open>Additivity of adjoints obviously implies their isotonicity.\<close>
lemma f_iso:
assumes "is_conjugation f g"
shows "x \<le> y \<longrightarrow> f x \<le> f y"
by (metis assms f_additive le_iff_sup)
lemma g_iso:
assumes "is_conjugation f g"
shows "x \<le> y \<longrightarrow> g x \<le> g y"
by (metis assms is_conjugation_sym f_iso)
lemma f_subdist:
assumes "is_conjugation f g"
shows "f (x \<cdot> y) \<le> f x"
by (metis assms f_iso inf_le1)
lemma g_subdist:
assumes "is_conjugation f g"
shows "g (x \<cdot> y) \<le> g x"
by (metis assms g_iso inf_le1)
text \<open>Next we prove cancellation and strictness laws.\<close>
lemma cancellation_1:
assumes "is_conjugation f g"
shows "f (-g x) \<le> -x"
by (metis assms conj_galois_1 double_compl eq_refl)
lemma cancellation_2:
assumes "is_conjugation f g"
shows "g (-f x) \<le> -x"
by (metis assms is_conjugation_sym cancellation_1)
lemma f_strict:
assumes "is_conjugation f g"
shows "f 0 = 0"
by (metis assms inf.idem inf_bot_left is_conjugation_def)
lemma g_strict:
assumes "is_conjugation f g"
shows "g 0 = 0"
by (metis assms is_conjugation_sym f_strict)
text \<open>The following variants of modular laws have more concrete counterparts
in relation algebra.\<close>
lemma modular_1_aux:
assumes "is_conjugation f g"
shows "f (x \<cdot> -g y) \<cdot> y = 0"
by (metis assms galois_aux inf_le2 is_conjugation_def)
lemma modular_2_aux:
assumes "is_conjugation f g"
shows "g (x \<cdot> -f y) \<cdot> y = 0"
by (metis assms is_conjugation_sym modular_1_aux)
lemma modular_1:
assumes "is_conjugation f g"
shows "f x \<cdot> y = f (x \<cdot> g y) \<cdot> y"
proof -
have "f x \<cdot> y = f (x \<cdot> g y + x \<cdot> -g y) \<cdot> y"
by (metis aux4)
hence "f x \<cdot> y = (f (x \<cdot> g y) + f (x \<cdot> -g y)) \<cdot> y"
by (metis assms f_additive)
hence "f x \<cdot> y = f (x \<cdot> g y) \<cdot> y + f (x \<cdot> -g y) \<cdot> y"
by (metis inf.commute inf_sup_distrib1)
thus ?thesis
by (metis assms modular_1_aux sup_bot_right)
qed
lemma modular_2:
assumes "is_conjugation f g"
shows "g x \<cdot> y = g (x \<cdot> f y) \<cdot> y"
by (metis assms is_conjugation_sym modular_1)
lemma conjugate_eq_aux:
"is_conjugation f g \<Longrightarrow> f (x \<cdot> -g y) \<le> f x \<cdot> -y"
by (metis f_subdist galois_aux le_inf_iff modular_1_aux)
lemma conjugate_eq:
"is_conjugation f g \<longleftrightarrow> (\<forall>x y. f (x \<cdot> -g y) \<le> f x \<cdot> -y \<and> g (y \<cdot> -f x) \<le> g y \<cdot> -x)"
(is "?l \<longleftrightarrow> ?r")
proof
assume ?l thus ?r
by (metis is_conjugation_sym conjugate_eq_aux)
next
assume r: ?r
have "\<forall>x y. f x \<cdot> y = 0 \<longrightarrow> x \<cdot> g y = 0"
by (metis aux4 inf.left_commute inf_absorb1 inf_compl_bot inf_left_idem sup_bot_left r)
hence "\<forall>x y. x \<cdot> g y = 0 \<longleftrightarrow> f x \<cdot> y = 0"
by (metis aux4 inf.commute inf.left_commute inf_absorb1 inf_compl_bot sup_commute sup_inf_absorb r)
thus "is_conjugation f g"
by (metis is_conjugation_def)
qed
lemma conjugation_prop1: "is_conjugation f g \<Longrightarrow> f y \<cdot> z \<le> f (y \<cdot> g z)"
by (metis le_infE modular_1 order_refl)
lemma conjugation_prop2: "is_conjugation f g \<Longrightarrow> g z \<cdot> y \<le> g (z \<cdot> f y)"
by (metis is_conjugation_sym conjugation_prop1)
end (* boolean_algebra *)
end
diff --git a/thys/Relation_Algebra/Relation_Algebra.thy b/thys/Relation_Algebra/Relation_Algebra.thy
--- a/thys/Relation_Algebra/Relation_Algebra.thy
+++ b/thys/Relation_Algebra/Relation_Algebra.thy
@@ -1,300 +1,300 @@
(* Title: Relation Algebra
Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Relation Algebra\<close>
theory Relation_Algebra
imports More_Boolean_Algebra Kleene_Algebra.Kleene_Algebra
begin
text \<open>We follow Tarski's original article and Maddux's book, in particular we
use their notation. In contrast to Schmidt and Str\"ohlein we do not assume
that the Boolean algebra is complete and we do not consider the Tarski rule in
this development.
A main reason for using complete Boolean algebras seems to be that the
Knaster-Tarski fixpoint theorem becomes available for defining notions of
iteration. In fact, several chapters of Schmidt and Str\"ohlein's book deal
with iteration.
We capture iteration in an alternative way by linking relation algebras with
Kleene algebras (cf.~\emph{relation-algebra-rtc}).\<close>
class relation_algebra = boolean_algebra +
fixes composition :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl ";" 75)
and converse :: "'a \<Rightarrow> 'a" ("(_\<^sup>\<smile>)" [1000] 999)
and unit :: "'a" ("1''")
assumes comp_assoc: "(x ; y) ; z = x ; (y ; z)"
and comp_unitr [simp]: "x ; 1' = x"
and comp_distr: "(x + y) ; z = x ; z + y ; z"
and conv_invol [simp]: "(x\<^sup>\<smile>)\<^sup>\<smile> = x"
and conv_add [simp]: "(x + y)\<^sup>\<smile> = x\<^sup>\<smile> + y\<^sup>\<smile>"
and conv_contrav [simp]: "(x ; y)\<^sup>\<smile> = y\<^sup>\<smile> ; x\<^sup>\<smile>"
and comp_res: "x\<^sup>\<smile> ; -(x ; y) \<le> -y"
text \<open>We first show that every relation algebra is a dioid. We do not yet
treat the zero (the minimal element of the boolean reduct) since the proof of
the annihilation laws is rather tricky to automate. Following Maddux we derive
them from properties of Boolean algebras with operators.\<close>
sublocale relation_algebra \<subseteq> dioid_one "(+)" "(;)" "(\<le>) " "(<)" "1'"
proof
fix x y z :: 'a
show "x ; y ; z = x ; (y ; z)"
by (fact comp_assoc)
show "x + y + z = x + (y + z)"
by (metis sup.commute sup.left_commute)
show "x + y = y + x"
by (fact sup.commute)
show "(x + y) ; z = x ; z + y ; z"
by (fact comp_distr)
show "x + x = x"
by (fact sup.idem)
show "x ; (y + z) = x ; y + x ; z"
by (metis conv_invol conv_add conv_contrav comp_distr)
show "x ; 1' = x"
by (fact comp_unitr)
show "1' ; x = x"
by (metis comp_unitr conv_contrav conv_invol)
show "x \<le> y \<longleftrightarrow> x + y = y"
by (metis sup.commute sup_absorb1 sup_ge1)
show "x < y \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
by (fact less_le)
qed
context relation_algebra
begin
text \<open>First we prove some basic facts about joins and meets.\<close>
lemma meet_interchange: "(w \<cdot> x) ; (y \<cdot> z) \<le> w ; y \<cdot> x ; z"
by (metis inf_le1 inf_le2 le_infI mult_isol_var)
lemma join_interchange: "w ; x + y ; z \<le> (w + y) ; (x + z)"
using local.mult_isol_var local.sup.bounded_iff local.sup.cobounded2 local.sup_ge1 by presburger
text \<open>We now prove some simple facts about conversion.\<close>
lemma conv_iso: "x \<le> y \<longleftrightarrow> x\<^sup>\<smile> \<le> y\<^sup>\<smile>"
by (metis conv_add conv_invol le_iff_sup)
lemma conv_zero [simp]: "0\<^sup>\<smile> = 0"
by (metis conv_add conv_invol sup_bot_right sup_eq_bot_iff)
lemma conv_one [simp]: "1\<^sup>\<smile> = 1"
by (metis conv_add conv_invol sup_top_left sup_top_right)
lemma conv_compl_aux: "(-x)\<^sup>\<smile> = (-x)\<^sup>\<smile> + (-x\<^sup>\<smile>)"
by (metis aux9 conv_add conv_one double_compl galois_aux4 inf.commute less_eq_def sup.commute sup_top_left)
lemma conv_compl: "(-x)\<^sup>\<smile> = -(x\<^sup>\<smile>)"
by (metis add_commute conv_add conv_compl_aux conv_invol)
lemma comp_res_aux [simp]: "x\<^sup>\<smile> ; -(x ; y) \<cdot> y = 0"
by (metis comp_res galois_aux)
lemma conv_e [simp]: "1'\<^sup>\<smile> = 1'"
by (metis comp_unitr conv_contrav conv_invol)
lemma conv_times [simp]: "(x \<cdot> y)\<^sup>\<smile> = x\<^sup>\<smile> \<cdot> y\<^sup>\<smile>"
by (metis compl_inf double_compl conv_add conv_compl)
text \<open>The next lemmas show that conversion is self-conjugate in the sense of
Boolean algebra with operators.\<close>
lemma conv_self_conjugate: "x\<^sup>\<smile> \<cdot> y = 0 \<longleftrightarrow> x \<cdot> y\<^sup>\<smile> = 0"
by (metis conv_invol conv_times conv_zero)
lemma conv_self_conjugate_var: "is_conjugation converse converse"
by (metis conv_self_conjugate is_conjugation_def)
text \<open>The following lemmas link the relative product and meet.\<close>
lemma one_idem_mult [simp]: "1 ; 1 = 1"
by (metis compl_eq_compl_iff galois_aux2 inf.commute inf_top_right mult_1_left mult_isor top_greatest)
lemma mult_subdistl: "x ; (y \<cdot> z) \<le> x ; y"
by (metis inf_le1 mult_isol)
lemma mult_subdistr: "(x \<cdot> y) ; z \<le> x ; z"
by (metis inf_le1 mult_isor)
lemma mult_subdistr_var: "(x \<cdot> y) ; z \<le> x ; z \<cdot> y ; z"
by (metis inf.commute le_inf_iff mult_subdistr)
text \<open>The following lemmas deal with variants of the Peirce law, the
Schr\"oder laws and the Dedekind law. Some of them are obtained from Boolean
algebras with operators by instantiation, using conjugation properties.
However, Isabelle does not always pick up this relationship.\<close>
lemma peirce_1: "x ; y \<cdot> z\<^sup>\<smile> = 0 \<Longrightarrow> y ; z \<cdot> x\<^sup>\<smile> = 0"
by (metis compl_le_swap1 conv_contrav conv_self_conjugate galois_aux comp_res conv_invol galois_aux mult_isol order_trans)
lemma peirce: "x ; y \<cdot> z\<^sup>\<smile> = 0 \<longleftrightarrow> y ; z \<cdot> x\<^sup>\<smile> = 0"
by (metis peirce_1)
lemma schroeder_1: "x ; y \<cdot> z = 0 \<longleftrightarrow> y \<cdot> x\<^sup>\<smile> ; z = 0"
by (metis conv_invol peirce conv_contrav conv_invol conv_self_conjugate inf.commute)
lemma schroeder_2: "y ; x \<cdot> z = 0 \<longleftrightarrow> y \<cdot> z ; x\<^sup>\<smile> = 0"
by (metis conv_invol peirce schroeder_1)
text \<open>The following two conjugation properties between multiplication with
elements and their converses are used for deriving modular laws of relation
algebra from those of Boolean algebras with operators.\<close>
lemma schroeder_1_var: "is_conjugation (composition x) (composition (x\<^sup>\<smile>))"
by (metis schroeder_1 is_conjugation_def)
lemma schroeder_2_var: "is_conjugation (\<lambda>x. x ; y) (\<lambda>x. x ; y\<^sup>\<smile>)"
by (unfold is_conjugation_def, metis schroeder_2)
text \<open>The following Galois connections define residuals. They link relation
algebras with action algebras. This could be further explored and formalised.
\<close>
lemma conv_galois_1: "x ; y \<le> z \<longleftrightarrow> y \<le> -(x\<^sup>\<smile> ; -z)"
by (metis galois_aux galois_aux2 schroeder_1)
lemma conv_galois_2: "y ; x \<le> z \<longleftrightarrow> y \<le> -(-z ; x\<^sup>\<smile>)"
by (metis galois_aux galois_aux2 schroeder_2)
text \<open>Variants of the modular law for relation algebras can now be
instantiated from Boolean algebras with operators.\<close>
lemma modular_1_aux': "x ; (y \<cdot> -(x\<^sup>\<smile> ; z)) \<cdot> z = 0"
by (metis schroeder_1_var modular_1_aux)
lemma modular_2_aux': "(y \<cdot> -(z ; x\<^sup>\<smile>)) ; x \<cdot> z = 0"
by (metis modular_1_aux schroeder_2_var)
lemma modular_1': "x ; y \<cdot> z = x ; (y \<cdot> x\<^sup>\<smile> ; z) \<cdot> z"
by (metis schroeder_1_var modular_1)
lemma modular_2': "y ; x \<cdot> z = (y \<cdot> z ; x\<^sup>\<smile>) ; x \<cdot> z"
proof -
have "y ; x \<cdot> z = (y \<cdot> z ; x\<^sup>\<smile>) ; x \<cdot> z + (y \<cdot> -(z ; x\<^sup>\<smile>)) ; x \<cdot> z"
by (metis aux4 distrib_right inf.commute inf_sup_distrib1)
thus ?thesis
by (metis sup_bot_right modular_2_aux')
qed
lemma modular_1_var: "x ; y \<cdot> z \<le> x ; (y \<cdot> x\<^sup>\<smile> ; z)"
by (metis inf.commute inf_le2 modular_1')
lemma modular_2_var: "x ; y \<cdot> z \<le> (x \<cdot> z ; y\<^sup>\<smile>) ; y"
by (metis inf.commute inf_le2 modular_2')
lemma modular_var_2: "x ; y \<le> x ; (y \<cdot> x\<^sup>\<smile> ; 1)"
by (metis inf_top_right modular_1_var)
lemma modular_var_3: "x ; y \<le> (x \<cdot> 1 ; y\<^sup>\<smile>) ; y"
by (metis inf_top_right modular_2_var)
text \<open>The modular laws are used to prove the Dedekind rule.\<close>
lemma dedekind: "x ; y \<cdot> z \<le> (x \<cdot> z ; y\<^sup>\<smile>) ; (y \<cdot> x\<^sup>\<smile> ; z)"
proof -
have "x ; y \<cdot> z \<le> (x \<cdot> z ; y\<^sup>\<smile>) ; (y \<cdot> ((x \<cdot> z ; y\<^sup>\<smile>)\<^sup>\<smile> ; z))"
by (metis modular_2' modular_1_var)
also have "\<dots> \<le> (x \<cdot> z ; y\<^sup>\<smile>) ; (y \<cdot> x\<^sup>\<smile> ; z)"
by (metis conv_iso inf_le1 inf_mono mult_isol mult_isor order_refl)
thus ?thesis
by (metis calculation order_trans)
qed
lemma dedekind_var_1: "x ; y \<le> (x \<cdot> 1 ; y\<^sup>\<smile>) ; (y \<cdot> x\<^sup>\<smile> ; 1)"
by (metis dedekind inf.commute inf_top_left)
end (* relation_algebra *)
text \<open>The Schr\"oder laws allow us, finally, to prove the annihilation laws
for zero. We formalise this by proving that relation algebras form dioids with
zero.\<close>
sublocale relation_algebra < dioid_one_zero "(+)" "(;)" "1'" 0 "(\<le>)" "(<)"
proof
fix x :: 'a
show "0 + x = x"
by (fact sup_bot_left)
show "0 ; x = 0"
by (metis f_strict schroeder_2_var)
show "x ; 0 = 0"
by (metis f_strict schroeder_1_var)
qed
context relation_algebra
begin
text \<open>Next we prove miscellaneous properties which we found in the books of
Maddux and Schmidt and Str\"ohlein. Most of them do not carry any meaningful
names.\<close>
lemma ra_1: "(x \<cdot> y ; 1) ; z = x ; z \<cdot> y ; 1"
-proof (rule antisym)
+proof (rule order.antisym)
show "x ; z \<cdot> y ; 1 \<le> (x \<cdot> y ; 1) ; z"
by (metis modular_2_var comp_assoc order_trans eq_refl inf_mono inf_top_left mult_isor mult_subdistl)
show "(x \<cdot> y ; 1) ; z \<le> x ; z \<cdot> y ; 1"
by (metis inf.commute inf_greatest inf_top_left mult.assoc mult_subdistl mult_subdistr order_trans)
qed
lemma ra_2: "x ; (z \<cdot> y ; 1) = (x \<cdot> (y ; 1)\<^sup>\<smile>) ; z"
-proof (rule antisym)
+proof (rule order.antisym)
have "(x \<cdot> 1 ; (z \<cdot> y ; 1)\<^sup>\<smile>) ; z \<le> (x \<cdot> (y ; 1)\<^sup>\<smile>) ; z"
by (metis conv_contrav conv_invol conv_one conv_times inf.idem inf.left_commute le_iff_inf meet_assoc mult_isor ra_1)
thus "x ; (z \<cdot> y ; 1) \<le> (x \<cdot> (y ; 1)\<^sup>\<smile>) ; z"
by (metis modular_var_3 mult_subdistl order_trans)
next
have "x ; (z \<cdot> ((x \<cdot> (y ; 1)\<^sup>\<smile>)\<^sup>\<smile> ; 1)) \<le> x ; (z \<cdot> y ; 1)"
by (metis conv_invol conv_times eq_refl inf_le2 inf_mono mult.assoc mult_isol mult_isor one_idem_mult)
thus "x ; (z \<cdot> y ; 1) \<ge> (x \<cdot> (y ; 1)\<^sup>\<smile>) ; z"
by (metis modular_var_2 mult_subdistr order_trans)
qed
lemma one_conv: "1' \<cdot> x ; 1 = 1' \<cdot> x ; x\<^sup>\<smile>"
by (metis inf.commute inf_top_left modular_1' mult.right_neutral)
lemma maddux_12: "-(y ; x) ; x\<^sup>\<smile> \<le> -y"
by (metis galois_aux inf.commute inf_compl_bot schroeder_2)
lemma maddux_141: "x ; y \<cdot> z = 0 \<longleftrightarrow> x\<^sup>\<smile> ; z \<cdot> y = 0"
by (metis inf.commute schroeder_1)
lemma maddux_142: "x\<^sup>\<smile> ; z \<cdot> y = 0 \<longleftrightarrow> z ; y\<^sup>\<smile> \<cdot> x = 0"
by (metis inf.commute schroeder_1 schroeder_2)
lemmas maddux_16 = modular_1_var
lemmas maddux_17 = modular_2_var
lemma maddux_20: "x \<le> x ; 1"
by (metis inf_top_left mult.right_neutral mult_subdistl)
lemma maddux_21: "x \<le> 1 ; x"
by (metis mult_isor mult_onel top_greatest)
lemma maddux_23: "x ; y \<cdot> -(x ; z) = x ; (y \<cdot> -z) \<cdot> -(x ; z)"
-apply (rule antisym)
+apply (rule order.antisym)
apply (metis local.aux6 local.aux6_var local.aux9 local.compl_inf_bot local.compl_sup_top local.compl_unique local.distrib_left local.galois_2 local.sup_ge2)
using local.meet_iso local.mult_subdistl by blast
lemma maddux_24: "-(x ; y) + x ; z = -(x ; (y \<cdot> -z)) + x ; z"
by (metis de_morgan_3 double_compl maddux_23)
lemma one_compl: "-(x ; 1) ; 1 = -(x ; 1)"
-by (metis antisym conv_one maddux_12 mult.assoc one_idem_mult maddux_20)
+by (metis order.antisym conv_one maddux_12 mult.assoc one_idem_mult maddux_20)
lemma ss_p18: "x ; 1 = 0 \<longleftrightarrow> x = 0"
by (metis annil le_bot maddux_20)
end (* relation_algebra *)
text \<open>This finishes our development of the basic laws of relation algebras.
The next sections are devoted to special elements such as vectors, test or
subidentities, and, in particular, functions.\<close>
end
diff --git a/thys/Relation_Algebra/Relation_Algebra_Functions.thy b/thys/Relation_Algebra/Relation_Algebra_Functions.thy
--- a/thys/Relation_Algebra/Relation_Algebra_Functions.thy
+++ b/thys/Relation_Algebra/Relation_Algebra_Functions.thy
@@ -1,359 +1,359 @@
(* Title: Relation Algebra
Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Functions\<close>
theory Relation_Algebra_Functions
imports Relation_Algebra_Vectors Relation_Algebra_Tests
begin
subsection \<open>Functions\<close>
text \<open>This section collects the most important properties of functions. Most
of them can be found in the books by Maddux and by Schmidt and Str\"ohlein. The
main material is on partial and total functions, injections, surjections,
bijections.\<close>
(* Perhaps this material should be reorganised so that related theorems are
grouped together ... *)
context relation_algebra
begin
definition is_p_fun :: "'a \<Rightarrow> bool"
where "is_p_fun x \<equiv> x\<^sup>\<smile> ; x \<le> 1'"
definition is_total :: "'a \<Rightarrow> bool"
where "is_total x \<equiv> 1' \<le> x ; x\<^sup>\<smile>"
definition is_map :: "'a \<Rightarrow> bool"
where "is_map x \<equiv> is_p_fun x \<and> is_total x"
definition is_inj :: "'a \<Rightarrow> bool"
where "is_inj x \<equiv> x ; x\<^sup>\<smile> \<le> 1'"
definition is_sur :: "'a \<Rightarrow> bool"
where "is_sur x \<equiv> 1' \<le> x\<^sup>\<smile> ; x"
text \<open>We distinguish between partial and total bijections. As usual we call
the latter just bijections.\<close>
definition is_p_bij :: "'a \<Rightarrow> bool"
where "is_p_bij x \<equiv> is_p_fun x \<and> is_inj x \<and> is_sur x"
definition is_bij :: "'a \<Rightarrow> bool"
where "is_bij x \<equiv> is_map x \<and> is_inj x \<and> is_sur x"
text \<open>Our first set of lemmas relates the various concepts.\<close>
lemma inj_p_fun: "is_inj x \<longleftrightarrow> is_p_fun (x\<^sup>\<smile>)"
by (metis conv_invol is_inj_def is_p_fun_def)
lemma p_fun_inj: "is_p_fun x \<longleftrightarrow> is_inj (x\<^sup>\<smile>)"
by (metis conv_invol inj_p_fun)
lemma sur_total: "is_sur x \<longleftrightarrow> is_total (x\<^sup>\<smile>)"
by (metis conv_invol is_sur_def is_total_def)
lemma total_sur: "is_total x \<longleftrightarrow> is_sur (x\<^sup>\<smile>)"
by (metis conv_invol sur_total)
lemma bij_conv: "is_bij x \<longleftrightarrow> is_bij (x\<^sup>\<smile>)"
by (metis is_bij_def inj_p_fun is_map_def p_fun_inj sur_total total_sur)
text \<open>Next we show that tests are partial injections.\<close>
lemma test_is_inj_fun: "is_test x \<Longrightarrow> (is_p_fun x \<and> is_inj x)"
by (metis is_inj_def p_fun_inj test_comp test_eq_conv is_test_def)
text \<open>Next we show composition properties.\<close>
lemma p_fun_comp:
assumes "is_p_fun x" and "is_p_fun y"
shows "is_p_fun (x ; y)"
proof (unfold is_p_fun_def)
have "(x ; y)\<^sup>\<smile> ; x ; y = y\<^sup>\<smile> ; x\<^sup>\<smile> ; x ; y"
by (metis conv_contrav mult.assoc)
also have "... \<le> y\<^sup>\<smile> ; y"
by (metis assms(1) is_p_fun_def mult_double_iso mult.right_neutral mult.assoc)
finally show "(x ; y)\<^sup>\<smile> ; (x ; y) \<le> 1'"
by (metis assms(2) order_trans is_p_fun_def mult.assoc)
qed
lemma p_fun_mult_var: "x\<^sup>\<smile> ; x \<le> 1' \<Longrightarrow> (x \<cdot> y)\<^sup>\<smile> ; (x \<cdot> y) \<le> 1'"
by (metis conv_times inf_le1 mult_isol_var order_trans)
lemma inj_compose:
assumes "is_inj x" and "is_inj y"
shows "is_inj (x ; y)"
by (metis assms conv_contrav inj_p_fun p_fun_comp)
lemma inj_mult_var: "x\<^sup>\<smile> ; x \<le> 1' \<Longrightarrow> (x \<cdot> y)\<^sup>\<smile> ; (x \<cdot> y) \<le> 1'"
by (metis p_fun_mult_var)
lemma total_comp:
assumes "is_total x" and "is_total y"
shows "is_total (x ; y)"
by (metis assms inf_top_left le_iff_inf mult.assoc mult.right_neutral one_conv ra_2 is_total_def)
lemma total_add_var: "1' \<le> x\<^sup>\<smile> ; x \<Longrightarrow> 1' \<le> (x + y)\<^sup>\<smile> ; (x + y)"
by (metis local.conv_add local.inf.absorb2 local.inf.coboundedI1 local.join_interchange local.le_sup_iff)
lemma sur_comp:
assumes "is_sur x" and "is_sur y"
shows "is_sur (x ; y)"
by (metis assms conv_contrav sur_total total_comp)
lemma sur_sum_var: "1' \<le> x\<^sup>\<smile> ; x \<Longrightarrow> 1' \<le> (x + y)\<^sup>\<smile> ; (x + y)"
by (metis total_add_var)
lemma map_comp:
assumes "is_map x" and "is_map y"
shows "is_map (x ; y)"
by (metis assms is_map_def p_fun_comp total_comp)
lemma bij_comp:
assumes "is_bij x" and "is_bij y"
shows "is_bij (x ; y)"
by (metis assms is_bij_def inj_compose map_comp sur_comp)
text \<open>We now show that (partial) functions, unlike relations, distribute over
meets from the left.\<close>
lemma p_fun_distl: "is_p_fun x \<Longrightarrow> x ; (y \<cdot> z) = x ; y \<cdot> x ; z"
proof -
assume "is_p_fun x"
hence "x ; (z \<cdot> ((x\<^sup>\<smile> ; x) ; y)) \<le> x ; (z \<cdot> y)"
by (metis is_p_fun_def inf_le1 le_infI le_infI2 mult_isol mult_isor mult_onel)
hence "x ; y \<cdot> x ; z \<le> x ; (z \<cdot> y)"
by (metis inf.commute mult.assoc order_trans modular_1_var)
thus "x ; (y \<cdot> z) = x ; y \<cdot> x ; z"
- by (metis eq_iff inf.commute le_infI mult_subdistl)
+ by (metis order.eq_iff inf.commute le_infI mult_subdistl)
qed
lemma map_distl: "is_map x \<Longrightarrow> x ; (y \<cdot> z) = x ; y \<cdot> x ; z"
by (metis is_map_def p_fun_distl)
text \<open>Next we prove simple properties of functions which arise in equivalent
definitions of those concepts.\<close>
lemma p_fun_zero: "is_p_fun x \<Longrightarrow> x ; y \<cdot> x ; -y = 0"
by (metis annir inf_compl_bot p_fun_distl)
lemma total_one: "is_total x \<Longrightarrow> x ; 1 = 1"
by (metis conv_invol conv_one inf_top_left le_iff_inf mult.right_neutral one_conv ra_2 is_total_def)
lemma total_1: "is_total x \<Longrightarrow> (\<forall>y. y ; x = 0 \<longrightarrow> y = 0)"
by (metis conv_invol conv_zero inf_bot_left inf_top_left peirce total_one)
lemma surj_one: "is_sur x \<Longrightarrow> 1 ; x = 1"
by (metis conv_contrav conv_invol conv_one sur_total total_one)
lemma surj_1: "is_sur x \<Longrightarrow> (\<forall>y. x ; y = 0 \<longrightarrow> y = 0)"
by (metis comp_res_aux compl_bot_eq conv_contrav conv_one inf.commute inf_top_right surj_one)
lemma bij_is_maprop:
assumes "is_bij x" and "is_map x"
shows "x\<^sup>\<smile> ; x = 1' \<and> x ; x\<^sup>\<smile> = 1'"
-by (metis assms is_bij_def eq_iff is_inj_def is_map_def is_p_fun_def is_sur_def is_total_def)
+by (metis assms is_bij_def order.eq_iff is_inj_def is_map_def is_p_fun_def is_sur_def is_total_def)
text\<open>We now provide alternative definitions for functions. These can be found
in Schmidt and Str\"ohlein's book.\<close>
lemma p_fun_def_var: "is_p_fun x \<longleftrightarrow> x ; -(1') \<le> -x"
by (metis conv_galois_1 double_compl galois_aux inf.commute is_p_fun_def)
lemma total_def_var_1: "is_total x \<longleftrightarrow> x ; 1 = 1"
by (metis inf_top_right le_iff_inf one_conv total_one is_total_def)
lemma total_def_var_2: "is_total x \<longleftrightarrow> -x \<le> x ; -(1')"
by (metis total_def_var_1 distrib_left sup_compl_top mult.right_neutral galois_aux3)
lemma sur_def_var1: "is_sur x \<longleftrightarrow> 1 ; x = 1"
by (metis conv_contrav conv_one sur_total surj_one total_def_var_1)
lemma sur_def_var2: "is_sur x \<longleftrightarrow> -x \<le> -(1') ; x"
by (metis sur_total total_def_var_2 conv_compl conv_contrav conv_e conv_iso)
lemma inj_def_var1: "is_inj x \<longleftrightarrow> -(1') ; x \<le> -x"
by (metis conv_galois_2 double_compl galois_aux inf.commute is_inj_def)
lemma is_maprop: "is_map x \<longleftrightarrow> x ; -(1') = -x"
-by (metis eq_iff is_map_def p_fun_def_var total_def_var_2)
+by (metis order.eq_iff is_map_def p_fun_def_var total_def_var_2)
text \<open>Finally we prove miscellaneous properties of functions.\<close>
lemma ss_422iii: "is_p_fun y \<Longrightarrow> (x \<cdot> z ; y\<^sup>\<smile>) ; y = x ; y \<cdot> z"
(* by (smt antisym comp_assoc inf_commute maddux_17 meet_iso mult_isol mult_oner mult_subdistr_var order_trans is_p_fun_def) *)
-proof (rule antisym)
+proof (rule order.antisym)
assume "is_p_fun y"
show "x ; y \<cdot> z \<le> (x \<cdot> z ; y\<^sup>\<smile>) ; y"
by (metis maddux_17)
have "(x \<cdot> z ; y\<^sup>\<smile>) ; y \<le> x ; y \<cdot> (z ; (y\<^sup>\<smile> ; y))"
by (metis mult_subdistr_var mult.assoc)
also have "\<dots> \<le> x ; y \<cdot> z ; 1'"
by (metis \<open>is_p_fun y\<close> inf_absorb2 inf_le1 le_infI le_infI2 mult_subdistl is_p_fun_def)
finally show "(x \<cdot> z ; y\<^sup>\<smile>) ; y \<le> x ; y \<cdot> z"
by (metis mult.right_neutral)
qed
lemma p_fun_compl: "is_p_fun x \<Longrightarrow> x ; -y \<le> -(x; y)"
by (metis annir galois_aux inf.commute inf_compl_bot p_fun_distl)
lemma ss_422v: "is_p_fun x \<Longrightarrow> x ; -y = x ; 1 \<cdot> -(x ; y)"
by (metis inf.commute inf_absorb2 inf_top_left maddux_23 p_fun_compl)
text \<open>The next property is a Galois connection.\<close>
lemma ss43iii: "is_map x \<longleftrightarrow> (\<forall>y. x ; -y = -(x ; y))"
by standard (metis inf_top_left is_map_def ss_422v total_one, metis is_maprop mult.right_neutral)
text \<open>Next we prove a lemma from Schmidt and Str\"ohlein's book and some of
its consequences. We show the proof in detail since the textbook proof uses
Tarski's rule which we omit.\<close>
lemma ss423: "is_map x \<Longrightarrow> y ; x \<le> z \<longleftrightarrow> y \<le> z ; x\<^sup>\<smile>"
proof
assume "is_map x" and "y ; x \<le> z"
hence "y \<le> y ; x ; x\<^sup>\<smile>"
by (metis is_map_def mult_1_right mult.assoc mult_isol is_total_def)
thus "y \<le> z ; x\<^sup>\<smile>"
by (metis \<open>y ; x \<le> z\<close> mult_isor order_trans)
next
assume "is_map x" and "y \<le> z ; x\<^sup>\<smile>"
hence "y ; x \<le> z ; x\<^sup>\<smile> ; x"
by (metis mult_isor)
also have "\<dots> \<le> z ; 1'"
by (metis \<open>is_map x\<close> is_map_def mult.assoc mult_isol is_p_fun_def)
finally show "y ; x \<le> z"
by (metis mult_1_right)
qed
lemma ss424i: "is_total x \<longleftrightarrow> (\<forall>y. -(x ; y) \<le> x ; -y)"
by (metis galois_aux3 distrib_left sup_compl_top total_def_var_1)
lemma ss434ii: "is_p_fun x \<longleftrightarrow> (\<forall>y. x ; -y \<le> -(x ; y))"
by (metis mult.right_neutral p_fun_compl p_fun_def_var)
lemma is_maprop1: "is_map x \<Longrightarrow> (y \<le> x ; z ; x\<^sup>\<smile> \<longleftrightarrow> y ; x \<le> x ; z)"
by (metis ss423)
lemma is_maprop2: "is_map x \<Longrightarrow> (y ; x \<le> x ; z \<longleftrightarrow> x\<^sup>\<smile> ; y; x \<le> z)"
by standard (metis galois_aux2 inf_commute mult.assoc schroeder_1 ss43iii, metis conv_contrav conv_invol conv_iso mult.assoc ss423)
lemma is_maprop3: "is_map x \<Longrightarrow> (x\<^sup>\<smile> ; y; x \<le> z \<longleftrightarrow> x\<^sup>\<smile> ; y \<le> z ; x\<^sup>\<smile>)"
by (metis ss423)
lemma p_fun_sur_id [simp]:
assumes "is_p_fun x" and "is_sur x"
shows "x\<^sup>\<smile> ; x = 1'"
-by (metis assms eq_iff is_p_fun_def is_sur_def)
+by (metis assms order.eq_iff is_p_fun_def is_sur_def)
lemma total_inj_id [simp]:
assumes "is_total x" and "is_inj x"
shows "x ; x\<^sup>\<smile> = 1'"
by (metis assms conv_invol inj_p_fun p_fun_sur_id sur_total)
lemma bij_inv_1 [simp]: "is_bij x \<Longrightarrow> x ; x\<^sup>\<smile> = 1'"
by (metis bij_is_maprop is_bij_def)
lemma bij_inv_2 [simp]: "is_bij x \<Longrightarrow> x\<^sup>\<smile> ; x = 1'"
by (metis bij_is_maprop is_bij_def)
lemma bij_inv_comm: "is_bij x \<Longrightarrow> x ; x\<^sup>\<smile> = x\<^sup>\<smile> ; x"
by (metis bij_inv_1 bij_inv_2)
lemma is_bijrop: "is_bij x \<Longrightarrow> (y = x ; z \<longleftrightarrow> z = x\<^sup>\<smile> ; y)"
by (metis bij_inv_1 bij_inv_2 mult.assoc mult.left_neutral)
lemma inj_map_monomorph: "\<lbrakk>is_inj x; is_map x\<rbrakk> \<Longrightarrow> (\<forall>y z. y ; x = z ; x \<longrightarrow> y = z)"
by (metis is_map_def mult.assoc mult.right_neutral total_inj_id)
lemma sur_map_epimorph: "\<lbrakk>is_sur x; is_map x\<rbrakk> \<Longrightarrow> (\<forall>y z. x ; y = x ; z \<longrightarrow> y = z)"
-by (metis eq_iff mult.assoc mult.left_neutral ss423 is_sur_def)
+by (metis order.eq_iff mult.assoc mult.left_neutral ss423 is_sur_def)
subsection \<open>Points and Rectangles\<close>
text \<open>Finally here is a section on points and rectangles. This is only a
beginning.\<close>
definition is_point :: "'a \<Rightarrow> bool"
where "is_point x \<equiv> is_vector x \<and> is_inj x \<and> x \<noteq> 0"
definition is_rectangle :: "'a \<Rightarrow> bool"
where "is_rectangle x \<equiv> x ; 1 ; x \<le> x"
lemma rectangle_eq [simp]: "is_rectangle x \<longleftrightarrow> x ; 1 ; x = x"
-by (metis conv_one dedekind eq_iff inf_top_left mult.assoc one_idem_mult is_rectangle_def)
+by (metis conv_one dedekind order.eq_iff inf_top_left mult.assoc one_idem_mult is_rectangle_def)
subsection \<open>Antidomain\<close>
text\<open>This section needs to be linked with domain semirings. We essentially
prove the antidomain semiring axioms. Then we have the abstract properties at
our disposition.\<close>
definition antidom :: "'a \<Rightarrow> 'a" ("a")
where "a x = 1' \<cdot> (-(x ; 1))"
definition dom :: "'a \<Rightarrow> 'a" ("d")
where "d x = a (a x)"
lemma antidom_test_comp [simp]: "a x = tc (x ; 1)"
by (metis antidom_def tc_def)
lemma dom_def_aux: "d x = 1' \<cdot> x ; 1"
by (metis antidom_test_comp dom_def double_compl inf_top_left mult.left_neutral one_compl ra_1 tc_def)
lemma dom_def_aux_var: "d x = 1' \<cdot> x ; x\<^sup>\<smile>"
by (metis dom_def_aux one_conv)
lemma antidom_dom [simp]: "a (d x) = a x"
by (metis antidom_test_comp dom_def_aux inf_top_left mult.left_neutral ra_1)
lemma dom_antidom [simp]: "d (a x) = a x"
by (metis antidom_dom dom_def)
lemma dom_verystrict: "d x = 0 \<longleftrightarrow> x = 0"
using dom_def_aux_var local.schroeder_2 by force
lemma a_1 [simp]: "a x ; x = 0"
by (metis antidom_test_comp galois_aux2 maddux_20 mult.left_neutral one_compl ra_1 tc_def)
lemma a_2: "a (x ; y) = a (x ; d y)"
by (metis antidom_test_comp dom_def_aux inf_top_left mult.assoc mult.left_neutral ra_1)
lemma a_3 [simp]: "a x + d x = 1'"
by (metis antidom_def aux4 dom_def_aux double_compl)
lemma test_domain: "x = d x \<longleftrightarrow> x \<le> 1'"
apply standard
apply (metis dom_def_aux inf_le1)
apply (metis dom_def_aux inf.commute mult.right_neutral test_1 is_test_def)
done
text \<open>At this point we have all the necessary ingredients to prove that
relation algebras form Boolean domain semirings. However, we omit a formal
proof since we haven't formalized the latter.\<close>
lemma dom_one: "x ; 1 = d x ; 1"
by (metis dom_def_aux inf_top_left mult.left_neutral ra_1)
lemma test_dom: "is_test (d x)"
by (metis dom_def_aux inf_le1 is_test_def)
lemma p_fun_dom: "is_p_fun (d x)"
by (metis test_dom test_is_inj_fun)
lemma inj_dom: "is_inj (d x)"
by (metis test_dom test_is_inj_fun)
lemma total_alt_def: "is_total x \<longleftrightarrow> (d x) = 1'"
by (metis dom_def_aux_var le_iff_inf is_total_def)
end (* relation_algebra *)
end
diff --git a/thys/Relation_Algebra/Relation_Algebra_RTC.thy b/thys/Relation_Algebra/Relation_Algebra_RTC.thy
--- a/thys/Relation_Algebra/Relation_Algebra_RTC.thy
+++ b/thys/Relation_Algebra/Relation_Algebra_RTC.thy
@@ -1,50 +1,50 @@
(* Title: Relation Algebra
Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Reflexive Transitive Closure\<close>
theory Relation_Algebra_RTC
imports Relation_Algebra
begin
text \<open>We impose the Kleene algebra axioms for the star on relation algebras.
This gives us a reflexive transitive closure operation.\<close>
class relation_algebra_rtc = relation_algebra + star_op +
assumes rtc_unfoldl: "1' + x ; x\<^sup>\<star> \<le> x\<^sup>\<star>"
and rtc_inductl: "z + x ; y \<le> y \<longrightarrow> x\<^sup>\<star> ; z \<le> y"
and rtc_inductr: "z + y ; x \<le> y \<longrightarrow> z ; x\<^sup>\<star> \<le> y"
sublocale relation_algebra_rtc \<subseteq> kleene_algebra "(+)" "(;)" "1'" 0 "(\<le>)" "(<)" star
apply (unfold_locales)
apply (rule rtc_unfoldl)
apply (simp add: local.rtc_inductl)
apply (simp add: rtc_inductr)
done
context relation_algebra_rtc
begin
text \<open>First, we prove that the obvious interaction between the star and
converse is captured by the axioms.\<close>
lemma star_conv: "(x\<^sup>\<star>)\<^sup>\<smile> = (x\<^sup>\<smile>)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "(x\<^sup>\<smile>)\<^sup>\<star> \<le> (x\<^sup>\<star>)\<^sup>\<smile>"
by (metis local.conv_add local.conv_contrav local.conv_e local.conv_iso local.star_rtc1 local.star_rtc_least)
show "(x\<^sup>\<star>)\<^sup>\<smile> \<le> (x\<^sup>\<smile>)\<^sup>\<star>"
by (metis boffa_var conv_add conv_contrav conv_e conv_invol conv_iso star_denest star_ext star_iso star_plus_one sup.idem)
qed
text \<open>Next we provide an example to show how facts from Kleene algebra are
picked up in relation algebra.\<close>
lemma rel_church_rosser: "(x\<^sup>\<smile>)\<^sup>\<star> ; x\<^sup>\<star> \<le> x\<^sup>\<star> ; (x\<^sup>\<smile>)\<^sup>\<star> \<Longrightarrow> (x + x\<^sup>\<smile>)\<^sup>\<star> = x\<^sup>\<star> ; (x\<^sup>\<smile>)\<^sup>\<star>"
by (fact church_rosser)
end (* relation_algebra_rtc *)
end
diff --git a/thys/Relation_Algebra/Relation_Algebra_Tests.thy b/thys/Relation_Algebra/Relation_Algebra_Tests.thy
--- a/thys/Relation_Algebra/Relation_Algebra_Tests.thy
+++ b/thys/Relation_Algebra/Relation_Algebra_Tests.thy
@@ -1,153 +1,153 @@
(* Title: Relation Algebra
Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber
Maintainer: Georg Struth <g.struth at sheffield.ac.uk>
Tjark Weber <tjark.weber at it.uu.se>
*)
section \<open>Tests\<close>
theory Relation_Algebra_Tests
imports Relation_Algebra
begin
subsection \<open>Tests\<close>
text \<open>Tests or subidentities provide another way of modelling sets. Once more
we prove the basic properties, most of which stem from Maddux's book.\<close>
context relation_algebra
begin
definition is_test :: "'a \<Rightarrow> bool"
where "is_test x \<equiv> x \<le> 1'"
lemma test_conv: "is_test x \<Longrightarrow> is_test (x\<^sup>\<smile>)"
by (metis conv_e conv_iso is_test_def)
lemma test_conv_var: "is_test x \<Longrightarrow> x\<^sup>\<smile> \<le> 1'"
by (metis test_conv is_test_def)
lemma test_eq_conv [simp]: "is_test x \<Longrightarrow> x\<^sup>\<smile> = x"
-proof (rule antisym)
+proof (rule order.antisym)
assume hyp: "is_test x"
hence "x \<le> x ; (1' \<cdot> x\<^sup>\<smile> ; 1')"
by (metis inf.commute inf_absorb2 inf_le2 modular_1' mult.right_neutral is_test_def)
thus "x \<le> x\<^sup>\<smile>"
- by (metis comp_unitr conv_contrav conv_invol eq_iff hyp inf_absorb2 mult_subdistl test_conv_var)
+ by (metis comp_unitr conv_contrav conv_invol order.eq_iff hyp inf_absorb2 mult_subdistl test_conv_var)
thus "x\<^sup>\<smile> \<le> x"
by (metis conv_invol conv_times le_iff_inf)
qed
lemma test_sum: "\<lbrakk>is_test x; is_test y\<rbrakk> \<Longrightarrow> is_test (x + y)"
by (simp add: is_test_def)
lemma test_prod: "\<lbrakk>is_test x; is_test y\<rbrakk> \<Longrightarrow> is_test (x \<cdot> y)"
by (metis le_infI2 is_test_def)
lemma test_comp: "\<lbrakk>is_test x; is_test y\<rbrakk> \<Longrightarrow> is_test (x ; y)"
by (metis mult_isol comp_unitr order_trans is_test_def)
lemma test_comp_eq_mult:
assumes "is_test x"
and "is_test y"
shows "x ; y = x \<cdot> y"
-proof (rule antisym)
+proof (rule order.antisym)
show "x ; y \<le> x \<cdot> y"
by (metis assms comp_unitr inf_absorb2 le_inf_iff mult_onel mult_subdistl mult_subdistr is_test_def)
next
have "x \<cdot> y \<le> x ; (1' \<cdot> x\<^sup>\<smile> ; y)"
by (metis comp_unitr modular_1_var)
thus "x \<cdot> y \<le> x ; y"
by (metis assms(1) inf_absorb2 le_infI2 mult.left_neutral mult_isol mult_subdistr order_trans test_eq_conv is_test_def)
qed
lemma test_1 [simp]: "is_test x \<Longrightarrow> x ; 1 \<cdot> y = x ; y"
by (metis inf.commute inf.idem inf_absorb2 mult.left_neutral one_conv ra_1 test_comp_eq_mult test_eq_conv is_test_def)
lemma maddux_32 [simp]: "is_test x \<Longrightarrow> -(x ; 1) \<cdot> 1' = -x \<cdot> 1'"
-proof (rule antisym)
+proof (rule order.antisym)
assume "is_test x"
show "-(x ; 1) \<cdot> 1' \<le> -x \<cdot> 1'"
by (metis maddux_20 comp_anti inf.commute meet_isor)
next
assume "is_test x"
have one: "x ; 1 \<cdot> (-x \<cdot> 1') \<le> x ; x\<^sup>\<smile> ; (-x \<cdot> 1')"
by (metis maddux_16 inf_top_left mult.assoc)
hence two: "x ; 1 \<cdot> (-x \<cdot> 1') \<le> -x"
by (metis inf.commute inf_le1 le_infE)
hence "x ; 1 \<cdot> (-x \<cdot> 1') \<le> x"
- by (metis one inf.commute le_infE meet_iso one_conv \<open>is_test x\<close> eq_iff test_1 test_eq_conv)
+ by (metis one inf.commute le_infE meet_iso one_conv \<open>is_test x\<close> order.eq_iff test_1 test_eq_conv)
hence "x ; 1 \<cdot> (-x \<cdot> 1') = 0"
by (metis two galois_aux2 le_iff_inf)
thus "-x \<cdot> 1' \<le> -(x ; 1) \<cdot> 1'"
by (metis double_compl galois_aux2 inf.commute inf_le1 le_inf_iff)
qed
lemma test_distr_1 :
assumes "is_test x"
and "is_test y"
shows "x ; z \<cdot> y ; z = (x \<cdot> y) ; z"
-proof (rule antisym)
+proof (rule order.antisym)
have "x ; z \<cdot> y ; z \<le> x ; 1 \<cdot> y ; z"
by (metis inf_top_left meet_iso mult_subdistl)
also have "\<dots> = x ; y ; z"
by (metis assms(1) mult.assoc test_1)
finally show "x ; z \<cdot> y ; z \<le> (x \<cdot> y) ; z"
by (metis assms test_comp_eq_mult)
next
show "(x \<cdot> y) ; z \<le> x ; z \<cdot> y ; z"
by (metis mult_subdistr_var)
qed
lemma maddux_35: "is_test x \<Longrightarrow> x ; y \<cdot> -z = x ; y \<cdot> -(x ; z)"
-proof (rule antisym)
+proof (rule order.antisym)
assume "is_test x"
show "x ; y \<cdot> -z \<le> x ; y \<cdot> -(x ; z)"
by (metis \<open>is_test x\<close> comp_anti mult_isor mult_onel is_test_def inf.commute inf_le2 le_infI le_infI1)
have one: "x ; y \<cdot> -(x ; z) \<le> x ; (y \<cdot> -z)"
- by (metis eq_iff le_infE maddux_23)
+ by (metis order.eq_iff le_infE maddux_23)
hence two: "x ; y \<cdot> -(x ; z) \<le> x ; y"
by (metis inf_le1)
have "x ; y \<cdot> -(x ; z) \<le> x ; -z"
using one
by (metis galois_1 le_iff_sup distrib_left sup_compl_top sup_top_right)
hence "x ; y \<cdot> -(x ; z) \<le> -z"
by (metis \<open>is_test x\<close> mult_isor mult_onel is_test_def order_trans)
thus "x ; y \<cdot> -(x ; z) \<le> x ; y \<cdot> -z"
using two
by (metis le_inf_iff)
qed
subsection \<open>Test Complements\<close>
text \<open>Text complements are complements of elements that are ``pushed below''
the multiplicative unit.\<close>
definition tc :: "'a \<Rightarrow> 'a"
where "tc x = 1' \<cdot> -x"
lemma test_compl_1 [simp]: "is_test x \<Longrightarrow> x + tc x = 1'"
by (metis is_test_def local.aux4 local.inf.absorb_iff1 local.inf_commute tc_def)
lemma test_compl_2 [simp]: "is_test x \<Longrightarrow> x \<cdot> tc x = 0"
by (metis galois_aux inf.commute inf_le2 tc_def)
lemma test_test_compl: "is_test x \<Longrightarrow> is_test (tc x)"
by (simp add: is_test_def tc_def)
lemma test_compl_de_morgan_1: "tc (x + y) = tc x \<cdot> tc y"
by (metis compl_sup inf.left_commute inf.left_idem meet_assoc tc_def)
lemma test_compl_de_morgan_2: "tc (x \<cdot> y) = tc x + tc y"
by (metis compl_inf inf_sup_distrib1 tc_def)
lemma test_compl_three [simp]: "tc (tc (tc x)) = tc x"
by (metis aux4 aux6 de_morgan_3 inf.commute inf_sup_absorb tc_def)
lemma test_compl_double [simp]: "is_test x \<Longrightarrow> tc (tc x) = x"
by (metis aux6_var compl_inf double_compl inf.commute le_iff_inf tc_def is_test_def)
end (* relation_algebra *)
end
diff --git a/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy b/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy
--- a/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy
+++ b/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy
@@ -1,2280 +1,2280 @@
(* Title: Disjoint-Set Forests
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
theory Disjoint_Set_Forests
imports
"HOL-Hoare.Hoare_Logic"
Stone_Kleene_Relation_Algebras.Kleene_Relation_Algebras
begin
no_notation
trancl ("(_\<^sup>+)" [1000] 999)
text \<open>
An arc in a Stone relation algebra corresponds to an atom in a relation algebra and represents a single edge in a graph.
A point represents a set of nodes.
A rectangle represents the Cartesian product of two sets of nodes \cite{BerghammerStruth2010}.
\<close>
context times_top
begin
abbreviation rectangle :: "'a \<Rightarrow> bool"
where "rectangle x \<equiv> x * top * x = x"
end
context stone_relation_algebra
begin
lemma arc_rectangle:
"arc x \<Longrightarrow> rectangle x"
using arc_top_arc by blast
section \<open>Relation-Algebraic Semantics of Associative Array Access\<close>
text \<open>
The following two operations model updating array $x$ at index $y$ to value $z$,
and reading the content of array $x$ at index $y$, respectively.
The read operation uses double brackets to avoid ambiguity with list syntax.
The remainder of this section shows basic properties of these operations.
\<close>
abbreviation rel_update :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("(_[_\<longmapsto>_])" [70, 65, 65] 61)
where "x[y\<longmapsto>z] \<equiv> (y \<sqinter> z\<^sup>T) \<squnion> (-y \<sqinter> x)"
abbreviation rel_access :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("(2_[[_]])" [70, 65] 65)
where "x[[y]] \<equiv> x\<^sup>T * y"
text \<open>Theorem 1.1\<close>
lemma update_univalent:
assumes "univalent x"
and "vector y"
and "injective z"
shows "univalent (x[y\<longmapsto>z])"
proof -
have 1: "univalent (y \<sqinter> z\<^sup>T)"
using assms(3) inf_commute univalent_inf_closed by force
have "(y \<sqinter> z\<^sup>T)\<^sup>T * (-y \<sqinter> x) = (y\<^sup>T \<sqinter> z) * (-y \<sqinter> x)"
by (simp add: conv_dist_inf)
also have "... = z * (y \<sqinter> -y \<sqinter> x)"
by (metis assms(2) covector_inf_comp_3 inf.sup_monoid.add_assoc inf.sup_monoid.add_commute)
finally have 2: "(y \<sqinter> z\<^sup>T)\<^sup>T * (-y \<sqinter> x) = bot"
by simp
have 3: "vector (-y)"
using assms(2) vector_complement_closed by simp
have "(-y \<sqinter> x)\<^sup>T * (y \<sqinter> z\<^sup>T) = (-y\<^sup>T \<sqinter> x\<^sup>T) * (y \<sqinter> z\<^sup>T)"
by (simp add: conv_complement conv_dist_inf)
also have "... = x\<^sup>T * (-y \<sqinter> y \<sqinter> z\<^sup>T)"
using 3 by (metis (mono_tags, hide_lams) conv_complement covector_inf_comp_3 inf.sup_monoid.add_assoc inf.sup_monoid.add_commute)
finally have 4: "(-y \<sqinter> x)\<^sup>T * (y \<sqinter> z\<^sup>T) = bot"
by simp
have 5: "univalent (-y \<sqinter> x)"
using assms(1) inf_commute univalent_inf_closed by fastforce
have "(x[y\<longmapsto>z])\<^sup>T * (x[y\<longmapsto>z]) = (y \<sqinter> z\<^sup>T)\<^sup>T * (x[y\<longmapsto>z]) \<squnion> (-y \<sqinter> x)\<^sup>T * (x[y\<longmapsto>z])"
by (simp add: conv_dist_sup mult_right_dist_sup)
also have "... = (y \<sqinter> z\<^sup>T)\<^sup>T * (y \<sqinter> z\<^sup>T) \<squnion> (y \<sqinter> z\<^sup>T)\<^sup>T * (-y \<sqinter> x) \<squnion> (-y \<sqinter> x)\<^sup>T * (y \<sqinter> z\<^sup>T) \<squnion> (-y \<sqinter> x)\<^sup>T * (-y \<sqinter> x)"
by (simp add: mult_left_dist_sup sup_assoc)
finally show ?thesis
using 1 2 4 5 by simp
qed
text \<open>Theorem 1.2\<close>
lemma update_total:
assumes "total x"
and "vector y"
and "regular y"
and "surjective z"
shows "total (x[y\<longmapsto>z])"
proof -
have "(x[y\<longmapsto>z]) * top = x*top[y\<longmapsto>top*z]"
by (simp add: assms(2) semiring.distrib_right vector_complement_closed vector_inf_comp conv_dist_comp)
also have "... = top[y\<longmapsto>top]"
using assms(1) assms(4) by simp
also have "... = top"
using assms(3) regular_complement_top by auto
finally show ?thesis
by simp
qed
text \<open>Theorem 1.3\<close>
lemma update_mapping:
assumes "mapping x"
and "vector y"
and "regular y"
and "bijective z"
shows "mapping (x[y\<longmapsto>z])"
using assms update_univalent update_total by simp
text \<open>Theorem 1.4\<close>
lemma read_injective:
assumes "injective y"
and "univalent x"
shows "injective (x[[y]])"
using assms injective_mult_closed univalent_conv_injective by blast
text \<open>Theorem 1.5\<close>
lemma read_surjective:
assumes "surjective y"
and "total x"
shows "surjective (x[[y]])"
using assms surjective_mult_closed total_conv_surjective by blast
text \<open>Theorem 1.6\<close>
lemma read_bijective:
assumes "bijective y"
and "mapping x"
shows "bijective (x[[y]])"
by (simp add: assms read_injective read_surjective)
text \<open>Theorem 1.7\<close>
lemma read_point:
assumes "point p"
and "mapping x"
shows "point (x[[p]])"
using assms comp_associative read_injective read_surjective by auto
text \<open>Theorem 1.8\<close>
lemma update_postcondition:
assumes "point x" "point y"
shows "x \<sqinter> p = x * y\<^sup>T \<longleftrightarrow> p[[x]] = y"
apply (rule iffI)
subgoal by (metis assms comp_associative conv_dist_comp conv_involutive covector_inf_comp_3 equivalence_top_closed vector_covector)
subgoal
- apply (rule antisym)
+ apply (rule order.antisym)
subgoal by (metis assms conv_dist_comp conv_involutive inf.boundedI inf.cobounded1 vector_covector vector_restrict_comp_conv)
subgoal by (smt assms comp_associative conv_dist_comp conv_involutive covector_restrict_comp_conv dense_conv_closed equivalence_top_closed inf.boundedI shunt_mapping vector_covector preorder_idempotent)
done
done
text \<open>Back and von Wright's array independence requirements \cite{BackWright1998},
later also lens laws \cite{FosterGreenwaldMoorePierceSchmitt2007}\<close>
lemma put_get:
assumes "vector y" "surjective y" "vector z"
shows "(x[y\<longmapsto>z])[[y]] = z"
proof -
have "(x[y\<longmapsto>z])[[y]] = (y\<^sup>T \<sqinter> z) * y \<squnion> (-y\<^sup>T \<sqinter> x\<^sup>T) * y"
by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup)
also have "... = z * y"
proof -
have "(-y\<^sup>T \<sqinter> x\<^sup>T) * y = bot"
by (metis assms(1) covector_inf_comp_3 inf_commute conv_complement mult_right_zero p_inf vector_complement_closed)
thus ?thesis
by (simp add: assms covector_inf_comp_3 inf_commute)
qed
also have "... = z"
by (metis assms(2,3) mult_assoc)
finally show ?thesis
.
qed
lemma put_put:
"(x[y\<longmapsto>z])[y\<longmapsto>w] = x[y\<longmapsto>w]"
by (metis inf_absorb2 inf_commute inf_le1 inf_sup_distrib1 maddux_3_13 sup_inf_absorb)
lemma get_put:
assumes "point y"
shows "x[y\<longmapsto>x[[y]]] = x"
proof -
have "x[y\<longmapsto>x[[y]]] = (y \<sqinter> y\<^sup>T * x) \<squnion> (-y \<sqinter> x)"
by (simp add: conv_dist_comp)
also have "... = (y \<sqinter> x) \<squnion> (-y \<sqinter> x)"
proof -
have "y \<sqinter> y\<^sup>T * x = y \<sqinter> x"
- proof (rule antisym)
+ proof (rule order.antisym)
have "y \<sqinter> y\<^sup>T * x = (y \<sqinter> y\<^sup>T) * x"
by (simp add: assms vector_inf_comp)
also have "(y \<sqinter> y\<^sup>T) * x = y * y\<^sup>T * x"
by (simp add: assms vector_covector)
also have "... \<le> x"
using assms comp_isotone by fastforce
finally show "y \<sqinter> y\<^sup>T * x \<le> y \<sqinter> x"
by simp
have "y \<sqinter> x \<le> y\<^sup>T * x"
by (simp add: assms vector_restrict_comp_conv)
thus "y \<sqinter> x \<le> y \<sqinter> y\<^sup>T * x"
by simp
qed
thus ?thesis
by simp
qed
also have "... = x"
proof -
have "regular y"
using assms bijective_regular by blast
thus ?thesis
by (metis inf.sup_monoid.add_commute maddux_3_11_pp)
qed
finally show ?thesis
.
qed
lemma update_inf:
"u \<le> y \<Longrightarrow> (x[y\<longmapsto>z]) \<sqinter> u = z\<^sup>T \<sqinter> u"
by (smt comp_inf.mult_right_dist_sup comp_inf.semiring.mult_zero_right inf.left_commute inf.sup_monoid.add_assoc inf_absorb2 p_inf sup_bot_right inf.sup_monoid.add_commute)
lemma update_inf_same:
"(x[y\<longmapsto>z]) \<sqinter> y = z\<^sup>T \<sqinter> y"
by (simp add: update_inf)
lemma update_inf_different:
"u \<le> -y \<Longrightarrow> (x[y\<longmapsto>z]) \<sqinter> u = x \<sqinter> u"
by (smt inf.right_idem inf.sup_monoid.add_commute inf.sup_relative_same_increasing inf_import_p maddux_3_13 sup.cobounded2 update_inf_same)
end
section \<open>Relation-Algebraic Semantics of Disjoint-Set Forests\<close>
text \<open>
A disjoint-set forest represents a partition of a set into equivalence classes.
We take the represented equivalence relation as the semantics of a forest.
It is obtained by operation \<open>fc\<close> below.
Additionally, operation \<open>wcc\<close> giving the weakly connected components of a graph will be used for the semantics of the union of two disjoint sets.
Finally, operation \<open>root\<close> yields the root of a component tree, that is, the representative of a set containing a given element.
This section defines these operations and derives their properties.
\<close>
context stone_kleene_relation_algebra
begin
text \<open>Theorem 4.2\<close>
lemma omit_redundant_points:
assumes "point p"
shows "p \<sqinter> x\<^sup>\<star> = (p \<sqinter> 1) \<squnion> (p \<sqinter> x) * (-p \<sqinter> x)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
let ?p = "p \<sqinter> 1"
have "?p * x * (-p \<sqinter> x)\<^sup>\<star> * ?p \<le> ?p * top * ?p"
by (metis comp_associative mult_left_isotone mult_right_isotone top.extremum)
also have "... \<le> ?p"
by (simp add: assms injective_codomain vector_inf_one_comp)
finally have "?p * x * (-p \<sqinter> x)\<^sup>\<star> * ?p * x \<le> ?p * x"
using mult_left_isotone by blast
hence "?p * x * (-p \<sqinter> x)\<^sup>\<star> * (p \<sqinter> x) \<le> ?p * x"
by (simp add: assms comp_associative vector_inf_one_comp)
also have 1: "... \<le> ?p * x * (-p \<sqinter> x)\<^sup>\<star>"
using mult_right_isotone star.circ_reflexive by fastforce
finally have "?p * x * (-p \<sqinter> x)\<^sup>\<star> * (p \<sqinter> x) \<squnion> ?p * x * (-p \<sqinter> x)\<^sup>\<star> * (-p \<sqinter> x) \<le> ?p * x * (-p \<sqinter> x)\<^sup>\<star>"
by (simp add: mult_right_isotone star.circ_plus_same star.left_plus_below_circ mult_assoc)
hence "?p * x * (-p \<sqinter> x)\<^sup>\<star> * ((p \<squnion> -p) \<sqinter> x) \<le> ?p * x * (-p \<sqinter> x)\<^sup>\<star>"
by (simp add: comp_inf.mult_right_dist_sup mult_left_dist_sup)
hence "?p * x * (-p \<sqinter> x)\<^sup>\<star> * x \<le> ?p * x * (-p \<sqinter> x)\<^sup>\<star>"
by (metis assms bijective_regular inf.absorb2 inf.cobounded1 inf.sup_monoid.add_commute shunting_p)
hence "?p * x * (-p \<sqinter> x)\<^sup>\<star> * x \<squnion> ?p * x \<le> ?p * x * (-p \<sqinter> x)\<^sup>\<star>"
using 1 by simp
hence "?p * (1 \<squnion> x * (-p \<sqinter> x)\<^sup>\<star>) * x \<le> ?p * x * (-p \<sqinter> x)\<^sup>\<star>"
by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup)
also have "... \<le> ?p * (1 \<squnion> x * (-p \<sqinter> x)\<^sup>\<star>)"
by (simp add: comp_associative mult_right_isotone)
finally have "?p * x\<^sup>\<star> \<le> ?p * (1 \<squnion> x * (-p \<sqinter> x)\<^sup>\<star>)"
using star_right_induct by (meson dual_order.trans le_supI mult_left_sub_dist_sup_left mult_sub_right_one)
also have "... = ?p \<squnion> ?p * x * (-p \<sqinter> x)\<^sup>\<star>"
by (simp add: comp_associative semiring.distrib_left)
finally show "p \<sqinter> x\<^sup>\<star> \<le> ?p \<squnion> (p \<sqinter> x) * (-p \<sqinter> x)\<^sup>\<star>"
by (simp add: assms vector_inf_one_comp)
show "?p \<squnion> (p \<sqinter> x) * (-p \<sqinter> x)\<^sup>\<star> \<le> p \<sqinter> x\<^sup>\<star>"
by (metis assms comp_isotone inf.boundedI inf.cobounded1 inf.coboundedI2 inf.sup_monoid.add_commute le_supI star.circ_increasing star.circ_transitive_equal star_isotone star_left_unfold_equal sup.cobounded1 vector_export_comp)
qed
text \<open>Weakly connected components\<close>
abbreviation "wcc x \<equiv> (x \<squnion> x\<^sup>T)\<^sup>\<star>"
text \<open>Theorem 5.1\<close>
lemma wcc_equivalence:
"equivalence (wcc x)"
apply (intro conjI)
subgoal by (simp add: star.circ_reflexive)
subgoal by (simp add: star.circ_transitive_equal)
subgoal by (simp add: conv_dist_sup conv_star_commute sup_commute)
done
text \<open>Theorem 5.2\<close>
lemma wcc_increasing:
"x \<le> wcc x"
by (simp add: star.circ_sub_dist_1)
lemma wcc_isotone:
"x \<le> y \<Longrightarrow> wcc x \<le> wcc y"
using conv_isotone star_isotone sup_mono by blast
lemma wcc_idempotent:
"wcc (wcc x) = wcc x"
using star_involutive wcc_equivalence by auto
text \<open>Theorem 5.3\<close>
lemma wcc_below_wcc:
"x \<le> wcc y \<Longrightarrow> wcc x \<le> wcc y"
using wcc_idempotent wcc_isotone by fastforce
text \<open>Theorem 5.4\<close>
lemma wcc_bot:
"wcc bot = 1"
by (simp add: star.circ_zero)
lemma wcc_one:
"wcc 1 = 1"
by (simp add: star_one)
text \<open>Theorem 5.5\<close>
lemma wcc_top:
"wcc top = top"
by (simp add: star.circ_top)
text \<open>Theorem 5.6\<close>
lemma wcc_with_loops:
"wcc x = wcc (x \<squnion> 1)"
by (metis conv_dist_sup star_decompose_1 star_sup_one sup_commute symmetric_one_closed)
lemma wcc_without_loops:
"wcc x = wcc (x \<sqinter> -1)"
by (metis conv_star_commute star_sum reachable_without_loops)
lemma forest_components_wcc:
"injective x \<Longrightarrow> wcc x = forest_components x"
by (simp add: cancel_separate_1)
text \<open>Components of a forest, which is represented using edges directed towards the roots\<close>
abbreviation "fc x \<equiv> x\<^sup>\<star> * x\<^sup>T\<^sup>\<star>"
text \<open>Theorem 2.1\<close>
lemma fc_equivalence:
"univalent x \<Longrightarrow> equivalence (fc x)"
apply (intro conjI)
subgoal by (simp add: reflexive_mult_closed star.circ_reflexive)
- subgoal by (metis cancel_separate_1 eq_iff star.circ_transitive_equal)
+ subgoal by (metis cancel_separate_1 order.eq_iff star.circ_transitive_equal)
subgoal by (simp add: conv_dist_comp conv_star_commute)
done
text \<open>Theorem 2.2\<close>
lemma fc_increasing:
"x \<le> fc x"
by (metis le_supE mult_left_isotone star.circ_back_loop_fixpoint star.circ_increasing)
text \<open>Theorem 2.3\<close>
lemma fc_isotone:
"x \<le> y \<Longrightarrow> fc x \<le> fc y"
by (simp add: comp_isotone conv_isotone star_isotone)
text \<open>Theorem 2.4\<close>
lemma fc_idempotent:
"univalent x \<Longrightarrow> fc (fc x) = fc x"
by (metis fc_equivalence cancel_separate_1 star.circ_transitive_equal star_involutive)
text \<open>Theorem 2.5\<close>
lemma fc_star:
"univalent x \<Longrightarrow> (fc x)\<^sup>\<star> = fc x"
using fc_equivalence fc_idempotent star.circ_transitive_equal by simp
lemma fc_plus:
"univalent x \<Longrightarrow> (fc x)\<^sup>+ = fc x"
by (metis fc_star star.circ_decompose_9)
text \<open>Theorem 2.6\<close>
lemma fc_bot:
"fc bot = 1"
by (simp add: star.circ_zero)
lemma fc_one:
"fc 1 = 1"
by (simp add: star_one)
text \<open>Theorem 2.7\<close>
lemma fc_top:
"fc top = top"
by (simp add: star.circ_top)
text \<open>Theorem 5.7\<close>
lemma fc_wcc:
"univalent x \<Longrightarrow> wcc x = fc x"
by (simp add: fc_star star_decompose_1)
lemma fc_via_root:
assumes "total (p\<^sup>\<star> * (p \<sqinter> 1))"
shows "fc p = p\<^sup>\<star> * (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "1 \<le> p\<^sup>\<star> * (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star>"
by (smt assms comp_associative conv_dist_comp conv_star_commute coreflexive_idempotent coreflexive_symmetric inf.cobounded2 total_var)
hence "fc p \<le> p\<^sup>\<star> * p\<^sup>\<star> * (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star> * p\<^sup>T\<^sup>\<star>"
by (metis comp_right_one mult_left_isotone mult_right_isotone mult_assoc)
thus "fc p \<le> p\<^sup>\<star> * (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star>"
by (simp add: star.circ_transitive_equal mult_assoc)
show "p\<^sup>\<star> * (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star> \<le> fc p"
by (metis comp_isotone inf.cobounded2 mult_1_right order.refl)
qed
text \<open>Theorem 4.1\<close>
lemma update_acyclic_1:
assumes "acyclic (p \<sqinter> -1)"
and "point y"
and "vector w"
and "w \<le> p\<^sup>\<star> * y"
shows "acyclic ((p[w\<longmapsto>y]) \<sqinter> -1)"
proof -
let ?p = "p[w\<longmapsto>y]"
have "w * y\<^sup>T \<le> p\<^sup>\<star>"
using assms(2,4) shunt_bijective by blast
hence "w * y\<^sup>T \<le> (p \<sqinter> -1)\<^sup>\<star>"
using reachable_without_loops by auto
hence "w * y\<^sup>T \<sqinter> -1 \<le> (p \<sqinter> -1)\<^sup>\<star> \<sqinter> -1"
by (simp add: inf.coboundedI2 inf.sup_monoid.add_commute)
also have "... \<le> (p \<sqinter> -1)\<^sup>+"
by (simp add: star_plus_without_loops)
finally have 1: "w \<sqinter> y\<^sup>T \<sqinter> -1 \<le> (p \<sqinter> -1)\<^sup>+"
using assms(2,3) vector_covector by auto
have "?p \<sqinter> -1 = (w \<sqinter> y\<^sup>T \<sqinter> -1) \<squnion> (-w \<sqinter> p \<sqinter> -1)"
by (simp add: inf_sup_distrib2)
also have "... \<le> (p \<sqinter> -1)\<^sup>+ \<squnion> (-w \<sqinter> p \<sqinter> -1)"
using 1 sup_left_isotone by blast
also have "... \<le> (p \<sqinter> -1)\<^sup>+ \<squnion> (p \<sqinter> -1)"
using comp_inf.mult_semi_associative sup_right_isotone by auto
also have "... = (p \<sqinter> -1)\<^sup>+"
by (metis star.circ_back_loop_fixpoint sup.right_idem)
finally have "(?p \<sqinter> -1)\<^sup>+ \<le> (p \<sqinter> -1)\<^sup>+"
by (metis comp_associative comp_isotone star.circ_transitive_equal star.left_plus_circ star_isotone)
also have "... \<le> -1"
using assms(1) by blast
finally show ?thesis
by simp
qed
lemma update_acyclic_2:
assumes "acyclic (p \<sqinter> -1)"
and "point y"
and "point x"
and "y \<le> p\<^sup>T\<^sup>\<star> * x"
and "univalent p"
and "p\<^sup>T * y \<le> y"
shows "acyclic ((p[p\<^sup>T\<^sup>\<star>*x\<longmapsto>y]) \<sqinter> -1)"
proof -
have "p\<^sup>T * p\<^sup>\<star> * y = p\<^sup>T * p * p\<^sup>\<star> * y \<squnion> p\<^sup>T * y"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint)
also have "... \<le> p\<^sup>\<star> * y"
by (metis assms(5,6) comp_right_one le_supI le_supI2 mult_left_isotone star.circ_loop_fixpoint star.circ_transitive_equal)
finally have "p\<^sup>T\<^sup>\<star> * x \<le> p\<^sup>\<star> * y"
by (simp add: assms(2-4) bijective_reverse conv_star_commute comp_associative star_left_induct)
thus ?thesis
by (simp add: assms(1-3) vector_mult_closed update_acyclic_1)
qed
lemma update_acyclic_3:
assumes "acyclic (p \<sqinter> -1)"
and "point y"
and "point w"
and "y \<le> p\<^sup>T\<^sup>\<star> * w"
shows "acyclic ((p[w\<longmapsto>y]) \<sqinter> -1)"
by (simp add: assms bijective_reverse conv_star_commute update_acyclic_1)
lemma rectangle_star_rectangle:
"rectangle a \<Longrightarrow> a * x\<^sup>\<star> * a \<le> a"
by (metis mult_left_isotone mult_right_isotone top.extremum)
lemma arc_star_arc:
"arc a \<Longrightarrow> a * x\<^sup>\<star> * a \<le> a"
using arc_top_arc rectangle_star_rectangle by blast
lemma star_rectangle_decompose:
assumes "rectangle a"
shows "(a \<squnion> x)\<^sup>\<star> = x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have 1: "1 \<le> x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
by (simp add: star.circ_reflexive sup.coboundedI1)
have "(a \<squnion> x) * (x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>) = a * x\<^sup>\<star> \<squnion> a * x\<^sup>\<star> * a * x\<^sup>\<star> \<squnion> x\<^sup>+ \<squnion> x\<^sup>+ * a * x\<^sup>\<star>"
by (metis comp_associative semiring.combine_common_factor semiring.distrib_left sup_commute)
also have "... = a * x\<^sup>\<star> \<squnion> x\<^sup>+ \<squnion> x\<^sup>+ * a * x\<^sup>\<star>"
using assms rectangle_star_rectangle by (simp add: mult_left_isotone sup_absorb1)
also have "... = x\<^sup>+ \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
by (metis comp_associative star.circ_loop_fixpoint sup_assoc sup_commute)
also have "... \<le> x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
using star.left_plus_below_circ sup_left_isotone by auto
finally show "(a \<squnion> x)\<^sup>\<star> \<le> x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
using 1 by (metis comp_right_one le_supI star_left_induct)
next
show "x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star> \<le> (a \<squnion> x)\<^sup>\<star>"
by (metis comp_isotone le_supE le_supI star.circ_increasing star.circ_transitive_equal star_isotone sup_ge2)
qed
lemma star_arc_decompose:
"arc a \<Longrightarrow> (a \<squnion> x)\<^sup>\<star> = x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
using arc_top_arc star_rectangle_decompose by blast
lemma plus_rectangle_decompose:
assumes "rectangle a"
shows "(a \<squnion> x)\<^sup>+ = x\<^sup>+ \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
proof -
have "(a \<squnion> x)\<^sup>+ = (a \<squnion> x) * (x\<^sup>\<star> \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>)"
by (simp add: assms star_rectangle_decompose)
also have "... = a * x\<^sup>\<star> \<squnion> a * x\<^sup>\<star> * a * x\<^sup>\<star> \<squnion> x\<^sup>+ \<squnion> x\<^sup>+ * a * x\<^sup>\<star>"
by (metis comp_associative semiring.combine_common_factor semiring.distrib_left sup_commute)
also have "... = a * x\<^sup>\<star> \<squnion> x\<^sup>+ \<squnion> x\<^sup>+ * a * x\<^sup>\<star>"
using assms rectangle_star_rectangle by (simp add: mult_left_isotone sup_absorb1)
also have "... = x\<^sup>+ \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
by (metis comp_associative star.circ_loop_fixpoint sup_assoc sup_commute)
finally show ?thesis
by simp
qed
text \<open>Theorem 6.1\<close>
lemma plus_arc_decompose:
"arc a \<Longrightarrow> (a \<squnion> x)\<^sup>+ = x\<^sup>+ \<squnion> x\<^sup>\<star> * a * x\<^sup>\<star>"
using arc_top_arc plus_rectangle_decompose by blast
text \<open>Theorem 6.2\<close>
lemma update_acyclic_4:
assumes "acyclic (p \<sqinter> -1)"
and "point y"
and "point w"
and "y \<sqinter> p\<^sup>\<star> * w = bot"
shows "acyclic ((p[w\<longmapsto>y]) \<sqinter> -1)"
proof -
let ?p = "p[w\<longmapsto>y]"
have "y\<^sup>T * p\<^sup>\<star> * w \<le> -1"
using assms(4) comp_associative pseudo_complement schroeder_3_p by auto
hence 1: "p\<^sup>\<star> * w * y\<^sup>T * p\<^sup>\<star> \<le> -1"
by (metis comp_associative comp_commute_below_diversity star.circ_transitive_equal)
have "?p \<sqinter> -1 \<le> (w \<sqinter> y\<^sup>T) \<squnion> (p \<sqinter> -1)"
by (metis comp_inf.mult_right_dist_sup dual_order.trans inf.cobounded1 inf.coboundedI2 inf.sup_monoid.add_assoc le_supI sup.cobounded1 sup_ge2)
also have "... = w * y\<^sup>T \<squnion> (p \<sqinter> -1)"
using assms(2,3) by (simp add: vector_covector)
finally have "(?p \<sqinter> -1)\<^sup>+ \<le> (w * y\<^sup>T \<squnion> (p \<sqinter> -1))\<^sup>+"
by (simp add: comp_isotone star_isotone)
also have "... = (p \<sqinter> -1)\<^sup>+ \<squnion> (p \<sqinter> -1)\<^sup>\<star> * w * y\<^sup>T * (p \<sqinter> -1)\<^sup>\<star>"
using assms(2,3) plus_arc_decompose points_arc by (simp add: comp_associative)
also have "... \<le> (p \<sqinter> -1)\<^sup>+ \<squnion> p\<^sup>\<star> * w * y\<^sup>T * p\<^sup>\<star>"
using reachable_without_loops by auto
also have "... \<le> -1"
using 1 assms(1) by simp
finally show ?thesis
by simp
qed
text \<open>Theorem 6.3\<close>
lemma update_acyclic_5:
assumes "acyclic (p \<sqinter> -1)"
and "point w"
shows "acyclic ((p[w\<longmapsto>w]) \<sqinter> -1)"
proof -
let ?p = "p[w\<longmapsto>w]"
have "?p \<sqinter> -1 \<le> (w \<sqinter> w\<^sup>T \<sqinter> -1) \<squnion> (p \<sqinter> -1)"
by (metis comp_inf.mult_right_dist_sup inf.cobounded2 inf.sup_monoid.add_assoc sup_right_isotone)
also have "... = p \<sqinter> -1"
using assms(2) by (metis comp_inf.covector_complement_closed equivalence_top_closed inf_top.right_neutral maddux_3_13 pseudo_complement regular_closed_top regular_one_closed vector_covector vector_top_closed)
finally show ?thesis
using assms(1) acyclic_down_closed by blast
qed
text \<open>Root of the tree containing point $x$ in the disjoint-set forest $p$\<close>
abbreviation "root p x \<equiv> p\<^sup>T\<^sup>\<star> * x \<sqinter> (p \<sqinter> 1) * top"
text \<open>Theorem 3.1\<close>
lemma root_var:
"root p x = (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star> * x"
by (simp add: coreflexive_comp_top_inf inf_commute mult_assoc)
text \<open>Theorem 3.2\<close>
lemma root_successor_loop:
"univalent p \<Longrightarrow> root p x = p[[root p x]]"
by (metis root_var injective_codomain comp_associative conv_dist_inf coreflexive_symmetric equivalence_one_closed inf.cobounded2 univalent_conv_injective)
lemma root_transitive_successor_loop:
"univalent p \<Longrightarrow> root p x = p\<^sup>T\<^sup>\<star> * (root p x)"
by (metis mult_1_right star_one star_simulation_right_equal root_successor_loop)
text \<open>The root of a tree of a node belongs to the same component as the node.\<close>
lemma root_same_component:
"injective x \<Longrightarrow> root p x * x\<^sup>T \<le> fc p"
by (metis comp_associative coreflexive_comp_top_inf eq_refl inf.sup_left_divisibility inf.sup_monoid.add_commute mult_isotone star.circ_circ_mult star.circ_right_top star.circ_transitive_equal star_one star_outer_increasing test_preserves_equation top_greatest)
lemma root_vector:
"vector x \<Longrightarrow> vector (root p x)"
by (simp add: vector_mult_closed root_var)
lemma root_vector_inf:
"vector x \<Longrightarrow> root p x * x\<^sup>T = root p x \<sqinter> x\<^sup>T"
by (simp add: vector_covector root_vector)
lemma root_same_component_vector:
"injective x \<Longrightarrow> vector x \<Longrightarrow> root p x \<sqinter> x\<^sup>T \<le> fc p"
using root_same_component root_vector_inf by fastforce
lemma univalent_root_successors:
assumes "univalent p"
shows "(p \<sqinter> 1) * p\<^sup>\<star> = p \<sqinter> 1"
-proof (rule antisym)
+proof (rule order.antisym)
have "(p \<sqinter> 1) * p \<le> p \<sqinter> 1"
by (smt assms(1) comp_inf.mult_semi_associative conv_dist_comp conv_dist_inf conv_order equivalence_one_closed inf.absorb1 inf.sup_monoid.add_assoc injective_codomain)
thus "(p \<sqinter> 1) * p\<^sup>\<star> \<le> p \<sqinter> 1"
using star_right_induct_mult by blast
show "p \<sqinter> 1 \<le> (p \<sqinter> 1) * p\<^sup>\<star>"
by (metis coreflexive_idempotent inf_le1 inf_le2 mult_right_isotone order_trans star.circ_increasing)
qed
lemma same_component_same_root_sub:
assumes "univalent p"
and "bijective y"
and "x * y\<^sup>T \<le> fc p"
shows "root p x \<le> root p y"
proof -
have "root p x * y\<^sup>T \<le> (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star>"
- by (smt assms(1,3) mult_isotone mult_assoc root_var fc_plus fc_star eq_iff univalent_root_successors)
+ by (smt assms(1,3) mult_isotone mult_assoc root_var fc_plus fc_star order.eq_iff univalent_root_successors)
thus ?thesis
by (simp add: assms(2) shunt_bijective root_var)
qed
lemma same_component_same_root:
assumes "univalent p"
and "bijective x"
and "bijective y"
and "x * y\<^sup>T \<le> fc p"
shows "root p x = root p y"
-proof (rule antisym)
+proof (rule order.antisym)
show "root p x \<le> root p y"
using assms(1,3,4) same_component_same_root_sub by blast
have "y * x\<^sup>T \<le> fc p"
using assms(1,4) fc_equivalence conv_dist_comp conv_isotone by fastforce
thus "root p y \<le> root p x"
using assms(1,2) same_component_same_root_sub by blast
qed
lemma same_roots_sub:
assumes "univalent q"
and "p \<sqinter> 1 \<le> q \<sqinter> 1"
and "fc p \<le> fc q"
shows "p\<^sup>\<star> * (p \<sqinter> 1) \<le> q\<^sup>\<star> * (q \<sqinter> 1)"
proof -
have "p\<^sup>\<star> * (p \<sqinter> 1) \<le> p\<^sup>\<star> * (q \<sqinter> 1)"
using assms(2) mult_right_isotone by auto
also have "... \<le> fc p * (q \<sqinter> 1)"
using mult_left_isotone mult_right_isotone star.circ_reflexive by fastforce
also have "... \<le> fc q * (q \<sqinter> 1)"
by (simp add: assms(3) mult_left_isotone)
also have "... = q\<^sup>\<star> * (q \<sqinter> 1)"
by (metis assms(1) conv_dist_comp conv_dist_inf conv_star_commute inf_commute one_inf_conv symmetric_one_closed mult_assoc univalent_root_successors)
finally show ?thesis
.
qed
lemma same_roots:
assumes "univalent p"
and "univalent q"
and "p \<sqinter> 1 = q \<sqinter> 1"
and "fc p = fc q"
shows "p\<^sup>\<star> * (p \<sqinter> 1) = q\<^sup>\<star> * (q \<sqinter> 1)"
by (smt assms conv_dist_comp conv_dist_inf conv_involutive conv_star_commute inf_commute one_inf_conv symmetric_one_closed root_var univalent_root_successors)
lemma same_root:
assumes "univalent p"
and "univalent q"
and "p \<sqinter> 1 = q \<sqinter> 1"
and "fc p = fc q"
shows "root p x = root q x"
by (metis assms mult_assoc root_var univalent_root_successors)
lemma loop_root:
assumes "injective x"
and "x = p[[x]]"
shows "x = root p x"
-proof (rule antisym)
+proof (rule order.antisym)
have "x \<le> p * x"
by (metis assms comp_associative comp_right_one conv_order equivalence_one_closed ex231c inf.orderE inf.sup_monoid.add_commute mult_left_isotone mult_right_isotone one_inf_conv)
hence "x = (p \<sqinter> 1) * x"
by (simp add: assms(1) inf_absorb2 injective_comp_right_dist_inf)
thus "x \<le> root p x"
by (metis assms(2) coreflexive_comp_top_inf inf.boundedI inf.cobounded1 inf.cobounded2 mult_isotone star.circ_increasing)
next
show "root p x \<le> x"
using assms(2) le_infI1 star_left_induct_mult by auto
qed
lemma one_loop:
assumes "acyclic (p \<sqinter> -1)"
and "univalent p"
shows "(p \<sqinter> 1) * (p\<^sup>T \<sqinter> -1)\<^sup>+ * (p \<sqinter> 1) = bot"
proof -
have "p\<^sup>T\<^sup>+ \<sqinter> (p \<sqinter> 1) * top * (p \<sqinter> 1) = (p \<sqinter> 1) * p\<^sup>T\<^sup>+ * (p \<sqinter> 1)"
by (simp add: test_comp_test_top)
also have "... \<le> p\<^sup>T\<^sup>\<star> * (p \<sqinter> 1)"
by (simp add: inf.coboundedI2 mult_left_isotone star.circ_mult_upper_bound star.circ_reflexive star.left_plus_below_circ)
also have "... = p \<sqinter> 1"
by (metis assms(2) conv_dist_comp conv_dist_inf conv_star_commute inf_commute one_inf_conv symmetric_one_closed univalent_root_successors)
also have "... \<le> 1"
by simp
finally have "(p \<sqinter> 1) * top * (p \<sqinter> 1) \<le> -(p\<^sup>T\<^sup>+ \<sqinter> -1)"
using p_antitone p_antitone_iff p_shunting_swap by blast
hence "(p \<sqinter> 1)\<^sup>T * (p\<^sup>T\<^sup>+ \<sqinter> -1) * (p \<sqinter> 1)\<^sup>T \<le> bot"
using triple_schroeder_p p_top by blast
hence "(p \<sqinter> 1) * (p\<^sup>T\<^sup>+ \<sqinter> -1) * (p \<sqinter> 1) = bot"
by (simp add: coreflexive_symmetric le_bot)
thus ?thesis
by (smt assms(1) conv_complement conv_dist_comp conv_dist_inf conv_star_commute inf_absorb1 star.circ_plus_same symmetric_one_closed reachable_without_loops star_plus_without_loops)
qed
lemma root_root:
"root p x = root p (root p x)"
by (smt comp_associative comp_inf.mult_right_sub_dist_sup_right dual_order.eq_iff inf.cobounded1 inf.cobounded2 inf.orderE mult_right_isotone star.circ_loop_fixpoint star.circ_transitive_equal root_var)
lemma loop_root_2:
assumes "acyclic (p \<sqinter> -1)"
and "univalent p"
and "injective x"
and "x \<le> p\<^sup>T\<^sup>+ * x"
shows "x = root p x"
-proof (rule antisym)
+proof (rule order.antisym)
have 1: "x = x \<sqinter> -(-1 * x)"
by (metis assms(3) comp_injective_below_complement inf.orderE mult_1_left regular_one_closed)
have "x \<le> (p\<^sup>T \<sqinter> -1)\<^sup>+ * x \<squnion> (p \<sqinter> 1) * x"
by (metis assms(4) inf_commute mult_right_dist_sup one_inf_conv plus_reachable_without_loops)
also have "... \<le> -1 * x \<squnion> (p \<sqinter> 1) * x"
by (metis assms(1) conv_complement conv_dist_inf conv_isotone conv_plus_commute mult_left_isotone semiring.add_right_mono symmetric_one_closed)
also have "... \<le> -1 * x \<squnion> root p x"
using comp_isotone inf.coboundedI2 star.circ_reflexive sup_right_isotone by auto
finally have "x \<le> (-1 * x \<squnion> root p x) \<sqinter> -(-1 * x)"
using 1 inf.boundedI inf.order_iff by blast
also have "... \<le> root p x"
using inf.sup_left_divisibility by auto
finally show 2: "x \<le> root p x"
.
have "root p x = (p \<sqinter> 1) * x \<squnion> (p \<sqinter> 1) * (p\<^sup>T \<sqinter> -1)\<^sup>+ * x"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute reachable_without_loops root_var)
also have "... \<le> x \<squnion> (p \<sqinter> 1) * (p\<^sup>T \<sqinter> -1)\<^sup>+ * root p x"
using 2 by (metis coreflexive_comp_top_inf inf.cobounded2 mult_right_isotone semiring.add_mono)
also have "... = x"
by (metis assms(1,2) one_loop root_var mult_assoc semiring.mult_not_zero sup_bot_right)
finally show "root p x \<le> x"
.
qed
end
context stone_relation_algebra_tarski
begin
text \<open>Theorem 4.3 \<open>distinct_points\<close> has been moved to theory \<open>Relation_Algebras\<close> in entry \<open>Stone_Relation_Algebras\<close>\<close>
text \<open>Back and von Wright's array independence requirements \cite{BackWright1998}\<close>
lemma put_get_different_vector:
assumes "vector y" "vector w" "w \<le> -y"
shows "(x[y\<longmapsto>z])[[w]] = x[[w]]"
proof -
have "(x[y\<longmapsto>z])[[w]] = (y\<^sup>T \<sqinter> z) * w \<squnion> (-y\<^sup>T \<sqinter> x\<^sup>T) * w"
by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup)
also have "... = z * (w \<sqinter> y) \<squnion> x\<^sup>T * (w \<sqinter> -y)"
by (metis assms(1) conv_complement covector_inf_comp_3 inf_commute vector_complement_closed)
also have "... = z * (w \<sqinter> y) \<squnion> x\<^sup>T * w"
by (simp add: assms(3) inf.absorb1)
also have "... = z * bot \<squnion> x\<^sup>T * w"
by (metis assms(3) comp_inf.semiring.mult_zero_right inf.absorb1 inf.sup_monoid.add_assoc p_inf)
also have "... = x\<^sup>T * w"
by simp
finally show ?thesis
.
qed
lemma put_get_different:
assumes "point y" "point w" "w \<noteq> y"
shows "(x[y\<longmapsto>z])[[w]] = x[[w]]"
proof -
have "w \<sqinter> y = bot"
using assms distinct_points by simp
hence "w \<le> -y"
using pseudo_complement by simp
thus ?thesis
by (simp add: assms(1) assms(2) put_get_different_vector)
qed
lemma put_put_different_vector:
assumes "vector y" "vector v" "v \<sqinter> y = bot"
shows "(x[y\<longmapsto>z])[v\<longmapsto>w] = (x[v\<longmapsto>w])[y\<longmapsto>z]"
proof -
have "(x[y\<longmapsto>z])[v\<longmapsto>w] = (v \<sqinter> w\<^sup>T) \<squnion> (-v \<sqinter> y \<sqinter> z\<^sup>T) \<squnion> (-v \<sqinter> -y \<sqinter> x)"
by (simp add: comp_inf.semiring.distrib_left inf_assoc sup_assoc)
also have "... = (v \<sqinter> w\<^sup>T) \<squnion> (y \<sqinter> z\<^sup>T) \<squnion> (-v \<sqinter> -y \<sqinter> x)"
by (metis assms(3) inf_commute inf_import_p p_inf selection_closed_id)
also have "... = (y \<sqinter> z\<^sup>T) \<squnion> (v \<sqinter> w\<^sup>T) \<squnion> (-y \<sqinter> -v \<sqinter> x)"
by (simp add: inf_commute sup_commute)
also have "... = (y \<sqinter> z\<^sup>T) \<squnion> (-y \<sqinter> v \<sqinter> w\<^sup>T) \<squnion> (-y \<sqinter> -v \<sqinter> x)"
using assms distinct_points pseudo_complement inf.absorb2 by simp
also have "... = (x[v\<longmapsto>w])[y\<longmapsto>z]"
by (simp add: comp_inf.semiring.distrib_left inf_assoc sup_assoc)
finally show ?thesis
.
qed
lemma put_put_different:
assumes "point y" "point v" "v \<noteq> y"
shows "(x[y\<longmapsto>z])[v\<longmapsto>w] = (x[v\<longmapsto>w])[y\<longmapsto>z]"
using assms distinct_points put_put_different_vector by blast
end
section \<open>Verifying Operations on Disjoint-Set Forests\<close>
text \<open>
In this section we verify the make-set, find-set and union-sets operations of disjoint-set forests.
We start by introducing syntax for updating arrays in programs.
Updating the value at a given array index means updating the whole array.
\<close>
syntax
"_rel_update" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b com" ("(2_[_] :=/ _)" [70, 65, 65] 61)
translations
"x[y] := z" => "(x := (y \<sqinter> z\<^sup>T) \<squnion> (CONST uminus y \<sqinter> x))"
text \<open>
The finiteness requirement in the following class is used for proving that the operations terminate.
\<close>
class finite_regular_p_algebra = p_algebra +
assumes finite_regular: "finite { x . regular x }"
class stone_kleene_relation_algebra_tarski = stone_kleene_relation_algebra + stone_relation_algebra_tarski
class stone_kleene_relation_algebra_tarski_finite_regular = stone_kleene_relation_algebra_tarski + finite_regular_p_algebra
begin
subsection \<open>Make-Set\<close>
text \<open>
We prove two correctness results about make-set.
The first shows that the forest changes only to the extent of making one node the root of a tree.
The second result adds that only singleton sets are created.
\<close>
definition "make_set_postcondition p x p0 \<equiv> x \<sqinter> p = x * x\<^sup>T \<and> -x \<sqinter> p = -x \<sqinter> p0"
theorem make_set:
"VARS p
[ point x \<and> p0 = p ]
p[x] := x
[ make_set_postcondition p x p0 ]"
apply vcg_tc_simp
by (simp add: make_set_postcondition_def inf_sup_distrib1 inf_assoc[THEN sym] vector_covector[THEN sym])
theorem make_set_2:
"VARS p
[ point x \<and> p0 = p \<and> p \<le> 1 ]
p[x] := x
[ make_set_postcondition p x p0 \<and> p \<le> 1 ]"
proof vcg_tc
fix p
assume 1: "point x \<and> p0 = p \<and> p \<le> 1"
show "make_set_postcondition (p[x\<longmapsto>x]) x p0 \<and> p[x\<longmapsto>x] \<le> 1"
proof (rule conjI)
show "make_set_postcondition (p[x\<longmapsto>x]) x p0"
using 1 by (simp add: make_set_postcondition_def inf_sup_distrib1 inf_assoc[THEN sym] vector_covector[THEN sym])
show "p[x\<longmapsto>x] \<le> 1"
using 1 by (metis coreflexive_sup_closed dual_order.trans inf.cobounded2 vector_covector)
qed
qed
text \<open>
The above total-correctness proof allows us to extract a function, which can be used in other implementations below.
This is a technique of \cite{Guttmann2018c}.
\<close>
lemma make_set_exists:
"point x \<Longrightarrow> \<exists>p' . make_set_postcondition p' x p"
using tc_extract_function make_set by blast
definition "make_set p x \<equiv> (SOME p' . make_set_postcondition p' x p)"
lemma make_set_function:
assumes "point x"
and "p' = make_set p x"
shows "make_set_postcondition p' x p"
proof -
let ?P = "\<lambda>p' . make_set_postcondition p' x p"
have "?P (SOME z . ?P z)"
using assms(1) make_set_exists by (meson someI)
thus ?thesis
using assms(2) make_set_def by auto
qed
subsection \<open>Find-Set\<close>
text \<open>
Disjoint-set forests are represented by their parent mapping.
It is a forest except each root of a component tree points to itself.
We prove that find-set returns the root of the component tree of the given node.
\<close>
abbreviation "disjoint_set_forest p \<equiv> mapping p \<and> acyclic (p \<sqinter> -1)"
definition "find_set_precondition p x \<equiv> disjoint_set_forest p \<and> point x"
definition "find_set_invariant p x y \<equiv> find_set_precondition p x \<and> point y \<and> y \<le> p\<^sup>T\<^sup>\<star> * x"
definition "find_set_postcondition p x y \<equiv> point y \<and> y = root p x"
lemma find_set_1:
"find_set_precondition p x \<Longrightarrow> find_set_invariant p x x"
apply (unfold find_set_invariant_def)
using mult_left_isotone star.circ_reflexive find_set_precondition_def by fastforce
lemma find_set_2:
"find_set_invariant p x y \<and> y \<noteq> p[[y]] \<Longrightarrow> find_set_invariant p x (p[[y]]) \<and> card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * (p[[y]]) } < card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * y }"
proof -
let ?s = "{ z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * y }"
let ?t = "{ z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * (p[[y]]) }"
assume 1: "find_set_invariant p x y \<and> y \<noteq> p[[y]]"
have 2: "point (p[[y]])"
using 1 read_point find_set_invariant_def find_set_precondition_def by simp
show "find_set_invariant p x (p[[y]]) \<and> card ?t < card ?s"
proof (unfold find_set_invariant_def, intro conjI)
show "find_set_precondition p x"
using 1 find_set_invariant_def by simp
show "vector (p[[y]])"
using 2 by simp
show "injective (p[[y]])"
using 2 by simp
show "surjective (p[[y]])"
using 2 by simp
show "p[[y]] \<le> p\<^sup>T\<^sup>\<star> * x"
using 1 by (metis (hide_lams) find_set_invariant_def comp_associative comp_isotone star.circ_increasing star.circ_transitive_equal)
show "card ?t < card ?s"
proof -
have 3: "(p\<^sup>T \<sqinter> -1) * (p\<^sup>T \<sqinter> -1)\<^sup>+ * y \<le> (p\<^sup>T \<sqinter> -1)\<^sup>+ * y"
by (simp add: mult_left_isotone mult_right_isotone star.left_plus_below_circ)
have "p[[y]] = (p\<^sup>T \<sqinter> 1) * y \<squnion> (p\<^sup>T \<sqinter> -1) * y"
by (metis maddux_3_11_pp mult_right_dist_sup regular_one_closed)
also have "... \<le> ((p[[y]]) \<sqinter> y) \<squnion> (p\<^sup>T \<sqinter> -1) * y"
by (metis comp_left_subdist_inf mult_1_left semiring.add_right_mono)
also have "... = (p\<^sup>T \<sqinter> -1) * y"
using 1 2 find_set_invariant_def distinct_points by auto
finally have 4: "(p\<^sup>T \<sqinter> -1)\<^sup>\<star> * (p[[y]]) \<le> (p\<^sup>T \<sqinter> -1)\<^sup>+ * y"
using 3 by (metis inf.antisym_conv inf.eq_refl inf_le1 mult_left_isotone star_plus mult_assoc)
hence "p\<^sup>T\<^sup>\<star> * (p[[y]]) \<le> p\<^sup>T\<^sup>\<star> * y"
by (metis mult_isotone order_refl star.left_plus_below_circ star_plus mult_assoc)
hence 5: "?t \<subseteq> ?s"
using order_trans by auto
have 6: "y \<in> ?s"
using 1 find_set_invariant_def bijective_regular mult_left_isotone star.circ_reflexive by fastforce
have 7: "\<not> y \<in> ?t"
proof
assume "y \<in> ?t"
hence "y \<le> (p\<^sup>T \<sqinter> -1)\<^sup>+ * y"
using 4 by (metis reachable_without_loops mem_Collect_eq order_trans)
hence "y * y\<^sup>T \<le> (p\<^sup>T \<sqinter> -1)\<^sup>+"
using 1 find_set_invariant_def shunt_bijective by simp
also have "... \<le> -1"
using 1 by (metis (mono_tags, lifting) find_set_invariant_def find_set_precondition_def conv_dist_comp conv_dist_inf conv_isotone conv_star_commute equivalence_one_closed star.circ_plus_same symmetric_complement_closed)
finally have "y \<le> -y"
using schroeder_4_p by auto
thus False
using 1 by (metis find_set_invariant_def comp_inf.coreflexive_idempotent conv_complement covector_vector_comp inf.absorb1 inf.sup_monoid.add_commute pseudo_complement surjective_conv_total top.extremum vector_top_closed regular_closed_top)
qed
show "card ?t < card ?s"
apply (rule psubset_card_mono)
subgoal using finite_regular by simp
subgoal using 5 6 7 by auto
done
qed
qed
qed
lemma find_set_3:
"find_set_invariant p x y \<and> y = p[[y]] \<Longrightarrow> find_set_postcondition p x y"
proof -
assume 1: "find_set_invariant p x y \<and> y = p[[y]]"
show "find_set_postcondition p x y"
proof (unfold find_set_postcondition_def, rule conjI)
show "point y"
using 1 find_set_invariant_def by simp
show "y = root p x"
- proof (rule antisym)
+ proof (rule order.antisym)
have "y * y\<^sup>T \<le> p"
using 1 by (metis find_set_invariant_def find_set_precondition_def shunt_bijective shunt_mapping top_right_mult_increasing)
hence "y * y\<^sup>T \<le> p \<sqinter> 1"
using 1 find_set_invariant_def le_infI by blast
hence "y \<le> (p \<sqinter> 1) * top"
using 1 by (metis find_set_invariant_def order_lesseq_imp shunt_bijective top_right_mult_increasing mult_assoc)
thus "y \<le> root p x"
using 1 find_set_invariant_def by simp
next
have 2: "x \<le> p\<^sup>\<star> * y"
using 1 find_set_invariant_def find_set_precondition_def bijective_reverse conv_star_commute by auto
have "p\<^sup>T * p\<^sup>\<star> * y = p\<^sup>T * p * p\<^sup>\<star> * y \<squnion> (p[[y]])"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint)
also have "... \<le> p\<^sup>\<star> * y \<squnion> y"
using 1 by (metis find_set_invariant_def find_set_precondition_def comp_isotone mult_left_sub_dist_sup semiring.add_right_mono star.circ_back_loop_fixpoint star.circ_circ_mult star.circ_top star.circ_transitive_equal star_involutive star_one)
also have "... = p\<^sup>\<star> * y"
by (metis star.circ_loop_fixpoint sup.left_idem sup_commute)
finally have 3: "p\<^sup>T\<^sup>\<star> * x \<le> p\<^sup>\<star> * y"
using 2 by (simp add: comp_associative star_left_induct)
have "p * y \<sqinter> (p \<sqinter> 1) * top = (p \<sqinter> 1) * p * y"
using comp_associative coreflexive_comp_top_inf inf_commute by auto
also have "... \<le> p\<^sup>T * p * y"
by (metis inf.cobounded2 inf.sup_monoid.add_commute mult_left_isotone one_inf_conv)
also have "... \<le> y"
using 1 find_set_invariant_def find_set_precondition_def mult_left_isotone by fastforce
finally have 4: "p * y \<le> y \<squnion> -((p \<sqinter> 1) * top)"
using 1 by (metis find_set_invariant_def shunting_p bijective_regular)
have "p\<^sup>T * (p \<sqinter> 1) \<le> p\<^sup>T \<sqinter> 1"
using 1 by (metis find_set_invariant_def find_set_precondition_def N_top comp_isotone coreflexive_idempotent inf.cobounded2 inf.sup_monoid.add_commute inf_assoc one_inf_conv shunt_mapping)
hence "p\<^sup>T * (p \<sqinter> 1) * top \<le> (p \<sqinter> 1) * top"
using inf_commute mult_isotone one_inf_conv by auto
hence "p * -((p \<sqinter> 1) * top) \<le> -((p \<sqinter> 1) * top)"
by (metis comp_associative inf.sup_monoid.add_commute p_antitone p_antitone_iff schroeder_3_p)
hence "p * y \<squnion> p * -((p \<sqinter> 1) * top) \<le> y \<squnion> -((p \<sqinter> 1) * top)"
using 4 dual_order.trans le_supI sup_ge2 by blast
hence "p * (y \<squnion> -((p \<sqinter> 1) * top)) \<le> y \<squnion> -((p \<sqinter> 1) * top)"
by (simp add: mult_left_dist_sup)
hence "p\<^sup>\<star> * y \<le> y \<squnion> -((p \<sqinter> 1) * top)"
by (simp add: star_left_induct)
hence "p\<^sup>T\<^sup>\<star> * x \<le> y \<squnion> -((p \<sqinter> 1) * top)"
using 3 dual_order.trans by blast
thus "root p x \<le> y"
using 1 by (metis find_set_invariant_def shunting_p bijective_regular)
qed
qed
qed
theorem find_set:
"VARS y
[ find_set_precondition p x ]
y := x;
WHILE y \<noteq> p[[y]]
INV { find_set_invariant p x y }
VAR { card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * y } }
DO y := p[[y]]
OD
[ find_set_postcondition p x y ]"
apply vcg_tc_simp
apply (fact find_set_1)
apply (fact find_set_2)
by (fact find_set_3)
lemma find_set_exists:
"find_set_precondition p x \<Longrightarrow> \<exists>y . find_set_postcondition p x y"
using tc_extract_function find_set by blast
text \<open>
The root of a component tree is a point, that is, represents a singleton set of nodes.
This could be proved from the definitions using Kleene-relation algebraic calculations.
But they can be avoided because the property directly follows from the postcondition of the previous correctness proof.
The corresponding algorithm shows how to obtain the root.
We therefore have an essentially constructive proof of the following result.
\<close>
text \<open>Theorem 3.3\<close>
lemma root_point:
"disjoint_set_forest p \<Longrightarrow> point x \<Longrightarrow> point (root p x)"
using find_set_exists find_set_precondition_def find_set_postcondition_def by simp
definition "find_set p x \<equiv> (SOME y . find_set_postcondition p x y)"
lemma find_set_function:
assumes "find_set_precondition p x"
and "y = find_set p x"
shows "find_set_postcondition p x y"
by (metis assms find_set_def find_set_exists someI)
subsection \<open>Path Compression\<close>
text \<open>
The path-compression technique is frequently implemented in recursive implementations of find-set
modifying the tree on the way out from recursive calls. Here we implement it using a second while-loop,
which iterates over the same path to the root and changes edges to point to the root of the component,
which is known after the while-loop in find-set completes. We prove that path compression preserves
the equivalence-relational semantics of the disjoint-set forest and also preserves the roots of the
component trees. Additionally we prove the exact effect of path compression.
\<close>
definition "path_compression_precondition p x y \<equiv> disjoint_set_forest p \<and> point x \<and> point y \<and> y = root p x"
definition "path_compression_invariant p x y p0 w \<equiv>
path_compression_precondition p x y \<and> point w \<and> y \<le> p\<^sup>T\<^sup>\<star> * w \<and>
(w \<noteq> x \<longrightarrow> p[[x]] = y \<and> y \<noteq> x \<and> p\<^sup>T\<^sup>+ * w \<le> -x) \<and> p \<sqinter> 1 = p0 \<sqinter> 1 \<and> fc p = fc p0 \<and>
root p w = y \<and> (w \<noteq> y \<longrightarrow> p[[w]] \<noteq> w \<and> p\<^sup>T\<^sup>+ * w \<le> -w) \<and> p[[w]] = p0[[w]] \<and> p0[p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)\<longmapsto>y] = p \<and>
disjoint_set_forest p0 \<and> y = root p0 x \<and> w \<le> p0\<^sup>T\<^sup>\<star> * x"
definition "path_compression_postcondition p x y p0 \<equiv>
path_compression_precondition p x y \<and> p \<sqinter> 1 = p0 \<sqinter> 1 \<and> fc p = fc p0 \<and>
p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y] = p"
text \<open>
We first consider a variant that achieves the effect as a single update.
The parents of all nodes reachable from x are simultaneously updated to the root of the component of x.
\<close>
lemma path_compression_exact:
assumes "path_compression_precondition p0 x y"
and "p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y] = p"
shows "p \<sqinter> 1 = p0 \<sqinter> 1" "fc p = fc p0"
proof -
have a1: "disjoint_set_forest p0" and a2: "point x" and a3: "point y" and a4: "y = root p0 x"
using path_compression_precondition_def assms(1) by auto
have 1: "regular (p0\<^sup>T\<^sup>\<star> * x)"
using a1 a2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
have "p \<sqinter> 1 = (p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T \<sqinter> 1) \<squnion> (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0 \<sqinter> 1)"
using assms(2) inf_sup_distrib2 by auto
also have "... = (p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 \<sqinter> 1) \<squnion> (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0 \<sqinter> 1)"
proof -
have "p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T \<sqinter> 1 = p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 \<sqinter> 1"
- proof (rule antisym)
+ proof (rule order.antisym)
have "(p0 \<sqinter> 1) * p0\<^sup>T\<^sup>\<star> * x \<sqinter> 1 \<le> p0"
by (smt coreflexive_comp_top_inf_one inf.absorb_iff2 inf.cobounded2 inf.sup_monoid.add_assoc root_var)
hence "p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T \<sqinter> 1 \<le> p0"
by (metis inf_le1 a4 conv_dist_inf coreflexive_symmetric inf.absorb2 inf.cobounded2 inf.sup_monoid.add_assoc root_var symmetric_one_closed)
thus "p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T \<sqinter> 1 \<le> p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 \<sqinter> 1"
by (meson inf.le_sup_iff order.refl)
have "p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 \<sqinter> 1 \<le> y"
by (metis a4 coreflexive_comp_top_inf_one inf.cobounded1 inf_assoc inf_le2)
thus "p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 \<sqinter> 1 \<le> p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T \<sqinter> 1"
by (smt conv_dist_inf coreflexive_symmetric inf.absorb_iff2 inf.cobounded2 inf.sup_monoid.add_assoc)
qed
thus ?thesis
by simp
qed
also have "... = p0 \<sqinter> 1"
using 1 by (metis inf.sup_monoid.add_commute inf_sup_distrib1 maddux_3_11_pp)
finally show "p \<sqinter> 1 = p0 \<sqinter> 1"
.
show "fc p = fc p0"
- proof (rule antisym)
+ proof (rule order.antisym)
have 2: "univalent (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])"
by (simp add: a1 a2 a3 update_univalent mult_assoc)
have 3: "-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0 \<le> (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>\<star> * (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>T\<^sup>\<star>"
using fc_increasing inf.order_trans sup.cobounded2 by blast
have "p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 \<le> (p0\<^sup>T\<^sup>\<star> \<sqinter> p0 * x\<^sup>T) * (x \<sqinter> p0\<^sup>\<star> * p0)"
by (metis conv_involutive conv_star_commute dedekind)
also have "... \<le> p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 * x\<^sup>T * p0\<^sup>\<star> * p0"
by (metis comp_associative inf.boundedI inf.cobounded2 inf_le1 mult_isotone)
also have "... \<le> p0\<^sup>T\<^sup>\<star> * x \<sqinter> top * x\<^sup>T * p0\<^sup>\<star>"
using comp_associative comp_inf.mult_right_isotone mult_isotone star.right_plus_below_circ by auto
also have "... = p0\<^sup>T\<^sup>\<star> * x * x\<^sup>T * p0\<^sup>\<star>"
by (metis a2 symmetric_top_closed vector_covector vector_inf_comp vector_mult_closed)
also have "... \<le> (p0\<^sup>T\<^sup>\<star> * x * y\<^sup>T) * (y * x\<^sup>T * p0\<^sup>\<star>)"
- by (metis a3 antisym comp_inf.top_right_mult_increasing conv_involutive dedekind_1 inf.sup_left_divisibility inf.sup_monoid.add_commute mult_right_isotone surjective_conv_total mult_assoc)
+ by (metis a3 order.antisym comp_inf.top_right_mult_increasing conv_involutive dedekind_1 inf.sup_left_divisibility inf.sup_monoid.add_commute mult_right_isotone surjective_conv_total mult_assoc)
also have "... = (p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T) * (y \<sqinter> x\<^sup>T * p0\<^sup>\<star>)"
by (metis a2 a3 vector_covector vector_inf_comp vector_mult_closed)
also have "... = (p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T) * (p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T)\<^sup>T"
by (simp add: conv_dist_comp conv_dist_inf conv_star_commute inf_commute)
also have "... \<le> (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>\<star> * (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>T\<^sup>\<star>"
by (meson conv_isotone dual_order.trans mult_isotone star.circ_increasing sup.cobounded1)
finally have "p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0 \<le> (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>\<star> * (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>T\<^sup>\<star>"
.
hence "(p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0) \<squnion> (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0) \<le> (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>\<star> * (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>T\<^sup>\<star>"
using 3 le_supI by blast
hence "p0 \<le> (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>\<star> * (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>T\<^sup>\<star>"
using 1 by (metis inf_commute maddux_3_11_pp)
hence "fc p0 \<le> (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>\<star> * (p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>y])\<^sup>T\<^sup>\<star>"
using 2 fc_idempotent fc_isotone by fastforce
thus "fc p0 \<le> fc p"
by (simp add: assms(2))
have "((p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T) \<squnion> (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0))\<^sup>\<star> = (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0)\<^sup>\<star> * ((p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T) \<squnion> 1)"
proof (rule star_sup_2)
have 4: "transitive (p0\<^sup>T\<^sup>\<star> * x)"
using a2 comp_associative mult_right_isotone rectangle_star_rectangle by auto
have "transitive (y\<^sup>T)"
by (metis a3 conv_dist_comp inf.eq_refl mult_assoc)
thus "transitive (p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T)"
using 4 transitive_inf_closed by auto
have 5: "p0\<^sup>T\<^sup>\<star> * x * (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0) \<le> p0\<^sup>T\<^sup>\<star> * x"
by (metis a2 mult_right_isotone top_greatest mult_assoc)
have "(-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0)\<^sup>T * y \<le> p0\<^sup>T * y"
by (simp add: conv_dist_inf mult_left_isotone)
also have "... \<le> y"
using a1 a4 root_successor_loop by auto
finally have "y\<^sup>T * (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0) \<le> y\<^sup>T"
using conv_dist_comp conv_isotone by fastforce
thus "(p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T) * (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0) \<le> p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T"
using 5 comp_left_subdist_inf inf_mono order_trans by blast
qed
hence "p\<^sup>\<star> = (-(p0\<^sup>T\<^sup>\<star> * x) \<sqinter> p0)\<^sup>\<star> * ((p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T) \<squnion> 1)"
by (simp add: assms(2))
also have "... \<le> p0\<^sup>\<star> * ((p0\<^sup>T\<^sup>\<star> * x \<sqinter> y\<^sup>T) \<squnion> 1)"
by (simp add: mult_left_isotone star_isotone)
also have "... = p0\<^sup>\<star> * (p0\<^sup>T\<^sup>\<star> * x * y\<^sup>T \<squnion> 1)"
by (simp add: a2 a3 vector_covector vector_mult_closed)
also have "... = p0\<^sup>\<star> * (p0\<^sup>T\<^sup>\<star> * (x * x\<^sup>T) * p0\<^sup>\<star> * (p0 \<sqinter> 1) \<squnion> 1)"
by (metis a4 coreflexive_symmetric inf.cobounded2 root_var comp_associative conv_dist_comp conv_involutive conv_star_commute)
also have "... \<le> p0\<^sup>\<star> * (p0\<^sup>T\<^sup>\<star> * 1 * p0\<^sup>\<star> * (p0 \<sqinter> 1) \<squnion> 1)"
by (metis a2 mult_left_isotone mult_right_isotone semiring.add_left_mono sup_commute)
also have "... = p0\<^sup>\<star> * (p0\<^sup>T\<^sup>\<star> * (p0 \<sqinter> 1) \<squnion> p0\<^sup>\<star> * (p0 \<sqinter> 1) \<squnion> 1)"
by (simp add: a1 cancel_separate_eq mult_right_dist_sup)
also have "... = p0\<^sup>\<star> * ((p0 \<sqinter> 1) \<squnion> p0\<^sup>\<star> * (p0 \<sqinter> 1) \<squnion> 1)"
by (smt univalent_root_successors a1 conv_dist_comp conv_dist_inf coreflexive_idempotent coreflexive_symmetric inf.cobounded2 injective_codomain loop_root root_transitive_successor_loop symmetric_one_closed)
also have "... = p0\<^sup>\<star> * (p0\<^sup>\<star> * (p0 \<sqinter> 1) \<squnion> 1)"
by (metis inf.sup_left_divisibility inf_commute sup.left_idem sup_commute sup_relative_same_increasing)
also have "... \<le> p0\<^sup>\<star> * p0\<^sup>\<star>"
by (metis inf.cobounded2 inf_commute order.refl order_lesseq_imp star.circ_mult_upper_bound star.circ_reflexive star.circ_transitive_equal sup.boundedI sup_monoid.add_commute)
also have "... = p0\<^sup>\<star>"
by (simp add: star.circ_transitive_equal)
finally show "fc p \<le> fc p0"
by (metis conv_order conv_star_commute mult_isotone)
qed
qed
lemma update_acyclic_6:
assumes "disjoint_set_forest p"
and "point x"
shows "acyclic ((p[p\<^sup>T\<^sup>\<star>*x\<longmapsto>root p x]) \<sqinter> -1)"
using assms root_point root_successor_loop update_acyclic_2 by auto
theorem path_compression_assign:
"VARS p t w
[ path_compression_precondition p x y \<and> p0 = p ]
p[p\<^sup>T\<^sup>\<star> * x] := y
[ path_compression_postcondition p x y p0 ]"
apply vcg_tc_simp
apply (unfold path_compression_precondition_def path_compression_postcondition_def)
apply (intro conjI)
subgoal using update_univalent mult_assoc by auto
subgoal using bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed update_mapping mult_assoc by auto
subgoal using update_acyclic_6 by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by (smt same_root path_compression_exact path_compression_precondition_def update_univalent vector_mult_closed)
subgoal using path_compression_exact(1) path_compression_precondition_def by blast
subgoal using path_compression_exact(2) path_compression_precondition_def by blast
by blast
text \<open>
We next look at implementing these updates using a loop.
\<close>
lemma path_compression_1a:
assumes "point x"
and "disjoint_set_forest p"
and "x \<noteq> root p x"
shows "p\<^sup>T\<^sup>+ * x \<le> - x"
by (meson assms bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed vector_mult_closed point_in_vector_or_complement_2 loop_root_2)
lemma path_compression_1b:
"x \<le> p\<^sup>T\<^sup>\<star> * x"
using mult_left_isotone star.circ_reflexive by fastforce
lemma path_compression_1:
"path_compression_precondition p x y \<Longrightarrow> path_compression_invariant p x y p x"
using path_compression_invariant_def path_compression_precondition_def loop_root path_compression_1a path_compression_1b by auto
lemma path_compression_2:
"path_compression_invariant p x y p0 w \<and> y \<noteq> p[[w]] \<Longrightarrow> path_compression_invariant (p[w\<longmapsto>y]) x y p0 (p[[w]]) \<and> card { z . regular z \<and> z \<le> (p[w\<longmapsto>y])\<^sup>T\<^sup>\<star> * (p[[w]]) } < card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * w }"
proof -
let ?p = "p[w\<longmapsto>y]"
let ?s = "{ z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * w }"
let ?t = "{ z . regular z \<and> z \<le> ?p\<^sup>T\<^sup>\<star> * (p[[w]]) }"
assume 1: "path_compression_invariant p x y p0 w \<and> y \<noteq> p[[w]]"
have i1: "disjoint_set_forest p" and i2: "point x" and i3: "point y" and i4: "y = root p x"
using 1 path_compression_invariant_def path_compression_precondition_def by meson+
have i5: "point w" and i6: "y \<le> p\<^sup>T\<^sup>\<star> * w"
and i7: "w \<noteq> x \<longrightarrow> p[[x]] = y \<and> y \<noteq> x \<and> p\<^sup>T\<^sup>+ * w \<le> -x" and i8: "p \<sqinter> 1 = p0 \<sqinter> 1" and i9: "fc p = fc p0"
and i10: "root p w = y" and i11: "w \<noteq> y \<longrightarrow> p[[w]] \<noteq> w" and i12: "p[[w]] = p0[[w]]" and i13: "p0[p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)\<longmapsto>y] = p"
using 1 path_compression_invariant_def by blast+
have i14: "disjoint_set_forest p0" and i15: "y = root p0 x" and i16: "w \<le> p0\<^sup>T\<^sup>\<star> * x"
using 1 path_compression_invariant_def by auto
have 2: "point (p[[w]])"
using i1 i5 read_point by blast
show "path_compression_invariant ?p x y p0 (p[[w]]) \<and> card ?t < card ?s"
proof (unfold path_compression_invariant_def, intro conjI)
have 3: "mapping ?p"
by (simp add: i1 i3 i5 bijective_regular update_total update_univalent)
have 4: "w \<noteq> y"
using 1 i1 i4 root_successor_loop by blast
hence 5: "w \<sqinter> y = bot"
by (simp add: i3 i5 distinct_points)
hence "y * w\<^sup>T \<le> -1"
using pseudo_complement schroeder_4_p by auto
hence "y * w\<^sup>T \<le> p\<^sup>T\<^sup>\<star> \<sqinter> -1"
using i5 i6 shunt_bijective by auto
also have "... \<le> p\<^sup>T\<^sup>+"
by (simp add: star_plus_without_loops)
finally have 6: "y \<le> p\<^sup>T\<^sup>+ * w"
using i5 shunt_bijective by auto
have 7: "w * w\<^sup>T \<le> -p\<^sup>T\<^sup>+"
proof (rule ccontr)
assume "\<not> w * w\<^sup>T \<le> -p\<^sup>T\<^sup>+"
hence "w * w\<^sup>T \<le> --p\<^sup>T\<^sup>+"
using i5 point_arc arc_in_partition by blast
hence "w * w\<^sup>T \<le> p\<^sup>T\<^sup>+ \<sqinter> 1"
using i1 i5 mapping_regular regular_conv_closed regular_closed_star regular_mult_closed by simp
also have "... = ((p\<^sup>T \<sqinter> 1) * p\<^sup>T\<^sup>\<star> \<sqinter> 1) \<squnion> ((p\<^sup>T \<sqinter> -1) * p\<^sup>T\<^sup>\<star> \<sqinter> 1)"
by (metis comp_inf.mult_right_dist_sup maddux_3_11_pp mult_right_dist_sup regular_one_closed)
also have "... = ((p\<^sup>T \<sqinter> 1) * p\<^sup>T\<^sup>\<star> \<sqinter> 1) \<squnion> ((p \<sqinter> -1)\<^sup>+ \<sqinter> 1)\<^sup>T"
by (metis conv_complement conv_dist_inf conv_plus_commute equivalence_one_closed reachable_without_loops)
also have "... \<le> ((p\<^sup>T \<sqinter> 1) * p\<^sup>T\<^sup>\<star> \<sqinter> 1) \<squnion> (-1 \<sqinter> 1)\<^sup>T"
by (metis (no_types, hide_lams) i1 sup_right_isotone inf.sup_left_isotone conv_isotone)
also have "... = (p\<^sup>T \<sqinter> 1) * p\<^sup>T\<^sup>\<star> \<sqinter> 1"
by simp
also have "... \<le> (p\<^sup>T \<sqinter> 1) * top \<sqinter> 1"
by (metis comp_inf.comp_isotone coreflexive_comp_top_inf equivalence_one_closed inf.cobounded1 inf.cobounded2)
also have "... \<le> p\<^sup>T"
by (simp add: coreflexive_comp_top_inf_one)
finally have "w * w\<^sup>T \<le> p\<^sup>T"
by simp
hence "w \<le> p[[w]]"
using i5 shunt_bijective by blast
hence "w = p[[w]]"
using 2 by (metis i5 epm_3 mult_semi_associative)
thus False
using 4 i11 by auto
qed
hence 8: "w \<sqinter> p\<^sup>T\<^sup>+ * w = bot"
using p_antitone_iff pseudo_complement schroeder_4_p by blast
show "y \<le> ?p\<^sup>T\<^sup>\<star> * (p[[w]])"
proof -
have "(w \<sqinter> y\<^sup>T)\<^sup>T * (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> w\<^sup>T * (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: conv_isotone mult_left_isotone)
also have "... \<le> w\<^sup>T * p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: conv_isotone mult_left_isotone star_isotone mult_right_isotone)
also have "... = w\<^sup>T * p\<^sup>T\<^sup>+ * w"
by (simp add: star_plus mult_assoc)
also have "... = bot"
using 8 by (smt i5 covector_inf_comp_3 mult_assoc conv_dist_comp conv_star_commute covector_bot_closed equivalence_top_closed inf.le_iff_sup mult_left_isotone)
finally have "((w \<sqinter> y\<^sup>T)\<^sup>T \<squnion> (-w \<sqinter> p)\<^sup>T) * (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> (-w \<sqinter> p)\<^sup>T * (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: bot_unique mult_right_dist_sup)
also have "... \<le> (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: mult_left_isotone star.left_plus_below_circ)
finally have "?p\<^sup>T * (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: conv_dist_sup)
hence "?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (metis comp_associative star.circ_loop_fixpoint star_left_induct sup_commute sup_least sup_left_divisibility)
hence "w \<sqinter> ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> w \<sqinter> (-w \<sqinter> p)\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
using inf.sup_right_isotone by blast
also have "... \<le> w \<sqinter> p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
using conv_isotone mult_left_isotone star_isotone inf.sup_right_isotone by simp
also have "... = bot"
using 8 by (simp add: star_plus)
finally have 9: "w\<^sup>T * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w = bot"
by (smt i5 covector_inf_comp_3 mult_assoc conv_dist_comp covector_bot_closed equivalence_top_closed inf.le_iff_sup mult_left_isotone bot_least inf.absorb1)
have "p\<^sup>T * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w = ((w \<sqinter> p)\<^sup>T \<squnion> (-w \<sqinter> p)\<^sup>T) * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (metis i5 bijective_regular conv_dist_sup inf.sup_monoid.add_commute maddux_3_11_pp)
also have "... = (w \<sqinter> p)\<^sup>T * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<squnion> (-w \<sqinter> p)\<^sup>T * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: mult_right_dist_sup)
also have "... \<le> w\<^sup>T * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<squnion> (-w \<sqinter> p)\<^sup>T * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
using semiring.add_right_mono comp_isotone conv_isotone by auto
also have "... = (-w \<sqinter> p)\<^sup>T * ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
using 9 by simp
also have "... \<le> ?p\<^sup>T\<^sup>+ * p\<^sup>T * w"
by (simp add: conv_isotone mult_left_isotone)
also have "... \<le> ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: comp_isotone star.left_plus_below_circ)
finally have "p\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> ?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (metis comp_associative star.circ_loop_fixpoint star_left_induct sup_commute sup_least sup_left_divisibility)
thus "y \<le> ?p\<^sup>T\<^sup>\<star> * (p[[w]])"
using 6 by (simp add: star_simulation_right_equal mult_assoc)
qed
have 10: "acyclic (?p \<sqinter> -1)"
using i1 i10 i3 i5 inf_le1 update_acyclic_3 by blast
have "?p[[p\<^sup>T\<^sup>+ * w]] \<le> p\<^sup>T\<^sup>+ * w"
proof -
have "(w\<^sup>T \<sqinter> y) * p\<^sup>T\<^sup>+ * w = y \<sqinter> w\<^sup>T * p\<^sup>T\<^sup>+ * w"
by (metis i3 inf_vector_comp vector_inf_comp)
hence "?p[[p\<^sup>T\<^sup>+ * w]] = (y \<sqinter> w\<^sup>T * p\<^sup>T\<^sup>+ * w) \<squnion> (-w\<^sup>T \<sqinter> p\<^sup>T) * p\<^sup>T\<^sup>+ * w"
by (simp add: comp_associative conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup)
also have "... \<le> y \<squnion> (-w\<^sup>T \<sqinter> p\<^sup>T) * p\<^sup>T\<^sup>+ * w"
using sup_left_isotone by auto
also have "... \<le> y \<squnion> p\<^sup>T * p\<^sup>T\<^sup>+ * w"
using mult_left_isotone sup_right_isotone by auto
also have "... \<le> y \<squnion> p\<^sup>T\<^sup>+ * w"
using semiring.add_left_mono mult_left_isotone mult_right_isotone star.left_plus_below_circ by auto
also have "... = p\<^sup>T\<^sup>+ * w"
using 6 by (simp add: sup_absorb2)
finally show ?thesis
by simp
qed
hence 11: "?p\<^sup>T\<^sup>\<star> * (p[[w]]) \<le> p\<^sup>T\<^sup>+ * w"
using star_left_induct by (simp add: mult_left_isotone star.circ_mult_increasing)
hence 12: "?p\<^sup>T\<^sup>+ * (p[[w]]) \<le> p\<^sup>T\<^sup>+ * w"
using dual_order.trans mult_left_isotone star.left_plus_below_circ by blast
have 13: "?p[[x]] = y \<and> y \<noteq> x \<and> ?p\<^sup>T\<^sup>+ * (p[[w]]) \<le> -x"
proof (cases "w = x")
case True
hence "?p[[x]] = (w\<^sup>T \<sqinter> y) * w \<squnion> (-w\<^sup>T \<sqinter> p\<^sup>T) * w"
by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup)
also have "... = (w\<^sup>T \<sqinter> y) * w \<squnion> p\<^sup>T * (-w \<sqinter> w)"
by (metis i5 conv_complement covector_inf_comp_3 inf.sup_monoid.add_commute vector_complement_closed)
also have "... = (w\<^sup>T \<sqinter> y) * w"
by simp
also have "... = y * w"
by (simp add: i5 covector_inf_comp_3 inf.sup_monoid.add_commute)
also have "... = y"
by (metis i3 i5 comp_associative)
finally show ?thesis
using 4 8 12 True pseudo_complement inf.sup_monoid.add_commute order.trans by blast
next
case False
have "?p[[x]] = (w\<^sup>T \<sqinter> y) * x \<squnion> (-w\<^sup>T \<sqinter> p\<^sup>T) * x"
by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup)
also have "... = y * (w \<sqinter> x) \<squnion> p\<^sup>T * (-w \<sqinter> x)"
by (metis i5 conv_complement covector_inf_comp_3 inf_commute vector_complement_closed)
also have "... = p\<^sup>T * (-w \<sqinter> x)"
using i2 i5 False distinct_points by auto
also have "... = y"
using i2 i5 i7 False distinct_points inf.absorb2 pseudo_complement by auto
finally show ?thesis
using 12 False i7 dual_order.trans by blast
qed
thus "p[[w]] \<noteq> x \<longrightarrow> ?p[[x]] = y \<and> y \<noteq> x \<and> ?p\<^sup>T\<^sup>+ * (p[[w]]) \<le> -x"
by simp
have 14: "?p\<^sup>T\<^sup>\<star> * x = x \<squnion> y"
- proof (rule antisym)
+ proof (rule order.antisym)
have "?p\<^sup>T * (x \<squnion> y) = y \<squnion> ?p\<^sup>T * y"
using 13 by (simp add: mult_left_dist_sup)
also have "... = y \<squnion> (w\<^sup>T \<sqinter> y) * y \<squnion> (-w\<^sup>T \<sqinter> p\<^sup>T) * y"
by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup sup_assoc)
also have "... \<le> y \<squnion> (w\<^sup>T \<sqinter> y) * y \<squnion> p\<^sup>T * y"
using mult_left_isotone sup_right_isotone by auto
also have "... = y \<squnion> (w\<^sup>T \<sqinter> y) * y"
using i1 i10 root_successor_loop sup_commute by auto
also have "... \<le> y \<squnion> y * y"
using mult_left_isotone sup_right_isotone by auto
also have "... = y"
by (metis i3 comp_associative sup.idem)
also have "... \<le> x \<squnion> y"
by simp
finally show "?p\<^sup>T\<^sup>\<star> * x \<le> x \<squnion> y"
by (simp add: star_left_induct)
next
show "x \<squnion> y \<le> ?p\<^sup>T\<^sup>\<star> * x"
using 13 by (metis mult_left_isotone star.circ_increasing star.circ_loop_fixpoint sup.boundedI sup_ge2)
qed
have 15: "y = root ?p x"
proof -
have "(p \<sqinter> 1) * y = (p \<sqinter> 1) * (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star> * x"
by (simp add: i4 comp_associative root_var)
also have "... = (p \<sqinter> 1) * p\<^sup>T\<^sup>\<star> * x"
using coreflexive_idempotent by auto
finally have 16: "(p \<sqinter> 1) * y = y"
by (simp add: i4 root_var)
have 17: "(p \<sqinter> 1) * x \<le> y"
by (metis (no_types, lifting) i4 comp_right_one mult_left_isotone mult_right_isotone star.circ_reflexive root_var)
have "root ?p x = (?p \<sqinter> 1) * (x \<squnion> y)"
using 14 by (metis mult_assoc root_var)
also have "... = (w \<sqinter> y\<^sup>T \<sqinter> 1) * (x \<squnion> y) \<squnion> (-w \<sqinter> p \<sqinter> 1) * (x \<squnion> y)"
by (simp add: inf_sup_distrib2 semiring.distrib_right)
also have "... = (w \<sqinter> 1 \<sqinter> y\<^sup>T) * (x \<squnion> y) \<squnion> (-w \<sqinter> p \<sqinter> 1) * (x \<squnion> y)"
by (simp add: inf.left_commute inf.sup_monoid.add_commute)
also have "... = (w \<sqinter> 1) * (y \<sqinter> (x \<squnion> y)) \<squnion> (-w \<sqinter> p \<sqinter> 1) * (x \<squnion> y)"
by (simp add: i3 covector_inf_comp_3)
also have "... = (w \<sqinter> 1) * y \<squnion> (-w \<sqinter> p \<sqinter> 1) * (x \<squnion> y)"
by (simp add: inf.absorb1)
also have "... = (w \<sqinter> 1 * y) \<squnion> (-w \<sqinter> (p \<sqinter> 1) * (x \<squnion> y))"
by (simp add: i5 inf_assoc vector_complement_closed vector_inf_comp)
also have "... = (w \<sqinter> y) \<squnion> (-w \<sqinter> ((p \<sqinter> 1) * x \<squnion> y))"
using 16 by (simp add: mult_left_dist_sup)
also have "... = (w \<sqinter> y) \<squnion> (-w \<sqinter> y)"
using 17 by (simp add: sup.absorb2)
also have "... = y"
using 5 inf.sup_monoid.add_commute le_iff_inf pseudo_complement sup_monoid.add_0_left by fastforce
finally show ?thesis
by simp
qed
show "path_compression_precondition ?p x y"
using 3 10 15 i2 i3 path_compression_precondition_def by blast
show "vector (p[[w]])"
using 2 by simp
show "injective (p[[w]])"
using 2 by simp
show "surjective (p[[w]])"
using 2 by simp
have "w \<sqinter> p \<sqinter> 1 \<le> w \<sqinter> w\<^sup>T \<sqinter> p"
by (metis inf.boundedE inf.boundedI inf.cobounded1 inf.cobounded2 one_inf_conv)
also have "... = w * w\<^sup>T \<sqinter> p"
by (simp add: i5 vector_covector)
also have "... \<le> -p\<^sup>T\<^sup>+ \<sqinter> p"
using 7 by (simp add: inf.coboundedI2 inf.sup_monoid.add_commute)
finally have "w \<sqinter> p \<sqinter> 1 = bot"
by (metis (no_types, hide_lams) conv_dist_inf coreflexive_symmetric inf.absorb1 inf.boundedE inf.cobounded2 pseudo_complement star.circ_mult_increasing)
also have "w \<sqinter> y\<^sup>T \<sqinter> 1 = bot"
using 5 antisymmetric_bot_closed asymmetric_bot_closed comp_inf.schroeder_2 inf.absorb1 one_inf_conv by fastforce
finally have "w \<sqinter> p \<sqinter> 1 = w \<sqinter> y\<^sup>T \<sqinter> 1"
by simp
thus 18: "?p \<sqinter> 1 = p0 \<sqinter> 1"
by (metis i5 i8 bijective_regular inf.sup_monoid.add_commute inf_sup_distrib2 maddux_3_11_pp)
show 19: "fc ?p = fc p0"
proof -
have "p[[w]] = p\<^sup>T * (w \<sqinter> p\<^sup>\<star> * y)"
by (metis i3 i5 i6 bijective_reverse conv_star_commute inf.absorb1)
also have "... = p\<^sup>T * (w \<sqinter> p\<^sup>\<star>) * y"
by (simp add: i5 vector_inf_comp mult_assoc)
also have "... = p\<^sup>T * ((w \<sqinter> 1) \<squnion> (w \<sqinter> p) * (-w \<sqinter> p)\<^sup>\<star>) * y"
by (simp add: i5 omit_redundant_points)
also have "... = p\<^sup>T * (w \<sqinter> 1) * y \<squnion> p\<^sup>T * (w \<sqinter> p) * (-w \<sqinter> p)\<^sup>\<star> * y"
by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup)
also have "... \<le> p\<^sup>T * y \<squnion> p\<^sup>T * (w \<sqinter> p) * (-w \<sqinter> p)\<^sup>\<star> * y"
- by (metis semiring.add_right_mono comp_isotone eq_iff inf.cobounded1 inf.sup_monoid.add_commute mult_1_right)
+ by (metis semiring.add_right_mono comp_isotone order.eq_iff inf.cobounded1 inf.sup_monoid.add_commute mult_1_right)
also have "... = y \<squnion> p\<^sup>T * (w \<sqinter> p) * (-w \<sqinter> p)\<^sup>\<star> * y"
using i1 i4 root_successor_loop by auto
also have "... \<le> y \<squnion> p\<^sup>T * p * (-w \<sqinter> p)\<^sup>\<star> * y"
using comp_isotone sup_right_isotone by auto
also have "... \<le> y \<squnion> (-w \<sqinter> p)\<^sup>\<star> * y"
by (metis i1 comp_associative eq_refl shunt_mapping sup_right_isotone)
also have "... = (-w \<sqinter> p)\<^sup>\<star> * y"
by (metis star.circ_loop_fixpoint sup.left_idem sup_commute)
finally have 20: "p[[w]] \<le> (-w \<sqinter> p)\<^sup>\<star> * y"
by simp
have "p\<^sup>T * (-w \<sqinter> p)\<^sup>\<star> * y = p\<^sup>T * y \<squnion> p\<^sup>T * (-w \<sqinter> p) * (-w \<sqinter> p)\<^sup>\<star> * y"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute)
also have "... = y \<squnion> p\<^sup>T * (-w \<sqinter> p) * (-w \<sqinter> p)\<^sup>\<star> * y"
using i1 i4 root_successor_loop by auto
also have "... \<le> y \<squnion> p\<^sup>T * p * (-w \<sqinter> p)\<^sup>\<star> * y"
using comp_isotone sup_right_isotone by auto
also have "... \<le> y \<squnion> (-w \<sqinter> p)\<^sup>\<star> * y"
by (metis i1 comp_associative eq_refl shunt_mapping sup_right_isotone)
also have "... = (-w \<sqinter> p)\<^sup>\<star> * y"
by (metis star.circ_loop_fixpoint sup.left_idem sup_commute)
finally have 21: "p\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> (-w \<sqinter> p)\<^sup>\<star> * y"
using 20 by (simp add: comp_associative star_left_induct)
have "w\<^sup>T \<sqinter> p\<^sup>T = p\<^sup>T * (w\<^sup>T \<sqinter> 1)"
by (metis i5 comp_right_one covector_inf_comp_3 inf.sup_monoid.add_commute one_inf_conv)
also have "... \<le> p[[w]]"
by (metis comp_right_subdist_inf inf.boundedE inf.sup_monoid.add_commute one_inf_conv)
also have "... \<le> p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
by (simp add: mult_left_isotone star.circ_mult_increasing_2)
also have "... \<le> (-w \<sqinter> p)\<^sup>\<star> * y"
using 21 by simp
finally have "w \<sqinter> p \<le> y\<^sup>T * (-w \<sqinter> p)\<^sup>T\<^sup>\<star>"
by (metis conv_dist_comp conv_dist_inf conv_involutive conv_isotone conv_star_commute)
hence "w \<sqinter> p \<le> (w \<sqinter> y\<^sup>T) * (-w \<sqinter> p)\<^sup>T\<^sup>\<star>"
by (simp add: i5 vector_inf_comp)
also have "... \<le> (w \<sqinter> y\<^sup>T) * ?p\<^sup>T\<^sup>\<star>"
by (simp add: conv_isotone mult_right_isotone star_isotone)
also have "... \<le> ?p * ?p\<^sup>T\<^sup>\<star>"
by (simp add: mult_left_isotone)
also have "... \<le> fc ?p"
by (simp add: mult_left_isotone star.circ_increasing)
finally have 22: "w \<sqinter> p \<le> fc ?p"
by simp
have "-w \<sqinter> p \<le> ?p"
by simp
also have "... \<le> fc ?p"
by (simp add: fc_increasing)
finally have "(w \<squnion> -w) \<sqinter> p \<le> fc ?p"
using 22 by (simp add: comp_inf.semiring.distrib_left inf.sup_monoid.add_commute)
hence "p \<le> fc ?p"
by (metis i5 bijective_regular inf.sup_monoid.add_commute inf_sup_distrib1 maddux_3_11_pp)
hence 23: "fc p \<le> fc ?p"
using 3 fc_idempotent fc_isotone by fastforce
have "?p \<le> (w \<sqinter> y\<^sup>T) \<squnion> p"
using sup_right_isotone by auto
also have "... = w * y\<^sup>T \<squnion> p"
by (simp add: i3 i5 vector_covector)
also have "... \<le> p\<^sup>\<star> \<squnion> p"
by (smt i5 i6 conv_dist_comp conv_involutive conv_isotone conv_star_commute le_supI shunt_bijective star.circ_increasing sup_absorb1)
also have "... \<le> fc p"
using fc_increasing star.circ_back_loop_prefixpoint by auto
finally have "fc ?p \<le> fc p"
using i1 fc_idempotent fc_isotone by fastforce
thus ?thesis
using 23 i9 by auto
qed
show "?p[[p[[w]]]] = p0[[p[[w]]]]"
proof -
have "?p[[p[[w]]]] = p[[p[[w]]]]"
using 2 4 i5 i11 by (simp add: put_get_different)
also have "... = p[[p0[[w]]]]"
by (simp add: i12)
also have "... = (p0[p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)\<longmapsto>y])[[p0[[w]]]]"
using i13 by auto
also have "... = p0[[p0[[w]]]]"
proof -
have "p0[[w]] \<le> -(p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w))"
by (meson inf.coboundedI2 mult_left_isotone p_antitone p_antitone_iff star.circ_increasing)
thus ?thesis
by (meson i2 i5 put_get_different_vector vector_complement_closed vector_inf_closed vector_mult_closed)
qed
also have "... = p0[[p[[w]]]]"
by (simp add: i12)
finally show ?thesis
.
qed
have 24: "root ?p (p[[w]]) = root p0 (p[[w]])"
using 3 18 19 i14 same_root by blast
also have "... = root p0 (p0[[w]])"
by (simp add: i12)
also have 25: "... = root p0 w"
by (metis i5 i14 conv_involutive forest_components_increasing mult_left_isotone shunt_bijective injective_mult_closed read_surjective same_component_same_root)
finally show 26: "root ?p (p[[w]]) = y"
by (metis i1 i10 i14 i8 i9 same_root)
show "p[[w]] \<noteq> y \<longrightarrow> ?p[[p[[w]]]] \<noteq> p[[w]] \<and> ?p\<^sup>T\<^sup>+ * (p[[w]]) \<le> -(p[[w]])"
proof
assume 27: "p[[w]] \<noteq> y"
show "?p[[p[[w]]]] \<noteq> p[[w]] \<and> ?p\<^sup>T\<^sup>+ * (p[[w]]) \<le> -(p[[w]])"
proof
show "?p[[p[[w]]]] \<noteq> p[[w]]"
using 2 26 27 loop_root by auto
show "?p\<^sup>T\<^sup>+ * (p[[w]]) \<le> -(p[[w]])"
using 2 3 10 26 27 by (simp add: path_compression_1a)
qed
qed
show "univalent p0" "total p0" "acyclic (p0 \<sqinter> -1)"
by (simp_all add: i14)
show "y = root p0 x"
by (simp add: i15)
show "p[[w]] \<le> p0\<^sup>T\<^sup>\<star> * x"
by (metis i12 i16 mult_isotone star.circ_increasing star.circ_transitive_equal mult_assoc)
let ?q = "p0[p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * (p[[w]]))\<longmapsto>y]"
show "?q = ?p"
proof -
have 28: "w \<squnion> p0\<^sup>T\<^sup>+ * w = p0\<^sup>T\<^sup>\<star> * w"
using comp_associative star.circ_loop_fixpoint sup_commute by auto
hence 29: "p0\<^sup>T\<^sup>+ * w = p0\<^sup>T\<^sup>\<star> * w \<sqinter> -w"
using 4 24 25 26 by (metis i12 i14 i5 inf.orderE maddux_3_13 path_compression_1a)
hence "p0\<^sup>T\<^sup>\<star> * (p[[w]]) \<le> -w"
by (metis i12 inf_le2 star_plus mult.assoc)
hence "w \<le> -(p0\<^sup>T\<^sup>\<star> * (p[[w]]))"
by (simp add: p_antitone_iff)
hence "w \<le> p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * (p[[w]]))"
by (simp add: i16)
hence 30: "?q \<sqinter> w = ?p \<sqinter> w"
by (metis update_inf update_inf_same)
have 31: "?q \<sqinter> p0\<^sup>T\<^sup>+ * w = ?p \<sqinter> p0\<^sup>T\<^sup>+ * w"
proof -
have "?q \<sqinter> p0\<^sup>T\<^sup>+ * w = p0 \<sqinter> p0\<^sup>T\<^sup>+ * w"
by (metis i12 comp_associative inf.cobounded2 p_antitone_iff star.circ_plus_same update_inf_different)
also have "... = p \<sqinter> p0\<^sup>T\<^sup>+ * w"
using 29 by (metis i13 inf.cobounded2 inf.sup_monoid.add_assoc p_antitone_iff update_inf_different)
also have "... = ?p \<sqinter> p0\<^sup>T\<^sup>+ * w"
using 29 by (simp add: update_inf_different)
finally show ?thesis
.
qed
have 32: "?q \<sqinter> p0\<^sup>T\<^sup>\<star> * w = ?p \<sqinter> p0\<^sup>T\<^sup>\<star> * w"
using 28 30 31 by (metis inf_sup_distrib1)
have 33: "?q \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)) = ?p \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w))"
proof -
have "p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w) \<le> p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * (p[[w]]))"
using 29 by (metis i12 inf.sup_right_isotone mult.semigroup_axioms p_antitone_inf star_plus semigroup.assoc)
hence "?q \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)) = y\<^sup>T \<sqinter> p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)"
by (metis inf_assoc update_inf)
also have "... = p \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w))"
by (metis i13 inf_assoc update_inf_same)
also have "... = ?p \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w))"
by (simp add: inf.coboundedI2 p_antitone path_compression_1b inf_assoc update_inf_different)
finally show ?thesis
.
qed
have "p0\<^sup>T\<^sup>\<star> * w \<squnion> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)) = p0\<^sup>T\<^sup>\<star> * x"
proof -
have 34: "regular (p0\<^sup>T\<^sup>\<star> * w)"
using i14 i5 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
have "p0\<^sup>T\<^sup>\<star> * w \<le> p0\<^sup>T\<^sup>\<star> * x"
by (metis i16 comp_associative mult_right_isotone star.circ_transitive_equal)
hence "p0\<^sup>T\<^sup>\<star> * w \<squnion> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)) = p0\<^sup>T\<^sup>\<star> * x \<sqinter> (p0\<^sup>T\<^sup>\<star> * w \<squnion> -(p0\<^sup>T\<^sup>\<star> * w))"
by (simp add: comp_inf.semiring.distrib_left inf.absorb2)
also have "... = p0\<^sup>T\<^sup>\<star> * x"
using 34 by (metis inf_sup_distrib1 maddux_3_11_pp)
finally show ?thesis
.
qed
hence 35: "?q \<sqinter> p0\<^sup>T\<^sup>\<star> * x = ?p \<sqinter> p0\<^sup>T\<^sup>\<star> * x"
using 32 33 by (metis inf_sup_distrib1)
have 36: "regular (p0\<^sup>T\<^sup>\<star> * x)"
using i14 i2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
have "-(p0\<^sup>T\<^sup>\<star> * x) \<le> -w"
by (simp add: i16 p_antitone)
hence "?q \<sqinter> -(p0\<^sup>T\<^sup>\<star> * x) = ?p \<sqinter> -(p0\<^sup>T\<^sup>\<star> * x)"
by (metis i13 p_antitone_inf update_inf_different)
thus ?thesis
using 35 36 by (metis maddux_3_11_pp)
qed
show "card ?t < card ?s"
proof -
have "?p\<^sup>T * p\<^sup>T\<^sup>\<star> * w = (w\<^sup>T \<sqinter> y) * p\<^sup>T\<^sup>\<star> * w \<squnion> (-w\<^sup>T \<sqinter> p\<^sup>T) * p\<^sup>T\<^sup>\<star> * w"
by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup)
also have "... \<le> (w\<^sup>T \<sqinter> y) * p\<^sup>T\<^sup>\<star> * w \<squnion> p\<^sup>T * p\<^sup>T\<^sup>\<star> * w"
using mult_left_isotone sup_right_isotone by auto
also have "... \<le> (w\<^sup>T \<sqinter> y) * p\<^sup>T\<^sup>\<star> * w \<squnion> p\<^sup>T\<^sup>\<star> * w"
using mult_left_isotone star.left_plus_below_circ sup_right_isotone by blast
also have "... \<le> y * p\<^sup>T\<^sup>\<star> * w \<squnion> p\<^sup>T\<^sup>\<star> * w"
using semiring.add_right_mono mult_left_isotone by auto
also have "... \<le> y * top \<squnion> p\<^sup>T\<^sup>\<star> * w"
by (simp add: comp_associative le_supI1 mult_right_isotone)
also have "... = p\<^sup>T\<^sup>\<star> * w"
by (simp add: i3 i6 sup_absorb2)
finally have "?p\<^sup>T\<^sup>\<star> * p\<^sup>T * w \<le> p\<^sup>T\<^sup>\<star> * w"
using 11 by (metis dual_order.trans star.circ_loop_fixpoint sup_commute sup_ge2 mult_assoc)
hence 37: "?t \<subseteq> ?s"
using order_lesseq_imp mult_assoc by auto
have 38: "w \<in> ?s"
by (simp add: i5 bijective_regular path_compression_1b)
have 39: "\<not> w \<in> ?t"
proof
assume "w \<in> ?t"
hence 40: "w \<le> (?p\<^sup>T \<sqinter> -1)\<^sup>\<star> * (p[[w]])"
using reachable_without_loops by auto
hence "p[[w]] \<le> (?p \<sqinter> -1)\<^sup>\<star> * w"
using 2 by (smt i5 bijective_reverse conv_star_commute reachable_without_loops)
also have "... \<le> p\<^sup>\<star> * w"
proof -
have "p\<^sup>T\<^sup>\<star> * y = y"
using i1 i4 root_transitive_successor_loop by auto
hence "y\<^sup>T * p\<^sup>\<star> * w = y\<^sup>T * w"
by (metis conv_dist_comp conv_involutive conv_star_commute)
also have "... = bot"
using 5 by (metis i5 inf.idem inf.sup_monoid.add_commute mult_left_zero schroeder_1 vector_inf_comp)
finally have 41: "y\<^sup>T * p\<^sup>\<star> * w = bot"
by simp
have "(?p \<sqinter> -1) * p\<^sup>\<star> * w = (w \<sqinter> y\<^sup>T \<sqinter> -1) * p\<^sup>\<star> * w \<squnion> (-w \<sqinter> p \<sqinter> -1) * p\<^sup>\<star> * w"
by (simp add: comp_inf.mult_right_dist_sup mult_right_dist_sup)
also have "... \<le> (w \<sqinter> y\<^sup>T \<sqinter> -1) * p\<^sup>\<star> * w \<squnion> p * p\<^sup>\<star> * w"
by (meson inf_le1 inf_le2 mult_left_isotone order_trans sup_right_isotone)
also have "... \<le> (w \<sqinter> y\<^sup>T \<sqinter> -1) * p\<^sup>\<star> * w \<squnion> p\<^sup>\<star> * w"
using mult_left_isotone star.left_plus_below_circ sup_right_isotone by blast
also have "... \<le> y\<^sup>T * p\<^sup>\<star> * w \<squnion> p\<^sup>\<star> * w"
by (meson inf_le1 inf_le2 mult_left_isotone order_trans sup_left_isotone)
also have "... = p\<^sup>\<star> * w"
using 41 by simp
finally show ?thesis
by (metis comp_associative le_supI star.circ_loop_fixpoint sup_ge2 star_left_induct)
qed
finally have "w \<le> p\<^sup>T\<^sup>\<star> * p\<^sup>T * w"
using 11 40 reachable_without_loops star_plus by auto
thus False
using 4 i1 i10 i5 loop_root_2 star.circ_plus_same by auto
qed
show "card ?t < card ?s"
apply (rule psubset_card_mono)
subgoal using finite_regular by simp
subgoal using 37 38 39 by auto
done
qed
qed
qed
lemma path_compression_3a:
assumes "path_compression_invariant p x (p[[w]]) p0 w"
shows "p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>p[[w]]] = p"
proof -
let ?y = "p[[w]]"
let ?p = "p0[p0\<^sup>T\<^sup>\<star> * x\<longmapsto>?y]"
have i1: "disjoint_set_forest p" and i2: "point x" and i3: "point ?y" and i4: "?y = root p x"
using assms path_compression_invariant_def path_compression_precondition_def by meson+
have i5: "point w" and i6: "?y \<le> p\<^sup>T\<^sup>\<star> * w"
and i7: "w \<noteq> x \<longrightarrow> p[[x]] = ?y \<and> ?y \<noteq> x \<and> p\<^sup>T\<^sup>+ * w \<le> -x" and i8: "p \<sqinter> 1 = p0 \<sqinter> 1" and i9: "fc p = fc p0"
and i10: "root p w = ?y" and i11: "w \<noteq> ?y \<longrightarrow> p[[w]] \<noteq> w" and i12: "p[[w]] = p0[[w]]" and i13: "p0[p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)\<longmapsto>?y] = p"
and i14: "disjoint_set_forest p0" and i15: "?y = root p0 x" and i16: "w \<le> p0\<^sup>T\<^sup>\<star> * x"
using assms path_compression_invariant_def by blast+
have 1: "?p \<sqinter> ?y = p \<sqinter> ?y"
by (metis i1 i15 i3 i4 get_put inf_le1 root_successor_loop update_inf update_inf_same)
have 2: "?p \<sqinter> w = p \<sqinter> w"
by (metis i5 i12 i16 get_put update_inf update_inf_same)
have "?y = root p0 w"
by (metis i1 i10 i14 i8 i9 same_root)
hence "p0\<^sup>T\<^sup>\<star> * w = w \<squnion> ?y"
by (metis i12 i14 root_transitive_successor_loop star.circ_loop_fixpoint star_plus sup_monoid.add_commute mult_assoc)
hence 3: "?p \<sqinter> p0\<^sup>T\<^sup>\<star> * w = p \<sqinter> p0\<^sup>T\<^sup>\<star> * w"
using 1 2 by (simp add: inf_sup_distrib1)
have "p0\<^sup>T\<^sup>\<star> * w \<le> p0\<^sup>T\<^sup>\<star> * x"
by (metis i16 comp_associative mult_right_isotone star.circ_transitive_equal)
hence 4: "?p \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0\<^sup>T\<^sup>\<star> * w) = p \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> p0\<^sup>T\<^sup>\<star> * w)"
using 3 by (simp add: inf.absorb2)
have 5: "?p \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w)) = p \<sqinter> (p0\<^sup>T\<^sup>\<star> * x \<sqinter> -(p0\<^sup>T\<^sup>\<star> * w))"
by (metis i13 inf_le1 update_inf update_inf_same)
have "regular (p0\<^sup>T\<^sup>\<star> * w)"
using i14 i5 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
hence 6: "?p \<sqinter> p0\<^sup>T\<^sup>\<star> * x = p \<sqinter> p0\<^sup>T\<^sup>\<star> * x"
using 4 5 by (smt inf_sup_distrib1 maddux_3_11_pp)
have 7: "?p \<sqinter> -(p0\<^sup>T\<^sup>\<star> * x) = p \<sqinter> -(p0\<^sup>T\<^sup>\<star> * x)"
by (smt i13 inf.sup_monoid.add_commute inf_import_p inf_sup_absorb le_iff_inf p_dist_inf update_inf_different inf.idem p_antitone_inf)
have "regular (p0\<^sup>T\<^sup>\<star> * x)"
using i14 i2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
thus "?p = p"
using 6 7 by (smt inf_sup_distrib1 maddux_3_11_pp)
qed
lemma path_compression_3:
"path_compression_invariant p x (p[[w]]) p0 w \<Longrightarrow> path_compression_postcondition p x (p[[w]]) p0"
using path_compression_invariant_def path_compression_postcondition_def path_compression_precondition_def path_compression_3a by blast
theorem path_compression:
"VARS p t w
[ path_compression_precondition p x y \<and> p0 = p ]
w := x;
WHILE y \<noteq> p[[w]]
INV { path_compression_invariant p x y p0 w }
VAR { card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * w } }
DO t := w;
w := p[[w]];
p[t] := y
OD
[ path_compression_postcondition p x y p0 ]"
apply vcg_tc_simp
apply (fact path_compression_1)
apply (fact path_compression_2)
using path_compression_3 by auto
lemma path_compression_exists:
"path_compression_precondition p x y \<Longrightarrow> \<exists>p' . path_compression_postcondition p' x y p"
using tc_extract_function path_compression by blast
definition "path_compression p x y \<equiv> (SOME p' . path_compression_postcondition p' x y p)"
lemma path_compression_function:
assumes "path_compression_precondition p x y"
and "p' = path_compression p x y"
shows "path_compression_postcondition p' x y p"
by (metis assms path_compression_def path_compression_exists someI)
subsection \<open>Find-Set with Path Compression\<close>
text \<open>
We sequentially combine find-set and path compression.
We consider implementations which use the previously derived functions and implementations which unfold their definitions.
\<close>
theorem find_set_path_compression:
"VARS p y
[ find_set_precondition p x \<and> p0 = p ]
y := find_set p x;
p := path_compression p x y
[ path_compression_postcondition p x y p0 ]"
apply vcg_tc_simp
using find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def by fastforce
theorem find_set_path_compression_1:
"VARS p t w y
[ find_set_precondition p x \<and> p0 = p ]
y := find_set p x;
w := x;
WHILE y \<noteq> p[[w]]
INV { path_compression_invariant p x y p0 w }
VAR { card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * w } }
DO t := w;
w := p[[w]];
p[t] := y
OD
[ path_compression_postcondition p x y p0 ]"
apply vcg_tc_simp
using find_set_function find_set_postcondition_def find_set_precondition_def path_compression_1 path_compression_precondition_def apply fastforce
apply (fact path_compression_2)
by (fact path_compression_3)
theorem find_set_path_compression_2:
"VARS p y
[ find_set_precondition p x \<and> p0 = p ]
y := x;
WHILE y \<noteq> p[[y]]
INV { find_set_invariant p x y \<and> p0 = p }
VAR { card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * y } }
DO y := p[[y]]
OD;
p := path_compression p x y
[ path_compression_postcondition p x y p0 ]"
apply vcg_tc_simp
apply (fact find_set_1)
apply (fact find_set_2)
by (smt find_set_3 find_set_invariant_def find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def)
theorem find_set_path_compression_3:
"VARS p t w y
[ find_set_precondition p x \<and> p0 = p ]
y := x;
WHILE y \<noteq> p[[y]]
INV { find_set_invariant p x y \<and> p0 = p }
VAR { card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * y } }
DO y := p[[y]]
OD;
w := x;
WHILE y \<noteq> p[[w]]
INV { path_compression_invariant p x y p0 w }
VAR { card { z . regular z \<and> z \<le> p\<^sup>T\<^sup>\<star> * w } }
DO t := w;
w := p[[w]];
p[t] := y
OD
[ path_compression_postcondition p x y p0 ]"
apply vcg_tc_simp
apply (simp add: find_set_1)
apply (fact find_set_2)
using find_set_3 find_set_invariant_def find_set_postcondition_def find_set_precondition_def path_compression_1 path_compression_precondition_def apply blast
apply (fact path_compression_2)
by (fact path_compression_3)
text \<open>
Find-set with path compression returns two results: the representative of the tree and the modified disjoint-set forest.
\<close>
lemma find_set_path_compression_exists:
"find_set_precondition p x \<Longrightarrow> \<exists>p' y . path_compression_postcondition p' x y p"
using tc_extract_function find_set_path_compression by blast
definition "find_set_path_compression p x \<equiv> (SOME (p',y) . path_compression_postcondition p' x y p)"
lemma find_set_path_compression_function:
assumes "find_set_precondition p x"
and "(p',y) = find_set_path_compression p x"
shows "path_compression_postcondition p' x y p"
proof -
let ?P = "\<lambda>(p',y) . path_compression_postcondition p' x y p"
have "?P (SOME z . ?P z)"
apply (unfold some_eq_ex)
using assms(1) find_set_path_compression_exists by simp
thus ?thesis
using assms(2) find_set_path_compression_def by auto
qed
text \<open>
We prove that \<open>find_set_path_compression\<close> returns the same representative as \<open>find_set\<close>.
\<close>
lemma find_set_path_compression_find_set:
assumes "find_set_precondition p x"
shows "find_set p x = snd (find_set_path_compression p x)"
proof -
let ?r = "find_set p x"
let ?p = "fst (find_set_path_compression p x)"
let ?y = "snd (find_set_path_compression p x)"
have 1: "find_set_postcondition p x ?r"
by (simp add: assms find_set_function)
have "path_compression_postcondition ?p x ?y p"
using assms find_set_path_compression_function prod.collapse by blast
thus "?r = ?y"
using 1 by (smt assms same_root find_set_precondition_def find_set_postcondition_def path_compression_precondition_def path_compression_postcondition_def)
qed
text \<open>
A weaker postcondition suffices to prove that the two forests have the same semantics; that is, they describe the same disjoint sets and have the same roots.
\<close>
lemma find_set_path_compression_path_compression_semantics:
assumes "find_set_precondition p x"
shows "fc (path_compression p x (find_set p x)) = fc (fst (find_set_path_compression p x))"
and "path_compression p x (find_set p x) \<sqinter> 1 = fst (find_set_path_compression p x) \<sqinter> 1"
proof -
let ?r = "find_set p x"
let ?q = "path_compression p x ?r"
let ?p = "fst (find_set_path_compression p x)"
let ?y = "snd (find_set_path_compression p x)"
have 1: "path_compression_postcondition (path_compression p x ?r) x ?r p"
using assms find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def by auto
have 2: "path_compression_postcondition ?p x ?y p"
using assms find_set_path_compression_function prod.collapse by blast
show "fc ?q = fc ?p"
using 1 2 by (simp add: path_compression_postcondition_def)
show "?q \<sqinter> 1 = ?p \<sqinter> 1"
using 1 2 by (simp add: path_compression_postcondition_def)
qed
text \<open>
With the current, stronger postcondition of path compression describing the precise effect of how links change, we can prove that the two forests are actually equal.
\<close>
lemma find_set_path_compression_find_set_pathcompression:
assumes "find_set_precondition p x"
shows "path_compression p x (find_set p x) = fst (find_set_path_compression p x)"
proof -
let ?r = "find_set p x"
let ?q = "path_compression p x ?r"
let ?p = "fst (find_set_path_compression p x)"
let ?y = "snd (find_set_path_compression p x)"
have 1: "path_compression_postcondition (path_compression p x ?r) x ?r p"
using assms find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def by auto
have 2: "path_compression_postcondition ?p x ?y p"
using assms find_set_path_compression_function prod.collapse by blast
have "?r = ?y"
by (simp add: assms find_set_path_compression_find_set)
thus "?q = ?p"
using 1 2 by (simp add: path_compression_postcondition_def)
qed
subsection \<open>Union-Sets\<close>
text \<open>
We only consider a naive union-sets operation (without ranks).
The semantics is the equivalence closure obtained after adding the link between the two given nodes,
which requires those two elements to be in the same set.
The implementation uses temporary variable \<open>t\<close> to store the two results returned by find-set with path compression.
The disjoint-set forest, which keeps being updated, is threaded through the sequence of operations.
\<close>
definition "union_sets_precondition p x y \<equiv> disjoint_set_forest p \<and> point x \<and> point y"
definition "union_sets_postcondition p x y p0 \<equiv> union_sets_precondition p x y \<and> fc p = wcc (p0 \<squnion> x * y\<^sup>T)"
lemma union_sets_1:
assumes "union_sets_precondition p0 x y"
and "path_compression_postcondition p1 x r p0"
and "path_compression_postcondition p2 y s p1"
shows "union_sets_postcondition (p2[r\<longmapsto>s]) x y p0"
proof (unfold union_sets_postcondition_def union_sets_precondition_def, intro conjI)
let ?p = "p2[r\<longmapsto>s]"
have 1: "disjoint_set_forest p1 \<and> point r \<and> r = root p1 x \<and> p1 \<sqinter> 1 = p0 \<sqinter> 1 \<and> fc p1 = fc p0"
using assms(2) path_compression_precondition_def path_compression_postcondition_def by auto
have 2: "disjoint_set_forest p2 \<and> point s \<and> s = root p2 y \<and> p2 \<sqinter> 1 = p1 \<sqinter> 1 \<and> fc p2 = fc p1"
using assms(3) path_compression_precondition_def path_compression_postcondition_def by auto
hence 3: "fc p2 = fc p0"
using 1 by simp
show 4: "univalent ?p"
using 1 2 update_univalent by blast
show "total ?p"
using 1 2 bijective_regular update_total by blast
show "acyclic (?p \<sqinter> -1)"
proof (cases "r = s")
case True
thus ?thesis
using 2 update_acyclic_5 by fastforce
next
case False
hence "bot = r \<sqinter> s"
using 1 2 distinct_points by blast
also have "... = r \<sqinter> p2\<^sup>T\<^sup>\<star> * s"
using 2 by (smt root_transitive_successor_loop)
finally have "s \<sqinter> p2\<^sup>\<star> * r = bot"
using schroeder_1 conv_star_commute inf.sup_monoid.add_commute by fastforce
thus ?thesis
using 1 2 update_acyclic_4 by blast
qed
show "vector x"
using assms(1) by (simp add: union_sets_precondition_def)
show "injective x"
using assms(1) by (simp add: union_sets_precondition_def)
show "surjective x"
using assms(1) by (simp add: union_sets_precondition_def)
show "vector y"
using assms(1) by (simp add: union_sets_precondition_def)
show "injective y"
using assms(1) by (simp add: union_sets_precondition_def)
show "surjective y"
using assms(1) by (simp add: union_sets_precondition_def)
show "fc ?p = wcc (p0 \<squnion> x * y\<^sup>T)"
- proof (rule antisym)
+ proof (rule order.antisym)
have "r = p1[[r]]"
using 1 by (metis root_successor_loop)
hence "r * r\<^sup>T \<le> p1\<^sup>T"
using 1 eq_refl shunt_bijective by blast
hence "r * r\<^sup>T \<le> p1"
using 1 conv_order coreflexive_symmetric by fastforce
hence "r * r\<^sup>T \<le> p1 \<sqinter> 1"
using 1 inf.boundedI by blast
also have "... = p2 \<sqinter> 1"
using 2 by simp
finally have "r * r\<^sup>T \<le> p2"
by simp
hence "r \<le> p2 * r"
using 1 shunt_bijective by blast
hence 5: "p2[[r]] \<le> r"
using 2 shunt_mapping by blast
have "r \<sqinter> p2 \<le> r * (top \<sqinter> r\<^sup>T * p2)"
using 1 by (metis dedekind_1)
also have "... = r * r\<^sup>T * p2"
by (simp add: mult_assoc)
also have "... \<le> r * r\<^sup>T"
using 5 by (metis comp_associative conv_dist_comp conv_involutive conv_order mult_right_isotone)
also have "... \<le> 1"
using 1 by blast
finally have 6: "r \<sqinter> p2 \<le> 1"
by simp
have "p0 \<le> wcc p0"
by (simp add: star.circ_sub_dist_1)
also have "... = wcc p2"
using 3 by (simp add: star_decompose_1)
also have 7: "... \<le> wcc ?p"
proof -
have "wcc p2 = wcc ((-r \<sqinter> p2) \<squnion> (r \<sqinter> p2))"
using 1 by (metis bijective_regular inf.sup_monoid.add_commute maddux_3_11_pp)
also have "... \<le> wcc ((-r \<sqinter> p2) \<squnion> 1)"
using 6 wcc_isotone sup_right_isotone by simp
also have "... = wcc (-r \<sqinter> p2)"
using wcc_with_loops by simp
also have "... \<le> wcc ?p"
using wcc_isotone sup_ge2 by blast
finally show ?thesis
by simp
qed
finally have 8: "p0 \<le> wcc ?p"
by force
have "r \<le> p1\<^sup>T\<^sup>\<star> * x"
using 1 by (metis inf_le1)
hence 9: "r * x\<^sup>T \<le> p1\<^sup>T\<^sup>\<star>"
using assms(1) shunt_bijective union_sets_precondition_def by blast
hence "x * r\<^sup>T \<le> p1\<^sup>\<star>"
using conv_dist_comp conv_order conv_star_commute by force
also have "... \<le> wcc p1"
by (simp add: star.circ_sub_dist)
also have "... = wcc p2"
using 1 2 by (simp add: fc_wcc)
also have "... \<le> wcc ?p"
using 7 by simp
finally have 10: "x * r\<^sup>T \<le> wcc ?p"
by simp
have 11: "r * s\<^sup>T \<le> wcc ?p"
using 1 2 star.circ_sub_dist_1 sup_assoc vector_covector by auto
have "s \<le> p2\<^sup>T\<^sup>\<star> * y"
using 2 by (metis inf_le1)
hence 12: "s * y\<^sup>T \<le> p2\<^sup>T\<^sup>\<star>"
using assms(1) shunt_bijective union_sets_precondition_def by blast
also have "... \<le> wcc p2"
using star_isotone sup_ge2 by blast
also have "... \<le> wcc ?p"
using 7 by simp
finally have 13: "s * y\<^sup>T \<le> wcc ?p"
by simp
have "x \<le> x * r\<^sup>T * r \<and> y \<le> y * s\<^sup>T * s"
using 1 2 shunt_bijective by blast
hence "x * y\<^sup>T \<le> x * r\<^sup>T * r * (y * s\<^sup>T * s)\<^sup>T"
using comp_isotone conv_isotone by blast
also have "... = x * r\<^sup>T * r * s\<^sup>T * s * y\<^sup>T"
by (simp add: comp_associative conv_dist_comp)
also have "... \<le> wcc ?p * (r * s\<^sup>T) * (s * y\<^sup>T)"
using 10 by (metis mult_left_isotone mult_assoc)
also have "... \<le> wcc ?p * wcc ?p * (s * y\<^sup>T)"
using 11 by (metis mult_left_isotone mult_right_isotone)
also have "... \<le> wcc ?p * wcc ?p * wcc ?p"
using 13 by (metis mult_right_isotone)
also have "... = wcc ?p"
by (simp add: star.circ_transitive_equal)
finally have "p0 \<squnion> x * y\<^sup>T \<le> wcc ?p"
using 8 by simp
hence "wcc (p0 \<squnion> x * y\<^sup>T) \<le> wcc ?p"
using wcc_below_wcc by simp
thus "wcc (p0 \<squnion> x * y\<^sup>T) \<le> fc ?p"
using 4 fc_wcc by simp
have "-r \<sqinter> p2 \<le> wcc p2"
by (simp add: inf.coboundedI2 star.circ_sub_dist_1)
also have "... = wcc p0"
using 3 by (simp add: star_decompose_1)
also have "... \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
by (simp add: wcc_isotone)
finally have 14: "-r \<sqinter> p2 \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
by simp
have "r * x\<^sup>T \<le> wcc p1"
using 9 inf.order_trans star.circ_sub_dist sup_commute by fastforce
also have "... = wcc p0"
using 1 by (simp add: star_decompose_1)
also have "... \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
by (simp add: wcc_isotone)
finally have 15: "r * x\<^sup>T \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
by simp
have 16: "x * y\<^sup>T \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
using le_supE star.circ_sub_dist_1 by blast
have "y * s\<^sup>T \<le> p2\<^sup>\<star>"
using 12 conv_dist_comp conv_order conv_star_commute by fastforce
also have "... \<le> wcc p2"
using star.circ_sub_dist sup_commute by fastforce
also have "... = wcc p0"
using 3 by (simp add: star_decompose_1)
also have "... \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
by (simp add: wcc_isotone)
finally have 17: "y * s\<^sup>T \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
by simp
have "r \<le> r * x\<^sup>T * x \<and> s \<le> s * y\<^sup>T * y"
using assms(1) shunt_bijective union_sets_precondition_def by blast
hence "r * s\<^sup>T \<le> r * x\<^sup>T * x * (s * y\<^sup>T * y)\<^sup>T"
using comp_isotone conv_isotone by blast
also have "... = r * x\<^sup>T * x * y\<^sup>T * y * s\<^sup>T"
by (simp add: comp_associative conv_dist_comp)
also have "... \<le> wcc (p0 \<squnion> x * y\<^sup>T) * (x * y\<^sup>T) * (y * s\<^sup>T)"
using 15 by (metis mult_left_isotone mult_assoc)
also have "... \<le> wcc (p0 \<squnion> x * y\<^sup>T) * wcc (p0 \<squnion> x * y\<^sup>T) * (y * s\<^sup>T)"
using 16 by (metis mult_left_isotone mult_right_isotone)
also have "... \<le> wcc (p0 \<squnion> x * y\<^sup>T) * wcc (p0 \<squnion> x * y\<^sup>T) * wcc (p0 \<squnion> x * y\<^sup>T)"
using 17 by (metis mult_right_isotone)
also have "... = wcc (p0 \<squnion> x * y\<^sup>T)"
by (simp add: star.circ_transitive_equal)
finally have "?p \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
using 1 2 14 vector_covector by auto
hence "wcc ?p \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
using wcc_below_wcc by blast
thus "fc ?p \<le> wcc (p0 \<squnion> x * y\<^sup>T)"
using 4 fc_wcc by simp
qed
qed
theorem union_sets:
"VARS p r s t
[ union_sets_precondition p x y \<and> p0 = p ]
t := find_set_path_compression p x;
p := fst t;
r := snd t;
t := find_set_path_compression p y;
p := fst t;
s := snd t;
p[r] := s
[ union_sets_postcondition p x y p0 ]"
proof vcg_tc_simp
let ?t1 = "find_set_path_compression p0 x"
let ?p1 = "fst ?t1"
let ?r = "snd ?t1"
let ?t2 = "find_set_path_compression ?p1 y"
let ?p2 = "fst ?t2"
let ?s = "snd ?t2"
let ?p = "?p2[?r\<longmapsto>?s]"
assume 1: "union_sets_precondition p0 x y"
hence 2: "path_compression_postcondition ?p1 x ?r p0"
by (simp add: find_set_precondition_def union_sets_precondition_def find_set_path_compression_function)
hence "path_compression_postcondition ?p2 y ?s ?p1"
using 1 by (meson find_set_precondition_def union_sets_precondition_def find_set_path_compression_function path_compression_postcondition_def path_compression_precondition_def prod.collapse)
thus "union_sets_postcondition (?p2[?r\<longmapsto>?s]) x y p0"
using 1 2 by (simp add: union_sets_1)
qed
lemma union_sets_exists:
"union_sets_precondition p x y \<Longrightarrow> \<exists>p' . union_sets_postcondition p' x y p"
using tc_extract_function union_sets by blast
definition "union_sets p x y \<equiv> (SOME p' . union_sets_postcondition p' x y p)"
lemma union_sets_function:
assumes "union_sets_precondition p x y"
and "p' = union_sets p x y"
shows "union_sets_postcondition p' x y p"
by (metis assms union_sets_def union_sets_exists someI)
theorem union_sets_2:
"VARS p r s t
[ union_sets_precondition p x y \<and> p0 = p ]
r := find_set p x;
p := path_compression p x r;
s := find_set p y;
p := path_compression p y s;
p[r] := s
[ union_sets_postcondition p x y p0 ]"
proof vcg_tc_simp
let ?r = "find_set p0 x"
let ?p1 = "path_compression p0 x ?r"
let ?s = "find_set ?p1 y"
let ?p2 = "path_compression ?p1 y ?s"
assume 1: "union_sets_precondition p0 x y"
hence 2: "path_compression_postcondition ?p1 x ?r p0"
using find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def union_sets_precondition_def by auto
hence "path_compression_postcondition ?p2 y ?s ?p1"
using 1 find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def union_sets_precondition_def path_compression_postcondition_def by meson
thus "union_sets_postcondition (?p2[?r\<longmapsto>?s]) x y p0"
using 1 2 by (simp add: union_sets_1)
qed
end
end
diff --git a/thys/Relational_Minimum_Spanning_Trees/Boruvka.thy b/thys/Relational_Minimum_Spanning_Trees/Boruvka.thy
--- a/thys/Relational_Minimum_Spanning_Trees/Boruvka.thy
+++ b/thys/Relational_Minimum_Spanning_Trees/Boruvka.thy
@@ -1,3620 +1,3619 @@
(* Title: Borůvka's Minimum Spanning Tree Algorithm
Author: Nicolas Robinson-O'Brien
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Bor\r{u}vka's Minimum Spanning Tree Algorithm\<close>
text \<open>
In this theory we prove partial correctness of Bor\r{u}vka's minimum spanning tree algorithm.
\<close>
theory Boruvka
imports
Relational_Disjoint_Set_Forests.Disjoint_Set_Forests
Kruskal
begin
subsection \<open>General results\<close>
text \<open>
The proof is carried out in $m$-$k$-Stone-Kleene relation algebras.
In this section we give results that hold more generally.
\<close>
context stone_kleene_relation_algebra
begin
definition "big_forest H d \<equiv>
equivalence H
\<and> d \<le> -H
\<and> univalent (H * d)
\<and> H \<sqinter> d * d\<^sup>T \<le> 1
\<and> (H * d)\<^sup>+ \<le> - H"
definition "bf_between_points p q H d \<equiv> point p \<and> point q \<and> p \<le> (H * d)\<^sup>\<star> * H * d"
definition "bf_between_arcs a b H d \<equiv> arc a \<and> arc b \<and> a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top"
text \<open>Theorem 3\<close>
lemma He_eq_He_THe_star:
assumes "arc e"
and "equivalence H"
shows "H * e = H * e * (top * H * e)\<^sup>\<star>"
proof -
let ?x = "H * e"
have 1: "H * e \<le> H * e * (top * H * e)\<^sup>\<star>"
using comp_isotone star.circ_reflexive by fastforce
have "H * e * (top * H * e)\<^sup>\<star> = H * e * (top * e)\<^sup>\<star>"
by (metis assms(2) preorder_idempotent surjective_var)
also have "... \<le> H * e * (1 \<squnion> top * (e * top)\<^sup>\<star> * e)"
by (metis eq_refl star.circ_mult_1)
also have "... \<le> H * e * (1 \<squnion> top * top * e)"
by (simp add: star.circ_left_top)
also have "... = H * e \<squnion> H * e * top * e"
by (simp add: mult.semigroup_axioms semiring.distrib_left semigroup.assoc)
finally have 2: "H * e * (top * H * e)\<^sup>\<star> \<le> H * e"
using assms arc_top_arc mult_assoc by auto
thus ?thesis
using 1 2 by simp
qed
lemma path_through_components:
assumes "equivalence H"
and "arc e"
shows "(H * (d \<squnion> e))\<^sup>\<star> = (H * d)\<^sup>\<star> \<squnion> (H * d)\<^sup>\<star> * H * e * (H * d)\<^sup>\<star>"
proof -
have "H * e * (H * d)\<^sup>\<star> * H * e \<le> H * e * top * H * e"
by (simp add: comp_isotone)
also have "... = H * e * top * e"
by (metis assms(1) preorder_idempotent surjective_var mult_assoc)
also have "... = H * e"
using assms(2) arc_top_arc mult_assoc by auto
finally have 1: "H * e * (H * d)\<^sup>\<star> * H * e \<le> H * e"
by simp
have "(H * (d \<squnion> e))\<^sup>\<star> = (H * d \<squnion> H * e)\<^sup>\<star>"
by (simp add: comp_left_dist_sup)
also have "... = (H * d)\<^sup>\<star> \<squnion> (H * d)\<^sup>\<star> * H * e * (H * d)\<^sup>\<star>"
using 1 star_separate_3 by (simp add: mult_assoc)
finally show ?thesis
by simp
qed
lemma simplify_f:
assumes "regular p"
and "regular e"
shows "(f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> - p)\<^sup>T \<squnion> e\<^sup>T \<squnion> e = f \<squnion> f\<^sup>T \<squnion> e \<squnion> e\<^sup>T"
proof -
have "(f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> - p)\<^sup>T \<squnion> e\<^sup>T \<squnion> e
= (f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p) \<squnion> (f\<^sup>T \<sqinter> - e \<sqinter> p\<^sup>T) \<squnion> (f\<^sup>T \<sqinter> - e \<sqinter> - p\<^sup>T) \<squnion> e\<^sup>T \<squnion> e"
by (simp add: conv_complement conv_dist_inf)
also have "... =
((f \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)) \<sqinter> (- e\<^sup>T \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)) \<sqinter> (- p \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)))
\<squnion> ((f\<^sup>T \<squnion> (f\<^sup>T \<sqinter> - e \<sqinter> - p\<^sup>T)) \<sqinter> (- e \<squnion> (f\<^sup>T \<sqinter> - e \<sqinter> - p\<^sup>T)) \<sqinter> (p\<^sup>T \<squnion> (f\<^sup>T \<sqinter> - e \<sqinter> - p\<^sup>T)))
\<squnion> e\<^sup>T \<squnion> e"
by (metis sup_inf_distrib2 sup_assoc)
also have "... =
((f \<squnion> f) \<sqinter> (f \<squnion> - e\<^sup>T) \<sqinter> (f \<squnion> p) \<sqinter> (- e\<^sup>T \<squnion> f) \<sqinter> (- e\<^sup>T \<squnion> - e\<^sup>T) \<sqinter> (- e\<^sup>T \<squnion> p) \<sqinter> (- p \<squnion> f) \<sqinter> (- p \<squnion> - e\<^sup>T) \<sqinter> (- p \<squnion> p))
\<squnion> ((f\<^sup>T \<squnion> f\<^sup>T) \<sqinter> (f\<^sup>T \<squnion> - e) \<sqinter> (f\<^sup>T \<squnion> - p\<^sup>T) \<sqinter> (- e \<squnion> f\<^sup>T) \<sqinter> (- e \<squnion> - e) \<sqinter> (- e \<squnion> - p\<^sup>T) \<sqinter> (p\<^sup>T \<squnion> f\<^sup>T) \<sqinter> (p\<^sup>T \<squnion> - e) \<sqinter> (p\<^sup>T \<squnion> - p\<^sup>T))
\<squnion> e\<^sup>T \<squnion> e"
using sup_inf_distrib1 sup_assoc inf_assoc sup_inf_distrib1 by simp
also have "... =
((f \<squnion> f) \<sqinter> (f \<squnion> - e\<^sup>T) \<sqinter> (f \<squnion> p) \<sqinter> (f \<squnion> - p) \<sqinter> (- e\<^sup>T \<squnion> f) \<sqinter> (- e\<^sup>T \<squnion> - e\<^sup>T) \<sqinter> (- e\<^sup>T \<squnion> p) \<sqinter> (- e\<^sup>T \<squnion> - p) \<sqinter> (- p \<squnion> p))
\<squnion> ((f\<^sup>T \<squnion> f\<^sup>T) \<sqinter> (f\<^sup>T \<squnion> - e) \<sqinter> (f\<^sup>T \<squnion> - p\<^sup>T) \<sqinter> (- e \<squnion> f\<^sup>T) \<sqinter> (f\<^sup>T \<squnion> p\<^sup>T) \<sqinter> (- e \<squnion> - e) \<sqinter> (- e \<squnion> - p\<^sup>T) \<sqinter> (- e \<squnion> p\<^sup>T) \<sqinter> (p\<^sup>T \<squnion> - p\<^sup>T))
\<squnion> e\<^sup>T \<squnion> e"
by (smt abel_semigroup.commute inf.abel_semigroup_axioms inf.left_commute sup.abel_semigroup_axioms)
also have "... = (f \<sqinter> - e\<^sup>T \<sqinter> (- p \<squnion> p)) \<squnion> (f\<^sup>T \<sqinter> - e \<sqinter> (p\<^sup>T \<squnion> - p\<^sup>T)) \<squnion> e\<^sup>T \<squnion> e"
by (smt inf.sup_monoid.add_assoc inf.sup_monoid.add_commute inf_sup_absorb sup.idem)
also have "... = (f \<sqinter> - e\<^sup>T) \<squnion> (f\<^sup>T \<sqinter> - e) \<squnion> e\<^sup>T \<squnion> e"
by (metis assms(1) conv_complement inf_top_right stone)
also have "... = (f \<squnion> e\<^sup>T) \<sqinter> (- e\<^sup>T \<squnion> e\<^sup>T) \<squnion> (f\<^sup>T \<squnion> e) \<sqinter> (- e \<squnion> e)"
by (metis sup.left_commute sup_assoc sup_inf_distrib2)
finally show ?thesis
by (metis abel_semigroup.commute assms(2) conv_complement inf_top_right stone sup.abel_semigroup_axioms sup_assoc)
qed
lemma simplify_forest_components_f:
assumes "regular p"
and "regular e"
and "injective (f \<sqinter> - e\<^sup>T \<sqinter> - p \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e)"
and "injective f"
shows "forest_components ((f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> -e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e) = (f \<squnion> f\<^sup>T \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
proof -
have "forest_components ((f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> -e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e) = wcc ((f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e)"
by (simp add: assms(3) forest_components_wcc)
also have "... = ((f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> - p)\<^sup>T \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p) \<squnion> e\<^sup>T)\<^sup>\<star>"
using conv_dist_sup sup_assoc by auto
also have "... = ((f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p) \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> - p)\<^sup>T \<squnion> e\<^sup>T \<squnion> e)\<^sup>\<star>"
using sup_assoc sup_commute by auto
also have "... = (f \<squnion> f\<^sup>T \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
using assms(1, 2, 3, 4) simplify_f by auto
finally show ?thesis
by simp
qed
lemma components_disj_increasing:
assumes "regular p"
and "regular e"
and "injective (f \<sqinter> - e\<^sup>T \<sqinter> - p \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e)"
and "injective f"
shows "forest_components f \<le> forest_components (f \<sqinter> - e\<^sup>T \<sqinter> - p \<squnion> (f \<sqinter> - e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e)"
proof -
have 1: "forest_components ((f \<sqinter> - e\<^sup>T \<sqinter> - p) \<squnion> (f \<sqinter> -e\<^sup>T \<sqinter> p)\<^sup>T \<squnion> e) = (f \<squnion> f\<^sup>T \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
using simplify_forest_components_f assms(1, 2, 3, 4) by blast
have "forest_components f = wcc f"
by (simp add: assms(4) forest_components_wcc)
also have "... \<le> (f \<squnion> f\<^sup>T \<squnion> e\<^sup>T \<squnion> e)\<^sup>\<star>"
by (simp add: le_supI2 star_isotone sup_commute)
finally show ?thesis
using 1 sup.left_commute sup_commute by simp
qed
lemma fch_equivalence:
assumes "forest h"
shows "equivalence (forest_components h)"
by (simp add: assms forest_components_equivalence)
lemma big_forest_path_split_1:
assumes "arc a"
and "equivalence H"
shows "(H * d)\<^sup>\<star> * H * a * top = (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top"
proof -
let ?H = "H"
let ?x = "?H * (d \<sqinter> -a)"
let ?y = "?H * a"
let ?a = "?H * a * top"
let ?d = "?H * d"
have 1: "?d\<^sup>\<star> * ?a \<le> ?x\<^sup>\<star> * ?a"
proof -
have "?x\<^sup>\<star> *?y * ?x\<^sup>\<star> * ?a \<le> ?x\<^sup>\<star> * ?a * ?a"
by (smt mult_left_isotone star.circ_right_top top_right_mult_increasing mult_assoc)
also have "... = ?x\<^sup>\<star> * ?a * a * top"
by (metis ex231e mult_assoc)
also have "... = ?x\<^sup>\<star> * ?a"
by (simp add: assms(1) mult_assoc)
finally have 11: "?x\<^sup>\<star> *?y * ?x\<^sup>\<star> * ?a \<le> ?x\<^sup>\<star> * ?a"
by simp
have "?d\<^sup>\<star> * ?a = (?H * (d \<sqinter> a) \<squnion> ?H * (d \<sqinter> -a))\<^sup>\<star> * ?a"
proof -
have 12: "regular a"
using assms(1) arc_regular by simp
have "?H * ((d \<sqinter> a) \<squnion> (d \<sqinter> - a)) = ?H * (d \<sqinter> top)"
using 12 by (metis inf_top_right maddux_3_11_pp)
thus ?thesis
using mult_left_dist_sup by auto
qed
also have "... \<le> (?y \<squnion> ?x)\<^sup>\<star> * ?a"
by (metis comp_inf.coreflexive_idempotent comp_isotone inf.cobounded1 inf.sup_monoid.add_commute semiring.add_mono star_isotone top.extremum)
also have "... = (?x \<squnion> ?y)\<^sup>\<star> * ?a"
by (simp add: sup_commute mult_assoc)
also have "... = ?x\<^sup>\<star> * ?a \<squnion> (?x\<^sup>\<star> * ?y * (?x\<^sup>\<star> * ?y)\<^sup>\<star> * ?x\<^sup>\<star>) * ?a"
by (smt mult_right_dist_sup star.circ_sup_9 star.circ_unfold_sum mult_assoc)
also have "... \<le> ?x\<^sup>\<star> * ?a \<squnion> (?x\<^sup>\<star> * ?y * (top * ?y)\<^sup>\<star> * ?x\<^sup>\<star>) * ?a"
proof -
have "(?x\<^sup>\<star> * ?y)\<^sup>\<star> \<le> (top * ?y)\<^sup>\<star>"
by (simp add: mult_left_isotone star_isotone)
thus ?thesis
by (metis comp_inf.coreflexive_idempotent comp_inf.transitive_star eq_refl mult_left_dist_sup top.extremum mult_assoc)
qed
also have "... = ?x\<^sup>\<star> * ?a \<squnion> (?x\<^sup>\<star> * ?y * ?x\<^sup>\<star>) * ?a"
using assms(1, 2) He_eq_He_THe_star arc_regular mult_assoc by auto
finally have 13: "(?H * d)\<^sup>\<star> * ?a \<le> ?x\<^sup>\<star> * ?a \<squnion> ?x\<^sup>\<star> * ?y * ?x\<^sup>\<star> * ?a"
by (simp add: mult_assoc)
have 14: "?x\<^sup>\<star> * ?y * ?x\<^sup>\<star> * ?a \<le> ?x\<^sup>\<star> * ?a"
using 11 mult_assoc by auto
thus ?thesis
using 13 14 sup.absorb1 by auto
qed
have 2: "?d\<^sup>\<star> * ?a \<ge> ?x\<^sup>\<star> *?a"
by (simp add: comp_isotone star_isotone)
thus ?thesis
using 1 2 antisym mult_assoc by simp
qed
lemma dTransHd_le_1:
assumes "equivalence H"
and "univalent (H * d)"
shows "d\<^sup>T * H * d \<le> 1"
proof -
have "d\<^sup>T * H\<^sup>T * H * d \<le> 1"
using assms(2) conv_dist_comp mult_assoc by auto
thus ?thesis
using assms(1) mult_assoc by (simp add: preorder_idempotent)
qed
lemma HcompaT_le_compHaT:
assumes "equivalence H"
and "injective (a * top)"
shows "-H * a * top \<le> - (H * a * top)"
proof -
have "a * top * a\<^sup>T \<le> 1"
by (metis assms(2) conv_dist_comp symmetric_top_closed vector_top_closed mult_assoc)
hence "a * top * a\<^sup>T * H \<le> H"
using assms(1) comp_isotone order_trans by blast
hence "a * top * top * a\<^sup>T * H \<le> H"
by (simp add: vector_mult_closed)
hence "a * top * (H * a * top)\<^sup>T \<le> H"
by (metis assms(1) conv_dist_comp symmetric_top_closed vector_top_closed mult_assoc)
thus ?thesis
using assms(2) comp_injective_below_complement mult_assoc by auto
qed
text \<open>Theorem 4\<close>
lemma expand_big_forest:
assumes "big_forest H d"
shows "(d\<^sup>T * H)\<^sup>\<star> * (H * d)\<^sup>\<star> = (d\<^sup>T * H)\<^sup>\<star> \<squnion> (H * d)\<^sup>\<star>"
proof -
have "(H * d)\<^sup>T * H * d \<le> 1"
using assms big_forest_def mult_assoc by auto
hence "d\<^sup>T * H * H * d \<le> 1"
using assms big_forest_def conv_dist_comp by auto
thus ?thesis
by (simp add: cancel_separate_eq comp_associative)
qed
lemma big_forest_path_bot:
assumes "arc a"
and "a \<le> d"
and "big_forest H d"
shows "(d \<sqinter> - a)\<^sup>T * (H * a * top) \<le> bot"
proof -
have 1: "d\<^sup>T * H * d \<le> 1"
using assms(3) big_forest_def dTransHd_le_1 by blast
hence "d * - 1 * d\<^sup>T \<le> - H"
using triple_schroeder_p by force
hence "d * - 1 * d\<^sup>T \<le> 1 \<squnion> - H"
by (simp add: le_supI2)
hence "d * d\<^sup>T \<squnion> d * - 1 * d\<^sup>T \<le> 1 \<squnion> - H"
by (metis assms(3) big_forest_def inf_commute regular_one_closed shunting_p le_supI)
hence "d * 1 * d\<^sup>T \<squnion> d * - 1 * d\<^sup>T \<le> 1 \<squnion> - H"
by simp
hence "d * (1 \<squnion> - 1) * d\<^sup>T \<le> 1 \<squnion> - H"
using comp_associative mult_right_dist_sup by (simp add: mult_left_dist_sup)
hence "d * top * d\<^sup>T \<le> 1 \<squnion> - H"
using regular_complement_top by auto
hence "d * top * a\<^sup>T \<le> 1 \<squnion> - H"
using assms(2) conv_isotone dual_order.trans mult_right_isotone by blast
hence "d * (a * top)\<^sup>T \<le> 1 \<squnion> - H"
by (simp add: comp_associative conv_dist_comp)
hence "d \<le> (1 \<squnion> - H) * (a * top)"
by (simp add: assms(1) shunt_bijective)
hence "d \<le> a * top \<squnion> - H * a * top"
by (simp add: comp_associative mult_right_dist_sup)
also have "... \<le> a * top \<squnion> - (H * a * top)"
using assms(1, 3) HcompaT_le_compHaT big_forest_def sup_right_isotone by auto
finally have "d \<le> a * top \<squnion> - (H * a * top)"
by simp
hence "d \<sqinter> --( H * a * top) \<le> a * top"
using shunting_var_p by auto
hence 2:"d \<sqinter> H * a * top \<le> a * top"
using inf.sup_right_isotone order.trans pp_increasing by blast
have 3:"d \<sqinter> H * a * top \<le> top * a"
proof -
have "d \<sqinter> H * a * top \<le> (H * a \<sqinter> d * top\<^sup>T) * (top \<sqinter> (H * a)\<^sup>T * d)"
by (metis dedekind inf_commute)
also have "... = d * top \<sqinter> H * a * a\<^sup>T * H\<^sup>T * d"
by (simp add: conv_dist_comp inf_vector_comp mult_assoc)
also have "... \<le> d * top \<sqinter> H * a * d\<^sup>T * H\<^sup>T * d"
using assms(2) mult_right_isotone mult_left_isotone conv_isotone inf.sup_right_isotone by auto
also have "... = d * top \<sqinter> H * a * d\<^sup>T * H * d"
using assms(3) big_forest_def by auto
also have "... \<le> d * top \<sqinter> H * a * 1"
using 1 by (metis inf.sup_right_isotone mult_right_isotone mult_assoc)
also have "... \<le> H * a"
by simp
also have "... \<le> top * a"
by (simp add: mult_left_isotone)
finally have "d \<sqinter> H * a * top \<le> top * a"
by simp
thus ?thesis
by simp
qed
have "d \<sqinter> H * a * top \<le> a * top \<sqinter> top * a"
using 2 3 by simp
also have "... = a * top * top * a"
by (metis comp_associative comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_inf_covector vector_inf_comp vector_top_closed)
also have "... = a * top * a"
by (simp add: vector_mult_closed)
finally have 4:"d \<sqinter> H * a * top \<le> a"
by (simp add: assms(1) arc_top_arc)
hence "d \<sqinter> - a \<le> -(H * a * top)"
using assms(1) arc_regular p_shunting_swap by fastforce
hence "(d \<sqinter> - a) * top \<le> -(H * a * top)"
using mult.semigroup_axioms p_antitone_iff schroeder_4_p semigroup.assoc by fastforce
thus ?thesis
by (simp add: schroeder_3_p)
qed
lemma big_forest_path_split_2:
assumes "arc a"
and "a \<le> d"
and "big_forest H d"
shows "(H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top = (H * ((d \<sqinter> - a) \<squnion> (d \<sqinter> - a)\<^sup>T))\<^sup>\<star> * H * a * top"
proof -
let ?lhs = "(H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top"
have 1: "d\<^sup>T * H * d \<le> 1"
using assms(3) big_forest_def dTransHd_le_1 by blast
have 2: "H * a * top \<le> ?lhs"
by (metis le_iff_sup star.circ_loop_fixpoint star.circ_transitive_equal star_involutive sup_commute mult_assoc)
have "(d \<sqinter> - a)\<^sup>T * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top) = (d \<sqinter> - a)\<^sup>T * H * (d \<sqinter> - a) * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top)"
proof -
have "(d \<sqinter> - a)\<^sup>T * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top) = (d \<sqinter> - a)\<^sup>T * (1 \<squnion> H * (d \<sqinter> - a) * (H * (d \<sqinter> - a))\<^sup>\<star>) * (H * a * top)"
by (simp add: star_left_unfold_equal)
also have "... = (d \<sqinter> - a)\<^sup>T * H * a * top \<squnion> (d \<sqinter> - a)\<^sup>T * H * (d \<sqinter> - a) * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top)"
by (smt mult_left_dist_sup star.circ_loop_fixpoint star.circ_mult_1 star_slide sup_commute mult_assoc)
also have "... = bot \<squnion> (d \<sqinter> - a)\<^sup>T * H * (d \<sqinter> - a) * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top)"
by (metis assms(1, 2, 3) big_forest_path_bot mult_assoc le_bot)
thus ?thesis
by (simp add: calculation)
qed
also have "... \<le> d\<^sup>T * H * d * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top)"
using conv_isotone inf.cobounded1 mult_isotone by auto
also have "... \<le> 1 * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top)"
using 1 by (metis le_iff_sup mult_right_dist_sup)
finally have 3: "(d \<sqinter> - a)\<^sup>T * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top) \<le> ?lhs"
using mult_assoc by auto
hence 4: "H * (d \<sqinter> - a)\<^sup>T * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top) \<le> ?lhs"
proof -
have "H * (d \<sqinter> - a)\<^sup>T * (H * (d \<sqinter> - a))\<^sup>\<star> * (H * a * top) \<le> H * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top"
using 3 mult_right_isotone mult_assoc by auto
also have "... = H * H * ((d \<sqinter> - a) * H)\<^sup>\<star> * H * a * top"
by (metis assms(3) big_forest_def star_slide mult_assoc preorder_idempotent)
also have "... = H * ((d \<sqinter> - a) * H)\<^sup>\<star> * H * a * top"
using assms(3) big_forest_def preorder_idempotent by fastforce
finally show ?thesis
by (metis assms(3) big_forest_def preorder_idempotent star_slide mult_assoc)
qed
have 5: "(H * (d \<sqinter> - a) \<squnion> H * (d \<sqinter> - a)\<^sup>T) * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top \<le> ?lhs"
proof -
have 51: "H * (d \<sqinter> - a) * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top \<le> (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top"
using star.left_plus_below_circ mult_left_isotone by simp
have 52: "(H * (d \<sqinter> - a) \<squnion> H * (d \<sqinter> - a)\<^sup>T) * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top = H * (d \<sqinter> - a) * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top \<squnion> H * (d \<sqinter> - a)\<^sup>T * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top"
using mult_right_dist_sup by auto
hence "... \<le> (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top \<squnion> H * (d \<sqinter> - a)\<^sup>T * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top"
using star.left_plus_below_circ mult_left_isotone sup_left_isotone by auto
thus ?thesis
using 4 51 52 mult_assoc by auto
qed
hence "(H * (d \<sqinter> - a) \<squnion> H * (d \<sqinter> - a)\<^sup>T)\<^sup>\<star> * H * a * top \<le> ?lhs"
proof -
have "(H * (d \<sqinter> - a) \<squnion> H * (d \<sqinter> - a)\<^sup>T)\<^sup>\<star> * (H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top \<le> ?lhs"
using 5 star_left_induct_mult_iff mult_assoc by auto
thus ?thesis
using star.circ_decompose_11 star_decompose_1 by auto
qed
hence 6: "(H * ((d \<sqinter> - a) \<squnion> (d \<sqinter> - a)\<^sup>T))\<^sup>\<star> * H * a * top \<le> ?lhs"
using mult_left_dist_sup by auto
have 7: "(H * (d \<sqinter> - a))\<^sup>\<star> * H * a * top \<le> (H * ((d \<sqinter> - a) \<squnion> (d \<sqinter> - a)\<^sup>T))\<^sup>\<star> * H * a * top"
by (simp add: mult_left_isotone semiring.distrib_left star_isotone)
thus ?thesis
using 6 7 by (simp add: mult_assoc)
qed
end
subsection \<open>An operation to select components\<close>
text \<open>
We introduce the operation \<open>choose_component\<close>.
\begin{itemize}
\item Axiom \<open>component_in_v\<close> expresses that the result of \<open>choose_component\<close> is contained in the set of vertices, $v$, we are selecting from, ignoring the weights.
\item Axiom \<open>component_is_vector\<close> states that the result of \<open>choose_component\<close> is a vector.
\item Axiom \<open>component_is_regular\<close> states that the result of \<open>choose_component\<close> is regular.
\item Axiom \<open>component_is_connected\<close> states that any two vertices from the result of \<open>choose_component\<close> are connected in $e$.
\item Axiom \<open>component_single\<close> states that the result of \<open>choose_component\<close> is closed under being connected in $e$.
\item Finally, axiom \<open>component_not_bot_when_v_bot_bot\<close> expresses that the operation \<open>choose_component\<close> returns a non-empty component if the input satisfies the given criteria.
\end{itemize}
\<close>
class choose_component =
fixes choose_component :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
class choose_component_algebra = choose_component + stone_relation_algebra +
assumes component_in_v: "choose_component e v \<le> --v"
assumes component_is_vector: "vector (choose_component e v)"
assumes component_is_regular: "regular (choose_component e v)"
assumes component_is_connected: "choose_component e v * (choose_component e v)\<^sup>T \<le> e"
assumes component_single: "choose_component e v = e * choose_component e v"
assumes component_not_bot_when_v_bot_bot: "
regular e
\<and> equivalence e
\<and> vector v
\<and> regular v
\<and> e * v = v
\<and> v \<noteq> bot \<longrightarrow> choose_component e v \<noteq> bot"
text \<open>Theorem 1\<close>
text \<open>
Every \<open>m_kleene_algebra\<close> is an instance of \<open>choose_component_algebra\<close> when the \<open>choose_component\<close> operation is defined as follows:
\<close>
context m_kleene_algebra
begin
definition "choose_component_input_condition e v \<equiv>
regular e
\<and> equivalence e
\<and> vector v
\<and> regular v
\<and> e * v = v"
definition "m_choose_component e v \<equiv>
if choose_component_input_condition e v then
e * minarc(v) * top
else
bot"
sublocale m_choose_component_algebra: choose_component_algebra where choose_component = m_choose_component
proof
fix e v
show "m_choose_component e v \<le> -- v"
proof (cases "choose_component_input_condition e v")
case True
hence "m_choose_component e v = e * minarc(v) * top"
by (simp add: m_choose_component_def)
also have "... \<le> e * --v * top"
by (simp add: comp_isotone minarc_below)
also have "... = e * v * top"
using True choose_component_input_condition_def by auto
also have "... = v * top"
using True choose_component_input_condition_def by auto
finally show ?thesis
using True choose_component_input_condition_def by auto
next
case False
hence "m_choose_component e v = bot"
using False m_choose_component_def by auto
thus ?thesis
by simp
qed
next
fix e v
show "vector (m_choose_component e v)"
proof (cases "choose_component_input_condition e v")
case True
thus ?thesis
by (simp add: mult_assoc m_choose_component_def)
next
case False
thus ?thesis
by (simp add: m_choose_component_def)
qed
next
fix e v
show "regular (m_choose_component e v)"
using choose_component_input_condition_def minarc_regular regular_closed_star regular_mult_closed m_choose_component_def by auto
next
fix e v
show "m_choose_component e v * (m_choose_component e v)\<^sup>T \<le> e"
proof (cases "choose_component_input_condition e v")
case True
assume 1: "choose_component_input_condition e v"
hence "m_choose_component e v * (m_choose_component e v)\<^sup>T = e * minarc(v) * top * (e * minarc(v) * top)\<^sup>T"
by (simp add: m_choose_component_def)
also have "... = e * minarc(v) * top * top\<^sup>T * minarc(v)\<^sup>T * e\<^sup>T"
by (metis comp_associative conv_dist_comp)
also have "... = e * minarc(v) * top * top * minarc(v)\<^sup>T * e"
using 1 choose_component_input_condition_def by auto
also have "... = e * minarc(v) * top * minarc(v)\<^sup>T * e"
by (simp add: comp_associative)
also have "... \<le> e"
proof (cases "v = bot")
case True
thus ?thesis
by (simp add: True minarc_bot)
next
case False
assume 3: "v \<noteq> bot"
hence "e * minarc(v) * top * minarc(v)\<^sup>T \<le> e * 1"
using 3 minarc_arc arc_expanded comp_associative mult_right_isotone by fastforce
hence "e * minarc(v) * top * minarc(v)\<^sup>T * e \<le> e * 1 * e"
using mult_left_isotone by auto
also have "... = e"
using 1 choose_component_input_condition_def preorder_idempotent by auto
thus ?thesis
using calculation by auto
qed
thus ?thesis
by (simp add: calculation)
next
case False
thus ?thesis
by (simp add: m_choose_component_def)
qed
next
fix e v
show "m_choose_component e v = e * m_choose_component e v"
proof (cases "choose_component_input_condition e v")
case True
thus ?thesis
by (metis choose_component_input_condition_def preorder_idempotent m_choose_component_def mult_assoc)
next
case False
thus ?thesis
by (simp add: m_choose_component_def)
qed
next
fix e v
show "regular e \<and> equivalence e \<and> vector v \<and> regular v \<and> e * v = v \<and> v \<noteq> bot \<longrightarrow> m_choose_component e v \<noteq> bot"
proof (cases "choose_component_input_condition e v")
case True
hence "m_choose_component e v \<ge> minarc(v) * top"
by (metis choose_component_input_condition_def mult_1_left mult_left_isotone m_choose_component_def)
also have "... \<ge> minarc(v)"
using calculation dual_order.trans top_right_mult_increasing by blast
thus ?thesis
using True bot_unique minarc_bot_iff by fastforce
next
case False
thus ?thesis
using choose_component_input_condition_def by blast
qed
qed
end
subsection \<open>m-k-Stone-Kleene relation algebras\<close>
text \<open>
$m$-$k$-Stone-Kleene relation algebras are an extension of $m$-Kleene algebras where the \<open>choose_component\<close> operation has been added.
\<close>
class m_kleene_algebra_choose_component =
m_kleene_algebra
+ choose_component_algebra
begin
text \<open>
A \<open>selected_edge\<close> is a minimum-weight edge whose source is in a component, with respect to $h$, $j$ and $g$, and whose target is not in that component.
\<close>
abbreviation "selected_edge h j g \<equiv> minarc (choose_component (forest_components h) j * - choose_component (forest_components h) j\<^sup>T \<sqinter> g)"
text \<open>
A \<open>path\<close> is any sequence of edges in the forest, $f$, of the graph, $g$, backwards from the target of the \<open>selected_edge\<close> to a root in $f$.
\<close>
abbreviation "path f h j g \<equiv> top * selected_edge h j g * (f \<sqinter> - selected_edge h j g\<^sup>T)\<^sup>T\<^sup>\<star>"
definition "boruvka_outer_invariant f g \<equiv>
symmetric g
\<and> forest f
\<and> f \<le> --g
\<and> regular f
\<and> (\<exists>w . minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T)"
definition "boruvka_inner_invariant j f h g d \<equiv>
boruvka_outer_invariant f g
\<and> g \<noteq> bot
\<and> vector j
\<and> regular j
\<and> boruvka_outer_invariant h g
\<and> forest h
\<and> forest_components h \<le> forest_components f
\<and> big_forest (forest_components h) d
\<and> d * top \<le> - j
\<and> forest_components h * j = j
\<and> forest_components f = (forest_components h * (d \<squnion> d\<^sup>T))\<^sup>\<star> * forest_components h
\<and> f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T
\<and> (\<forall> a b . bf_between_arcs a b (forest_components h) d \<and> a \<le> -(forest_components h) \<sqinter> -- g \<and> b \<le> d
\<longrightarrow> sum(b \<sqinter> g) \<le> sum(a \<sqinter> g))
\<and> regular d"
lemma expression_equivalent_without_e_complement:
assumes "selected_edge h j g \<le> - forest_components f"
shows "f \<sqinter> - (selected_edge h j g)\<^sup>T \<sqinter> - (path f h j g) \<squnion> (f \<sqinter> - (selected_edge h j g)\<^sup>T \<sqinter> (path f h j g))\<^sup>T \<squnion> (selected_edge h j g)
= f \<sqinter> - (path f h j g) \<squnion> (f \<sqinter> (path f h j g))\<^sup>T \<squnion> (selected_edge h j g)"
proof -
let ?p = "path f h j g"
let ?e = "selected_edge h j g"
let ?F = "forest_components f"
have 1: "?e \<le> - ?F"
by (simp add: assms)
have "f\<^sup>T \<le> ?F"
by (metis conv_dist_comp conv_involutive conv_order conv_star_commute forest_components_increasing)
hence "- ?F \<le> - f\<^sup>T"
using p_antitone by auto
hence "?e \<le> - f\<^sup>T"
using 1 dual_order.trans by blast
hence "f\<^sup>T \<le> - ?e"
by (simp add: p_antitone_iff)
hence "f\<^sup>T\<^sup>T \<le> - ?e\<^sup>T"
by (metis conv_complement conv_dist_inf inf.orderE inf.orderI)
hence "f \<le> - ?e\<^sup>T"
by auto
hence "f = f \<sqinter> - ?e\<^sup>T"
using inf.orderE by blast
hence "f \<sqinter> - ?e\<^sup>T \<sqinter> - ?p \<squnion> (f \<sqinter> - ?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> ?e = f \<sqinter> - ?p \<squnion> (f \<sqinter> ?p)\<^sup>T \<squnion> ?e"
by auto
thus ?thesis by auto
qed
text \<open>Theorem 2\<close>
text \<open>
The source of the \<open>selected_edge\<close> is contained in $j$, the vector describing those vertices still to be processed in the inner loop of Bor\r{u}vka's algorithm.
\<close>
lemma et_below_j:
assumes "vector j"
and "regular j"
and "j \<noteq> bot"
shows "selected_edge h j g * top \<le> j"
proof -
let ?e = "selected_edge h j g"
let ?c = "choose_component (forest_components h) j"
have "?e * top \<le> --(?c * -?c\<^sup>T \<sqinter> g) * top"
using comp_isotone minarc_below by blast
also have "... = (--(?c * -?c\<^sup>T) \<sqinter> --g) * top"
by simp
also have "... = (?c * -?c\<^sup>T \<sqinter> --g) * top"
using component_is_regular regular_mult_closed by auto
also have "... = (?c \<sqinter> -?c\<^sup>T \<sqinter> --g) * top"
by (metis component_is_vector conv_complement vector_complement_closed vector_covector)
also have "... \<le> ?c * top"
using inf.cobounded1 mult_left_isotone order_trans by blast
also have "... \<le> j * top"
by (metis assms(2) comp_inf.star.circ_sup_2 comp_isotone component_in_v)
also have "... = j"
by (simp add: assms(1))
finally show ?thesis
by simp
qed
subsubsection \<open>Components of forests and big forests\<close>
text \<open>
We prove a number of properties about \<open>big_forest\<close> and \<open>forest_components\<close>.
\<close>
lemma fc_j_eq_j_inv:
assumes "forest h"
and "forest_components h * j = j"
shows "forest_components h * (j \<sqinter> - choose_component (forest_components h) j) = j \<sqinter> - choose_component (forest_components h) j"
proof -
let ?c = "choose_component (forest_components h) j"
let ?H = "forest_components h"
have 1:"equivalence ?H"
by (simp add: assms(1) forest_components_equivalence)
have "?H * (j \<sqinter> - ?c) = ?H * j \<sqinter> ?H * - ?c"
using 1 by (metis assms(2) equivalence_comp_dist_inf inf.sup_monoid.add_commute)
hence 2: "?H * (j \<sqinter> - ?c) = j \<sqinter> ?H * - ?c"
by (simp add: assms(2))
have 3: "j \<sqinter> - ?c \<le> ?H * - ?c"
using 1 by (metis assms(2) dedekind_1 dual_order.trans equivalence_comp_dist_inf inf.cobounded2)
have "?H * ?c \<le> ?c"
using component_single by auto
hence "?H\<^sup>T * ?c \<le> ?c"
using 1 by simp
hence "?H * - ?c \<le> - ?c"
using component_is_regular schroeder_3_p by force
hence "j \<sqinter> ?H * - ?c \<le> j \<sqinter> - ?c"
using inf.sup_right_isotone by auto
thus ?thesis
- using 2 3 antisym by simp
+ using 2 3 order.antisym by simp
qed
text \<open>Theorem 5\<close>
text \<open>
There is a path in the \<open>big_forest\<close> between edges $a$ and $b$ if and only if there is either a path in the \<open>big_forest\<close> from $a$ to $b$ or one from $a$ to $c$ and one from $c$ to $b$.
\<close>
lemma big_forest_path_split_disj:
assumes "equivalence H"
and "arc c"
and "regular a \<and> regular b \<and> regular c \<and> regular d \<and> regular H"
shows "bf_between_arcs a b H (d \<squnion> c) \<longleftrightarrow> bf_between_arcs a b H d \<or> (bf_between_arcs a c H d \<and> bf_between_arcs c b H d)"
proof -
have 1: "bf_between_arcs a b H (d \<squnion> c) \<longrightarrow> bf_between_arcs a b H d \<or> (bf_between_arcs a c H d \<and> bf_between_arcs c b H d)"
proof (rule impI)
assume 11: "bf_between_arcs a b H (d \<squnion> c)"
hence "a\<^sup>T * top \<le> (H * (d \<squnion> c))\<^sup>\<star> * H * b * top"
by (simp add: bf_between_arcs_def)
also have "... = ((H * d)\<^sup>\<star> \<squnion> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star>) * H * b * top"
using assms(1, 2) path_through_components by simp
also have "... = (H * d)\<^sup>\<star> * H * b * top \<squnion> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star> * H * b * top"
by (simp add: mult_right_dist_sup)
finally have 12:"a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top \<squnion> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star> * H * b * top"
by simp
have 13: "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top \<or> a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star> * H * b * top"
proof (rule point_in_vector_sup)
show "point (a\<^sup>T * top)"
using 11 bf_between_arcs_def mult_assoc by auto
next
show "vector ((H * d)\<^sup>\<star> * H * b * top)"
using vector_mult_closed by simp
next
show "regular ((H * d)\<^sup>\<star> * H * b * top)"
using assms(3) pp_dist_comp pp_dist_star by auto
next
show "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top \<squnion> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star> * H * b * top"
using 12 by simp
qed
thus "bf_between_arcs a b H d \<or> (bf_between_arcs a c H d \<and> bf_between_arcs c b H d)"
proof (cases "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top")
case True
assume "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top"
hence "bf_between_arcs a b H d"
using 11 bf_between_arcs_def by auto
thus ?thesis
by simp
next
case False
have 14: "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star> * H * b * top"
using 13 by (simp add: False)
hence 15: "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * c * top"
by (metis mult_right_isotone order_lesseq_imp top_greatest mult_assoc)
have "c\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top"
proof (rule ccontr)
assume "\<not> c\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top"
hence "c\<^sup>T * top \<le> -((H * d)\<^sup>\<star> * H * b * top)"
by (meson assms(2, 3) point_in_vector_or_complement regular_closed_star regular_closed_top regular_mult_closed vector_mult_closed vector_top_closed)
hence "c * (H * d)\<^sup>\<star> * H * b * top \<le> bot"
using schroeder_3_p mult_assoc by auto
thus "False"
using 13 False sup.absorb_iff1 mult_assoc by auto
qed
hence "bf_between_arcs a c H d \<and> bf_between_arcs c b H d"
using 11 15 assms(2) bf_between_arcs_def by auto
thus ?thesis
by simp
qed
qed
have 2: "bf_between_arcs a b H d \<or> (bf_between_arcs a c H d \<and> bf_between_arcs c b H d) \<longrightarrow> bf_between_arcs a b H (d \<squnion> c)"
proof -
have 21: "bf_between_arcs a b H d \<longrightarrow> bf_between_arcs a b H (d \<squnion> c)"
proof (rule impI)
assume 22:"bf_between_arcs a b H d"
hence "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * b * top"
using bf_between_arcs_def by blast
hence "a\<^sup>T * top \<le> (H * (d \<squnion> c))\<^sup>\<star> * H * b * top"
by (simp add: mult_left_isotone mult_right_dist_sup mult_right_isotone order.trans star_isotone star_slide)
thus "bf_between_arcs a b H (d \<squnion> c)"
using 22 bf_between_arcs_def by blast
qed
have "bf_between_arcs a c H d \<and> bf_between_arcs c b H d \<longrightarrow> bf_between_arcs a b H (d \<squnion> c)"
proof (rule impI)
assume 23: "bf_between_arcs a c H d \<and> bf_between_arcs c b H d"
hence "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * c * top"
using bf_between_arcs_def by blast
also have "... \<le> (H * d)\<^sup>\<star> * H * c * c\<^sup>T * c * top"
by (metis ex231c comp_inf.star.circ_sup_2 mult_isotone mult_right_isotone mult_assoc)
also have "... \<le> (H * d)\<^sup>\<star> * H * c * c\<^sup>T * top"
by (simp add: mult_right_isotone mult_assoc)
also have "... \<le> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star> * H * b * top"
using 23 mult_right_isotone mult_assoc by (simp add: bf_between_arcs_def)
also have "... \<le> (H * d)\<^sup>\<star> * H * b * top \<squnion> (H * d)\<^sup>\<star> * H * c * (H * d)\<^sup>\<star> * H * b * top"
by (simp add: bf_between_arcs_def)
finally have "a\<^sup>T * top \<le> (H * (d \<squnion> c))\<^sup>\<star> * H * b * top"
using assms(1, 2) path_through_components mult_right_dist_sup by simp
thus "bf_between_arcs a b H (d \<squnion> c)"
using 23 bf_between_arcs_def by blast
qed
thus ?thesis
using 21 by auto
qed
thus ?thesis
using 1 2 by blast
qed
lemma dT_He_eq_bot:
assumes "vector j"
and "regular j"
and "d * top \<le> - j"
and "forest_components h * j = j"
and "j \<noteq> bot"
shows "d\<^sup>T * forest_components h * selected_edge h j g \<le> bot"
proof -
let ?e = "selected_edge h j g"
let ?H = "forest_components h"
have 1: "?e * top \<le> j"
using assms(1, 2, 5) et_below_j by auto
have "d\<^sup>T * ?H * ?e \<le> (d * top)\<^sup>T * ?H * (?e * top)"
by (simp add: comp_isotone conv_isotone top_right_mult_increasing)
also have "... \<le> (d * top)\<^sup>T * ?H * j"
using 1 mult_right_isotone by auto
also have "... \<le> (- j)\<^sup>T * ?H * j"
by (simp add: assms(3) conv_isotone mult_left_isotone)
also have "... = (- j)\<^sup>T * j"
using assms(4) comp_associative by auto
also have "... = bot"
by (simp add: assms(1) conv_complement covector_vector_comp)
finally show ?thesis
using coreflexive_bot_closed le_bot by blast
qed
lemma big_forest_d_U_e:
assumes "forest f"
and "vector j"
and "regular j"
and "forest h"
and "forest_components h \<le> forest_components f"
and "big_forest (forest_components h) d"
and "d * top \<le> - j"
and "forest_components h * j = j"
and "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
and "selected_edge h j g \<le> - forest_components f"
and "selected_edge h j g \<noteq> bot"
and "j \<noteq> bot"
shows "big_forest (forest_components h) (d \<squnion> selected_edge h j g)"
proof (unfold big_forest_def, intro conjI)
let ?H = "forest_components h"
let ?F = "forest_components f"
let ?e = "selected_edge h j g"
let ?d' = "d \<squnion> ?e"
show 01: "reflexive ?H"
by (simp add: assms(4) forest_components_equivalence)
show 02: "transitive ?H"
by (simp add: assms(4) forest_components_equivalence)
show 03: "symmetric ?H"
by (simp add: assms(4) forest_components_equivalence)
have 04: "equivalence ?H"
by (simp add: 01 02 03)
show 1: "?d' \<le> - ?H"
proof -
have "?H \<le> ?F"
by (simp add: assms(5))
hence 11: "?e \<le> - ?H"
using assms(10) order_lesseq_imp p_antitone by blast
have "d \<le> - ?H"
using assms(6) big_forest_def by auto
thus ?thesis
by (simp add: 11)
qed
show "univalent (?H * ?d')"
proof -
have "(?H * ?d')\<^sup>T * (?H * ?d') = ?d'\<^sup>T * ?H\<^sup>T * ?H * ?d'"
using conv_dist_comp mult_assoc by auto
also have "... = ?d'\<^sup>T * ?H * ?H * ?d'"
by (simp add: conv_dist_comp conv_star_commute)
also have "... = ?d'\<^sup>T * ?H * ?d'"
using 01 02 by (metis preorder_idempotent mult_assoc)
finally have 21: "univalent (?H * ?d') \<longleftrightarrow> ?e\<^sup>T * ?H * d \<squnion> d\<^sup>T * ?H * ?e \<squnion> ?e\<^sup>T * ?H * ?e \<squnion> d\<^sup>T * ?H * d \<le> 1"
using conv_dist_sup semiring.distrib_left semiring.distrib_right by auto
have 22: "?e\<^sup>T * ?H * ?e \<le> 1"
proof -
have 221: "?e\<^sup>T * ?H * ?e \<le> ?e\<^sup>T * top * ?e"
by (simp add: mult_left_isotone mult_right_isotone)
have "arc ?e"
using assms(11) minarc_arc minarc_bot_iff by blast
hence "?e\<^sup>T * top * ?e \<le> 1"
using arc_expanded by blast
thus ?thesis
using 221 dual_order.trans by blast
qed
have 24: "d\<^sup>T * ?H * ?e \<le> 1"
by (metis assms(2, 3, 7, 8, 12) dT_He_eq_bot coreflexive_bot_closed le_bot)
hence "(d\<^sup>T * ?H * ?e)\<^sup>T \<le> 1\<^sup>T"
using conv_isotone by blast
hence "?e\<^sup>T * ?H\<^sup>T * d\<^sup>T\<^sup>T \<le> 1"
by (simp add: conv_dist_comp mult_assoc)
hence 25: "?e\<^sup>T * ?H * d \<le> 1"
using assms(4) fch_equivalence by auto
have 8: "d\<^sup>T * ?H * d \<le> 1"
using 04 assms(6) dTransHd_le_1 big_forest_def by blast
thus ?thesis
using 21 22 24 25 by simp
qed
show "coreflexive (?H \<sqinter> ?d' * ?d'\<^sup>T)"
proof -
have "coreflexive (?H \<sqinter> ?d' * ?d'\<^sup>T) \<longleftrightarrow> ?H \<sqinter> (d \<squnion> ?e) * (d\<^sup>T \<squnion> ?e\<^sup>T) \<le> 1"
by (simp add: conv_dist_sup)
also have "... \<longleftrightarrow> ?H \<sqinter> (d * d\<^sup>T \<squnion> d * ?e\<^sup>T \<squnion> ?e * d\<^sup>T \<squnion> ?e * ?e\<^sup>T) \<le> 1"
by (metis mult_left_dist_sup mult_right_dist_sup sup.left_commute sup_commute)
finally have 1: "coreflexive (?H \<sqinter> ?d' * ?d'\<^sup>T) \<longleftrightarrow> ?H \<sqinter> d * d\<^sup>T \<squnion> ?H \<sqinter> d * ?e\<^sup>T \<squnion> ?H \<sqinter> ?e * d\<^sup>T \<squnion> ?H \<sqinter> ?e * ?e\<^sup>T \<le> 1"
by (simp add: inf_sup_distrib1)
have 31: "?H \<sqinter> d * d\<^sup>T \<le> 1"
using assms(6) big_forest_def by blast
have 32: "?H \<sqinter> ?e * d\<^sup>T \<le> 1"
proof -
have "?e * d\<^sup>T \<le> ?e * top * (d * top)\<^sup>T"
by (simp add: conv_isotone mult_isotone top_right_mult_increasing)
also have "... \<le> ?e * top * - j\<^sup>T"
by (metis assms(7) conv_complement conv_isotone mult_right_isotone)
also have "... \<le> j * - j\<^sup>T"
using assms(2, 3, 12) et_below_j mult_left_isotone by auto
also have "... \<le> - ?H"
using 03 by (metis assms(2, 3, 8) conv_complement conv_dist_comp equivalence_top_closed mult_left_isotone schroeder_3_p vector_top_closed)
finally have "?e * d\<^sup>T \<le> - ?H"
by simp
thus ?thesis
by (metis inf.coboundedI1 p_antitone_iff p_shunting_swap regular_one_closed)
qed
have 33: "?H \<sqinter> d * ?e\<^sup>T \<le> 1"
proof -
have 331: "injective h"
by (simp add: assms(4))
have "(?H \<sqinter> ?e * d\<^sup>T)\<^sup>T \<le> 1"
using 32 coreflexive_conv_closed by auto
hence "?H \<sqinter> (?e * d\<^sup>T)\<^sup>T \<le> 1"
using 331 conv_dist_inf forest_components_equivalence by auto
thus ?thesis
using conv_dist_comp by auto
qed
have 34: "?H \<sqinter> ?e * ?e\<^sup>T \<le> 1"
proof -
have 341:"arc ?e \<and> arc (?e\<^sup>T)"
using assms(11) minarc_arc minarc_bot_iff by auto
have "?H \<sqinter> ?e * ?e\<^sup>T \<le> ?e * ?e\<^sup>T"
by auto
thus ?thesis
using 341 arc_injective le_infI2 by blast
qed
thus ?thesis
using 1 31 32 33 34 by simp
qed
show 4:"(?H * (d \<squnion> ?e))\<^sup>+ \<le> - ?H"
proof -
have "?e \<le> - ?F"
by (simp add: assms(10))
hence "?F \<le> - ?e"
by (simp add: p_antitone_iff)
hence "?F\<^sup>T * ?F \<le> - ?e"
using assms(1) fch_equivalence by fastforce
hence "?F\<^sup>T * ?F * ?F\<^sup>T \<le> - ?e"
by (metis assms(1) fch_equivalence forest_components_star star.circ_decompose_9)
hence 41: "?F * ?e * ?F \<le> - ?F"
using triple_schroeder_p by blast
hence 42:"(?F * ?F)\<^sup>\<star> * ?F * ?e * (?F * ?F)\<^sup>\<star> \<le> - ?F"
proof -
have 43: "?F * ?F = ?F"
using assms(1) forest_components_equivalence preorder_idempotent by auto
hence "?F * ?e * ?F = ?F * ?F * ?e * ?F"
by simp
also have "... = (?F)\<^sup>\<star> * ?F * ?e * (?F)\<^sup>\<star>"
by (simp add: assms(1) forest_components_star)
also have "... = (?F * ?F)\<^sup>\<star> * ?F * ?e * (?F * ?F)\<^sup>\<star>"
using 43 by simp
finally show ?thesis
using 41 by simp
qed
hence 44: "(?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> \<le> - ?H"
proof -
have 45: "?H \<le> ?F"
by (simp add: assms(5))
hence 46:"?H * ?e \<le> ?F * ?e"
by (simp add: mult_left_isotone)
have "d \<le> f \<squnion> f\<^sup>T"
using assms(9) sup.left_commute sup_commute by auto
also have "... \<le> ?F"
by (metis forest_components_increasing le_supI2 star.circ_back_loop_fixpoint star.circ_increasing sup.bounded_iff)
finally have "d \<le> ?F"
by simp
hence "?H * d \<le> ?F * ?F"
using 45 mult_isotone by auto
hence 47: "(?H * d)\<^sup>\<star> \<le> (?F * ?F)\<^sup>\<star>"
by (simp add: star_isotone)
hence "(?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> \<le> (?H * d)\<^sup>\<star> * ?F * ?e * (?H * d)\<^sup>\<star>"
using 46 by (metis mult_left_isotone mult_right_isotone mult_assoc)
also have "... \<le> (?F * ?F)\<^sup>\<star> * ?F * ?e * (?F * ?F)\<^sup>\<star>"
using 47 mult_left_isotone mult_right_isotone by (simp add: comp_isotone)
also have "... \<le> - ?F"
using 42 by simp
also have "... \<le> - ?H"
using 45 by (simp add: p_antitone)
finally show ?thesis
by simp
qed
have "(?H * (d \<squnion> ?e))\<^sup>+ = (?H * (d \<squnion> ?e))\<^sup>\<star> * (?H * (d \<squnion> ?e))"
using star.circ_plus_same by auto
also have "... = ((?H * d)\<^sup>\<star> \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star>) * (?H * (d \<squnion> ?e))"
using assms(4, 11) forest_components_equivalence minarc_arc minarc_bot_iff path_through_components by auto
also have "... = (?H * d)\<^sup>\<star> * (?H * (d \<squnion> ?e)) \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * (?H * (d \<squnion> ?e))"
using mult_right_dist_sup by auto
also have "... = (?H * d)\<^sup>\<star> * (?H * d \<squnion> ?H * ?e) \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * (?H * d \<squnion> ?H * ?e)"
by (simp add: mult_left_dist_sup)
also have "... = (?H * d)\<^sup>\<star> * ?H * d \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * (?H * d \<squnion> ?H * ?e)"
using mult_left_dist_sup mult_assoc by auto
also have "... = (?H * d)\<^sup>+ \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * (?H * d \<squnion> ?H * ?e)"
by (simp add: star.circ_plus_same mult_assoc)
also have "... = (?H * d)\<^sup>+ \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * ?H * d \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * ?H * ?e"
by (simp add: mult.semigroup_axioms semiring.distrib_left sup.semigroup_axioms semigroup.assoc)
also have "... \<le> (?H * d)\<^sup>+ \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * ?H * d \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e"
proof -
have "?e * (?H * d)\<^sup>\<star> * ?H * ?e \<le> ?e * top * ?e"
by (metis comp_associative comp_inf.coreflexive_idempotent comp_inf.coreflexive_transitive comp_isotone top.extremum)
also have "... \<le> ?e"
using assms(11) arc_top_arc minarc_arc minarc_bot_iff by auto
finally have "?e * (?H * d)\<^sup>\<star> * ?H * ?e \<le> ?e"
by simp
hence "(?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (simp add: comp_associative comp_isotone)
thus ?thesis
using sup_right_isotone by blast
qed
also have "... = (?H * d)\<^sup>+ \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star> * ?H * d"
- by (smt eq_iff sup.left_commute sup.orderE sup_commute)
+ by (simp add: order.eq_iff ac_simps)
also have "... = (?H * d)\<^sup>+ \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>+"
using star.circ_plus_same mult_assoc by auto
also have "... = (?H * d)\<^sup>+ \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (1 \<squnion> (?H * d)\<^sup>+)"
by (simp add: mult_left_dist_sup sup_assoc)
also have "... = (?H * d)\<^sup>+ \<squnion> (?H * d)\<^sup>\<star> * ?H * ?e * (?H * d)\<^sup>\<star>"
by (simp add: star_left_unfold_equal)
also have "... \<le> - ?H"
using 44 assms(6) big_forest_def by auto
finally show ?thesis
by simp
qed
qed
subsubsection \<open>Identifying arcs\<close>
text \<open>
The expression $d \sqcap \top e^\top H \sqcap (H d^\top)^* H a^\top \top$ identifies the edge incoming to the component that the \<open>selected_edge\<close>, $e$, is outgoing from and which is on the path from edge $a$ to $e$.
Here, we prove this expression is an \<open>arc\<close>.
\<close>
lemma shows_arc_x:
assumes "big_forest H d"
and "bf_between_arcs a e H d"
and "H * d * (H * d)\<^sup>\<star> \<le> - H"
and "\<not> a\<^sup>T * top \<le> H * e * top"
and "regular a"
and "regular e"
and "regular H"
and "regular d"
shows "arc (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)"
proof -
let ?x = "d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top"
have 1:"regular ?x"
using assms(5, 6, 7, 8) regular_closed_star regular_conv_closed regular_mult_closed by auto
have 2: "a\<^sup>T * top * a \<le> 1"
using arc_expanded assms(2) bf_between_arcs_def by auto
have 3: "e * top * e\<^sup>T \<le> 1"
using assms(2) bf_between_arcs_def arc_expanded by blast
have 4: "top * ?x * top = top"
proof -
have "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * e * top"
using assms(2) bf_between_arcs_def by blast
also have "... = H * e * top \<squnion> (H * d)\<^sup>\<star> * H * d * H * e * top"
by (metis star.circ_loop_fixpoint star.circ_plus_same sup_commute mult_assoc)
finally have "a\<^sup>T * top \<le> H * e * top \<squnion> (H * d)\<^sup>\<star> * H * d * H * e * top"
by simp
hence "a\<^sup>T * top \<le> H * e * top \<or> a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * d * H * e * top"
using assms(2, 6, 7) point_in_vector_sup bf_between_arcs_def regular_mult_closed vector_mult_closed by auto
hence "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * d * H * e * top"
using assms(4) by blast
also have "... = (H * d)\<^sup>\<star> * H * d * (H * e * top \<sqinter> H * e * top)"
by (simp add: mult_assoc)
also have "... = (H * d)\<^sup>\<star> * H * (d \<sqinter> (H * e * top)\<^sup>T) * H * e * top"
by (metis comp_associative covector_inf_comp_3 star.circ_left_top star.circ_top)
also have "... = (H * d)\<^sup>\<star> * H * (d \<sqinter> top\<^sup>T * e\<^sup>T * H\<^sup>T) * H * e * top"
using conv_dist_comp mult_assoc by auto
also have "... = (H * d)\<^sup>\<star> * H * (d \<sqinter> top * e\<^sup>T * H) * H * e * top"
using assms(1) by (simp add: big_forest_def)
finally have 2: "a\<^sup>T * top \<le> (H * d)\<^sup>\<star> * H * (d \<sqinter> top * e\<^sup>T * H) * H * e * top"
by simp
hence "e * top \<le> ((H * d)\<^sup>\<star> * H * (d \<sqinter> top * e\<^sup>T * H) * H)\<^sup>T * a\<^sup>T * top"
proof -
have "bijective (e * top) \<and> bijective (a\<^sup>T * top)"
using assms(2) bf_between_arcs_def by auto
thus ?thesis
using 2 by (metis bijective_reverse mult_assoc)
qed
also have "... = H\<^sup>T * (d \<sqinter> top * e\<^sup>T * H)\<^sup>T * H\<^sup>T * (H * d)\<^sup>\<star>\<^sup>T * a\<^sup>T * top"
by (simp add: conv_dist_comp mult_assoc)
also have "... = H * (d \<sqinter> top * e\<^sup>T * H)\<^sup>T * H * (H * d)\<^sup>\<star>\<^sup>T * a\<^sup>T * top"
using assms(1) big_forest_def by auto
also have "... = H * (d \<sqinter> top * e\<^sup>T * H)\<^sup>T * H * (d\<^sup>T * H)\<^sup>\<star> * a\<^sup>T * top"
using assms(1) big_forest_def conv_dist_comp conv_star_commute by auto
also have "... = H * (d\<^sup>T \<sqinter> H * e * top) * H * (d\<^sup>T * H)\<^sup>\<star> * a\<^sup>T * top"
using assms(1) conv_dist_comp big_forest_def comp_associative conv_dist_inf by auto
also have "... = H * (d\<^sup>T \<sqinter> H * e * top) * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top"
by (simp add: comp_associative star_slide)
also have "... = H * (d\<^sup>T \<sqinter> H * e * top) * ((H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)"
using mult_assoc by auto
also have "... = H * (d\<^sup>T \<sqinter> H * e * top \<sqinter> ((H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)\<^sup>T) * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top"
by (smt comp_inf_vector covector_comp_inf vector_conv_covector vector_top_closed mult_assoc)
also have "... = H * (d\<^sup>T \<sqinter> (top * e\<^sup>T * H)\<^sup>T \<sqinter> ((H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)\<^sup>T) * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top"
using assms(1) big_forest_def conv_dist_comp mult_assoc by auto
also have "... = H * (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top"
by (simp add: conv_dist_inf)
finally have 3: "e * top \<le> H * ?x\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top"
by auto
have "?x \<noteq> bot"
proof (rule ccontr)
assume "\<not> ?x \<noteq> bot"
hence "e * top = bot"
using 3 le_bot by auto
thus "False"
using assms(2, 4) bf_between_arcs_def mult_assoc semiring.mult_zero_right by auto
qed
thus ?thesis
using 1 using tarski by blast
qed
have 5: "?x * top * ?x\<^sup>T \<le> 1"
proof -
have 51: "H * (d * H)\<^sup>\<star> \<sqinter> d * H * d\<^sup>T \<le> 1"
proof -
have 511: "d * (H * d)\<^sup>\<star> \<le> - H"
using assms(1, 3) big_forest_def preorder_idempotent schroeder_4_p triple_schroeder_p by fastforce
hence "(d * H)\<^sup>\<star> * d \<le> - H"
using star_slide by auto
hence "H * (d\<^sup>T * H)\<^sup>\<star> \<le> - d"
by (smt assms(1) big_forest_def conv_dist_comp conv_star_commute schroeder_4_p star_slide)
hence "H * (d * H)\<^sup>\<star> \<le> - d\<^sup>T"
using 511 by (metis assms(1) big_forest_def schroeder_5_p star_slide)
hence "H * (d * H)\<^sup>\<star> \<le> - (H * d\<^sup>T)"
by (metis assms(3) p_antitone_iff schroeder_4_p star_slide mult_assoc)
hence "H * (d * H)\<^sup>\<star> \<sqinter> H * d\<^sup>T \<le> bot"
by (simp add: bot_unique pseudo_complement)
hence "H * d * (H * (d * H)\<^sup>\<star> \<sqinter> H * d\<^sup>T) \<le> 1"
by (simp add: bot_unique)
hence 512: "H * d * H * (d * H)\<^sup>\<star> \<sqinter> H * d * H * d\<^sup>T \<le> 1"
using univalent_comp_left_dist_inf assms(1) big_forest_def mult_assoc by fastforce
hence 513: "H * d * H * (d * H)\<^sup>\<star> \<sqinter> d * H * d\<^sup>T \<le> 1"
proof -
have "d * H * d\<^sup>T \<le> H * d * H * d\<^sup>T"
by (metis assms(1) big_forest_def conv_dist_comp conv_involutive mult_1_right mult_left_isotone)
thus ?thesis
using 512 by (smt dual_order.trans p_antitone p_shunting_swap regular_one_closed)
qed
have "d\<^sup>T * H * d \<le> 1 \<squnion> - H"
using assms(1) big_forest_def dTransHd_le_1 le_supI1 by blast
hence "(- 1 \<sqinter> H) * d\<^sup>T * H \<le> - d\<^sup>T"
by (metis assms(1) big_forest_def dTransHd_le_1 inf.sup_monoid.add_commute le_infI2 p_antitone_iff regular_one_closed schroeder_4_p mult_assoc)
hence "d * (- 1 \<sqinter> H) * d\<^sup>T \<le> - H"
by (metis assms(1) big_forest_def conv_dist_comp schroeder_3_p triple_schroeder_p)
hence "H \<sqinter> d * (- 1 \<sqinter> H) * d\<^sup>T \<le> 1"
by (metis inf.coboundedI1 p_antitone_iff p_shunting_swap regular_one_closed)
hence "H \<sqinter> d * d\<^sup>T \<squnion> H \<sqinter> d * (- 1 \<sqinter> H) * d\<^sup>T \<le> 1"
using assms(1) big_forest_def le_supI by blast
hence "H \<sqinter> (d * 1 * d\<^sup>T \<squnion> d * (- 1 \<sqinter> H) * d\<^sup>T) \<le> 1"
using comp_inf.semiring.distrib_left by auto
hence "H \<sqinter> (d * (1 \<squnion> (- 1 \<sqinter> H)) * d\<^sup>T) \<le> 1"
by (simp add: mult_left_dist_sup mult_right_dist_sup)
hence 514: "H \<sqinter> d * H * d\<^sup>T \<le> 1"
by (metis assms(1) big_forest_def comp_inf.semiring.distrib_left inf.le_iff_sup inf.sup_monoid.add_commute inf_top_right regular_one_closed stone)
thus ?thesis
proof -
have "H \<sqinter> d * H * d\<^sup>T \<squnion> H * d * H * (d * H)\<^sup>\<star> \<sqinter> d * H * d\<^sup>T \<le> 1"
using 513 514 by simp
hence "d * H * d\<^sup>T \<sqinter> (H \<squnion> H * d * H * (d * H)\<^sup>\<star>) \<le> 1"
by (simp add: comp_inf.semiring.distrib_left inf.sup_monoid.add_commute)
hence "d * H * d\<^sup>T \<sqinter> H * (1 \<squnion> d * H * (d * H)\<^sup>\<star>) \<le> 1"
by (simp add: mult_left_dist_sup mult_assoc)
thus ?thesis
by (simp add: inf.sup_monoid.add_commute star_left_unfold_equal)
qed
qed
have "?x * top * ?x\<^sup>T = (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top) * top * (d\<^sup>T \<sqinter> H\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T \<sqinter> top\<^sup>T * a\<^sup>T\<^sup>T * H\<^sup>T * (d\<^sup>T\<^sup>T * H\<^sup>T)\<^sup>\<star>)"
by (simp add: conv_dist_comp conv_dist_inf conv_star_commute mult_assoc)
also have "... = (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top) * top * (d\<^sup>T \<sqinter> H * e * top \<sqinter> top * a * H * (d * H)\<^sup>\<star>)"
using assms(1) big_forest_def by auto
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top \<sqinter> (d \<sqinter> top * e\<^sup>T * H) * top * (d\<^sup>T \<sqinter> H * e * top \<sqinter> top * a * H * (d * H)\<^sup>\<star>)"
by (metis inf_vector_comp vector_export_comp)
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top \<sqinter> (d \<sqinter> top * e\<^sup>T * H) * top * top * (d\<^sup>T \<sqinter> H * e * top \<sqinter> top * a * H * (d * H)\<^sup>\<star>)"
by (simp add: vector_mult_closed)
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top \<sqinter> d * ((top * e\<^sup>T * H)\<^sup>T \<sqinter> top) * top * (d\<^sup>T \<sqinter> H * e * top \<sqinter> top * a * H * (d * H)\<^sup>\<star>)"
by (simp add: covector_comp_inf_1 covector_mult_closed)
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top \<sqinter> d * ((top * e\<^sup>T * H)\<^sup>T \<sqinter> (H * e * top)\<^sup>T) * d\<^sup>T \<sqinter> top * a * H * (d * H)\<^sup>\<star>"
by (smt comp_associative comp_inf.star_star_absorb comp_inf_vector conv_star_commute covector_comp_inf covector_conv_vector fc_top star.circ_top total_conv_surjective vector_conv_covector vector_inf_comp)
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top \<sqinter> top * a * H * (d * H)\<^sup>\<star> \<sqinter> d * ((top * e\<^sup>T * H)\<^sup>T \<sqinter> (H * e * top)\<^sup>T) * d\<^sup>T"
using inf.sup_monoid.add_assoc inf.sup_monoid.add_commute by auto
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * top * a * H * (d * H)\<^sup>\<star> \<sqinter> d * ((top * e\<^sup>T * H)\<^sup>T \<sqinter> (H * e * top)\<^sup>T) * d\<^sup>T"
by (smt comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_inf_covector fc_top star.circ_decompose_11 star.circ_top vector_export_comp)
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\<star> \<sqinter> d * (H * e * top \<sqinter> top * e\<^sup>T * H) * d\<^sup>T"
using assms(1) big_forest_def conv_dist_comp mult_assoc by auto
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\<star> \<sqinter> d * H * e * top * e\<^sup>T * H * d\<^sup>T"
by (metis comp_inf_covector inf_top.left_neutral mult_assoc)
also have "... \<le> (H * d\<^sup>T)\<^sup>\<star> * (H * d)\<^sup>\<star> * H \<sqinter> d * H * e * top * e\<^sup>T * H * d\<^sup>T"
proof -
have "(H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\<star> \<le> (H * d\<^sup>T)\<^sup>\<star> * H * 1 * H * (d * H)\<^sup>\<star>"
using 2 by (metis comp_associative comp_isotone mult_left_isotone mult_semi_associative star.circ_transitive_equal)
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H * (d * H)\<^sup>\<star>"
using assms(1) big_forest_def mult.semigroup_axioms preorder_idempotent semigroup.assoc by fastforce
also have "... = (H * d\<^sup>T)\<^sup>\<star> * (H * d)\<^sup>\<star> * H"
by (metis star_slide mult_assoc)
finally show ?thesis
using inf.sup_left_isotone by auto
qed
also have "... \<le> (H * d\<^sup>T)\<^sup>\<star> * (H * d)\<^sup>\<star> * H \<sqinter> d * H * d\<^sup>T"
proof -
have "d * H * e * top * e\<^sup>T * H * d\<^sup>T \<le> d * H * 1 * H * d\<^sup>T"
using 3 by (metis comp_isotone idempotent_one_closed mult_left_isotone mult_sub_right_one mult_assoc)
also have "... \<le> d * H * d\<^sup>T"
by (metis assms(1) big_forest_def mult_left_isotone mult_one_associative mult_semi_associative preorder_idempotent)
finally show ?thesis
using inf.sup_right_isotone by auto
qed
also have "... = H * (d\<^sup>T * H)\<^sup>\<star> * (H * d)\<^sup>\<star> * H \<sqinter> d * H * d\<^sup>T"
by (metis assms(1) big_forest_def comp_associative preorder_idempotent star_slide)
also have "... = H * ((d\<^sup>T * H)\<^sup>\<star> \<squnion> (H * d)\<^sup>\<star>) * H \<sqinter> d * H * d\<^sup>T"
by (simp add: assms(1) expand_big_forest mult.semigroup_axioms semigroup.assoc)
also have "... = (H * (d\<^sup>T * H)\<^sup>\<star> * H \<squnion> H * (H * d)\<^sup>\<star> * H) \<sqinter> d * H * d\<^sup>T"
by (simp add: mult_left_dist_sup mult_right_dist_sup)
also have "... = (H * d\<^sup>T)\<^sup>\<star> * H \<sqinter> d * H * d\<^sup>T \<squnion> H * (d * H)\<^sup>\<star> \<sqinter> d * H * d\<^sup>T"
by (smt assms(1) big_forest_def inf_sup_distrib2 mult.semigroup_axioms preorder_idempotent star_slide semigroup.assoc)
also have "... \<le> (H * d\<^sup>T)\<^sup>\<star> * H \<sqinter> d * H * d\<^sup>T \<squnion> 1"
using 51 comp_inf.semiring.add_left_mono by blast
finally have "?x * top * ?x\<^sup>T \<le> 1"
using 51 by (smt assms(1) big_forest_def conv_dist_comp conv_dist_inf conv_dist_sup conv_involutive conv_star_commute equivalence_one_closed mult.semigroup_axioms sup.absorb2 semigroup.assoc conv_isotone conv_order)
thus ?thesis
by simp
qed
have 6: "?x\<^sup>T * top * ?x \<le> 1"
proof -
have "?x\<^sup>T * top * ?x = (d\<^sup>T \<sqinter> H\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T \<sqinter> top\<^sup>T * a\<^sup>T\<^sup>T * H\<^sup>T * (d\<^sup>T\<^sup>T * H\<^sup>T)\<^sup>\<star>) * top * (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)"
by (simp add: conv_dist_comp conv_dist_inf conv_star_commute mult_assoc)
also have "... = (d\<^sup>T \<sqinter> H * e * top \<sqinter> top * a * H * (d * H)\<^sup>\<star>) * top * (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)"
using assms(1) big_forest_def by auto
also have "... = H * e * top \<sqinter> (d\<^sup>T \<sqinter> top * a * H * (d * H)\<^sup>\<star>) * top * (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)"
by (smt comp_associative inf.sup_monoid.add_assoc inf.sup_monoid.add_commute star.circ_left_top star.circ_top vector_inf_comp)
also have "... = H * e * top \<sqinter> d\<^sup>T * ((top * a * H * (d * H)\<^sup>\<star>)\<^sup>T \<sqinter> top) * (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)"
by (simp add: covector_comp_inf_1 covector_mult_closed)
also have "... = H * e * top \<sqinter> d\<^sup>T * (d * H)\<^sup>\<star>\<^sup>T * H * a\<^sup>T * top * (d \<sqinter> top * e\<^sup>T * H \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)"
using assms(1) big_forest_def comp_associative conv_dist_comp by auto
also have "... = H * e * top \<sqinter> d\<^sup>T * (d * H)\<^sup>\<star>\<^sup>T * H * a\<^sup>T * top * (d \<sqinter> (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top) \<sqinter> top * e\<^sup>T * H"
by (smt comp_associative comp_inf_covector inf.sup_monoid.add_assoc inf.sup_monoid.add_commute)
also have "... = H * e * top \<sqinter> d\<^sup>T * (d * H)\<^sup>\<star>\<^sup>T * H * a\<^sup>T * (top \<sqinter> ((H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)\<^sup>T) * d \<sqinter> top * e\<^sup>T * H"
by (metis comp_associative comp_inf_vector vector_conv_covector vector_top_closed)
also have "... = H * e * top \<sqinter> (H * e * top)\<^sup>T \<sqinter> d\<^sup>T * (d * H)\<^sup>\<star>\<^sup>T * H * a\<^sup>T * ((H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)\<^sup>T * d"
by (smt assms(1) big_forest_def conv_dist_comp inf.left_commute inf.sup_monoid.add_commute symmetric_top_closed mult_assoc inf_top.left_neutral)
also have "... = H * e * top * (H * e * top)\<^sup>T \<sqinter> d\<^sup>T * (d * H)\<^sup>\<star>\<^sup>T * H * a\<^sup>T * ((H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top)\<^sup>T * d"
using vector_covector vector_mult_closed by auto
also have "... = H * e * top * top\<^sup>T * e\<^sup>T * H\<^sup>T \<sqinter> d\<^sup>T * (d * H)\<^sup>\<star>\<^sup>T * H * a\<^sup>T * top\<^sup>T * a\<^sup>T\<^sup>T * H\<^sup>T * (H * d\<^sup>T)\<^sup>\<star>\<^sup>T * d"
by (smt conv_dist_comp mult.semigroup_axioms symmetric_top_closed semigroup.assoc)
also have "... = H * e * top * top * e\<^sup>T * H \<sqinter> d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\<star> * d"
using assms(1) big_forest_def conv_dist_comp conv_star_commute by auto
also have "... = H * e * top * e\<^sup>T * H \<sqinter> d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\<star> * d"
using vector_top_closed mult_assoc by auto
also have "... \<le> H \<sqinter> d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * (d * H)\<^sup>\<star> * d"
proof -
have "H * e * top * e\<^sup>T * H \<le> H * 1 * H"
using 3 by (metis comp_associative mult_left_isotone mult_right_isotone)
also have "... = H"
using assms(1) big_forest_def preorder_idempotent by auto
finally have 611: "H * e * top * e\<^sup>T * H \<le> H"
by simp
have "d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\<star> * d \<le> d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * 1 * H * (d * H)\<^sup>\<star> * d"
using 2 by (metis comp_associative mult_left_isotone mult_right_isotone)
also have "... = d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * (d * H)\<^sup>\<star> * d"
using assms(1) big_forest_def mult.semigroup_axioms preorder_idempotent semigroup.assoc by fastforce
finally have "d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\<star> * d \<le> d\<^sup>T * (H * d\<^sup>T)\<^sup>\<star> * H * (d * H)\<^sup>\<star> * d"
by simp
thus ?thesis
using 611 comp_inf.comp_isotone by blast
qed
also have "... = H \<sqinter> (d\<^sup>T * H)\<^sup>\<star> * d\<^sup>T * H * d * (H * d)\<^sup>\<star>"
using star_slide mult_assoc by auto
also have "... \<le> H \<sqinter> (d\<^sup>T * H)\<^sup>\<star> * (H * d)\<^sup>\<star>"
proof -
have "(d\<^sup>T * H)\<^sup>\<star> * d\<^sup>T * H * d * (H * d)\<^sup>\<star> \<le> (d\<^sup>T * H)\<^sup>\<star> * 1 * (H * d)\<^sup>\<star>"
by (smt assms(1) big_forest_def conv_dist_comp mult_left_isotone mult_right_isotone preorder_idempotent mult_assoc)
also have "... = (d\<^sup>T * H)\<^sup>\<star> * (H * d)\<^sup>\<star>"
by simp
finally show ?thesis
using inf.sup_right_isotone by blast
qed
also have "... = H \<sqinter> ((d\<^sup>T * H)\<^sup>\<star> \<squnion> (H * d)\<^sup>\<star>)"
by (simp add: assms(1) expand_big_forest)
also have "... = H \<sqinter> (d\<^sup>T * H)\<^sup>\<star> \<squnion> H \<sqinter> (H * d)\<^sup>\<star>"
by (simp add: comp_inf.semiring.distrib_left)
also have "... = 1 \<squnion> H \<sqinter> (d\<^sup>T * H)\<^sup>+ \<squnion> H \<sqinter> (H * d)\<^sup>+"
proof -
have 612: "H \<sqinter> (H * d)\<^sup>\<star> = 1 \<squnion> H \<sqinter> (H * d)\<^sup>+"
using assms(1) big_forest_def reflexive_inf_star by blast
have "H \<sqinter> (d\<^sup>T * H)\<^sup>\<star> = 1 \<squnion> H \<sqinter> (d\<^sup>T * H)\<^sup>+"
using assms(1) big_forest_def reflexive_inf_star by auto
thus ?thesis
using 612 sup_assoc sup_commute by auto
qed
also have "... \<le> 1"
proof -
have 613: "H \<sqinter> (H * d)\<^sup>+ \<le> 1"
by (metis assms(3) inf.coboundedI1 p_antitone_iff p_shunting_swap regular_one_closed)
hence "H \<sqinter> (d\<^sup>T * H)\<^sup>+ \<le> 1"
by (metis assms(1) big_forest_def conv_dist_comp conv_dist_inf conv_plus_commute coreflexive_symmetric)
thus ?thesis
by (simp add: 613)
qed
finally show ?thesis
by simp
qed
have 7:"bijective (?x * top)"
using 4 5 6 arc_expanded by blast
have "bijective (?x\<^sup>T * top)"
using 4 5 6 arc_expanded by blast
thus ?thesis
using 7 by simp
qed
text \<open>
To maintain that $f$ can be extended to a minimum spanning forest we identify an edge, $i = v \sqcap \overline{F}e\top \sqcap \top e^\top F$, that may be exchanged with the \<open>selected_edge\<close>, $e$.
Here, we show that $i$ is an \<open>arc\<close>.
\<close>
lemma boruvka_edge_arc:
assumes "equivalence F"
and "forest v"
and "arc e"
and "regular F"
and "F \<le> forest_components (F \<sqinter> v)"
and "regular v"
and "v * e\<^sup>T = bot"
and "e * F * e = bot"
and "e\<^sup>T \<le> v\<^sup>\<star>"
and "e \<noteq> bot"
shows "arc (v \<sqinter> -F * e * top \<sqinter> top * e\<^sup>T * F)"
proof -
let ?i = "v \<sqinter> -F * e * top \<sqinter> top * e\<^sup>T * F"
have 1: "?i\<^sup>T * top * ?i \<le> 1"
proof -
have "?i\<^sup>T * top * ?i = (v\<^sup>T \<sqinter> top * e\<^sup>T * -F \<sqinter> F * e * top) * top * (v \<sqinter> -F * e * top \<sqinter> top * e\<^sup>T * F)"
using assms(1) conv_complement conv_dist_comp conv_dist_inf mult.semigroup_axioms semigroup.assoc by fastforce
also have "... = F * e * top \<sqinter> (v\<^sup>T \<sqinter> top * e\<^sup>T * -F) * top * (v \<sqinter> -F * e * top) \<sqinter> top * e\<^sup>T * F"
by (smt covector_comp_inf covector_mult_closed inf_vector_comp vector_export_comp vector_top_closed)
also have "... = F * e * top \<sqinter> (v\<^sup>T \<sqinter> top * e\<^sup>T * -F) * top * top * (v \<sqinter> -F * e * top) \<sqinter> top * e\<^sup>T * F"
by (simp add: comp_associative)
also have "... = F * e * top \<sqinter> v\<^sup>T * (top \<sqinter> (top * e\<^sup>T * -F)\<^sup>T) * top * (v \<sqinter> -F * e * top) \<sqinter> top * e\<^sup>T * F"
using comp_associative comp_inf_vector_1 by auto
also have "... = F * e * top \<sqinter> v\<^sup>T * (top \<sqinter> (top * e\<^sup>T * -F)\<^sup>T) * (top \<sqinter> (-F * e * top)\<^sup>T) * v \<sqinter> top * e\<^sup>T * F"
by (smt comp_inf_vector conv_dist_comp mult.semigroup_axioms symmetric_top_closed semigroup.assoc)
also have "... = F * e * top \<sqinter> v\<^sup>T * (top * e\<^sup>T * -F)\<^sup>T * (-F * e * top)\<^sup>T * v \<sqinter> top * e\<^sup>T * F"
by simp
also have "... = F * e * top \<sqinter> v\<^sup>T * -F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T * top\<^sup>T * e\<^sup>T * -F\<^sup>T * v \<sqinter> top * e\<^sup>T * F"
by (metis comp_associative conv_complement conv_dist_comp)
also have "... = F * e * top \<sqinter> v\<^sup>T * -F * e * top * top * e\<^sup>T * -F * v \<sqinter> top * e\<^sup>T * F"
by (simp add: assms(1))
also have "... = F * e * top \<sqinter> v\<^sup>T * -F * e * top \<sqinter> top * e\<^sup>T * -F * v \<sqinter> top * e\<^sup>T * F"
by (metis comp_associative comp_inf_covector inf.sup_monoid.add_assoc inf_top.left_neutral vector_top_closed)
also have "... = (F \<sqinter> v\<^sup>T * -F) * e * top \<sqinter> top * e\<^sup>T * -F * v \<sqinter> top * e\<^sup>T * F"
using assms(3) injective_comp_right_dist_inf mult_assoc by auto
also have "... = (F \<sqinter> v\<^sup>T * -F) * e * top \<sqinter> top * e\<^sup>T * (F \<sqinter> -F * v)"
using assms(3) conv_dist_comp inf.sup_monoid.add_assoc inf.sup_monoid.add_commute mult.semigroup_axioms univalent_comp_left_dist_inf semigroup.assoc by fastforce
also have "... = (F \<sqinter> v\<^sup>T * -F) * e * top * top * e\<^sup>T * (F \<sqinter> -F * v)"
by (metis comp_associative comp_inf_covector inf_top.left_neutral vector_top_closed)
also have "... = (F \<sqinter> v\<^sup>T * -F) * e * top * e\<^sup>T * (F \<sqinter> -F * v)"
by (simp add: comp_associative)
also have "... \<le> (F \<sqinter> v\<^sup>T * -F) * (F \<sqinter> -F * v)"
by (smt assms(3) conv_dist_comp mult_left_isotone shunt_bijective symmetric_top_closed top_right_mult_increasing mult_assoc)
also have "... \<le> (F \<sqinter> v\<^sup>T * -F) * (F \<sqinter> -F * v) \<sqinter> F"
by (metis assms(1) inf.absorb1 inf.cobounded1 mult_isotone preorder_idempotent)
also have "... \<le> (F \<sqinter> v\<^sup>T * -F) * (F \<sqinter> -F * v) \<sqinter> (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star>"
using assms(5) comp_inf.mult_right_isotone by auto
also have "... \<le> (-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star>"
proof -
have "F \<sqinter> v\<^sup>T * -F \<le> (v\<^sup>T \<sqinter> F * -F\<^sup>T) * -F"
by (metis conv_complement dedekind_2 inf_commute)
also have "... = (v\<^sup>T \<sqinter> -F\<^sup>T) * -F"
using assms(1) equivalence_comp_left_complement by simp
finally have "F \<sqinter> v\<^sup>T * -F \<le> F \<sqinter> (v\<^sup>T \<sqinter> -F) * -F"
using assms(1) by auto
hence 11: "F \<sqinter> v\<^sup>T * -F = F \<sqinter> (-F \<sqinter> v\<^sup>T) * -F"
by (metis inf.antisym_conv inf.sup_monoid.add_commute comp_left_subdist_inf inf.boundedE inf.sup_right_isotone)
hence "F\<^sup>T \<sqinter> -F\<^sup>T * v\<^sup>T\<^sup>T = F\<^sup>T \<sqinter> -F\<^sup>T * (-F\<^sup>T \<sqinter> v\<^sup>T\<^sup>T)"
by (metis (full_types) assms(1) conv_complement conv_dist_comp conv_dist_inf)
hence 12: "F \<sqinter> -F * v = F \<sqinter> -F * (-F \<sqinter> v)"
using assms(1) by (simp add: abel_semigroup.commute inf.abel_semigroup_axioms)
have "(F \<sqinter> v\<^sup>T * -F) * (F \<sqinter> -F * v) = (F \<sqinter> (-F \<sqinter> v\<^sup>T) * -F) * (F \<sqinter> -F * (-F \<sqinter> v))"
using 11 12 by auto
also have "... \<le> (-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v)"
by (metis comp_associative comp_isotone inf.cobounded2)
finally show ?thesis
using comp_inf.mult_left_isotone by blast
qed
also have "... = ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v)\<^sup>T * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star>) \<squnion> ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v)\<^sup>\<star>)"
by (metis comp_associative inf_sup_distrib1 star.circ_loop_fixpoint)
also have "... = ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v\<^sup>T) * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star>) \<squnion> ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v)\<^sup>\<star>)"
using assms(1) conv_dist_inf by auto
also have "... = bot \<squnion> ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v)\<^sup>\<star>)"
proof -
have "(-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v\<^sup>T) * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star> \<le> bot"
using assms(1, 2) forests_bot_2 by (simp add: comp_associative)
thus ?thesis
using le_bot by blast
qed
also have "... = (-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (1 \<squnion> (F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v))"
by (simp add: star.circ_plus_same star_left_unfold_equal)
also have "... = ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> 1) \<squnion> ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v))"
by (simp add: comp_inf.semiring.distrib_left)
also have "... \<le> 1 \<squnion> ((-F \<sqinter> v\<^sup>T) * -F * -F * (-F \<sqinter> v) \<sqinter> (F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v))"
using sup_left_isotone by auto
also have "... \<le> 1 \<squnion> bot"
using assms(1, 2) forests_bot_3 comp_inf.semiring.add_left_mono by simp
finally show ?thesis
by simp
qed
have 2: "?i * top * ?i\<^sup>T \<le> 1"
proof -
have "?i * top * ?i\<^sup>T = (v \<sqinter> -F * e * top \<sqinter> top * e\<^sup>T * F) * top * (v\<^sup>T \<sqinter> (-F * e * top)\<^sup>T \<sqinter> (top * e\<^sup>T * F)\<^sup>T)"
by (simp add: conv_dist_inf)
also have "... = (v \<sqinter> -F * e * top \<sqinter> top * e\<^sup>T * F) * top * (v\<^sup>T \<sqinter> top\<^sup>T * e\<^sup>T * -F\<^sup>T \<sqinter> F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T)"
by (simp add: conv_complement conv_dist_comp mult_assoc)
also have "... = (v \<sqinter> -F * e * top \<sqinter> top * e\<^sup>T * F) * top * (v\<^sup>T \<sqinter> top * e\<^sup>T * -F \<sqinter> F * e * top)"
by (simp add: assms(1))
also have "... = -F * e * top \<sqinter> (v \<sqinter> top * e\<^sup>T * F) * top * (v\<^sup>T \<sqinter> top * e\<^sup>T * -F \<sqinter> F * e * top)"
by (smt inf.left_commute inf.sup_monoid.add_assoc vector_export_comp)
also have "... = -F * e * top \<sqinter> (v \<sqinter> top * e\<^sup>T * F) * top * (v\<^sup>T \<sqinter> F * e * top) \<sqinter> top * e\<^sup>T * -F"
by (smt comp_inf_covector inf.sup_monoid.add_assoc inf.sup_monoid.add_commute mult_assoc)
also have "... = -F * e * top \<sqinter> (v \<sqinter> top * e\<^sup>T * F) * top * top * (v\<^sup>T \<sqinter> F * e * top) \<sqinter> top * e\<^sup>T * -F"
by (simp add: mult_assoc)
also have "... = -F * e * top \<sqinter> v * ((top * e\<^sup>T * F)\<^sup>T \<sqinter> top) * top * (v\<^sup>T \<sqinter> F * e * top) \<sqinter> top * e\<^sup>T * -F"
by (simp add: comp_inf_vector_1 mult.semigroup_axioms semigroup.assoc)
also have "... = -F * e * top \<sqinter> v * ((top * e\<^sup>T * F)\<^sup>T \<sqinter> top) * (top \<sqinter> (F * e * top)\<^sup>T) * v\<^sup>T \<sqinter> top * e\<^sup>T * -F"
by (smt comp_inf_vector covector_comp_inf vector_conv_covector vector_mult_closed vector_top_closed)
also have "... = -F * e * top \<sqinter> v * (top * e\<^sup>T * F)\<^sup>T * (F * e * top)\<^sup>T * v\<^sup>T \<sqinter> top * e\<^sup>T * -F"
by simp
also have "... = -F * e * top \<sqinter> v * F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T * top\<^sup>T * e\<^sup>T * F\<^sup>T * v\<^sup>T \<sqinter> top * e\<^sup>T * -F"
by (metis comp_associative conv_dist_comp)
also have "... = -F * e * top \<sqinter> v * F * e * top * top * e\<^sup>T * F * v\<^sup>T \<sqinter> top * e\<^sup>T * -F"
using assms(1) by auto
also have "... = -F * e * top \<sqinter> v * F * e * top \<sqinter> top * e\<^sup>T * F * v\<^sup>T \<sqinter> top * e\<^sup>T * -F"
by (smt comp_associative comp_inf_covector inf.sup_monoid.add_assoc inf_top.left_neutral vector_top_closed)
also have "... = (-F \<sqinter> v * F) * e * top \<sqinter> top * e\<^sup>T * F * v\<^sup>T \<sqinter> top * e\<^sup>T * -F"
using injective_comp_right_dist_inf assms(3) mult.semigroup_axioms semigroup.assoc by fastforce
also have "... = (-F \<sqinter> v * F) * e * top \<sqinter> top * e\<^sup>T * (F * v\<^sup>T \<sqinter> -F)"
using injective_comp_right_dist_inf assms(3) conv_dist_comp inf.sup_monoid.add_assoc mult.semigroup_axioms univalent_comp_left_dist_inf semigroup.assoc by fastforce
also have "... = (-F \<sqinter> v * F) * e * top * top * e\<^sup>T * (F * v\<^sup>T \<sqinter> -F)"
by (metis inf_top_right vector_export_comp vector_top_closed)
also have "... = (-F \<sqinter> v * F) * e * top * e\<^sup>T * (F * v\<^sup>T \<sqinter> -F)"
by (simp add: comp_associative)
also have "... \<le> (-F \<sqinter> v * F) * (F * v\<^sup>T \<sqinter> -F)"
by (smt assms(3) conv_dist_comp mult.semigroup_axioms mult_left_isotone shunt_bijective symmetric_top_closed top_right_mult_increasing semigroup.assoc)
also have "... = (-F \<sqinter> v * F) * ((v * F)\<^sup>T \<sqinter> -F)"
by (simp add: assms(1) conv_dist_comp)
also have "... = (-F \<sqinter> v * F) * (-F \<sqinter> v * F)\<^sup>T"
using assms(1) conv_complement conv_dist_inf by (simp add: inf.sup_monoid.add_commute)
also have "... \<le> (-F \<sqinter> v) * (F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (-F \<sqinter> v)\<^sup>T"
proof -
let ?Fv = "F \<sqinter> v"
have "-F \<sqinter> v * F \<le> -F \<sqinter> v * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star>"
using assms(5) inf.sup_right_isotone mult_right_isotone comp_associative by auto
also have "... \<le> -F \<sqinter> v * (F \<sqinter> v)\<^sup>\<star>"
proof -
have "v * v\<^sup>T \<le> 1"
by (simp add: assms(2))
hence "v * v\<^sup>T * F \<le> F"
using assms(1) dual_order.trans mult_left_isotone by blast
hence "v * v\<^sup>T * F\<^sup>T\<^sup>\<star> * F\<^sup>\<star> \<le> F"
by (metis assms(1) mult_1_right preorder_idempotent star.circ_sup_one_right_unfold star.circ_transitive_equal star_one star_simulation_right_equal mult_assoc)
hence "v * (F \<sqinter> v)\<^sup>T * F\<^sup>T\<^sup>\<star> * F\<^sup>\<star> \<le> F"
by (meson conv_isotone dual_order.trans inf.cobounded2 inf.sup_monoid.add_commute mult_left_isotone mult_right_isotone)
hence "v * (F \<sqinter> v)\<^sup>T * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star> \<le> F"
by (meson conv_isotone dual_order.trans inf.cobounded2 inf.sup_monoid.add_commute mult_left_isotone mult_right_isotone comp_isotone conv_dist_inf inf.cobounded1 star_isotone)
hence "-F \<sqinter> v * (F \<sqinter> v)\<^sup>T * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star> \<le> bot"
- using eq_iff p_antitone pseudo_complement by auto
+ using order.eq_iff p_antitone pseudo_complement by auto
hence "(-F \<sqinter> v * (F \<sqinter> v)\<^sup>T * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star>) \<squnion> v * (v \<sqinter> F)\<^sup>\<star> \<le> v * (v \<sqinter> F)\<^sup>\<star>"
using bot_least le_bot by fastforce
hence "(-F \<squnion> v * (v \<sqinter> F)\<^sup>\<star>) \<sqinter> (v * (F \<sqinter> v)\<^sup>T * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star> \<squnion> v * (v \<sqinter> F)\<^sup>\<star>) \<le> v * (v \<sqinter> F)\<^sup>\<star>"
by (simp add: sup_inf_distrib2)
hence "(-F \<squnion> v * (v \<sqinter> F)\<^sup>\<star>) \<sqinter> v * ((F \<sqinter> v)\<^sup>T * (F \<sqinter> v)\<^sup>T\<^sup>\<star> \<squnion> 1) * (v \<sqinter> F)\<^sup>\<star> \<le> v * (v \<sqinter> F)\<^sup>\<star>"
by (simp add: inf.sup_monoid.add_commute mult.semigroup_axioms mult_left_dist_sup mult_right_dist_sup semigroup.assoc)
hence "(-F \<squnion> v * (v \<sqinter> F)\<^sup>\<star>) \<sqinter> v * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (v \<sqinter> F)\<^sup>\<star> \<le> v * (v \<sqinter> F)\<^sup>\<star>"
by (simp add: star_left_unfold_equal sup_commute)
hence "-F \<sqinter> v * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (v \<sqinter> F)\<^sup>\<star> \<le> v * (v \<sqinter> F)\<^sup>\<star>"
using comp_inf.mult_right_sub_dist_sup_left inf.order_lesseq_imp by blast
thus ?thesis
by (simp add: inf.sup_monoid.add_commute)
qed
also have "... \<le> (v \<sqinter> -F * (F \<sqinter> v)\<^sup>T\<^sup>\<star>) * (F \<sqinter> v)\<^sup>\<star>"
by (metis dedekind_2 conv_star_commute inf.sup_monoid.add_commute)
also have "... \<le> (v \<sqinter> -F * F\<^sup>T\<^sup>\<star>) * (F \<sqinter> v)\<^sup>\<star>"
using conv_isotone inf.sup_right_isotone mult_left_isotone mult_right_isotone star_isotone by auto
also have "... = (v \<sqinter> -F * F) * (F \<sqinter> v)\<^sup>\<star>"
by (metis assms(1) equivalence_comp_right_complement mult_left_one star_one star_simulation_right_equal)
also have "... = (-F \<sqinter> v) * (F \<sqinter> v)\<^sup>\<star>"
using assms(1) equivalence_comp_right_complement inf.sup_monoid.add_commute by auto
finally have "-F \<sqinter> v * F \<le> (-F \<sqinter> v) * (F \<sqinter> v)\<^sup>\<star>"
by simp
hence "(-F \<sqinter> v * F) * (-F \<sqinter> v * F)\<^sup>T \<le> (-F \<sqinter> v) * (F \<sqinter> v)\<^sup>\<star> * ((-F \<sqinter> v) * (F \<sqinter> v)\<^sup>\<star>)\<^sup>T"
by (simp add: comp_isotone conv_isotone)
also have "... = (-F \<sqinter> v) * (F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (-F \<sqinter> v)\<^sup>T"
by (simp add: comp_associative conv_dist_comp conv_star_commute)
finally show ?thesis
by simp
qed
also have "... \<le> (-F \<sqinter> v) * ((F \<sqinter> v\<^sup>\<star>) \<squnion> (F \<sqinter> v\<^sup>T\<^sup>\<star>)) * (-F \<sqinter> v)\<^sup>T"
proof -
have "(F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v)\<^sup>T\<^sup>\<star> \<le> F\<^sup>\<star> * F\<^sup>T\<^sup>\<star>"
using fc_isotone by auto
also have "... \<le> F * F"
by (metis assms(1) preorder_idempotent star.circ_sup_one_left_unfold star.circ_transitive_equal star_right_induct_mult)
finally have 21: "(F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v)\<^sup>T\<^sup>\<star> \<le> F"
using assms(1) dual_order.trans by blast
have "(F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v)\<^sup>T\<^sup>\<star> \<le> v\<^sup>\<star> * v\<^sup>T\<^sup>\<star>"
by (simp add: fc_isotone)
hence "(F \<sqinter> v)\<^sup>\<star> * (F \<sqinter> v)\<^sup>T\<^sup>\<star> \<le> F \<sqinter> v\<^sup>\<star> * v\<^sup>T\<^sup>\<star>"
using 21 by simp
also have "... = F \<sqinter> (v\<^sup>\<star> \<squnion> v\<^sup>T\<^sup>\<star>)"
by (simp add: assms(2) cancel_separate_eq)
finally show ?thesis
by (metis assms(4, 6) comp_associative comp_inf.semiring.distrib_left comp_isotone inf_pp_semi_commute mult_left_isotone regular_closed_inf)
qed
also have "... \<le> (-F \<sqinter> v) * (F \<sqinter> v\<^sup>T\<^sup>\<star>) * (-F \<sqinter> v)\<^sup>T \<squnion> (-F \<sqinter> v) * (F \<sqinter> v\<^sup>\<star>) * (-F \<sqinter> v)\<^sup>T"
by (simp add: mult_left_dist_sup mult_right_dist_sup)
also have "... \<le> (-F \<sqinter> v) * (-F \<sqinter> v)\<^sup>T \<squnion> (-F \<sqinter> v) * (-F \<sqinter> v)\<^sup>T"
proof -
have "(-F \<sqinter> v) * (F \<sqinter> v\<^sup>T\<^sup>\<star>) \<le> (-F \<sqinter> v) * ((F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star> \<sqinter> v\<^sup>T\<^sup>\<star>)"
by (simp add: assms(5) inf.coboundedI1 mult_right_isotone)
also have "... = (-F \<sqinter> v) * ((F \<sqinter> v)\<^sup>T * (F \<sqinter> v)\<^sup>T\<^sup>\<star> * (F \<sqinter> v)\<^sup>\<star> \<sqinter> v\<^sup>T\<^sup>\<star>) \<squnion> (-F \<sqinter> v) * ((F \<sqinter> v)\<^sup>\<star> \<sqinter> v\<^sup>T\<^sup>\<star>)"
by (metis comp_associative comp_inf.mult_right_dist_sup mult_left_dist_sup star.circ_loop_fixpoint)
also have "... \<le> (-F \<sqinter> v) * (F \<sqinter> v)\<^sup>T * top \<squnion> (-F \<sqinter> v) * ((F \<sqinter> v)\<^sup>\<star> \<sqinter> v\<^sup>T\<^sup>\<star>)"
by (simp add: comp_associative comp_isotone inf.coboundedI2 inf.sup_monoid.add_commute le_supI1)
also have "... \<le> (-F \<sqinter> v) * (F \<sqinter> v)\<^sup>T * top \<squnion> (-F \<sqinter> v) * (v\<^sup>\<star> \<sqinter> v\<^sup>T\<^sup>\<star>)"
- by (smt comp_inf.mult_right_isotone comp_inf.semiring.add_mono eq_iff inf.cobounded2 inf.sup_monoid.add_commute mult_right_isotone star_isotone)
+ by (smt comp_inf.mult_right_isotone comp_inf.semiring.add_mono order.eq_iff inf.cobounded2 inf.sup_monoid.add_commute mult_right_isotone star_isotone)
also have "... \<le> bot \<squnion> (-F \<sqinter> v) * (v\<^sup>\<star> \<sqinter> v\<^sup>T\<^sup>\<star>)"
by (metis assms(1, 2) forests_bot_1 comp_associative comp_inf.semiring.add_right_mono mult_semi_associative vector_bot_closed)
also have "... \<le> -F \<sqinter> v"
by (simp add: assms(2) acyclic_star_inf_conv)
finally have 22: "(-F \<sqinter> v) * (F \<sqinter> v\<^sup>T\<^sup>\<star>) \<le> -F \<sqinter> v"
by simp
have "((-F \<sqinter> v) * (F \<sqinter> v\<^sup>T\<^sup>\<star>))\<^sup>T = (F \<sqinter> v\<^sup>\<star>) * (-F \<sqinter> v)\<^sup>T"
by (simp add: assms(1) conv_dist_inf conv_star_commute conv_dist_comp)
hence "(F \<sqinter> v\<^sup>\<star>) * (-F \<sqinter> v)\<^sup>T \<le> (-F \<sqinter> v)\<^sup>T"
using 22 conv_isotone by fastforce
thus ?thesis
using 22 by (metis assms(4, 6) comp_associative comp_inf.pp_comp_semi_commute comp_inf.semiring.add_mono comp_isotone inf_pp_commute mult_left_isotone)
qed
also have "... = (-F \<sqinter> v) * (-F \<sqinter> v)\<^sup>T"
by simp
also have "... \<le> v * v\<^sup>T"
by (simp add: comp_isotone conv_isotone)
also have "... \<le> 1"
by (simp add: assms(2))
thus ?thesis
using calculation dual_order.trans by blast
qed
have 3: "top * ?i * top = top"
proof -
have 31: "regular (e\<^sup>T * -F * v * F * e)"
using assms(3, 4, 6) arc_regular regular_mult_closed by auto
have 32: "bijective ((top * e\<^sup>T)\<^sup>T)"
using assms(3) by (simp add: conv_dist_comp)
have "top * ?i * top = top * (v \<sqinter> -F * e * top) * ((top * e\<^sup>T * F)\<^sup>T \<sqinter> top)"
by (simp add: comp_associative comp_inf_vector_1)
also have "... = (top \<sqinter> (-F * e * top)\<^sup>T) * v * ((top * e\<^sup>T * F)\<^sup>T \<sqinter> top)"
using comp_inf_vector conv_dist_comp by auto
also have "... = (-F * e * top)\<^sup>T * v * (top * e\<^sup>T * F)\<^sup>T"
by simp
also have "... = top\<^sup>T * e\<^sup>T * -F\<^sup>T * v * F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T"
by (simp add: comp_associative conv_complement conv_dist_comp)
finally have 33: "top * ?i * top = top * e\<^sup>T * -F * v * F * e * top"
by (simp add: assms(1))
have "top * ?i * top \<noteq> bot"
proof (rule ccontr)
assume "\<not> top * (v \<sqinter> - F * e * top \<sqinter> top * e\<^sup>T * F) * top \<noteq> bot"
hence "top * e\<^sup>T * -F * v * F * e * top = bot"
using 33 by auto
hence "e\<^sup>T * -F * v * F * e = bot"
using 31 tarski comp_associative le_bot by fastforce
hence "top * (-F * v * F * e)\<^sup>T \<le> -(e\<^sup>T)"
by (metis comp_associative conv_complement_sub_leq conv_involutive p_bot schroeder_5_p)
hence "top * e\<^sup>T * F\<^sup>T * v\<^sup>T * -F\<^sup>T \<le> -(e\<^sup>T)"
by (simp add: comp_associative conv_complement conv_dist_comp)
hence "v * F * e * top * e\<^sup>T \<le> F"
by (metis assms(1, 4) comp_associative conv_dist_comp schroeder_3_p symmetric_top_closed)
hence "v * F * e * top * top * e\<^sup>T \<le> F"
by (simp add: comp_associative)
hence "v * F * e * top \<le> F * (top * e\<^sup>T)\<^sup>T"
using 32 by (metis shunt_bijective comp_associative conv_involutive)
hence "v * F * e * top \<le> F * e * top"
using comp_associative conv_dist_comp by auto
hence "v\<^sup>\<star> * F * e * top \<le> F * e * top"
using comp_associative star_left_induct_mult_iff by auto
hence "e\<^sup>T * F * e * top \<le> F * e * top"
by (meson assms(9) mult_left_isotone order_trans)
hence "e\<^sup>T * F * e * top * (e * top)\<^sup>T \<le> F"
using 32 shunt_bijective assms(3) mult_assoc by auto
hence 34: "e\<^sup>T * F * e * top * top * e\<^sup>T \<le> F"
by (metis conv_dist_comp mult.semigroup_axioms symmetric_top_closed semigroup.assoc)
hence "e\<^sup>T \<le> F"
proof -
have "e\<^sup>T \<le> e\<^sup>T * e * e\<^sup>T"
by (metis conv_involutive ex231c)
also have "... \<le> e\<^sup>T * F * e * e\<^sup>T"
using assms(1) comp_associative mult_left_isotone mult_right_isotone by fastforce
also have "... \<le> e\<^sup>T * F * e * top * top * e\<^sup>T"
by (simp add: mult_left_isotone top_right_mult_increasing vector_mult_closed)
finally show ?thesis
using 34 by simp
qed
hence 35: "e \<le> F"
using assms(1) conv_order by fastforce
have "top * (F * e)\<^sup>T \<le> - e"
using assms(8) comp_associative schroeder_4_p by auto
hence "top * e\<^sup>T * F \<le> - e"
by (simp add: assms(1) comp_associative conv_dist_comp)
hence "(top * e\<^sup>T)\<^sup>T * e \<le> - F"
using schroeder_3_p by auto
hence "e * top * e \<le> - F"
by (simp add: conv_dist_comp)
hence "e \<le> - F"
by (simp add: assms(3) arc_top_arc)
hence "e \<le> F \<sqinter> - F"
using 35 inf.boundedI by blast
hence "e = bot"
using bot_unique by auto
thus "False"
using assms(10) by auto
qed
thus ?thesis
by (metis assms(3, 4, 6) arc_regular regular_closed_inf regular_closed_top regular_conv_closed regular_mult_closed semiring.mult_not_zero tarski)
qed
have "bijective (?i * top) \<and> bijective (?i\<^sup>T * top)"
using 1 2 3 arc_expanded by blast
thus ?thesis
by blast
qed
subsubsection \<open>Comparison of edge weights\<close>
text \<open>
In this section we compare the weight of the \<open>selected_edge\<close> with other edges of interest.
Theorems 8, 9, 10 and 11 are supporting lemmas.
For example, Theorem 8 is used to show that the \<open>selected_edge\<close> has its source inside and its target outside the component it is chosen for.
\<close>
text \<open>Theorem 8\<close>
lemma e_leq_c_c_complement_transpose_general:
assumes "e = minarc (c * -(c)\<^sup>T \<sqinter> g)"
and "regular c"
shows "e \<le> c * -(c)\<^sup>T"
proof -
have "e \<le> -- (c * - c\<^sup>T \<sqinter> g)"
using assms(1) minarc_below order_trans by blast
also have "... \<le> -- (c * - c\<^sup>T)"
using order_lesseq_imp pp_isotone_inf by blast
also have "... = c * - c\<^sup>T"
using assms(2) regular_mult_closed by auto
finally show ?thesis
by simp
qed
text \<open>Theorem 9\<close>
lemma x_leq_c_transpose_general:
assumes "forest h"
and "vector c"
and "x\<^sup>T * top \<le> forest_components(h) * e * top"
and "e \<le> c * -c\<^sup>T"
and "c = forest_components(h) * c"
shows "x \<le> c\<^sup>T"
proof -
let ?H = "forest_components h"
have "x \<le> top * x"
using top_left_mult_increasing by blast
also have "... \<le> (?H * e * top)\<^sup>T"
using assms(3) conv_dist_comp conv_order by force
also have "... = top * e\<^sup>T * ?H"
using assms(1) comp_associative conv_dist_comp forest_components_equivalence by auto
also have "... \<le> top * (c * - c\<^sup>T)\<^sup>T * ?H"
by (simp add: assms(4) conv_isotone mult_left_isotone mult_right_isotone)
also have "... = top * (- c * c\<^sup>T) * ?H"
by (simp add: conv_complement conv_dist_comp)
also have "... \<le> top * c\<^sup>T * ?H"
by (metis mult_left_isotone top.extremum mult_assoc)
also have "... = c\<^sup>T * ?H"
using assms(1, 2) component_is_vector vector_conv_covector by auto
also have "... = c\<^sup>T"
by (metis assms(1, 5) fch_equivalence conv_dist_comp)
finally show ?thesis
by simp
qed
text \<open>Theorem 10\<close>
lemma x_leq_c_complement_general:
assumes "vector c"
and "c * c\<^sup>T \<le> forest_components h"
and "x \<le> c\<^sup>T"
and "x \<le> -forest_components h"
shows "x \<le> -c"
proof -
let ?H = "forest_components h"
have "x \<le> - ?H \<sqinter> c\<^sup>T"
using assms(3, 4) by auto
also have "... \<le> - c"
proof -
have "c \<sqinter> c\<^sup>T \<le> ?H"
using assms(1, 2) vector_covector by auto
hence "-?H \<sqinter> c \<sqinter> c\<^sup>T \<le> bot"
using inf.sup_monoid.add_assoc p_antitone pseudo_complement by fastforce
thus ?thesis
using le_bot p_shunting_swap pseudo_complement by blast
qed
finally show ?thesis
by simp
qed
text \<open>Theorem 11\<close>
lemma sum_e_below_sum_x_when_outgoing_same_component_general:
assumes "e = minarc (c * -(c)\<^sup>T \<sqinter> g)"
and "regular c"
and "forest h"
and "vector c"
and "x\<^sup>T * top \<le> (forest_components h) * e * top"
and "c = (forest_components h) * c"
and "c * c\<^sup>T \<le> forest_components h"
and "x \<le> - forest_components h \<sqinter> -- g"
and "symmetric g"
and "arc x"
and "c \<noteq> bot"
shows "sum (e \<sqinter> g) \<le> sum (x \<sqinter> g)"
proof -
let ?H = "forest_components h"
have 1:"e \<le> c * - c\<^sup>T"
using assms(1, 2) e_leq_c_c_complement_transpose_general by auto
have 2: "x \<le> c\<^sup>T"
using 1 assms(3, 4, 5, 6) x_leq_c_transpose_general by auto
hence "x \<le> -c"
using assms(4, 7, 8) x_leq_c_complement_general inf.boundedE by blast
hence "x \<le> - c \<sqinter> c\<^sup>T"
using 2 by simp
hence "x \<le> - c * c\<^sup>T"
using assms(4) by (simp add: vector_complement_closed vector_covector)
hence "x\<^sup>T \<le> c\<^sup>T\<^sup>T * - c\<^sup>T"
by (metis conv_complement conv_dist_comp conv_isotone)
hence 3: "x\<^sup>T \<le> c * - c\<^sup>T"
by simp
hence "x \<le> -- g"
using assms(8) by auto
hence "x\<^sup>T \<le> -- g"
using assms(9) conv_complement conv_isotone by fastforce
hence "x\<^sup>T \<sqinter> c * - c\<^sup>T \<sqinter> -- g \<noteq> bot"
using 3 by (metis assms(10, 11) comp_inf.semiring.mult_not_zero conv_dist_comp
conv_involutive inf.orderE mult_right_zero top.extremum)
hence "x\<^sup>T \<sqinter> c * - c\<^sup>T \<sqinter> g \<noteq> bot"
using inf.sup_monoid.add_commute pp_inf_bot_iff by auto
hence "sum (minarc (c * - c\<^sup>T \<sqinter> g) \<sqinter> (c * - c\<^sup>T \<sqinter> g)) \<le> sum (x\<^sup>T \<sqinter> c * - c\<^sup>T \<sqinter> g)"
using assms(10) minarc_min inf.sup_monoid.add_assoc by auto
hence "sum (e \<sqinter> c * - c\<^sup>T \<sqinter> g) \<le> sum (x\<^sup>T \<sqinter> c * - c\<^sup>T \<sqinter> g)"
using assms(1) inf.sup_monoid.add_assoc by auto
hence "sum (e \<sqinter> g) \<le> sum (x\<^sup>T \<sqinter> g)"
using 1 3 by (metis inf.orderE)
hence "sum (e \<sqinter> g) \<le> sum (x \<sqinter> g)"
using assms(9) sum_symmetric by auto
thus ?thesis
by simp
qed
lemma sum_e_below_sum_x_when_outgoing_same_component:
assumes "symmetric g"
and "vector j"
and "forest h"
and "x \<le> - forest_components h \<sqinter> -- g"
and "x\<^sup>T * top \<le> forest_components h * selected_edge h j g * top"
and "j \<noteq> bot"
and "arc x"
shows "sum (selected_edge h j g \<sqinter> g) \<le> sum (x \<sqinter> g)"
proof -
let ?e = "selected_edge h j g"
let ?c = "choose_component (forest_components h) j"
let ?H = "forest_components h"
show ?thesis
proof (rule sum_e_below_sum_x_when_outgoing_same_component_general)
next
show "?e = minarc (?c * - ?c\<^sup>T \<sqinter> g)"
by simp
next
show "regular ?c"
using component_is_regular by auto
next
show "forest h"
by (simp add: assms(3))
next
show "vector ?c"
by (simp add: assms(2, 6) component_is_vector)
next
show "x\<^sup>T * top \<le> ?H * ?e * top"
by (simp add: assms(5))
next
show "?c = ?H * ?c"
using component_single by auto
next
show "?c * ?c\<^sup>T \<le> ?H"
by (simp add: component_is_connected)
next
show "x \<le> -?H \<sqinter> -- g"
using assms(4) by auto
next
show "symmetric g"
by (simp add: assms(1))
next
show "arc x"
by (simp add: assms(7))
next
show "?c \<noteq> bot"
using assms(2, 5 , 6, 7) inf_bot_left le_bot minarc_bot mult_left_zero mult_right_zero by fastforce
qed
qed
text \<open>
If there is a path in the \<open>big_forest\<close> from an edge between components, $a$, to the \<open>selected_edge\<close>, $e$, then the weight of $e$ is no greater than the weight of $a$.
This is because either,
\begin{itemize}
\item the edges $a$ and $e$ are adjacent the same component so that we can use \<open>sum_e_below_sum_x_when_outgoing_same_component\<close>, or
\item there is at least one edge between $a$ and $e$, namely $x$, the edge incoming to the component that $e$ is outgoing from.
The path from $a$ to $e$ is split on $x$ using \<open>big_forest_path_split_disj\<close>.
We show that the weight of $e$ is no greater than the weight of $x$ by making use of lemma \<open>sum_e_below_sum_x_when_outgoing_same_component\<close>.
We define $x$ in a way that we can show that the weight of $x$ is no greater than the weight of $a$ using the invariant.
Then, it follows that the weight of $e$ is no greater than the weight of $a$ owing to transitivity.
\end{itemize}
\<close>
lemma a_to_e_in_bigforest:
assumes "symmetric g"
and "f \<le> --g"
and "vector j"
and "forest h"
and "big_forest (forest_components h) d"
and "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
and "(\<forall> a b . bf_between_arcs a b (forest_components h) d \<and> a \<le> -(forest_components h) \<sqinter> -- g \<and> b \<le> d \<longrightarrow> sum(b \<sqinter> g) \<le> sum(a \<sqinter> g))"
and "regular d"
and "j \<noteq> bot"
and "b = selected_edge h j g"
and "arc a"
and "bf_between_arcs a b (forest_components h) (d \<squnion> selected_edge h j g)"
and "a \<le> - forest_components h \<sqinter> -- g"
and "regular h"
shows "sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof -
let ?p = "path f h j g"
let ?e = "selected_edge h j g"
let ?F = "forest_components f"
let ?H = "forest_components h"
have "sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof (cases "a\<^sup>T * top \<le> ?H * ?e * top")
case True
show "a\<^sup>T * top \<le> ?H * ?e * top \<Longrightarrow> sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof-
have "sum (?e \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof (rule sum_e_below_sum_x_when_outgoing_same_component)
show "symmetric g"
using assms(1) by auto
next
show "vector j"
using assms(3) by blast
next
show "forest h"
by (simp add: assms(4))
next
show "a \<le> - ?H \<sqinter> -- g"
using assms(13) by auto
next
show "a\<^sup>T * top \<le> ?H * ?e * top"
using True by auto
next
show "j \<noteq> bot"
by (simp add: assms(9))
next
show "arc a"
by (simp add: assms(11))
qed
thus ?thesis
using assms(10) by auto
qed
next
case False
show "\<not> a\<^sup>T * top \<le> ?H * ?e * top \<Longrightarrow> sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof -
let ?d' = "d \<squnion> ?e"
let ?x = "d \<sqinter> top * ?e\<^sup>T * ?H \<sqinter> (?H * d\<^sup>T)\<^sup>\<star> * ?H * a\<^sup>T * top"
have 61: "arc (?x)"
proof (rule shows_arc_x)
show "big_forest ?H d"
by (simp add: assms(5))
next
show "bf_between_arcs a ?e ?H d"
proof -
have 611: "bf_between_arcs a b ?H (d \<squnion> b)"
using assms(10, 12) by auto
have 616: "regular h"
using assms(14) by auto
have "regular a"
using 611 bf_between_arcs_def arc_regular by fastforce
thus ?thesis
using 616 by (smt big_forest_path_split_disj assms(4, 8, 10, 12) bf_between_arcs_def fch_equivalence minarc_regular regular_closed_star regular_conv_closed regular_mult_closed)
qed
next
show "(?H * d)\<^sup>+ \<le> - ?H"
using assms(5) big_forest_def by blast
next
show "\<not> a\<^sup>T * top \<le> ?H * ?e * top"
by (simp add: False)
next
show "regular a"
using assms(12) bf_between_arcs_def arc_regular by auto
next
show "regular ?e"
using minarc_regular by auto
next
show "regular ?H"
using assms(14) pp_dist_star regular_conv_closed regular_mult_closed by auto
next
show "regular d"
using assms(8) by auto
qed
have 62: "bijective (a\<^sup>T * top)"
by (simp add: assms(11))
have 63: "bijective (?x * top)"
using 61 by simp
have 64: "?x \<le> (?H * d\<^sup>T)\<^sup>\<star> * ?H * a\<^sup>T * top"
by simp
hence "?x * top \<le> (?H * d\<^sup>T)\<^sup>\<star> * ?H * a\<^sup>T * top"
using mult_left_isotone inf_vector_comp by auto
hence "a\<^sup>T * top \<le> ((?H * d\<^sup>T)\<^sup>\<star> * ?H)\<^sup>T * ?x * top"
using 62 63 64 by (smt bijective_reverse mult_assoc)
also have "... = ?H * (d * ?H)\<^sup>\<star> * ?x * top"
using conv_dist_comp conv_star_commute by auto
also have "... = (?H * d)\<^sup>\<star> * ?H * ?x * top"
by (simp add: star_slide)
finally have "a\<^sup>T * top \<le> (?H * d)\<^sup>\<star> * ?H * ?x * top"
by simp
hence 65: "bf_between_arcs a ?x ?H d"
using 61 assms(12) bf_between_arcs_def by blast
have 66: "?x \<le> d"
by (simp add: inf.sup_monoid.add_assoc)
hence x_below_a: "sum (?x \<sqinter> g) \<le> sum (a \<sqinter> g)"
using 65 bf_between_arcs_def assms(7, 13) by blast
have "sum (?e \<sqinter> g) \<le> sum (?x \<sqinter> g)"
proof (rule sum_e_below_sum_x_when_outgoing_same_component)
show "symmetric g"
using assms(1) by auto
next
show "vector j"
using assms(3) by blast
next
show "forest h"
by (simp add: assms(4))
next
show "?x \<le> - ?H \<sqinter> -- g"
proof -
have 67: "?x \<le> - ?H"
using 66 assms(5) big_forest_def order_lesseq_imp by blast
have "?x \<le> d"
by (simp add: conv_isotone inf.sup_monoid.add_assoc)
also have "... \<le> f \<squnion> f\<^sup>T"
proof -
have "h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T = f \<squnion> f\<^sup>T"
by (simp add: assms(6))
thus ?thesis
by (metis (no_types) le_supE sup.absorb_iff2 sup.idem)
qed
also have "... \<le> -- g"
using assms(1, 2) conv_complement conv_isotone by fastforce
finally have "?x \<le> -- g"
by simp
thus ?thesis
by (simp add: 67)
qed
next
show "?x\<^sup>T * top \<le> ?H * ?e * top"
proof -
have "?x \<le> top * ?e\<^sup>T * ?H"
using inf.coboundedI1 by auto
hence "?x\<^sup>T \<le> ?H * ?e * top"
using conv_dist_comp conv_dist_inf conv_star_commute inf.orderI inf.sup_monoid.add_assoc inf.sup_monoid.add_commute mult_assoc by auto
hence "?x\<^sup>T * top \<le> ?H * ?e * top * top"
by (simp add: mult_left_isotone)
thus ?thesis
by (simp add: mult_assoc)
qed
next
show "j \<noteq> bot"
by (simp add: assms(9))
next
show "arc (?x)"
using 61 by blast
qed
hence "sum (?e \<sqinter> g) \<le> sum (a \<sqinter> g)"
using x_below_a order.trans by blast
thus ?thesis
by (simp add: assms(10))
qed
qed
thus ?thesis
by simp
qed
subsubsection \<open>Maintenance of algorithm invariants\<close>
text \<open>
In this section, most of the work is done to maintain the invariants of the inner and outer loops of the algorithm.
In particular, we use \<open>exists_a_w\<close> to maintain that $f$ can be extended to a minimum spanning forest.
\<close>
lemma boruvka_exchange_spanning_inv:
assumes "forest v"
and "v\<^sup>\<star> * e\<^sup>T = e\<^sup>T"
and "i \<le> v \<sqinter> top * e\<^sup>T * w\<^sup>T\<^sup>\<star>"
and "arc i"
and "arc e"
and "v \<le> --g"
and "w \<le> --g"
and "e \<le> --g"
and "components g \<le> forest_components v"
shows "i \<le> (v \<sqinter> -i)\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
proof -
have 1: "(v \<sqinter> -i \<sqinter> -i\<^sup>T) * (v\<^sup>T \<sqinter> -i \<sqinter> -i\<^sup>T) \<le> 1"
using assms(1) comp_isotone order.trans inf.cobounded1 by blast
have 2: "bijective (i * top) \<and> bijective (e\<^sup>T * top)"
using assms(4, 5) mult_assoc by auto
have "i \<le> v * (top * e\<^sup>T * w\<^sup>T\<^sup>\<star>)\<^sup>T"
using assms(3) covector_mult_closed covector_restrict_comp_conv order_lesseq_imp vector_top_closed by blast
also have "... \<le> v * w\<^sup>T\<^sup>\<star>\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T"
by (simp add: comp_associative conv_dist_comp)
also have "... \<le> v * w\<^sup>\<star> * e * top"
by (simp add: conv_star_commute)
also have "... = v * w\<^sup>\<star> * e * e\<^sup>T * e * top"
using assms(5) arc_eq_1 by (simp add: comp_associative)
also have "... \<le> v * w\<^sup>\<star> * e * e\<^sup>T * top"
by (simp add: comp_associative mult_right_isotone)
also have "... \<le> (--g) * (--g)\<^sup>\<star> * (--g) * e\<^sup>T * top"
using assms(6, 7, 8) by (simp add: comp_isotone star_isotone)
also have "... \<le> (--g)\<^sup>\<star> * e\<^sup>T * top"
by (metis comp_isotone mult_left_isotone star.circ_increasing star.circ_transitive_equal)
also have "... \<le> v\<^sup>T\<^sup>\<star> * v\<^sup>\<star> * e\<^sup>T * top"
by (simp add: assms(9) mult_left_isotone)
also have "... \<le> v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
by (simp add: assms(2) comp_associative)
finally have "i \<le> v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
by simp
hence "i * top \<le> v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
by (metis comp_associative mult_left_isotone vector_top_closed)
hence "e\<^sup>T * top \<le> v\<^sup>T\<^sup>\<star>\<^sup>T * i * top"
using 2 by (metis bijective_reverse mult_assoc)
also have "... = v\<^sup>\<star> * i * top"
by (simp add: conv_star_commute)
also have "... \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
proof -
have 3: "i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
using star.circ_loop_fixpoint sup_right_divisibility mult_assoc by auto
have "(v \<sqinter> i) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> i * top * i * top"
by (metis comp_isotone inf.cobounded1 inf.sup_monoid.add_commute mult_left_isotone top.extremum)
also have "... \<le> i * top"
by simp
finally have 4: "(v \<sqinter> i) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
using 3 dual_order.trans by blast
have 5: "(v \<sqinter> -i \<sqinter> -i\<^sup>T) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
by (metis mult_left_isotone star.circ_increasing star.left_plus_circ)
have "v\<^sup>+ \<le> -1"
by (simp add: assms(1))
hence "v * v \<le> -1"
by (metis mult_left_isotone order_trans star.circ_increasing star.circ_plus_same)
hence "v * 1 \<le> -v\<^sup>T"
by (simp add: schroeder_5_p)
hence "v \<le> -v\<^sup>T"
by simp
hence "v \<sqinter> v\<^sup>T \<le> bot"
by (simp add: bot_unique pseudo_complement)
hence 7: "v \<sqinter> i\<^sup>T \<le> bot"
by (metis assms(3) comp_inf.mult_right_isotone conv_dist_inf inf.boundedE inf.le_iff_sup le_bot)
hence "(v \<sqinter> i\<^sup>T) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> bot"
using le_bot semiring.mult_zero_left by fastforce
hence 6: "(v \<sqinter> i\<^sup>T) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
using bot_least le_bot by blast
have 8: "v = (v \<sqinter> i) \<squnion> (v \<sqinter> i\<^sup>T) \<squnion> (v \<sqinter> -i \<sqinter> -i\<^sup>T)"
proof -
have 81: "regular i"
by (simp add: assms(4) arc_regular)
have "(v \<sqinter> i\<^sup>T) \<squnion> (v \<sqinter> -i \<sqinter> -i\<^sup>T) = (v \<sqinter> -i)"
using 7 by (metis comp_inf.coreflexive_comp_inf_complement inf_import_p inf_p le_bot maddux_3_11_pp top.extremum)
hence "(v \<sqinter> i) \<squnion> (v \<sqinter> i\<^sup>T) \<squnion> (v \<sqinter> -i \<sqinter> -i\<^sup>T) = (v \<sqinter> i) \<squnion> (v \<sqinter> -i)"
by (simp add: sup.semigroup_axioms semigroup.assoc)
also have "... = v"
using 81 by (metis maddux_3_11_pp)
finally show ?thesis
by simp
qed
have "(v \<sqinter> i) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<squnion> (v \<sqinter> i\<^sup>T) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<squnion> (v \<sqinter> -i \<sqinter> -i\<^sup>T) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
using 4 5 6 by simp
hence "((v \<sqinter> i) \<squnion> (v \<sqinter> i\<^sup>T) \<squnion> (v \<sqinter> -i \<sqinter> -i\<^sup>T)) * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
by (simp add: mult_right_dist_sup)
hence "v * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
using 8 by auto
hence "i * top \<squnion> v * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
using 3 by auto
hence 9:"v\<^sup>\<star> * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
by (simp add: star_left_induct_mult mult_assoc)
have "v\<^sup>\<star> * i * top \<le> v\<^sup>\<star> * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
using 3 mult_right_isotone mult_assoc by auto
thus ?thesis
using 9 order.trans by blast
qed
finally have "e\<^sup>T * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * i * top"
by simp
hence "i * top \<le> (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star>\<^sup>T * e\<^sup>T * top"
using 2 by (metis bijective_reverse mult_assoc)
also have "... = (v\<^sup>T \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
using comp_inf.inf_vector_comp conv_complement conv_dist_inf conv_star_commute inf.sup_monoid.add_commute by auto
also have "... \<le> ((v \<sqinter> -i \<sqinter> -i\<^sup>T) \<squnion> (v\<^sup>T \<sqinter> -i \<sqinter> -i\<^sup>T))\<^sup>\<star> * e\<^sup>T * top"
by (simp add: mult_left_isotone star_isotone)
finally have "i \<le> ((v\<^sup>T \<sqinter> -i \<sqinter> -i\<^sup>T) \<squnion> (v \<sqinter> -i \<sqinter> -i\<^sup>T))\<^sup>\<star> * e\<^sup>T * top"
using dual_order.trans top_right_mult_increasing sup_commute by auto
also have "... = (v\<^sup>T \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * (v \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
using 1 cancel_separate_1 by (simp add: sup_commute)
also have "... \<le> (v\<^sup>T \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * v\<^sup>\<star> * e\<^sup>T * top"
by (simp add: inf_assoc mult_left_isotone mult_right_isotone star_isotone)
also have "... = (v\<^sup>T \<sqinter> -i \<sqinter> -i\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
using assms(2) mult_assoc by simp
also have "... \<le> (v\<^sup>T \<sqinter> -i\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
by (metis mult_left_isotone star_isotone inf.cobounded2 inf.left_commute inf.sup_monoid.add_commute)
also have "... = (v \<sqinter> -i)\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
using conv_complement conv_dist_inf by auto
finally show ?thesis
by simp
qed
lemma exists_a_w:
assumes "symmetric g"
and "forest f"
and "f \<le> --g"
and "regular f"
and "(\<exists>w . minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T)"
and "vector j"
and "regular j"
and "forest h"
and "forest_components h \<le> forest_components f"
and "big_forest (forest_components h) d"
and "d * top \<le> - j"
and "forest_components h * j = j"
and "forest_components f = (forest_components h * (d \<squnion> d\<^sup>T))\<^sup>\<star> * forest_components h"
and "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
and "(\<forall> a b . bf_between_arcs a b (forest_components h) d \<and> a \<le> -(forest_components h) \<sqinter> -- g \<and> b \<le> d \<longrightarrow> sum(b \<sqinter> g) \<le> sum(a \<sqinter> g))"
and "regular d"
and "selected_edge h j g \<le> - forest_components f"
and "selected_edge h j g \<noteq> bot"
and "j \<noteq> bot"
and "regular h"
and "h \<le> --g"
shows "\<exists>w. minimum_spanning_forest w g \<and>
f \<sqinter> - (selected_edge h j g)\<^sup>T \<sqinter> - (path f h j g) \<squnion> (f \<sqinter> - (selected_edge h j g)\<^sup>T \<sqinter> (path f h j g))\<^sup>T \<squnion> (selected_edge h j g) \<le> w \<squnion> w\<^sup>T"
proof -
let ?p = "path f h j g"
let ?e = "selected_edge h j g"
let ?f = "(f \<sqinter> -?e\<^sup>T \<sqinter> -?p) \<squnion> (f \<sqinter> -?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> ?e"
let ?F = "forest_components f"
let ?H = "forest_components h"
let ?ec = "choose_component (forest_components h) j * - choose_component (forest_components h) j\<^sup>T \<sqinter> g"
from assms(4) obtain w where 2: "minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T"
using assms(5) by blast
hence 3: "regular w \<and> regular f \<and> regular ?e"
by (metis assms(4) minarc_regular minimum_spanning_forest_def spanning_forest_def)
have 5: "equivalence ?F"
using assms(2) forest_components_equivalence by auto
have "?e\<^sup>T * top * ?e\<^sup>T = ?e\<^sup>T"
by (metis arc_conv_closed arc_top_arc coreflexive_bot_closed coreflexive_symmetric minarc_arc minarc_bot_iff semiring.mult_not_zero)
hence "?e\<^sup>T * top * ?e\<^sup>T \<le> -?F"
using 5 assms(17) conv_complement conv_isotone by fastforce
hence 6: "?e * ?F * ?e = bot"
using assms(2) le_bot triple_schroeder_p by simp
let ?q = "w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?v = "(w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>)) \<squnion> ?q\<^sup>T"
have 7: "regular ?q"
using 3 regular_closed_star regular_conv_closed regular_mult_closed by auto
have 8: "injective ?v"
proof (rule kruskal_exchange_injective_inv_1)
show "injective w"
using 2 minimum_spanning_forest_def spanning_forest_def by blast
next
show "covector (top * ?e * w\<^sup>T\<^sup>\<star>)"
by (simp add: covector_mult_closed)
next
show "top * ?e * w\<^sup>T\<^sup>\<star> * w\<^sup>T \<le> top * ?e * w\<^sup>T\<^sup>\<star>"
by (simp add: mult_right_isotone star.right_plus_below_circ mult_assoc)
next
show "coreflexive ((top * ?e * w\<^sup>T\<^sup>\<star>)\<^sup>T * (top * ?e * w\<^sup>T\<^sup>\<star>) \<sqinter> w\<^sup>T * w)"
using 2 by (metis comp_inf.semiring.mult_not_zero forest_bot kruskal_injective_inv_3 minarc_arc minarc_bot_iff minimum_spanning_forest_def semiring.mult_not_zero spanning_forest_def)
qed
have 9: "components g \<le> forest_components ?v"
proof (rule kruskal_exchange_spanning_inv_1)
show "injective (w \<sqinter> - (top *?e * w\<^sup>T\<^sup>\<star>) \<squnion> (w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)\<^sup>T)"
using 8 by simp
next
show "regular (w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)"
using 7 by simp
next
show "components g \<le> forest_components w"
using 2 minimum_spanning_forest_def spanning_forest_def by blast
qed
have 10: "spanning_forest ?v g"
proof (unfold spanning_forest_def, intro conjI)
show "injective ?v"
using 8 by auto
next
show "acyclic ?v"
proof (rule kruskal_exchange_acyclic_inv_1)
show "pd_kleene_allegory_class.acyclic w"
using 2 minimum_spanning_forest_def spanning_forest_def by blast
next
show "covector (top * ?e * w\<^sup>T\<^sup>\<star>)"
by (simp add: covector_mult_closed)
qed
next
show "?v \<le> --g"
proof (rule sup_least)
show "w \<sqinter> - (top * ?e * w\<^sup>T\<^sup>\<star>) \<le> - - g"
using 7 inf.coboundedI1 minimum_spanning_forest_def spanning_forest_def 2 by blast
next
show "(w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)\<^sup>T \<le> - - g"
using 2 by (metis assms(1) conv_complement conv_isotone inf.coboundedI1 minimum_spanning_forest_def spanning_forest_def)
qed
next
show "components g \<le> forest_components ?v"
using 9 by simp
next
show "regular ?v"
using 3 regular_closed_star regular_conv_closed regular_mult_closed by auto
qed
have 11: "sum (?v \<sqinter> g) = sum (w \<sqinter> g)"
proof -
have "sum (?v \<sqinter> g) = sum (w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>) \<sqinter> g) + sum (?q\<^sup>T \<sqinter> g)"
using 2 by (smt conv_complement conv_top epm_8 inf_import_p inf_top_right regular_closed_top vector_top_closed minimum_spanning_forest_def spanning_forest_def sum_disjoint)
also have "... = sum (w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>) \<sqinter> g) + sum (?q \<sqinter> g)"
by (simp add: assms(1) sum_symmetric)
also have "... = sum (((w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>)) \<squnion> ?q) \<sqinter> g)"
using inf_commute inf_left_commute sum_disjoint by simp
also have "... = sum (w \<sqinter> g)"
using 3 7 8 maddux_3_11_pp by auto
finally show ?thesis
by simp
qed
have 12: "?v \<squnion> ?v\<^sup>T = w \<squnion> w\<^sup>T"
proof -
have "?v \<squnion> ?v\<^sup>T = (w \<sqinter> -?q) \<squnion> ?q\<^sup>T \<squnion> (w\<^sup>T \<sqinter> -?q\<^sup>T) \<squnion> ?q"
using conv_complement conv_dist_inf conv_dist_sup inf_import_p sup_assoc by simp
also have "... = w \<squnion> w\<^sup>T"
using 3 7 conv_complement conv_dist_inf inf_import_p maddux_3_11_pp sup_monoid.add_assoc sup_monoid.add_commute by auto
finally show ?thesis
by simp
qed
have 13: "?v * ?e\<^sup>T = bot"
proof (rule kruskal_reroot_edge)
show "injective (?e\<^sup>T * top)"
using assms(18) minarc_arc minarc_bot_iff by blast
next
show "pd_kleene_allegory_class.acyclic w"
using 2 minimum_spanning_forest_def spanning_forest_def by simp
qed
have "?v \<sqinter> ?e \<le> ?v \<sqinter> top * ?e"
using inf.sup_right_isotone top_left_mult_increasing by simp
also have "... \<le> ?v * (top * ?e)\<^sup>T"
using covector_restrict_comp_conv covector_mult_closed vector_top_closed by simp
finally have 14: "?v \<sqinter> ?e = bot"
using 13 by (metis conv_dist_comp mult_assoc le_bot mult_left_zero)
let ?i = "?v \<sqinter> (- ?F) * ?e * top \<sqinter> top * ?e\<^sup>T * ?F"
let ?w = "(?v \<sqinter> -?i) \<squnion> ?e"
have 15: "regular ?i"
using 3 regular_closed_star regular_conv_closed regular_mult_closed by simp
have 16: "?F \<le> -?i"
proof -
have 161: "bijective (?e * top)"
using assms(18) minarc_arc minarc_bot_iff by auto
have "?i \<le> - ?F * ?e * top"
using inf.cobounded2 inf.coboundedI1 by blast
also have "... = - (?F * ?e * top)"
using 161 comp_bijective_complement by (simp add: mult_assoc)
finally have "?i \<le> - (?F * ?e * top)"
by blast
hence 162: "?i \<sqinter> ?F \<le> - (?F * ?e * top)"
using inf.coboundedI1 by blast
have "?i \<sqinter> ?F \<le> ?F \<sqinter> (top * ?e\<^sup>T * ?F)"
by (meson inf_le1 inf_le2 le_infI order_trans)
also have "... \<le> ?F * (top * ?e\<^sup>T * ?F)\<^sup>T"
by (simp add: covector_mult_closed covector_restrict_comp_conv)
also have "... = ?F * ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T"
by (simp add: conv_dist_comp mult_assoc)
also have "... = ?F * ?F * ?e * top"
by (simp add: conv_dist_comp conv_star_commute)
also have "... = ?F * ?e * top"
by (simp add: 5 preorder_idempotent)
finally have "?i \<sqinter> ?F \<le> ?F * ?e * top"
by simp
hence "?i \<sqinter> ?F \<le> ?F * ?e * top \<sqinter> - (?F * ?e * top)"
using 162 inf.bounded_iff by blast
also have "... = bot"
by simp
finally show ?thesis
using le_bot p_antitone_iff pseudo_complement by blast
qed
have 17: "?i \<le> top * ?e\<^sup>T * (?F \<sqinter> ?v \<sqinter> -?i)\<^sup>T\<^sup>\<star>"
proof -
have "?i \<le> ?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> * (?F \<sqinter> ?v)\<^sup>\<star>"
using 2 8 12 by (smt inf.sup_right_isotone kruskal_forest_components_inf mult_right_isotone mult_assoc)
also have "... = ?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> * (1 \<squnion> (?F \<sqinter> ?v)\<^sup>\<star> * (?F \<sqinter> ?v))"
using star_left_unfold_equal star.circ_right_unfold_1 by auto
also have "... = ?v \<sqinter> - ?F * ?e * top \<sqinter> (top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> \<squnion> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> * (?F \<sqinter> ?v)\<^sup>\<star> * (?F \<sqinter> ?v))"
by (simp add: mult_left_dist_sup mult_assoc)
also have "... = (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star>) \<squnion> (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> * (?F \<sqinter> ?v)\<^sup>\<star> * (?F \<sqinter> ?v))"
using comp_inf.semiring.distrib_left by blast
also have "... \<le> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> \<squnion> (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> * (?F \<sqinter> ?v)\<^sup>\<star> * (?F \<sqinter> ?v))"
using comp_inf.semiring.add_right_mono inf_le2 by blast
also have "... \<le> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> \<squnion> (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * (?F\<^sup>T \<sqinter> ?v\<^sup>T)\<^sup>\<star> * (?F \<sqinter> ?v)\<^sup>\<star> * (?F \<sqinter> ?v))"
by (simp add: conv_dist_inf)
also have "... \<le> top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> \<squnion> (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F\<^sup>T\<^sup>\<star> * ?F\<^sup>\<star> * (?F \<sqinter> ?v))"
proof -
have "top * ?e\<^sup>T * (?F\<^sup>T \<sqinter> ?v\<^sup>T)\<^sup>\<star> * (?F \<sqinter> ?v)\<^sup>\<star> * (?F \<sqinter> ?v) \<le> top * ?e\<^sup>T * ?F\<^sup>T\<^sup>\<star> * ?F\<^sup>\<star> * (?F \<sqinter> ?v)"
using star_isotone by (simp add: comp_isotone)
hence "?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * (?F\<^sup>T \<sqinter> ?v\<^sup>T)\<^sup>\<star> * (?F \<sqinter> ?v)\<^sup>\<star> * (?F \<sqinter> ?v) \<le> ?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F\<^sup>T\<^sup>\<star> * ?F\<^sup>\<star> * (?F \<sqinter> ?v)"
using inf.sup_right_isotone by blast
thus ?thesis
using sup_right_isotone by blast
qed
also have "... = top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> \<squnion> (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F\<^sup>\<star> * ?F\<^sup>\<star> * (?F \<sqinter> ?v))"
using 5 by auto
also have "... = top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> \<squnion> (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F * ?F * (?F \<sqinter> ?v))"
by (simp add: assms(2) forest_components_star)
also have "... = top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star> \<squnion> (?v \<sqinter> - ?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F * (?F \<sqinter> ?v))"
using 5 mult.semigroup_axioms preorder_idempotent semigroup.assoc by fastforce
also have "... = top * ?e\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>\<star>"
proof -
have "?e * top * ?e\<^sup>T \<le> 1"
using assms(18) arc_expanded minarc_arc minarc_bot_iff by auto
hence "?F * ?e * top * ?e\<^sup>T \<le> ?F * 1"
by (metis comp_associative comp_isotone mult_semi_associative star.circ_transitive_equal)
hence "?v * ?v\<^sup>T * ?F * ?e * top * ?e\<^sup>T \<le> 1 * ?F * 1"
using 8 by (smt comp_isotone mult_assoc)
hence 171: "?v * ?v\<^sup>T * ?F * ?e * top * ?e\<^sup>T \<le> ?F"
by simp
hence "?v * (?F \<sqinter> ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T \<le> ?F"
proof -
have "?v * (?F \<sqinter> ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T \<le> ?v * ?v\<^sup>T * ?F * ?e * top * ?e\<^sup>T"
by (simp add: conv_dist_inf mult_left_isotone mult_right_isotone)
thus ?thesis
using 171 order_trans by blast
qed
hence 172: "-?F * ((?F \<sqinter> ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T)\<^sup>T \<le> -?v"
by (smt schroeder_4_p comp_associative order_lesseq_imp pp_increasing)
have "-?F * ((?F \<sqinter> ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T)\<^sup>T = -?F * ?e\<^sup>T\<^sup>T * top\<^sup>T * ?e\<^sup>T * ?F\<^sup>T * (?F \<sqinter> ?v)\<^sup>T\<^sup>T"
by (simp add: comp_associative conv_dist_comp)
also have "... = -?F * ?e * top * ?e\<^sup>T * ?F * (?F \<sqinter> ?v)"
using 5 by auto
also have "... = -?F * ?e * top * top * ?e\<^sup>T * ?F * (?F \<sqinter> ?v)"
using comp_associative by auto
also have "... = -?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F * (?F \<sqinter> ?v)"
by (smt comp_associative comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_inf_covector inf_vector_comp vector_top_closed)
finally have "-?F * ((?F \<sqinter> ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T)\<^sup>T = -?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F * (?F \<sqinter> ?v)"
by simp
hence "-?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F * (?F \<sqinter> ?v) \<le> -?v"
using 172 by auto
hence "?v \<sqinter> -?F * ?e * top \<sqinter> top * ?e\<^sup>T * ?F * (?F \<sqinter> ?v) \<le> bot"
by (smt bot_unique inf.sup_monoid.add_commute p_shunting_swap pseudo_complement)
thus ?thesis
using le_bot sup_monoid.add_0_right by blast
qed
also have "... = top * ?e\<^sup>T * (?F \<sqinter> ?v \<sqinter> -?i)\<^sup>T\<^sup>\<star>"
using 16 by (smt comp_inf.coreflexive_comp_inf_complement inf_top_right p_bot pseudo_complement top.extremum)
finally show ?thesis
by blast
qed
have 18: "?i \<le> top * ?e\<^sup>T * ?w\<^sup>T\<^sup>\<star>"
proof -
have "?i \<le> top * ?e\<^sup>T * (?F \<sqinter> ?v \<sqinter> -?i)\<^sup>T\<^sup>\<star>"
using 17 by simp
also have "... \<le> top * ?e\<^sup>T * (?v \<sqinter> -?i)\<^sup>T\<^sup>\<star>"
- using mult_right_isotone conv_isotone star_isotone inf.cobounded2 inf.sup_monoid.add_assoc by (simp add: inf.sup_monoid.add_assoc eq_iff inf.sup_monoid.add_commute)
+ using mult_right_isotone conv_isotone star_isotone inf.cobounded2 inf.sup_monoid.add_assoc by (simp add: inf.sup_monoid.add_assoc order.eq_iff inf.sup_monoid.add_commute)
also have "... \<le> top * ?e\<^sup>T * ((?v \<sqinter> -?i) \<squnion> ?e)\<^sup>T\<^sup>\<star>"
using mult_right_isotone conv_isotone star_isotone sup_ge1 by simp
finally show ?thesis
by blast
qed
have 19: "?i \<le> top * ?e\<^sup>T * ?v\<^sup>T\<^sup>\<star>"
proof -
have "?i \<le> top * ?e\<^sup>T * (?F \<sqinter> ?v \<sqinter> -?i)\<^sup>T\<^sup>\<star>"
using 17 by simp
also have "... \<le> top * ?e\<^sup>T * (?v \<sqinter> -?i)\<^sup>T\<^sup>\<star>"
- using mult_right_isotone conv_isotone star_isotone inf.cobounded2 inf.sup_monoid.add_assoc by (simp add: inf.sup_monoid.add_assoc eq_iff inf.sup_monoid.add_commute)
+ using mult_right_isotone conv_isotone star_isotone inf.cobounded2 inf.sup_monoid.add_assoc by (simp add: inf.sup_monoid.add_assoc order.eq_iff inf.sup_monoid.add_commute)
also have "... \<le> top * ?e\<^sup>T * (?v)\<^sup>T\<^sup>\<star>"
using mult_right_isotone conv_isotone star_isotone by auto
finally show ?thesis
by blast
qed
have 20: "f \<squnion> f\<^sup>T \<le> (?v \<sqinter> -?i \<sqinter> -?i\<^sup>T) \<squnion> (?v\<^sup>T \<sqinter> -?i \<sqinter> -?i\<^sup>T)"
proof (rule kruskal_edge_between_components_2)
show "?F \<le> - ?i"
using 16 by simp
next
show "injective f"
by (simp add: assms(2))
next
show "f \<squnion> f\<^sup>T \<le> w \<sqinter> - (top * ?e * w\<^sup>T\<^sup>\<star>) \<squnion> (w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)\<^sup>T \<squnion> (w \<sqinter> - (top * ?e * w\<^sup>T\<^sup>\<star>) \<squnion> (w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)\<^sup>T)\<^sup>T"
using 2 12 by (metis conv_dist_sup conv_involutive conv_isotone le_supI sup_commute)
qed
have "minimum_spanning_forest ?w g \<and> ?f \<le> ?w \<squnion> ?w\<^sup>T"
proof (intro conjI)
have 211: "?e\<^sup>T \<le> ?v\<^sup>\<star>"
proof (rule kruskal_edge_arc_1[where g=g and h="?ec"])
show "?e \<le> -- ?ec"
using minarc_below by blast
next
show "?ec \<le> g"
using assms(4) inf.cobounded2 by (simp add: boruvka_inner_invariant_def boruvka_outer_invariant_def conv_dist_inf)
next
show "symmetric g"
by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def)
next
show "components g \<le> forest_components (w \<sqinter> - (top * ?e * w\<^sup>T\<^sup>\<star>) \<squnion> (w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)\<^sup>T)"
using 9 by simp
next
show "(w \<sqinter> - (top * ?e * w\<^sup>T\<^sup>\<star>) \<squnion> (w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)\<^sup>T) * ?e\<^sup>T = bot"
using 13 by blast
qed
have 212: "arc ?i"
proof (rule boruvka_edge_arc)
show "equivalence ?F"
by (simp add: 5)
next
show "forest ?v"
using 10 spanning_forest_def by blast
next
show "arc ?e"
using assms(18) minarc_arc minarc_bot_iff by blast
next
show "regular ?F"
using 3 regular_closed_star regular_conv_closed regular_mult_closed by auto
next
show "?F \<le> forest_components (?F \<sqinter> ?v)"
by (simp add: 12 2 8 kruskal_forest_components_inf)
next
show "regular ?v"
using 10 spanning_forest_def by blast
next
show "?v * ?e\<^sup>T = bot"
using 13 by auto
next
show "?e * ?F * ?e = bot"
by (simp add: 6)
next
show "?e\<^sup>T \<le> ?v\<^sup>\<star>"
using 211 by auto
next
show "?e \<noteq> bot"
by (simp add: assms(18))
qed
show "minimum_spanning_forest ?w g"
proof (unfold minimum_spanning_forest_def, intro conjI)
have "(?v \<sqinter> -?i) * ?e\<^sup>T \<le> ?v * ?e\<^sup>T"
using inf_le1 mult_left_isotone by simp
hence "(?v \<sqinter> -?i) * ?e\<^sup>T = bot"
using 13 le_bot by simp
hence 221: "?e * (?v \<sqinter> -?i)\<^sup>T = bot"
using conv_dist_comp conv_involutive conv_bot by force
have 222: "injective ?w"
proof (rule injective_sup)
show "injective (?v \<sqinter> -?i)"
using 8 by (simp add: injective_inf_closed)
next
show "coreflexive (?e * (?v \<sqinter> -?i)\<^sup>T)"
using 221 by simp
next
show "injective ?e"
by (metis arc_injective minarc_arc coreflexive_bot_closed coreflexive_injective minarc_bot_iff)
qed
show "spanning_forest ?w g"
proof (unfold spanning_forest_def, intro conjI)
show "injective ?w"
using 222 by simp
next
show "acyclic ?w"
proof (rule kruskal_exchange_acyclic_inv_2)
show "acyclic ?v"
using 10 spanning_forest_def by blast
next
show "injective ?v"
using 8 by simp
next
show "?i \<le>?v"
using inf.coboundedI1 by simp
next
show "bijective (?i\<^sup>T * top)"
using 212 by simp
next
show "bijective (?e * top)"
using 14 212 by (smt assms(4) comp_inf.idempotent_bot_closed conv_complement minarc_arc minarc_bot_iff p_bot regular_closed_bot semiring.mult_not_zero symmetric_top_closed)
next
show "?i \<le> top * ?e\<^sup>T *?v\<^sup>T\<^sup>\<star>"
using 19 by simp
next
show "?v * ?e\<^sup>T * top = bot"
using 13 by simp
qed
next
have "?w \<le> ?v \<squnion> ?e"
using inf_le1 sup_left_isotone by simp
also have "... \<le> --g \<squnion> ?e"
using 10 sup_left_isotone spanning_forest_def by blast
also have "... \<le> --g \<squnion> --h"
proof -
have 1: "--g \<le> --g \<squnion> --h"
by simp
have 2: "?e \<le> --g \<squnion> --h"
by (metis inf.coboundedI1 inf.sup_monoid.add_commute minarc_below order.trans p_dist_inf p_dist_sup sup.cobounded1)
thus ?thesis
using 1 2 by simp
qed
also have "... \<le> --g"
using assms(20, 21) by auto
finally show "?w \<le> --g"
by simp
next
have 223: "?i \<le> (?v \<sqinter> -?i)\<^sup>T\<^sup>\<star> * ?e\<^sup>T * top"
proof (rule boruvka_exchange_spanning_inv)
show "forest ?v"
using 10 spanning_forest_def by blast
next
show "?v\<^sup>\<star> * ?e\<^sup>T = ?e\<^sup>T"
using 13 by (smt conv_complement conv_dist_comp conv_involutive conv_star_commute dense_pp fc_top regular_closed_top star_absorb)
next
show "?i \<le> ?v \<sqinter> top * ?e\<^sup>T * ?w\<^sup>T\<^sup>\<star>"
using 18 inf.sup_monoid.add_assoc by auto
next
show "arc ?i"
using 212 by blast
next
show "arc ?e"
using assms(18) minarc_arc minarc_bot_iff by auto
next
show "?v \<le> --g"
using 10 spanning_forest_def by blast
next
show "?w \<le> --g"
proof -
have 2231: "?e \<le> --g"
by (metis inf.boundedE minarc_below pp_dist_inf)
have "?w \<le> ?v \<squnion> ?e"
using inf_le1 sup_left_isotone by simp
also have "... \<le> --g"
using 2231 10 spanning_forest_def sup_least by blast
finally show ?thesis
by blast
qed
next
show "?e \<le> --g"
by (metis inf.boundedE minarc_below pp_dist_inf)
next
show "components g \<le> forest_components ?v"
by (simp add: 9)
qed
have "components g \<le> forest_components ?v"
using 10 spanning_forest_def by auto
also have "... \<le> forest_components ?w"
proof (rule kruskal_exchange_forest_components_inv)
next
show "injective ((?v \<sqinter> -?i) \<squnion> ?e)"
using 222 by simp
next
show "regular ?i"
using 15 by simp
next
show "?e * top * ?e = ?e"
by (metis arc_top_arc minarc_arc minarc_bot_iff semiring.mult_not_zero)
next
show "?i \<le> top * ?e\<^sup>T * ?v\<^sup>T\<^sup>\<star>"
using 19 by blast
next
show "?v * ?e\<^sup>T * top = bot"
using 13 by simp
next
show "injective ?v"
using 8 by simp
next
show "?i \<le> ?v"
by (simp add: le_infI1)
next
show "?i \<le> (?v \<sqinter> -?i)\<^sup>T\<^sup>\<star> * ?e\<^sup>T * top"
using 223 by blast
qed
finally show "components g \<le> forest_components ?w"
by simp
next
show "regular ?w"
using 3 7 regular_conv_closed by simp
qed
next
have 224: "?e \<sqinter> g \<noteq> bot"
using assms(18) inf.left_commute inf_bot_right minarc_meet_bot by fastforce
have 225: "sum (?e \<sqinter> g) \<le> sum (?i \<sqinter> g)"
proof (rule a_to_e_in_bigforest)
show "symmetric g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "j \<noteq> bot"
by (simp add: assms(19))
next
show "f \<le> -- g"
by (simp add: assms(3))
next
show "vector j"
using assms(6) boruvka_inner_invariant_def by blast
next
show "forest h"
by (simp add: assms(8))
next
show " big_forest (forest_components h) d"
by (simp add: assms(10))
next
show "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
by (simp add: assms(14))
next
show "\<forall>a b. bf_between_arcs a b (?H) d \<and> a \<le> - ?H \<sqinter> - - g \<and> b \<le> d \<longrightarrow> sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
by (simp add: assms(15))
next
show "regular d"
using assms(16) by auto
next
show "?e = ?e"
by simp
next
show "arc ?i"
using 212 by blast
next
show "bf_between_arcs ?i ?e ?H (d \<squnion> ?e)"
proof -
have "d\<^sup>T * ?H * ?e = bot"
using assms(6, 7, 11, 12, 19) dT_He_eq_bot le_bot by blast
hence 251: "d\<^sup>T * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by simp
hence "d\<^sup>T * ?H * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (metis assms(8) forest_components_star star.circ_decompose_9 mult_assoc)
hence "d\<^sup>T * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
proof -
have "d\<^sup>T * ?H * d \<le> 1"
using assms(10) big_forest_def dTransHd_le_1 by blast
hence "d\<^sup>T * ?H * d * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (metis mult_left_isotone star.circ_circ_mult star_involutive star_one)
hence "d\<^sup>T * ?H * ?e \<squnion> d\<^sup>T * ?H * d * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
using 251 by simp
hence "d\<^sup>T * (1 \<squnion> ?H * d * (?H * d)\<^sup>\<star>) * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (simp add: comp_associative comp_left_dist_sup semiring.distrib_right)
thus ?thesis
by (simp add: star_left_unfold_equal)
qed
hence "?H * d\<^sup>T * (?H * d)\<^sup>\<star> * ?H * ?e \<le> ?H * (?H * d)\<^sup>\<star> * ?H * ?e"
by (simp add: mult_right_isotone mult_assoc)
hence "?H * d\<^sup>T * (?H * d)\<^sup>\<star> * ?H * ?e \<le> ?H * ?H * (d * ?H)\<^sup>\<star> * ?e"
by (smt star_slide mult_assoc)
hence "?H * d\<^sup>T * (?H * d)\<^sup>\<star> * ?H * ?e \<le> ?H * (d * ?H)\<^sup>\<star> * ?e"
by (metis assms(8) forest_components_star star.circ_decompose_9)
hence "?H * d\<^sup>T * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
using star_slide by auto
hence "?H * d * (?H * d)\<^sup>\<star> * ?H * ?e \<squnion> ?H * d\<^sup>T * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (smt le_supI star.circ_loop_fixpoint sup.cobounded2 sup_commute mult_assoc)
hence "(?H * (d \<squnion> d\<^sup>T)) * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (simp add: semiring.distrib_left semiring.distrib_right)
hence "(?H * (d \<squnion> d\<^sup>T))\<^sup>\<star> * (?H * d)\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (simp add: star_left_induct_mult mult_assoc)
hence 252: "(?H * (d \<squnion> d\<^sup>T))\<^sup>\<star> * ?H * ?e \<le> (?H * d)\<^sup>\<star> * ?H * ?e"
by (smt mult_left_dist_sup star.circ_transitive_equal star_slide star_sup_1 mult_assoc)
have "?i \<le> top * ?e\<^sup>T * ?F"
by auto
hence "?i\<^sup>T \<le> ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T"
by (simp add: conv_dist_comp conv_dist_inf mult_assoc)
hence "?i\<^sup>T * top \<le> ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T * top"
using comp_isotone by blast
also have "... = ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T"
by (simp add: vector_mult_closed)
also have "... = ?F * ?e\<^sup>T\<^sup>T * top\<^sup>T"
by (simp add: conv_dist_comp conv_star_commute)
also have "... = ?F * ?e * top"
by simp
also have "... = (?H * (d \<squnion> d\<^sup>T))\<^sup>\<star> * ?H * ?e * top"
by (simp add: assms(13))
also have "... \<le> (?H * d)\<^sup>\<star> * ?H * ?e * top"
by (simp add: 252 comp_isotone)
also have "... \<le> (?H * (d \<squnion> ?e))\<^sup>\<star> * ?H * ?e * top"
by (simp add: comp_isotone star_isotone)
finally have "?i\<^sup>T * top \<le> (?H * (d \<squnion> ?e))\<^sup>\<star> * ?H * ?e * top"
by blast
thus ?thesis
using 212 assms(18) bf_between_arcs_def minarc_arc minarc_bot_iff by blast
qed
next
show "?i \<le> - ?H \<sqinter> -- g"
proof -
have 241: "?i \<le> -?H"
using 16 assms(9) inf.order_lesseq_imp p_antitone_iff by blast
have "?i \<le> -- g"
using 10 inf.coboundedI1 spanning_forest_def by blast
thus ?thesis
using 241 inf_greatest by blast
qed
next
show "regular h"
using assms(20) by auto
qed
have "?v \<sqinter> ?e \<sqinter> -?i = bot"
using 14 by simp
hence "sum (?w \<sqinter> g) = sum (?v \<sqinter> -?i \<sqinter> g) + sum (?e \<sqinter> g)"
using sum_disjoint inf_commute inf_assoc by simp
also have "... \<le> sum (?v \<sqinter> -?i \<sqinter> g) + sum (?i \<sqinter> g)"
using 224 225 sum_plus_right_isotone by simp
also have "... = sum (((?v \<sqinter> -?i) \<squnion> ?i) \<sqinter> g)"
using sum_disjoint inf_le2 pseudo_complement by simp
also have "... = sum ((?v \<squnion> ?i) \<sqinter> (-?i \<squnion> ?i) \<sqinter> g)"
by (simp add: sup_inf_distrib2)
also have "... = sum ((?v \<squnion> ?i) \<sqinter> g)"
using 15 by (metis inf_top_right stone)
also have "... = sum (?v \<sqinter> g)"
by (simp add: inf.sup_monoid.add_assoc)
finally have "sum (?w \<sqinter> g) \<le> sum (?v \<sqinter> g)"
by simp
thus "\<forall>u . spanning_forest u g \<longrightarrow> sum (?w \<sqinter> g) \<le> sum (u \<sqinter> g)"
using 2 11 minimum_spanning_forest_def by auto
qed
next
have "?f \<le> f \<squnion> f\<^sup>T \<squnion> ?e"
by (smt conv_dist_inf inf_le1 sup_left_isotone sup_mono inf.order_lesseq_imp)
also have "... \<le> (?v \<sqinter> -?i \<sqinter> -?i\<^sup>T) \<squnion> (?v\<^sup>T \<sqinter> -?i \<sqinter> -?i\<^sup>T) \<squnion> ?e"
using 20 sup_left_isotone by simp
also have "... \<le> (?v \<sqinter> -?i) \<squnion> (?v\<^sup>T \<sqinter> -?i \<sqinter> -?i\<^sup>T) \<squnion> ?e"
by (metis inf.cobounded1 sup_inf_distrib2)
also have "... = ?w \<squnion> (?v\<^sup>T \<sqinter> -?i \<sqinter> -?i\<^sup>T)"
by (simp add: sup_assoc sup_commute)
also have "... \<le> ?w \<squnion> (?v\<^sup>T \<sqinter> -?i\<^sup>T)"
using inf.sup_right_isotone inf_assoc sup_right_isotone by simp
also have "... \<le> ?w \<squnion> ?w\<^sup>T"
using conv_complement conv_dist_inf conv_dist_sup sup_right_isotone by simp
finally show "?f \<le> ?w \<squnion> ?w\<^sup>T"
by simp
qed
thus ?thesis by auto
qed
lemma boruvka_outer_invariant_when_e_not_bot:
assumes "boruvka_inner_invariant j f h g d"
and "j \<noteq> bot"
and "selected_edge h j g \<le> - forest_components f"
and "selected_edge h j g \<noteq> bot"
shows "boruvka_outer_invariant (f \<sqinter> - selected_edge h j g\<^sup>T \<sqinter> - path f h j g \<squnion> (f \<sqinter> - selected_edge h j g\<^sup>T \<sqinter> path f h j g)\<^sup>T \<squnion> selected_edge h j g) g"
proof -
let ?c = "choose_component (forest_components h) j"
let ?p = "path f h j g"
let ?F = "forest_components f"
let ?H = "forest_components h"
let ?e = "selected_edge h j g"
let ?f' = "f \<sqinter> -?e\<^sup>T \<sqinter> -?p \<squnion> (f \<sqinter> -?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> ?e"
let ?d' = "d \<squnion> ?e"
let ?j' = "j \<sqinter> -?c"
show "boruvka_outer_invariant ?f' g"
proof (unfold boruvka_outer_invariant_def, intro conjI)
show "symmetric g"
by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def)
next
show "injective ?f'"
proof (rule kruskal_injective_inv)
show "injective (f \<sqinter> - ?e\<^sup>T)"
by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def injective_inf_closed)
show "covector (?p)"
using covector_mult_closed by simp
show "?p * (f \<sqinter> - ?e\<^sup>T)\<^sup>T \<le> ?p"
by (simp add: mult_right_isotone star.left_plus_below_circ star_plus mult_assoc)
show "?e \<le> ?p"
by (meson mult_left_isotone order.trans star_outer_increasing top.extremum)
show "?p * (f \<sqinter> - ?e\<^sup>T)\<^sup>T \<le> - ?e"
proof -
have "?p * (f \<sqinter> - ?e\<^sup>T)\<^sup>T \<le> ?p * f\<^sup>T"
by (simp add: conv_dist_inf mult_right_isotone)
also have "... \<le> top * ?e * (f)\<^sup>T\<^sup>\<star> * f\<^sup>T"
using conv_dist_inf star_isotone comp_isotone by simp
also have "... \<le> - ?e"
using assms(1, 4) boruvka_inner_invariant_def boruvka_outer_invariant_def kruskal_injective_inv_2 minarc_arc minarc_bot_iff by auto
finally show ?thesis .
qed
show "injective (?e)"
by (metis arc_injective coreflexive_bot_closed minarc_arc minarc_bot_iff semiring.mult_not_zero)
show "coreflexive (?p\<^sup>T * ?p \<sqinter> (f \<sqinter> - ?e\<^sup>T)\<^sup>T * (f \<sqinter> - ?e\<^sup>T))"
proof -
have "(?p\<^sup>T * ?p \<sqinter> (f \<sqinter> - ?e\<^sup>T)\<^sup>T * (f \<sqinter> - ?e\<^sup>T)) \<le> ?p\<^sup>T * ?p \<sqinter> f\<^sup>T * f"
using conv_dist_inf inf.sup_right_isotone mult_isotone by simp
also have "... \<le> (top * ?e * f\<^sup>T\<^sup>\<star>)\<^sup>T * (top * ?e * f\<^sup>T\<^sup>\<star>) \<sqinter> f\<^sup>T * f"
by (metis comp_associative comp_inf.coreflexive_transitive comp_inf.mult_right_isotone comp_isotone conv_isotone inf.cobounded1 inf.idem inf.sup_monoid.add_commute star_isotone top.extremum)
also have "... \<le> 1"
using assms(1, 4) boruvka_inner_invariant_def boruvka_outer_invariant_def kruskal_injective_inv_3 minarc_arc minarc_bot_iff by auto
finally show ?thesis
by simp
qed
qed
next
show "acyclic ?f'"
proof (rule kruskal_acyclic_inv)
show "acyclic (f \<sqinter> - ?e\<^sup>T)"
proof -
have f_intersect_below: "(f \<sqinter> - ?e\<^sup>T) \<le> f" by simp
have "acyclic f"
by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def)
thus ?thesis
using comp_isotone dual_order.trans star_isotone f_intersect_below by blast
qed
next
show "covector ?p"
by (metis comp_associative vector_top_closed)
next
show "(f \<sqinter> - ?e\<^sup>T \<sqinter> ?p)\<^sup>T * (f \<sqinter> - ?e\<^sup>T)\<^sup>\<star> * ?e = bot"
proof -
have "?e \<le> - (f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>)"
by (simp add: assms(3))
hence "?e * top * ?e \<le> - (f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>)"
by (metis arc_top_arc minarc_arc minarc_bot_iff semiring.mult_not_zero)
hence "?e\<^sup>T * top * ?e\<^sup>T \<le> - (f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>)\<^sup>T"
by (metis comp_associative conv_complement conv_dist_comp conv_isotone symmetric_top_closed)
hence "?e\<^sup>T * top * ?e\<^sup>T \<le> - (f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>)"
by (simp add: conv_dist_comp conv_star_commute)
hence "?e * (f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>) * ?e \<le> bot"
using triple_schroeder_p by auto
hence 1: "?e * f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> * ?e \<le> bot"
using mult_assoc by auto
have 2: "(f \<sqinter> - ?e\<^sup>T)\<^sup>T\<^sup>\<star> \<le> f\<^sup>T\<^sup>\<star>"
by (simp add: conv_dist_inf star_isotone)
have "(f \<sqinter> - ?e\<^sup>T \<sqinter> ?p)\<^sup>T * (f \<sqinter> - ?e\<^sup>T)\<^sup>\<star> * ?e \<le> (f \<sqinter> ?p)\<^sup>T * (f \<sqinter> - ?e\<^sup>T)\<^sup>\<star> * ?e"
by (simp add: comp_isotone conv_dist_inf inf.orderI inf.sup_monoid.add_assoc)
also have "... \<le> (f \<sqinter> ?p)\<^sup>T * f\<^sup>\<star> * ?e"
by (simp add: comp_isotone star_isotone)
also have "... \<le> (f \<sqinter> top * ?e * (f)\<^sup>T\<^sup>\<star>)\<^sup>T * f\<^sup>\<star> * ?e"
using 2 by (metis comp_inf.comp_isotone comp_inf.coreflexive_transitive comp_isotone conv_isotone inf.idem top.extremum)
also have "... = (f\<^sup>T \<sqinter> (top * ?e * f\<^sup>T\<^sup>\<star>)\<^sup>T) * f\<^sup>\<star> * ?e"
by (simp add: conv_dist_inf)
also have "... \<le> top * (f\<^sup>T \<sqinter> (top * ?e * f\<^sup>T\<^sup>\<star>)\<^sup>T) * f\<^sup>\<star> * ?e"
using top_left_mult_increasing mult_assoc by auto
also have "... = (top \<sqinter> top * ?e * f\<^sup>T\<^sup>\<star>) * f\<^sup>T * f\<^sup>\<star> * ?e"
- by (smt covector_comp_inf_1 covector_mult_closed eq_iff inf.sup_monoid.add_commute vector_top_closed)
+ by (smt covector_comp_inf_1 covector_mult_closed order.eq_iff inf.sup_monoid.add_commute vector_top_closed)
also have "... = top * ?e * f\<^sup>T\<^sup>\<star> * f\<^sup>T * f\<^sup>\<star> * ?e"
by simp
also have "... \<le> top * ?e * f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> * ?e"
by (smt conv_dist_comp conv_isotone conv_star_commute mult_left_isotone mult_right_isotone star.left_plus_below_circ mult_assoc)
also have "... \<le> bot"
using 1 covector_bot_closed le_bot mult_assoc by fastforce
finally show ?thesis
using le_bot by auto
qed
next
show "?e * (f \<sqinter> - ?e\<^sup>T)\<^sup>\<star> * ?e = bot"
proof -
have 1: "?e \<le> - ?F"
by (simp add: assms(3))
have 2: "injective f"
by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def)
have 3: "equivalence ?F"
using 2 forest_components_equivalence by simp
hence 4: "?e\<^sup>T = ?e\<^sup>T * top * ?e\<^sup>T"
by (smt arc_conv_closed arc_top_arc covector_complement_closed covector_conv_vector ex231e minarc_arc minarc_bot_iff pp_surjective regular_closed_top vector_mult_closed vector_top_closed)
also have "... \<le> - ?F" using 1 3 conv_isotone conv_complement calculation by fastforce
finally have 5: "?e * ?F * ?e = bot"
using 4 by (smt triple_schroeder_p le_bot pp_total regular_closed_top vector_top_closed)
have "(f \<sqinter> - ?e\<^sup>T)\<^sup>\<star> \<le> f\<^sup>\<star>"
by (simp add: star_isotone)
hence "?e * (f \<sqinter> - ?e\<^sup>T)\<^sup>\<star> * ?e \<le> ?e * f\<^sup>\<star> * ?e"
using mult_left_isotone mult_right_isotone by blast
also have "... \<le> ?e * ?F * ?e"
by (metis conv_star_commute forest_components_increasing mult_left_isotone mult_right_isotone star_involutive)
also have 6: "... = bot"
using 5 by simp
finally show ?thesis using 6 le_bot by blast
qed
next
show "forest_components (f \<sqinter> -?e\<^sup>T) \<le> - ?e"
proof -
have 1: "?e \<le> - ?F"
by (simp add: assms(3))
have "f \<sqinter> - ?e\<^sup>T \<le> f"
by simp
hence "forest_components (f \<sqinter> - ?e\<^sup>T) \<le> ?F"
using forest_components_isotone by blast
thus ?thesis
using 1 order_lesseq_imp p_antitone_iff by blast
qed
qed
next
show "?f' \<le> --g"
proof -
have 1: "(f \<sqinter> - ?e\<^sup>T \<sqinter> - ?p) \<le> --g"
by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def inf.coboundedI1)
have 2: "(f \<sqinter> - ?e\<^sup>T \<sqinter> ?p)\<^sup>T \<le> --g"
proof -
have "(f \<sqinter> - ?e\<^sup>T \<sqinter> ?p)\<^sup>T \<le> f\<^sup>T"
by (simp add: conv_isotone inf.sup_monoid.add_assoc)
also have "... \<le> --g"
by (metis assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def conv_complement conv_isotone)
finally show ?thesis
by simp
qed
have 3: "?e \<le> --g"
by (metis inf.boundedE minarc_below pp_dist_inf)
show ?thesis using 1 2 3
by simp
qed
next
show "regular ?f'"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
next
show "\<exists>w. minimum_spanning_forest w g \<and> ?f' \<le> w \<squnion> w\<^sup>T"
proof (rule exists_a_w)
show "symmetric g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "forest f"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "f \<le> --g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "regular f"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "(\<exists>w . minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T)"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "vector j"
using assms(1) boruvka_inner_invariant_def by blast
next
show "regular j"
using assms(1) boruvka_inner_invariant_def by blast
next
show "forest h"
using assms(1) boruvka_inner_invariant_def by blast
next
show "forest_components h \<le> forest_components f"
using assms(1) boruvka_inner_invariant_def by blast
next
show "big_forest (forest_components h) d"
using assms(1) boruvka_inner_invariant_def by blast
next
show "d * top \<le> - j"
using assms(1) boruvka_inner_invariant_def by blast
next
show "forest_components h * j = j"
using assms(1) boruvka_inner_invariant_def by blast
next
show "forest_components f = (forest_components h * (d \<squnion> d\<^sup>T))\<^sup>\<star> * forest_components h"
using assms(1) boruvka_inner_invariant_def by blast
next
show "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
using assms(1) boruvka_inner_invariant_def by blast
next
show "(\<forall> a b . bf_between_arcs a b (forest_components h) d \<and> a \<le> -(forest_components h) \<sqinter> -- g \<and> b \<le> d \<longrightarrow> sum(b \<sqinter> g) \<le> sum(a \<sqinter> g))"
using assms(1) boruvka_inner_invariant_def by blast
next
show "regular d"
using assms(1) boruvka_inner_invariant_def by blast
next
show "selected_edge h j g \<le> - forest_components f"
by (simp add: assms(3))
next
show "selected_edge h j g \<noteq> bot"
by (simp add: assms(4))
next
show "j \<noteq> bot"
by (simp add: assms(2))
next
show "regular h"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "h \<le> --g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
qed
qed
qed
lemma second_inner_invariant_when_e_not_bot:
assumes "boruvka_inner_invariant j f h g d"
and "j \<noteq> bot"
and "selected_edge h j g \<le> - forest_components f"
and "selected_edge h j g \<noteq> bot"
shows "boruvka_inner_invariant
(j \<sqinter> - choose_component (forest_components h) j)
(f \<sqinter> - selected_edge h j g\<^sup>T \<sqinter> - path f h j g \<squnion>
(f \<sqinter> - selected_edge h j g\<^sup>T \<sqinter> path f h j g)\<^sup>T \<squnion>
selected_edge h j g)
h g (d \<squnion> selected_edge h j g)"
proof -
let ?c = "choose_component (forest_components h) j"
let ?p = "path f h j g"
let ?F = "forest_components f"
let ?H = "forest_components h"
let ?e = "selected_edge h j g"
let ?f' = "f \<sqinter> -?e\<^sup>T \<sqinter> -?p \<squnion> (f \<sqinter> -?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> ?e"
let ?d' = "d \<squnion> ?e"
let ?j' = "j \<sqinter> -?c"
show "boruvka_inner_invariant ?j' ?f' h g ?d'"
proof (unfold boruvka_inner_invariant_def, intro conjI)
have 1: "boruvka_outer_invariant ?f' g"
using assms(1, 2, 3, 4) boruvka_outer_invariant_when_e_not_bot by blast
show "boruvka_outer_invariant ?f' g"
using assms(1, 2, 3, 4) boruvka_outer_invariant_when_e_not_bot by blast
show "g \<noteq> bot"
using assms(1) boruvka_inner_invariant_def by force
show "vector ?j'"
using assms(1, 2) boruvka_inner_invariant_def component_is_vector vector_complement_closed vector_inf_closed by simp
show "regular ?j'"
using assms(1) boruvka_inner_invariant_def by auto
show "boruvka_outer_invariant h g"
by (meson assms(1) boruvka_inner_invariant_def)
show "injective h"
by (meson assms(1) boruvka_inner_invariant_def)
show "pd_kleene_allegory_class.acyclic h"
by (meson assms(1) boruvka_inner_invariant_def)
show "?H \<le> forest_components ?f'"
proof -
have 2: "?F \<le> forest_components ?f'"
proof (rule components_disj_increasing)
show "regular ?p"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto[1]
next
show "regular ?e"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto[1]
next
show "injective ?f'"
using 1 boruvka_outer_invariant_def by blast
next
show "injective f"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by blast
qed
thus ?thesis
using assms(1) boruvka_inner_invariant_def dual_order.trans by blast
qed
show "big_forest ?H ?d'"
using assms(1, 2, 3, 4) big_forest_d_U_e boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "?d' * top \<le> -?j'"
proof -
have 31: "?d' * top = d * top \<squnion> ?e * top"
by (simp add: mult_right_dist_sup)
have 32: "d * top \<le> -?j'"
by (meson assms(1) boruvka_inner_invariant_def inf.coboundedI1 p_antitone_iff)
have "regular (?c * - ?c\<^sup>T)"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def component_is_regular regular_conv_closed regular_mult_closed by auto
hence "minarc(?c * - ?c\<^sup>T \<sqinter> g) = minarc(?c \<sqinter> - ?c\<^sup>T \<sqinter> g)"
by (metis component_is_vector covector_comp_inf inf_top.left_neutral vector_conv_compl)
also have "... \<le> -- (?c \<sqinter> - ?c\<^sup>T \<sqinter> g)"
using minarc_below by blast
also have "... \<le> -- ?c"
by (simp add: inf.sup_monoid.add_assoc)
also have "... = ?c"
using component_is_regular by auto
finally have "?e \<le> ?c"
by simp
hence "?e * top \<le> ?c"
by (metis component_is_vector mult_left_isotone)
also have "... \<le> -j \<squnion> ?c"
by simp
also have "... = - (j \<sqinter> - ?c)"
using component_is_regular by auto
finally have 33: "?e * top \<le> - (j \<sqinter> - ?c)"
by simp
show ?thesis
using 31 32 33 by auto
qed
next
show "?H * ?j' = ?j'"
using fc_j_eq_j_inv assms(1) boruvka_inner_invariant_def by blast
next
show "forest_components ?f' = (?H * (?d' \<squnion> ?d'\<^sup>T))\<^sup>\<star> * ?H"
proof -
have "forest_components ?f' = (f \<squnion> f\<^sup>T \<squnion> ?e \<squnion> ?e\<^sup>T)\<^sup>\<star>"
proof (rule simplify_forest_components_f)
show "regular ?p"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
next
show "regular ?e"
using minarc_regular by auto
next
show "injective ?f'"
using assms(1, 2, 3, 4) boruvka_outer_invariant_def boruvka_outer_invariant_when_e_not_bot by blast
next
show "injective f"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by blast
qed
also have "... = (h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T \<squnion> ?e \<squnion> ?e\<^sup>T)\<^sup>\<star>"
using assms(1) boruvka_inner_invariant_def by simp
also have "... = (h \<squnion> h\<^sup>T \<squnion> ?d' \<squnion> ?d'\<^sup>T)\<^sup>\<star>"
by (smt conv_dist_sup sup_monoid.add_assoc sup_monoid.add_commute)
also have "... = ((h \<squnion> h\<^sup>T)\<^sup>\<star> * (?d' \<squnion> ?d'\<^sup>T))\<^sup>\<star> * (h \<squnion> h\<^sup>T)\<^sup>\<star>"
by (metis star.circ_sup_9 sup_assoc)
finally show ?thesis
using assms(1) boruvka_inner_invariant_def forest_components_wcc by simp
qed
next
show "?f' \<squnion> ?f'\<^sup>T = h \<squnion> h\<^sup>T \<squnion> ?d' \<squnion> ?d'\<^sup>T"
proof -
have "?f' \<squnion> ?f'\<^sup>T = f \<sqinter> - ?e\<^sup>T \<sqinter> - ?p \<squnion> (f \<sqinter> - ?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> ?e \<squnion> (f \<sqinter> - ?e\<^sup>T \<sqinter> - ?p)\<^sup>T \<squnion> (f \<sqinter> - ?e\<^sup>T \<sqinter> ?p) \<squnion> ?e\<^sup>T"
by (simp add: conv_dist_sup sup_monoid.add_assoc)
also have "... = (f \<sqinter> - ?e\<^sup>T \<sqinter> - ?p) \<squnion> (f \<sqinter> - ?e\<^sup>T \<sqinter> ?p) \<squnion> (f \<sqinter> - ?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> (f \<sqinter> - ?e\<^sup>T \<sqinter> - ?p)\<^sup>T \<squnion> ?e\<^sup>T \<squnion> ?e"
by (simp add: sup.left_commute sup_commute)
also have "... = f \<squnion> f\<^sup>T \<squnion> ?e \<squnion> ?e\<^sup>T"
proof (rule simplify_f)
show "regular ?p"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
next
show "regular ?e"
using minarc_regular by blast
qed
also have "... = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T \<squnion> ?e \<squnion> ?e\<^sup>T"
using assms(1) boruvka_inner_invariant_def by auto
finally show ?thesis
by (smt conv_dist_sup sup.left_commute sup_commute)
qed
next
show "\<forall> a b . bf_between_arcs a b ?H ?d' \<and> a \<le> - ?H \<sqinter> -- g \<and> b \<le> ?d' \<longrightarrow> sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof (intro allI, rule impI, unfold bf_between_arcs_def)
fix a b
assume 1: "(arc a \<and> arc b \<and> a\<^sup>T * top \<le> (?H * ?d')\<^sup>\<star> * ?H * b * top) \<and> a \<le> - ?H \<sqinter> -- g \<and> b \<le> ?d'"
thus "sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof (cases "b = ?e")
case b_equals_e: True
thus ?thesis
proof (cases "a = ?e")
case True
thus ?thesis
using b_equals_e by auto
next
case a_ne_e: False
have "sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof (rule a_to_e_in_bigforest)
show "symmetric g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "j \<noteq> bot"
by (simp add: assms(2))
next
show "f \<le> -- g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "vector j"
using assms(1) boruvka_inner_invariant_def by blast
next
show "forest h"
using assms(1) boruvka_inner_invariant_def by blast
next
show " big_forest (forest_components h) d"
using assms(1) boruvka_inner_invariant_def by blast
next
show "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
using assms(1) boruvka_inner_invariant_def by blast
next
show "\<forall>a b. bf_between_arcs a b (?H) d \<and> a \<le> - ?H \<sqinter> - - g \<and> b \<le> d \<longrightarrow> sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
using assms(1) boruvka_inner_invariant_def by blast
next
show "regular d"
using assms(1) boruvka_inner_invariant_def by blast
next
show "b = ?e"
using b_equals_e by simp
next
show "arc a"
using 1 by simp
next
show "bf_between_arcs a b ?H ?d'"
using 1 bf_between_arcs_def by simp
next
show "a \<le> - ?H \<sqinter> -- g"
using 1 by simp
next
show "regular h"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
qed
thus ?thesis
by simp
qed
next
case b_not_equal_e: False
hence b_below_d: "b \<le> d"
using 1 by (metis assms(4) different_arc_in_sup_arc minarc_arc minarc_bot_iff)
thus ?thesis
proof (cases "?e \<le> d")
case True
hence "bf_between_arcs a b ?H d \<and> b \<le> d"
using 1 bf_between_arcs_def sup.absorb1 by auto
thus ?thesis
using 1 assms(1) boruvka_inner_invariant_def by blast
next
case e_not_less_than_d: False
have 71:"equivalence ?H"
using assms(1) fch_equivalence boruvka_inner_invariant_def by auto
hence 72: "bf_between_arcs a b ?H ?d' \<longleftrightarrow> bf_between_arcs a b ?H d \<or> (bf_between_arcs a ?e ?H d \<and> bf_between_arcs ?e b ?H d)"
proof (rule big_forest_path_split_disj)
show "arc ?e"
using assms(4) minarc_arc minarc_bot_iff by blast
next
show "regular a \<and> regular b \<and> regular ?e \<and> regular d \<and> regular ?H"
using assms(1) 1 boruvka_inner_invariant_def boruvka_outer_invariant_def arc_regular minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto
qed
thus ?thesis
proof (cases "bf_between_arcs a b ?H d")
case True
have "bf_between_arcs a b ?H d \<and> b \<le> d"
using 1 by (metis assms(4) True b_not_equal_e minarc_arc minarc_bot_iff different_arc_in_sup_arc)
thus ?thesis
using 1 assms(1) b_below_d boruvka_inner_invariant_def by auto
next
case False
have 73:"bf_between_arcs a ?e ?H d \<and> bf_between_arcs ?e b ?H d"
using 1 72 False bf_between_arcs_def by blast
have 74: "?e \<le> --g"
by (metis inf.boundedE minarc_below pp_dist_inf)
have "?e \<le> - ?H"
by (meson assms(1, 3) boruvka_inner_invariant_def dual_order.trans p_antitone_iff)
hence "?e \<le> - ?H \<sqinter> --g"
using 74 by simp
hence 75: "sum (b \<sqinter> g) \<le> sum (?e \<sqinter> g)"
using assms(1) b_below_d 73 boruvka_inner_invariant_def by blast
have 76: "bf_between_arcs a ?e ?H ?d'"
using 73 by (meson big_forest_path_split_disj assms(1) bf_between_arcs_def boruvka_inner_invariant_def boruvka_outer_invariant_def fch_equivalence arc_regular regular_closed_star regular_conv_closed regular_mult_closed)
have 77: "sum (?e \<sqinter> g) \<le> sum (a \<sqinter> g)"
proof (rule a_to_e_in_bigforest)
show "symmetric g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "j \<noteq> bot"
by (simp add: assms(2))
next
show "f \<le> -- g"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
next
show "vector j"
using assms(1) boruvka_inner_invariant_def by blast
next
show "forest h"
using assms(1) boruvka_inner_invariant_def by blast
next
show " big_forest (forest_components h) d"
using assms(1) boruvka_inner_invariant_def by blast
next
show "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
using assms(1) boruvka_inner_invariant_def by blast
next
show "\<forall>a b. bf_between_arcs a b (?H) d \<and> a \<le> - ?H \<sqinter> - - g \<and> b \<le> d \<longrightarrow> sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
using assms(1) boruvka_inner_invariant_def by blast
next
show "regular d"
using assms(1) boruvka_inner_invariant_def by blast
next
show "?e = ?e"
by simp
next
show "arc a"
using 1 by simp
next
show "bf_between_arcs a ?e ?H ?d'"
by (simp add: 76)
next
show "a \<le> - ?H \<sqinter> --g"
using 1 by simp
next
show "regular h"
using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto
qed
thus ?thesis
using 75 order.trans by blast
qed
qed
qed
qed
next
show "regular ?d'"
using assms(1) boruvka_inner_invariant_def minarc_regular by auto
qed
qed
lemma second_inner_invariant_when_e_bot:
assumes "selected_edge h j g = bot"
and "selected_edge h j g \<le> - forest_components f"
and "boruvka_inner_invariant j f h g d"
shows "boruvka_inner_invariant
(j \<sqinter> - choose_component (forest_components h) j)
(f \<sqinter> - selected_edge h j g\<^sup>T \<sqinter> - path f h j g \<squnion>
(f \<sqinter> - selected_edge h j g\<^sup>T \<sqinter> path f h j g)\<^sup>T \<squnion>
selected_edge h j g)
h g (d \<squnion> selected_edge h j g)"
proof -
let ?c = "choose_component (forest_components h) j"
let ?p = "path f h j g"
let ?F = "forest_components f"
let ?H = "forest_components h"
let ?e = "selected_edge h j g"
let ?f' = "f \<sqinter> -?e\<^sup>T \<sqinter> -?p \<squnion> (f \<sqinter> -?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> ?e"
let ?d' = "d \<squnion> ?e"
let ?j' = "j \<sqinter> -?c"
show "boruvka_inner_invariant ?j' ?f' h g ?d'"
proof (unfold boruvka_inner_invariant_def, intro conjI)
next
show "boruvka_outer_invariant ?f' g"
using assms(1, 3) boruvka_inner_invariant_def by auto
next
show "g \<noteq> bot"
using assms(3) boruvka_inner_invariant_def by blast
next
show "vector ?j'"
by (metis assms(3) boruvka_inner_invariant_def component_is_vector vector_complement_closed vector_inf_closed)
next
show "regular ?j'"
using assms(3) boruvka_inner_invariant_def by auto
next
show "boruvka_outer_invariant h g"
using assms(3) boruvka_inner_invariant_def by blast
next
show "injective h"
using assms(3) boruvka_inner_invariant_def by blast
next
show "pd_kleene_allegory_class.acyclic h"
using assms(3) boruvka_inner_invariant_def by blast
next
show "?H \<le> forest_components ?f'"
using assms(1, 3) boruvka_inner_invariant_def by auto
next
show " big_forest ?H ?d'"
using assms(1, 3) boruvka_inner_invariant_def by auto
next
show "?d' * top \<le> -?j'"
by (metis assms(1, 3) boruvka_inner_invariant_def order.trans p_antitone_inf sup_monoid.add_0_right)
next
show "?H * ?j' = ?j'"
using assms(3) fc_j_eq_j_inv boruvka_inner_invariant_def by blast
next
show "forest_components ?f' = (?H * (?d' \<squnion> ?d'\<^sup>T))\<^sup>\<star> *?H"
using assms(1, 3) boruvka_inner_invariant_def by auto
next
show "?f' \<squnion> ?f'\<^sup>T = h \<squnion> h\<^sup>T \<squnion> ?d' \<squnion> ?d'\<^sup>T"
using assms(1, 3) boruvka_inner_invariant_def by auto
next
show "\<forall>a b. bf_between_arcs a b ?H ?d' \<and> a \<le> -?H \<sqinter> --g \<and> b \<le> ?d' \<longrightarrow> sum(b \<sqinter> g) \<le> sum(a \<sqinter> g)"
using assms(1, 3) boruvka_inner_invariant_def by auto
next
show "regular ?d'"
using assms(1, 3) boruvka_inner_invariant_def by auto
qed
qed
subsection \<open>Formalization and correctness proof\<close>
text \<open>
The following result shows that Bor\r{u}vka's algorithm constructs a minimum spanning forest.
We have the same postcondition as the proof of Kruskal's minimum spanning tree algorithm.
We show only partial correctness.
\<close>
theorem boruvka_mst:
"VARS f j h c e d
{ symmetric g }
f := bot;
WHILE -(forest_components f) \<sqinter> g \<noteq> bot
INV { boruvka_outer_invariant f g }
DO
j := top;
h := f;
d := bot;
WHILE j \<noteq> bot
INV { boruvka_inner_invariant j f h g d }
DO
c := choose_component (forest_components h) j;
e := minarc(c * -c\<^sup>T \<sqinter> g);
IF e \<le> -(forest_components f) THEN
f := f \<sqinter> -e\<^sup>T;
f := (f \<sqinter> -(top * e * f\<^sup>T\<^sup>\<star>)) \<squnion> (f \<sqinter> top * e * f\<^sup>T\<^sup>\<star>)\<^sup>T \<squnion> e;
d := d \<squnion> e
ELSE
SKIP
FI;
j := j \<sqinter> -c
OD
OD
{ minimum_spanning_forest f g }"
proof vcg_simp
assume 1: "symmetric g"
show "boruvka_outer_invariant bot g"
using 1 boruvka_outer_invariant_def kruskal_exists_minimal_spanning by auto
next
fix f
let ?F = "forest_components f"
assume 1: "boruvka_outer_invariant f g \<and> - ?F \<sqinter> g \<noteq> bot"
have 2: "equivalence ?F"
using 1 boruvka_outer_invariant_def forest_components_equivalence by auto
show "boruvka_inner_invariant top f f g bot"
proof (unfold boruvka_inner_invariant_def, intro conjI)
show "boruvka_outer_invariant f g"
by (simp add: 1)
next
show "g \<noteq> bot"
using 1 by auto
next
show "surjective top"
by simp
next
show "regular top"
by simp
next
show "boruvka_outer_invariant f g"
using 1 by auto
next
show "injective f"
using 1 boruvka_outer_invariant_def by blast
next
show "pd_kleene_allegory_class.acyclic f"
using 1 boruvka_outer_invariant_def by blast
next
show "?F \<le> ?F"
by simp
next
show "big_forest ?F bot"
by (simp add: 2 big_forest_def)
next
show "bot * top \<le> - top"
by simp
next
show "times_top_class.total (?F)"
by (simp add: star.circ_right_top mult_assoc)
next
show "?F = (?F * (bot \<squnion> bot\<^sup>T))\<^sup>\<star> * ?F"
by (metis mult_right_zero semiring.mult_zero_left star.circ_loop_fixpoint sup_commute sup_monoid.add_0_right symmetric_bot_closed)
next
show "f \<squnion> f\<^sup>T = f \<squnion> f\<^sup>T \<squnion> bot \<squnion> bot\<^sup>T"
by simp
next
show "\<forall> a b. bf_between_arcs a b ?F bot \<and> a \<le> - ?F \<sqinter> -- g \<and> b \<le> bot \<longrightarrow> sum (b \<sqinter> g) \<le> sum (a \<sqinter> g)"
by (metis (full_types) bf_between_arcs_def bot_unique mult_left_zero mult_right_zero top.extremum)
next
show "regular bot"
by auto
qed
next
fix f j h d
let ?c = "choose_component (forest_components h) j"
let ?p = "path f h j g"
let ?F = "forest_components f"
let ?H = "forest_components h"
let ?e = "selected_edge h j g"
let ?f' = "f \<sqinter> -?e\<^sup>T \<sqinter> -?p \<squnion> (f \<sqinter> -?e\<^sup>T \<sqinter> ?p)\<^sup>T \<squnion> ?e"
let ?d' = "d \<squnion> ?e"
let ?j' = "j \<sqinter> -?c"
assume 1: "boruvka_inner_invariant j f h g d \<and> j \<noteq> bot"
show "(?e \<le> -?F \<longrightarrow> boruvka_inner_invariant ?j' ?f' h g ?d') \<and> (\<not> ?e \<le> -?F \<longrightarrow> boruvka_inner_invariant ?j' f h g d)"
proof (intro conjI)
show "?e \<le> -?F \<longrightarrow> boruvka_inner_invariant ?j' ?f' h g ?d'"
proof (cases "?e = bot")
case True
thus ?thesis
using 1 second_inner_invariant_when_e_bot by simp
next
case False
thus ?thesis
using 1 second_inner_invariant_when_e_not_bot by simp
qed
next
show "\<not> ?e \<le> -?F \<longrightarrow> boruvka_inner_invariant ?j' f h g d"
proof (rule impI, unfold boruvka_inner_invariant_def, intro conjI)
show "boruvka_outer_invariant f g"
using 1 boruvka_inner_invariant_def by blast
next
show "g \<noteq> bot"
using 1 boruvka_inner_invariant_def by blast
next
show "vector ?j'"
using 1 boruvka_inner_invariant_def component_is_vector vector_complement_closed vector_inf_closed by auto
next
show "regular ?j'"
using 1 boruvka_inner_invariant_def by auto
next
show "boruvka_outer_invariant h g"
using 1 boruvka_inner_invariant_def by auto
next
show "injective h"
using 1 boruvka_inner_invariant_def by blast
next
show "pd_kleene_allegory_class.acyclic h"
using 1 boruvka_inner_invariant_def by blast
next
show "?H \<le> ?F"
using 1 boruvka_inner_invariant_def by blast
next
show "big_forest ?H d"
using 1 boruvka_inner_invariant_def by blast
next
show "d * top \<le> -?j'"
using 1 by (meson boruvka_inner_invariant_def dual_order.trans p_antitone_inf)
next
show "?H * ?j' = ?j'"
using 1 fc_j_eq_j_inv boruvka_inner_invariant_def by blast
next
show "?F = (?H * (d \<squnion> d\<^sup>T))\<^sup>\<star> * ?H"
using 1 boruvka_inner_invariant_def by blast
next
show "f \<squnion> f\<^sup>T = h \<squnion> h\<^sup>T \<squnion> d \<squnion> d\<^sup>T"
using 1 boruvka_inner_invariant_def by blast
next
show "\<not> ?e \<le> -?F \<Longrightarrow> \<forall>a b. bf_between_arcs a b ?H d \<and> a \<le> -?H \<sqinter> --g \<and> b \<le> d \<longrightarrow> sum(b \<sqinter> g) \<le> sum(a \<sqinter> g)"
using 1 boruvka_inner_invariant_def by blast
next
show "\<not> ?e \<le> -?F \<Longrightarrow> regular d"
using 1 boruvka_inner_invariant_def by blast
qed
qed
next
fix f h d
assume "boruvka_inner_invariant bot f h g d"
thus "boruvka_outer_invariant f g"
by (meson boruvka_inner_invariant_def)
next
fix f
assume 1: "boruvka_outer_invariant f g \<and> - forest_components f \<sqinter> g = bot"
hence 2:"spanning_forest f g"
proof (unfold spanning_forest_def, intro conjI)
show "injective f"
using 1 boruvka_outer_invariant_def by blast
next
show "acyclic f"
using 1 boruvka_outer_invariant_def by blast
next
show "f \<le> --g"
using 1 boruvka_outer_invariant_def by blast
next
show "components g \<le> forest_components f"
proof -
let ?F = "forest_components f"
have "-?F \<sqinter> g \<le> bot"
by (simp add: 1)
hence "--g \<le> bot \<squnion> --?F"
using 1 shunting_p p_antitone pseudo_complement by auto
hence "--g \<le> ?F"
using 1 boruvka_outer_invariant_def pp_dist_comp pp_dist_star regular_conv_closed by auto
hence "(--g)\<^sup>\<star> \<le> ?F\<^sup>\<star>"
by (simp add: star_isotone)
thus ?thesis
using 1 boruvka_outer_invariant_def forest_components_star by auto
qed
next
show "regular f"
using 1 boruvka_outer_invariant_def by auto
qed
from 1 obtain w where 3: "minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T"
using boruvka_outer_invariant_def by blast
hence "w = w \<sqinter> --g"
by (simp add: inf.absorb1 minimum_spanning_forest_def spanning_forest_def)
also have "... \<le> w \<sqinter> components g"
by (metis inf.sup_right_isotone star.circ_increasing)
also have "... \<le> w \<sqinter> f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>"
using 2 spanning_forest_def inf.sup_right_isotone by simp
also have "... \<le> f \<squnion> f\<^sup>T"
proof (rule cancel_separate_6[where z=w and y="w\<^sup>T"])
show "injective w"
using 3 minimum_spanning_forest_def spanning_forest_def by simp
next
show "f\<^sup>T \<le> w\<^sup>T \<squnion> w"
using 3 by (metis conv_dist_inf conv_dist_sup conv_involutive inf.cobounded2 inf.orderE)
next
show "f \<le> w\<^sup>T \<squnion> w"
using 3 by (simp add: sup_commute)
next
show "injective w"
using 3 minimum_spanning_forest_def spanning_forest_def by simp
next
show "w \<sqinter> w\<^sup>T\<^sup>\<star> = bot"
using 3 by (metis acyclic_star_below_complement comp_inf.mult_right_isotone inf_p le_bot minimum_spanning_forest_def spanning_forest_def)
qed
finally have 4: "w \<le> f \<squnion> f\<^sup>T"
by simp
have "sum (f \<sqinter> g) = sum ((w \<squnion> w\<^sup>T) \<sqinter> (f \<sqinter> g))"
using 3 by (metis inf_absorb2 inf.assoc)
also have "... = sum (w \<sqinter> (f \<sqinter> g)) + sum (w\<^sup>T \<sqinter> (f \<sqinter> g))"
using 3 inf.commute acyclic_asymmetric sum_disjoint minimum_spanning_forest_def spanning_forest_def by simp
also have "... = sum (w \<sqinter> (f \<sqinter> g)) + sum (w \<sqinter> (f\<^sup>T \<sqinter> g\<^sup>T))"
by (metis conv_dist_inf conv_involutive sum_conv)
also have "... = sum (f \<sqinter> (w \<sqinter> g)) + sum (f\<^sup>T \<sqinter> (w \<sqinter> g))"
proof -
have 51:"f\<^sup>T \<sqinter> (w \<sqinter> g) = f\<^sup>T \<sqinter> (w \<sqinter> g\<^sup>T)"
using 1 boruvka_outer_invariant_def by auto
have 52:"f \<sqinter> (w \<sqinter> g) = w \<sqinter> (f \<sqinter> g)"
by (simp add: inf.left_commute)
thus ?thesis
using 51 52 abel_semigroup.left_commute inf.abel_semigroup_axioms by fastforce
qed
also have "... = sum ((f \<squnion> f\<^sup>T) \<sqinter> (w \<sqinter> g))"
using 2 acyclic_asymmetric inf.sup_monoid.add_commute sum_disjoint spanning_forest_def by simp
also have "... = sum (w \<sqinter> g)"
using 4 by (metis inf_absorb2 inf.assoc)
finally show "minimum_spanning_forest f g"
using 2 3 minimum_spanning_forest_def by simp
qed
end
end
-
diff --git a/thys/Relational_Minimum_Spanning_Trees/Prim.thy b/thys/Relational_Minimum_Spanning_Trees/Prim.thy
--- a/thys/Relational_Minimum_Spanning_Trees/Prim.thy
+++ b/thys/Relational_Minimum_Spanning_Trees/Prim.thy
@@ -1,477 +1,477 @@
(* Title: Prim's Minimum Spanning Tree Algorithm
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Prim's Minimum Spanning Tree Algorithm\<close>
text \<open>
In this theory we prove total correctness of Prim's minimum spanning tree algorithm.
The proof has the same overall structure as the total-correctness proof of Kruskal's algorithm \cite{Guttmann2018c}.
The partial-correctness proof of Prim's algorithm is discussed in \cite{Guttmann2016c,Guttmann2018b}.
\<close>
theory Prim
imports "HOL-Hoare.Hoare_Logic" Aggregation_Algebras.Aggregation_Algebras
begin
context m_kleene_algebra
begin
abbreviation "component g r \<equiv> r\<^sup>T * (--g)\<^sup>\<star>"
definition "spanning_tree t g r \<equiv> forest t \<and> t \<le> (component g r)\<^sup>T * (component g r) \<sqinter> --g \<and> component g r \<le> r\<^sup>T * t\<^sup>\<star> \<and> regular t"
definition "minimum_spanning_tree t g r \<equiv> spanning_tree t g r \<and> (\<forall>u . spanning_tree u g r \<longrightarrow> sum (t \<sqinter> g) \<le> sum (u \<sqinter> g))"
definition "prim_precondition g r \<equiv> g = g\<^sup>T \<and> injective r \<and> vector r \<and> regular r"
definition "prim_spanning_invariant t v g r \<equiv> prim_precondition g r \<and> v\<^sup>T = r\<^sup>T * t\<^sup>\<star> \<and> spanning_tree t (v * v\<^sup>T \<sqinter> g) r"
definition "prim_invariant t v g r \<equiv> prim_spanning_invariant t v g r \<and> (\<exists>w . minimum_spanning_tree w g r \<and> t \<le> w)"
lemma span_tree_split:
assumes "vector r"
shows "t \<le> (component g r)\<^sup>T * (component g r) \<sqinter> --g \<longleftrightarrow> (t \<le> (component g r)\<^sup>T \<and> t \<le> component g r \<and> t \<le> --g)"
proof -
have "(component g r)\<^sup>T * (component g r) = (component g r)\<^sup>T \<sqinter> component g r"
by (metis assms conv_involutive covector_mult_closed vector_conv_covector vector_covector)
thus ?thesis
by simp
qed
lemma span_tree_component:
assumes "spanning_tree t g r"
shows "component g r = component t r"
- using assms by (simp add: antisym mult_right_isotone star_isotone spanning_tree_def)
+ using assms by (simp add: order.antisym mult_right_isotone star_isotone spanning_tree_def)
text \<open>
We first show three verification conditions which are used in both correctness proofs.
\<close>
lemma prim_vc_1:
assumes "prim_precondition g r"
shows "prim_spanning_invariant bot r g r"
proof (unfold prim_spanning_invariant_def, intro conjI)
show "prim_precondition g r"
using assms by simp
next
show "r\<^sup>T = r\<^sup>T * bot\<^sup>\<star>"
by (simp add: star_absorb)
next
let ?ss = "r * r\<^sup>T \<sqinter> g"
show "spanning_tree bot ?ss r"
proof (unfold spanning_tree_def, intro conjI)
show "injective bot"
by simp
next
show "acyclic bot"
by simp
next
show "bot \<le> (component ?ss r)\<^sup>T * (component ?ss r) \<sqinter> --?ss"
by simp
next
have "component ?ss r \<le> component (r * r\<^sup>T) r"
by (simp add: mult_right_isotone star_isotone)
also have "... \<le> r\<^sup>T * 1\<^sup>\<star>"
- using assms by (metis inf.eq_iff p_antitone regular_one_closed star_sub_one prim_precondition_def)
+ using assms by (metis order.eq_iff p_antitone regular_one_closed star_sub_one prim_precondition_def)
also have "... = r\<^sup>T * bot\<^sup>\<star>"
by (simp add: star.circ_zero star_one)
finally show "component ?ss r \<le> r\<^sup>T * bot\<^sup>\<star>"
.
next
show "regular bot"
by simp
qed
qed
lemma prim_vc_2:
assumes "prim_spanning_invariant t v g r"
and "v * -v\<^sup>T \<sqinter> g \<noteq> bot"
shows "prim_spanning_invariant (t \<squnion> minarc (v * -v\<^sup>T \<sqinter> g)) (v \<squnion> minarc (v * -v\<^sup>T \<sqinter> g)\<^sup>T * top) g r \<and> card { x . regular x \<and> x \<le> component g r \<and> x \<le> -(v \<squnion> minarc (v * -v\<^sup>T \<sqinter> g)\<^sup>T * top)\<^sup>T } < card { x . regular x \<and> x \<le> component g r \<and> x \<le> -v\<^sup>T }"
proof -
let ?vcv = "v * -v\<^sup>T \<sqinter> g"
let ?e = "minarc ?vcv"
let ?t = "t \<squnion> ?e"
let ?v = "v \<squnion> ?e\<^sup>T * top"
let ?c = "component g r"
let ?g = "--g"
let ?n1 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -v\<^sup>T }"
let ?n2 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -?v\<^sup>T }"
have 1: "regular v \<and> regular (v * v\<^sup>T) \<and> regular (?v * ?v\<^sup>T) \<and> regular (top * ?e)"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive regular_closed_top regular_closed_sup minarc_regular)
hence 2: "t \<le> v * v\<^sup>T \<sqinter> ?g"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def inf_pp_commute inf.boundedE)
hence 3: "t \<le> v * v\<^sup>T"
by simp
have 4: "t \<le> ?g"
using 2 by simp
have 5: "?e \<le> v * -v\<^sup>T \<sqinter> ?g"
using 1 by (metis minarc_below pp_dist_inf regular_mult_closed regular_closed_p)
hence 6: "?e \<le> v * -v\<^sup>T"
by simp
have 7: "vector v"
using assms(1) prim_spanning_invariant_def prim_precondition_def by (simp add: covector_mult_closed vector_conv_covector)
hence 8: "?e \<le> v"
using 6 by (metis conv_complement inf.boundedE vector_complement_closed vector_covector)
have 9: "?e * t = bot"
using 3 6 7 et(1) by blast
have 10: "?e * t\<^sup>T = bot"
using 3 6 7 et(2) by simp
have 11: "arc ?e"
using assms(2) minarc_arc by simp
have "r\<^sup>T \<le> r\<^sup>T * t\<^sup>\<star>"
by (metis mult_right_isotone order_refl semiring.mult_not_zero star.circ_separate_mult_1 star_absorb)
hence 12: "r\<^sup>T \<le> v\<^sup>T"
using assms(1) by (simp add: prim_spanning_invariant_def)
have 13: "vector r \<and> injective r \<and> v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
using assms(1) prim_spanning_invariant_def prim_precondition_def minimum_spanning_tree_def spanning_tree_def reachable_restrict by simp
have "g = g\<^sup>T"
using assms(1) prim_invariant_def prim_spanning_invariant_def prim_precondition_def by simp
hence 14: "?g\<^sup>T = ?g"
using conv_complement by simp
show "prim_spanning_invariant ?t ?v g r \<and> ?n2 < ?n1"
proof (rule conjI)
show "prim_spanning_invariant ?t ?v g r"
proof (unfold prim_spanning_invariant_def, intro conjI)
show "prim_precondition g r"
using assms(1) prim_spanning_invariant_def by simp
next
show "?v\<^sup>T = r\<^sup>T * ?t\<^sup>\<star>"
using assms(1) 6 7 9 by (simp add: reachable_inv prim_spanning_invariant_def prim_precondition_def spanning_tree_def)
next
let ?G = "?v * ?v\<^sup>T \<sqinter> g"
show "spanning_tree ?t ?G r"
proof (unfold spanning_tree_def, intro conjI)
show "injective ?t"
using assms(1) 10 11 by (simp add: injective_inv prim_spanning_invariant_def spanning_tree_def)
next
show "acyclic ?t"
using assms(1) 3 6 7 acyclic_inv prim_spanning_invariant_def spanning_tree_def by simp
next
show "?t \<le> (component ?G r)\<^sup>T * (component ?G r) \<sqinter> --?G"
using 1 2 5 7 13 prim_subgraph_inv inf_pp_commute mst_subgraph_inv_2 by auto
next
show "component (?v * ?v\<^sup>T \<sqinter> g) r \<le> r\<^sup>T * ?t\<^sup>\<star>"
proof -
have 15: "r\<^sup>T * (v * v\<^sup>T \<sqinter> ?g)\<^sup>\<star> \<le> r\<^sup>T * t\<^sup>\<star>"
using assms(1) 1 by (metis prim_spanning_invariant_def spanning_tree_def inf_pp_commute)
have "component (?v * ?v\<^sup>T \<sqinter> g) r = r\<^sup>T * (?v * ?v\<^sup>T \<sqinter> ?g)\<^sup>\<star>"
using 1 by simp
also have "... \<le> r\<^sup>T * ?t\<^sup>\<star>"
using 2 6 7 11 12 13 14 15 by (metis span_inv)
finally show ?thesis
.
qed
next
show "regular ?t"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def regular_closed_sup minarc_regular)
qed
qed
next
have 16: "top * ?e \<le> ?c"
proof -
have "top * ?e = top * ?e\<^sup>T * ?e"
using 11 by (metis arc_top_edge mult_assoc)
also have "... \<le> v\<^sup>T * ?e"
using 7 8 by (metis conv_dist_comp conv_isotone mult_left_isotone symmetric_top_closed)
also have "... \<le> v\<^sup>T * ?g"
using 5 mult_right_isotone by auto
also have "... = r\<^sup>T * t\<^sup>\<star> * ?g"
using 13 by simp
also have "... \<le> r\<^sup>T * ?g\<^sup>\<star> * ?g"
using 4 by (simp add: mult_left_isotone mult_right_isotone star_isotone)
also have "... \<le> ?c"
by (simp add: comp_associative mult_right_isotone star.right_plus_below_circ)
finally show ?thesis
by simp
qed
have 17: "top * ?e \<le> -v\<^sup>T"
using 6 7 by (simp add: schroeder_4_p vTeT)
have 18: "\<not> top * ?e \<le> -(top * ?e)"
by (metis assms(2) inf.orderE minarc_bot_iff conv_complement_sub_inf inf_p inf_top.left_neutral p_bot symmetric_top_closed vector_top_closed)
have 19: "-?v\<^sup>T = -v\<^sup>T \<sqinter> -(top * ?e)"
by (simp add: conv_dist_comp conv_dist_sup)
hence 20: "\<not> top * ?e \<le> -?v\<^sup>T"
using 18 by simp
show "?n2 < ?n1"
apply (rule psubset_card_mono)
using finite_regular apply simp
using 1 16 17 19 20 by auto
qed
qed
lemma prim_vc_3:
assumes "prim_spanning_invariant t v g r"
and "v * -v\<^sup>T \<sqinter> g = bot"
shows "spanning_tree t g r"
proof -
let ?g = "--g"
have 1: "regular v \<and> regular (v * v\<^sup>T)"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
have 2: "v * -v\<^sup>T \<sqinter> ?g = bot"
using assms(2) pp_inf_bot_iff pp_pp_inf_bot_iff by simp
have 3: "v\<^sup>T = r\<^sup>T * t\<^sup>\<star> \<and> vector v"
using assms(1) by (simp add: covector_mult_closed prim_invariant_def prim_spanning_invariant_def vector_conv_covector prim_precondition_def)
have 4: "t \<le> v * v\<^sup>T \<sqinter> ?g"
using assms(1) 1 by (metis prim_spanning_invariant_def inf_pp_commute spanning_tree_def inf.boundedE)
have "r\<^sup>T * (v * v\<^sup>T \<sqinter> ?g)\<^sup>\<star> \<le> r\<^sup>T * t\<^sup>\<star>"
using assms(1) 1 by (metis prim_spanning_invariant_def inf_pp_commute spanning_tree_def)
hence 5: "component g r = v\<^sup>T"
using 1 2 3 4 by (metis span_post)
have "regular (v * v\<^sup>T)"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
hence 6: "t \<le> v * v\<^sup>T \<sqinter> ?g"
by (metis assms(1) prim_spanning_invariant_def spanning_tree_def inf_pp_commute inf.boundedE)
show "spanning_tree t g r"
apply (unfold spanning_tree_def, intro conjI)
using assms(1) prim_spanning_invariant_def spanning_tree_def apply simp
using assms(1) prim_spanning_invariant_def spanning_tree_def apply simp
using 5 6 apply simp
using assms(1) 5 prim_spanning_invariant_def apply simp
using assms(1) prim_spanning_invariant_def spanning_tree_def by simp
qed
text \<open>
The following result shows that Prim's algorithm terminates and constructs a spanning tree.
We cannot yet show that this is a minimum spanning tree.
\<close>
theorem prim_spanning:
"VARS t v e
[ prim_precondition g r ]
t := bot;
v := r;
WHILE v * -v\<^sup>T \<sqinter> g \<noteq> bot
INV { prim_spanning_invariant t v g r }
VAR { card { x . regular x \<and> x \<le> component g r \<sqinter> -v\<^sup>T } }
DO e := minarc (v * -v\<^sup>T \<sqinter> g);
t := t \<squnion> e;
v := v \<squnion> e\<^sup>T * top
OD
[ spanning_tree t g r ]"
apply vcg_tc_simp
apply (simp add: prim_vc_1)
using prim_vc_2 apply blast
using prim_vc_3 by auto
text \<open>
Because we have shown total correctness, we conclude that a spanning tree exists.
\<close>
lemma prim_exists_spanning:
"prim_precondition g r \<Longrightarrow> \<exists>t . spanning_tree t g r"
using tc_extract_function prim_spanning by blast
text \<open>
This implies that a minimum spanning tree exists, which is used in the subsequent correctness proof.
\<close>
lemma prim_exists_minimal_spanning:
assumes "prim_precondition g r"
shows "\<exists>t . minimum_spanning_tree t g r"
proof -
let ?s = "{ t . spanning_tree t g r }"
have "\<exists>m\<in>?s . \<forall>z\<in>?s . sum (m \<sqinter> g) \<le> sum (z \<sqinter> g)"
apply (rule finite_set_minimal)
using finite_regular spanning_tree_def apply simp
using assms prim_exists_spanning apply simp
using sum_linear by simp
thus ?thesis
using minimum_spanning_tree_def by simp
qed
text \<open>
Prim's minimum spanning tree algorithm terminates and is correct.
This is the same algorithm that is used in the previous correctness proof, with the same precondition and variant, but with a different invariant and postcondition.
\<close>
theorem prim:
"VARS t v e
[ prim_precondition g r \<and> (\<exists>w . minimum_spanning_tree w g r) ]
t := bot;
v := r;
WHILE v * -v\<^sup>T \<sqinter> g \<noteq> bot
INV { prim_invariant t v g r }
VAR { card { x . regular x \<and> x \<le> component g r \<sqinter> -v\<^sup>T } }
DO e := minarc (v * -v\<^sup>T \<sqinter> g);
t := t \<squnion> e;
v := v \<squnion> e\<^sup>T * top
OD
[ minimum_spanning_tree t g r ]"
proof vcg_tc_simp
assume "prim_precondition g r \<and> (\<exists>w . minimum_spanning_tree w g r)"
thus "prim_invariant bot r g r"
using prim_invariant_def prim_vc_1 by simp
next
fix t v
let ?vcv = "v * -v\<^sup>T \<sqinter> g"
let ?vv = "v * v\<^sup>T \<sqinter> g"
let ?e = "minarc ?vcv"
let ?t = "t \<squnion> ?e"
let ?v = "v \<squnion> ?e\<^sup>T * top"
let ?c = "component g r"
let ?g = "--g"
let ?n1 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -v\<^sup>T }"
let ?n2 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -?v\<^sup>T }"
assume 1: "prim_invariant t v g r \<and> ?vcv \<noteq> bot"
hence 2: "regular v \<and> regular (v * v\<^sup>T)"
by (metis (no_types, hide_lams) prim_invariant_def prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
have 3: "t \<le> v * v\<^sup>T \<sqinter> ?g"
using 1 2 by (metis (no_types, hide_lams) prim_invariant_def prim_spanning_invariant_def spanning_tree_def inf_pp_commute inf.boundedE)
hence 4: "t \<le> v * v\<^sup>T"
by simp
have 5: "t \<le> ?g"
using 3 by simp
have 6: "?e \<le> v * -v\<^sup>T \<sqinter> ?g"
using 2 by (metis minarc_below pp_dist_inf regular_mult_closed regular_closed_p)
hence 7: "?e \<le> v * -v\<^sup>T"
by simp
have 8: "vector v"
using 1 prim_invariant_def prim_spanning_invariant_def prim_precondition_def by (simp add: covector_mult_closed vector_conv_covector)
have 9: "arc ?e"
using 1 minarc_arc by simp
from 1 obtain w where 10: "minimum_spanning_tree w g r \<and> t \<le> w"
by (metis prim_invariant_def)
hence 11: "vector r \<and> injective r \<and> v\<^sup>T = r\<^sup>T * t\<^sup>\<star> \<and> forest w \<and> t \<le> w \<and> w \<le> ?c\<^sup>T * ?c \<sqinter> ?g \<and> r\<^sup>T * (?c\<^sup>T * ?c \<sqinter> ?g)\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
using 1 2 prim_invariant_def prim_spanning_invariant_def prim_precondition_def minimum_spanning_tree_def spanning_tree_def reachable_restrict by simp
hence 12: "w * v \<le> v"
using predecessors_reachable reachable_restrict by auto
have 13: "g = g\<^sup>T"
using 1 prim_invariant_def prim_spanning_invariant_def prim_precondition_def by simp
hence 14: "?g\<^sup>T = ?g"
using conv_complement by simp
show "prim_invariant ?t ?v g r \<and> ?n2 < ?n1"
proof (unfold prim_invariant_def, intro conjI)
show "prim_spanning_invariant ?t ?v g r"
using 1 prim_invariant_def prim_vc_2 by blast
next
show "\<exists>w . minimum_spanning_tree w g r \<and> ?t \<le> w"
proof
let ?f = "w \<sqinter> v * -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?p = "w \<sqinter> -v * -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?fp = "w \<sqinter> -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?w = "(w \<sqinter> -?fp) \<squnion> ?p\<^sup>T \<squnion> ?e"
have 15: "regular ?f \<and> regular ?fp \<and> regular ?w"
using 2 10 by (metis regular_conv_closed regular_closed_star regular_mult_closed regular_closed_top regular_closed_inf regular_closed_sup minarc_regular minimum_spanning_tree_def spanning_tree_def)
show "minimum_spanning_tree ?w g r \<and> ?t \<le> ?w"
proof (intro conjI)
show "minimum_spanning_tree ?w g r"
proof (unfold minimum_spanning_tree_def, intro conjI)
show "spanning_tree ?w g r"
proof (unfold spanning_tree_def, intro conjI)
show "injective ?w"
using 7 8 9 11 exchange_injective by blast
next
show "acyclic ?w"
using 7 8 11 12 exchange_acyclic by blast
next
show "?w \<le> ?c\<^sup>T * ?c \<sqinter> --g"
proof -
have 16: "w \<sqinter> -?fp \<le> ?c\<^sup>T * ?c \<sqinter> --g"
using 10 by (simp add: le_infI1 minimum_spanning_tree_def spanning_tree_def)
have "?p\<^sup>T \<le> w\<^sup>T"
by (simp add: conv_isotone inf.sup_monoid.add_assoc)
also have "... \<le> (?c\<^sup>T * ?c \<sqinter> --g)\<^sup>T"
using 11 conv_order by simp
also have "... = ?c\<^sup>T * ?c \<sqinter> --g"
using 2 14 conv_dist_comp conv_dist_inf by simp
finally have 17: "?p\<^sup>T \<le> ?c\<^sup>T * ?c \<sqinter> --g"
.
have "?e \<le> ?c\<^sup>T * ?c \<sqinter> ?g"
using 5 6 11 mst_subgraph_inv by auto
thus ?thesis
using 16 17 by simp
qed
next
show "?c \<le> r\<^sup>T * ?w\<^sup>\<star>"
proof -
have "?c \<le> r\<^sup>T * w\<^sup>\<star>"
using 10 minimum_spanning_tree_def spanning_tree_def by simp
also have "... \<le> r\<^sup>T * ?w\<^sup>\<star>"
using 4 7 8 10 11 12 15 by (metis mst_reachable_inv)
finally show ?thesis
.
qed
next
show "regular ?w"
using 15 by simp
qed
next
have 18: "?f \<squnion> ?p = ?fp"
using 2 8 epm_1 by fastforce
have "arc (w \<sqinter> --v * -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)"
using 5 6 8 9 11 12 reachable_restrict arc_edge by auto
hence 19: "arc ?f"
using 2 by simp
hence "?f = bot \<longrightarrow> top = bot"
by (metis mult_left_zero mult_right_zero)
hence "?f \<noteq> bot"
using 1 le_bot by auto
hence "?f \<sqinter> v * -v\<^sup>T \<sqinter> ?g \<noteq> bot"
using 2 11 by (simp add: inf.absorb1 le_infI1)
hence "g \<sqinter> (?f \<sqinter> v * -v\<^sup>T) \<noteq> bot"
using inf_commute pp_inf_bot_iff by simp
hence 20: "?f \<sqinter> ?vcv \<noteq> bot"
by (simp add: inf_assoc inf_commute)
hence 21: "?f \<sqinter> g = ?f \<sqinter> ?vcv"
using 2 by (simp add: inf_assoc inf_commute inf_left_commute)
have 22: "?e \<sqinter> g = minarc ?vcv \<sqinter> ?vcv"
using 7 by (simp add: inf.absorb2 inf.assoc inf.commute)
hence 23: "sum (?e \<sqinter> g) \<le> sum (?f \<sqinter> g)"
using 15 19 20 21 by (simp add: minarc_min)
have "?e \<noteq> bot"
using 20 comp_inf.semiring.mult_not_zero semiring.mult_not_zero by blast
hence 24: "?e \<sqinter> g \<noteq> bot"
using 22 minarc_meet_bot by auto
have "sum (?w \<sqinter> g) = sum (w \<sqinter> -?fp \<sqinter> g) + sum (?p\<^sup>T \<sqinter> g) + sum (?e \<sqinter> g)"
using 7 8 10 by (metis sum_disjoint_3 epm_8 epm_9 epm_10 minimum_spanning_tree_def spanning_tree_def)
also have "... = sum (((w \<sqinter> -?fp) \<squnion> ?p\<^sup>T) \<sqinter> g) + sum (?e \<sqinter> g)"
using 11 by (metis epm_8 sum_disjoint)
also have "... \<le> sum (((w \<sqinter> -?fp) \<squnion> ?p\<^sup>T) \<sqinter> g) + sum (?f \<sqinter> g)"
using 23 24 by (simp add: sum_plus_right_isotone)
also have "... = sum (w \<sqinter> -?fp \<sqinter> g) + sum (?p\<^sup>T \<sqinter> g) + sum (?f \<sqinter> g)"
using 11 by (metis epm_8 sum_disjoint)
also have "... = sum (w \<sqinter> -?fp \<sqinter> g) + sum (?p \<sqinter> g) + sum (?f \<sqinter> g)"
using 13 sum_symmetric by auto
also have "... = sum (((w \<sqinter> -?fp) \<squnion> ?p \<squnion> ?f) \<sqinter> g)"
using 2 8 by (metis sum_disjoint_3 epm_11 epm_12 epm_13)
also have "... = sum (w \<sqinter> g)"
using 2 8 15 18 epm_2 by force
finally have "sum (?w \<sqinter> g) \<le> sum (w \<sqinter> g)"
.
thus "\<forall>u . spanning_tree u g r \<longrightarrow> sum (?w \<sqinter> g) \<le> sum (u \<sqinter> g)"
using 10 order_lesseq_imp minimum_spanning_tree_def by auto
qed
next
show "?t \<le> ?w"
using 4 8 10 mst_extends_new_tree by simp
qed
qed
next
show "?n2 < ?n1"
using 1 prim_invariant_def prim_vc_2 by auto
qed
next
fix t v
let ?g = "--g"
assume 25: "prim_invariant t v g r \<and> v * -v\<^sup>T \<sqinter> g = bot"
hence 26: "regular v"
by (metis prim_invariant_def prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
from 25 obtain w where 27: "minimum_spanning_tree w g r \<and> t \<le> w"
by (metis prim_invariant_def)
have "spanning_tree t g r"
using 25 prim_invariant_def prim_vc_3 by blast
hence "component g r = v\<^sup>T"
by (metis 25 prim_invariant_def span_tree_component prim_spanning_invariant_def spanning_tree_def)
hence 28: "w \<le> v * v\<^sup>T"
using 26 27 by (simp add: minimum_spanning_tree_def spanning_tree_def inf_pp_commute)
have "vector r \<and> injective r \<and> forest w"
using 25 27 by (simp add: prim_invariant_def prim_spanning_invariant_def prim_precondition_def minimum_spanning_tree_def spanning_tree_def)
hence "w = t"
using 25 27 28 prim_invariant_def prim_spanning_invariant_def mst_post by blast
thus "minimum_spanning_tree t g r"
using 27 by simp
qed
end
end
diff --git a/thys/Relational_Paths/More_Relation_Algebra.thy b/thys/Relational_Paths/More_Relation_Algebra.thy
--- a/thys/Relational_Paths/More_Relation_Algebra.thy
+++ b/thys/Relational_Paths/More_Relation_Algebra.thy
@@ -1,1130 +1,1130 @@
(* Title: (More) Relation Algebra
Author: Walter Guttmann, Peter Hoefner
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
Peter Hoefner <peter at hoefner-online.de>
*)
section \<open>(More) Relation Algebra\<close>
text \<open>
This theory presents fundamental properties of relation algebras, which are not present in the AFP entry on relation algebras but could be integrated there \cite{ArmstrongFosterStruthWeber2014}.
Many theorems concern vectors and points.
\<close>
theory More_Relation_Algebra
imports Relation_Algebra.Relation_Algebra_RTC Relation_Algebra.Relation_Algebra_Functions
begin
no_notation
trancl ("(_\<^sup>+)" [1000] 999)
context relation_algebra
begin
notation
converse ("(_\<^sup>T)" [102] 101)
abbreviation bijective
where "bijective x \<equiv> is_inj x \<and> is_sur x"
abbreviation reflexive
where "reflexive R \<equiv> 1' \<le> R"
abbreviation symmetric
where "symmetric R \<equiv> R = R\<^sup>T"
abbreviation transitive
where "transitive R \<equiv> R;R \<le> R"
text \<open>General theorems\<close>
lemma x_leq_triple_x:
"x \<le> x;x\<^sup>T;x"
proof -
have "x = x;1' \<cdot> 1"
by simp
also have "... \<le> (x \<cdot> 1;1'\<^sup>T);(1' \<cdot> x\<^sup>T;1)"
by (rule dedekind)
also have "... = x;(x\<^sup>T;1 \<cdot> 1')"
by (simp add: inf.commute)
also have "... \<le> x;(x\<^sup>T \<cdot> 1';1\<^sup>T);(1 \<cdot> (x\<^sup>T)\<^sup>T;1')"
by (metis comp_assoc dedekind mult_isol)
also have "... \<le> x;x\<^sup>T;x"
by simp
finally show ?thesis .
qed
lemma inj_triple:
assumes "is_inj x"
shows "x = x;x\<^sup>T;x"
-by (metis assms eq_iff inf_absorb2 is_inj_def mult_1_left mult_subdistr x_leq_triple_x)
+by (metis assms order.eq_iff inf_absorb2 is_inj_def mult_1_left mult_subdistr x_leq_triple_x)
lemma p_fun_triple:
assumes "is_p_fun x"
shows "x = x;x\<^sup>T;x"
-by (metis assms comp_assoc eq_iff is_p_fun_def mult_isol mult_oner x_leq_triple_x)
+by (metis assms comp_assoc order.eq_iff is_p_fun_def mult_isol mult_oner x_leq_triple_x)
lemma loop_backward_forward:
"x\<^sup>T \<le> -(1') + x"
by (metis conv_e conv_times inf.cobounded2 test_dom test_domain test_eq_conv galois_2 inf.commute
sup.commute)
lemma inj_sur_semi_swap:
assumes "is_sur z"
and "is_inj x"
shows "z \<le> y;x \<Longrightarrow> x \<le> y\<^sup>T;z"
proof -
assume "z \<le> y;x"
hence "z;x\<^sup>T \<le> y;(x;x\<^sup>T)"
by (metis mult_isor mult_assoc)
hence "z;x\<^sup>T \<le> y"
using \<open>is_inj x\<close> unfolding is_inj_def
by (metis mult_isol order.trans mult_1_right)
hence "(z\<^sup>T;z);x\<^sup>T \<le> z\<^sup>T;y"
by (metis mult_isol mult_assoc)
hence "x\<^sup>T \<le> z\<^sup>T;y"
using \<open>is_sur z\<close> unfolding is_sur_def
by (metis mult_isor order.trans mult_1_left)
thus ?thesis
using conv_iso by fastforce
qed
lemma inj_sur_semi_swap_short:
assumes "is_sur z"
and "is_inj x"
shows "z \<le> y\<^sup>T;x \<Longrightarrow> x \<le> y;z"
proof -
assume as: "z \<le> y\<^sup>T;x"
hence "z;x\<^sup>T \<le> y\<^sup>T"
using \<open>z \<le> y\<^sup>T;x\<close> \<open>is_inj x\<close> unfolding is_inj_def
by (metis assms(2) conv_invol inf.orderI inf_absorb1 inj_p_fun ss_422iii)
hence "x\<^sup>T \<le> z\<^sup>T;y\<^sup>T"
using \<open>is_sur z\<close> unfolding is_sur_def
by (metis as assms inj_sur_semi_swap conv_contrav conv_invol conv_iso)
thus "x \<le> y;z"
using conv_iso by fastforce
qed
lemma bij_swap:
assumes "bijective z"
and "bijective x"
shows "z \<le> y\<^sup>T;x \<longleftrightarrow> x \<le> y;z"
by (metis assms inj_sur_semi_swap conv_invol)
text \<open>The following result is \cite[Proposition 4.2.2(iv)]{SchmidtStroehlein1993}.\<close>
lemma ss422iv:
assumes "is_p_fun y"
and "x \<le> y"
and "y;1 \<le> x;1"
shows "x = y"
proof -
have "y \<le> (x;1)\<cdot>y"
using assms(3) le_infI maddux_20 order_trans by blast
also have "... \<le> x;x\<^sup>T;y"
by (metis inf_top_left modular_1_var comp_assoc)
also have "... \<le> x;y\<^sup>T;y"
using assms(2) conv_iso mult_double_iso by blast
also have "... \<le> x"
using assms(1) comp_assoc is_p_fun_def mult_isol mult_1_right
by fastforce
finally show ?thesis
- by (simp add: assms(2) antisym)
+ by (simp add: assms(2) order.antisym)
qed
text \<open>The following results are variants of \cite[Proposition 4.2.3]{SchmidtStroehlein1993}.\<close>
lemma ss423conv:
assumes "bijective x"
shows "x ; y \<le> z \<longleftrightarrow> y \<le> x\<^sup>T ; z"
by (metis assms conv_contrav conv_iso inj_p_fun is_map_def ss423 sur_total)
lemma ss423bij:
assumes "bijective x"
shows "y ; x\<^sup>T \<le> z \<longleftrightarrow> y \<le> z ; x"
by (simp add: assms is_map_def p_fun_inj ss423 total_sur)
lemma inj_distr:
assumes "is_inj z"
shows "(x\<cdot>y);z = (x;z)\<cdot>(y;z)"
-apply (rule antisym)
+apply (rule order.antisym)
using mult_subdistr_var apply blast
using assms conv_iso inj_p_fun p_fun_distl by fastforce
lemma test_converse:
"x \<cdot> 1' = x\<^sup>T \<cdot> 1'"
by (metis conv_e conv_times inf_le2 is_test_def test_eq_conv)
lemma injective_down_closed:
assumes "is_inj x"
and "y \<le> x"
shows "is_inj y"
by (meson assms conv_iso dual_order.trans is_inj_def mult_isol_var)
lemma injective_sup:
assumes "is_inj t"
and "e;t\<^sup>T \<le> 1'"
and "is_inj e"
shows "is_inj (t + e)"
proof -
have 1: "t;e\<^sup>T \<le> 1'"
using assms(2) conv_contrav conv_e conv_invol conv_iso by fastforce
have "(t + e);(t + e)\<^sup>T = t;t\<^sup>T + t;e\<^sup>T + e;t\<^sup>T + e;e\<^sup>T"
by (metis conv_add distrib_left distrib_right' sup_assoc)
also have "... \<le> 1'"
using 1 assms by (simp add: is_inj_def le_supI)
finally show ?thesis
unfolding is_inj_def .
qed
text \<open>Some (more) results about vectors\<close>
lemma vector_meet_comp:
assumes "is_vector v"
and "is_vector w"
shows "v;w\<^sup>T = v\<cdot>w\<^sup>T"
by (metis assms conv_contrav conv_one inf_top_right is_vector_def vector_1)
lemma vector_meet_comp':
assumes "is_vector v"
shows "v;v\<^sup>T = v\<cdot>v\<^sup>T"
using assms vector_meet_comp by blast
lemma vector_meet_comp_x:
"x;1;x\<^sup>T = x;1\<cdot>1;x\<^sup>T"
by (metis comp_assoc inf_top.right_neutral is_vector_def one_idem_mult vector_1)
lemma vector_meet_comp_x':
"x;1;x = x;1\<cdot>1;x"
by (metis inf_commute inf_top.right_neutral ra_1)
lemma vector_prop1:
assumes "is_vector v"
shows "-v\<^sup>T;v = 0"
by (metis assms compl_inf_bot inf_top.right_neutral one_compl one_idem_mult vector_2)
text \<open>The following results and a number of others in this theory are from \cite{Guttmann2017a}.\<close>
lemma ee:
assumes "is_vector v"
and "e \<le> v;-v\<^sup>T"
shows "e;e = 0"
proof -
have "e;v \<le> 0"
by (metis assms annir mult_isor vector_prop1 comp_assoc)
thus ?thesis
- by (metis assms(2) annil antisym bot_least comp_assoc mult_isol)
+ by (metis assms(2) annil order.antisym bot_least comp_assoc mult_isol)
qed
lemma et:
assumes "is_vector v"
and "e \<le> v;-v\<^sup>T"
and "t \<le> v;v\<^sup>T"
shows "e;t = 0"
and "e;t\<^sup>T = 0"
proof -
have "e;t \<le> v;-v\<^sup>T;v;v\<^sup>T"
by (metis assms(2-3) mult_isol_var comp_assoc)
thus "e;t = 0"
by (simp add: assms(1) comp_assoc le_bot vector_prop1)
next
have "t\<^sup>T \<le> v;v\<^sup>T"
using assms(3) conv_iso by fastforce
hence "e;t\<^sup>T \<le> v;-v\<^sup>T;v;v\<^sup>T"
by (metis assms(2) mult_isol_var comp_assoc)
thus "e;t\<^sup>T = 0"
by (simp add: assms(1) comp_assoc le_bot vector_prop1)
qed
text \<open>Some (more) results about points\<close>
definition point
where "point x \<equiv> is_vector x \<and> bijective x"
lemma point_swap:
assumes "point p"
and "point q"
shows "p \<le> x;q \<longleftrightarrow> q \<le> x\<^sup>T;p"
by (metis assms conv_invol inj_sur_semi_swap point_def)
text \<open>Some (more) results about singletons\<close>
abbreviation singleton
where "singleton x \<equiv> bijective (x;1) \<and> bijective (x\<^sup>T;1)"
lemma singleton_injective:
assumes "singleton x"
shows "is_inj x"
using assms injective_down_closed maddux_20 by blast
lemma injective_inv:
assumes "is_vector v"
and "singleton e"
and "e \<le> v;-v\<^sup>T"
and "t \<le> v;v\<^sup>T"
and "is_inj t"
shows "is_inj (t + e)"
by (metis assms singleton_injective injective_sup bot_least et(2))
lemma singleton_is_point:
assumes "singleton p"
shows "point (p;1)"
by (simp add: assms comp_assoc is_vector_def point_def)
lemma singleton_transp:
assumes "singleton p"
shows "singleton (p\<^sup>T)"
by (simp add: assms)
lemma point_to_singleton:
assumes "singleton p"
shows "singleton (1'\<cdot>p;p\<^sup>T)"
using assms dom_def_aux_var dom_one is_vector_def point_def by fastforce
lemma singleton_singletonT:
assumes "singleton p"
shows "p;p\<^sup>T \<le> 1'"
using assms singleton_injective is_inj_def by blast
text \<open>Minimality\<close>
abbreviation minimum
where "minimum x v \<equiv> v \<cdot> -(x\<^sup>T;v)"
text \<open>Regressively finite\<close>
abbreviation regressively_finite
where "regressively_finite x \<equiv> \<forall>v . is_vector v \<and> v \<le> x\<^sup>T;v \<longrightarrow> v = 0"
lemma regressively_finite_minimum:
"regressively_finite R \<Longrightarrow> is_vector v \<Longrightarrow> v \<noteq> 0 \<Longrightarrow> minimum R v \<noteq> 0"
using galois_aux2 by blast
lemma regressively_finite_irreflexive:
assumes "regressively_finite x"
shows "x \<le> -1'"
proof -
have 1: "is_vector ((x\<^sup>T \<cdot> 1');1)"
by (simp add: is_vector_def mult_assoc)
have "(x\<^sup>T \<cdot> 1');1 = (x\<^sup>T \<cdot> 1');(x\<^sup>T \<cdot> 1');1"
by (simp add: is_test_def test_comp_eq_mult)
with 1 have "(x\<^sup>T \<cdot> 1');1 = 0"
by (metis assms comp_assoc mult_subdistr)
thus ?thesis
by (metis conv_e conv_invol conv_times conv_zero galois_aux ss_p18)
qed
end (* relation_algebra *)
subsection \<open>Relation algebras satisfying the Tarski rule\<close>
class relation_algebra_tarski = relation_algebra +
assumes tarski: "x \<noteq> 0 \<longleftrightarrow> 1;x;1 = 1"
begin
text \<open>Some (more) results about points\<close>
lemma point_equations:
assumes "is_point p"
shows "p;1=p"
and "1;p=1"
and "p\<^sup>T;1=1"
and "1;p\<^sup>T=p\<^sup>T"
apply (metis assms is_point_def is_vector_def)
using assms is_point_def is_vector_def tarski vector_comp apply fastforce
apply (metis assms conv_contrav conv_one conv_zero is_point_def is_vector_def tarski)
by (metis assms conv_contrav conv_one is_point_def is_vector_def)
text \<open>The following result is \cite[Proposition 2.4.5(i)]{SchmidtStroehlein1993}.\<close>
lemma point_singleton:
assumes "is_point p"
and "is_vector v"
and "v \<noteq> 0"
and "v \<le> p"
shows "v = p"
proof -
have "1;v = 1"
using assms(2,3) comp_assoc is_vector_def tarski by fastforce
hence "p = 1;v \<cdot> p"
by simp
also have "... \<le> (1 \<cdot> p;v\<^sup>T);(v \<cdot> 1\<^sup>T;p)"
using dedekind by blast
also have "... \<le> p;v\<^sup>T;v"
by (simp add: mult_subdistl)
also have "... \<le> p;p\<^sup>T;v"
using assms(4) conv_iso mult_double_iso by blast
also have "... \<le> v"
by (metis assms(1) is_inj_def is_point_def mult_isor mult_onel)
finally show ?thesis
using assms(4) by simp
qed
lemma point_not_equal_aux:
assumes "is_point p"
and "is_point q"
shows "p\<noteq>q \<longleftrightarrow> p \<cdot> -q \<noteq> 0"
proof
show "p \<noteq> q \<Longrightarrow> p \<cdot> - q \<noteq> 0"
proof (rule contrapos_nn)
assume "p \<cdot> -q = 0"
thus "p = q"
using assms galois_aux2 is_point_def point_singleton by fastforce
qed
next
show "p \<cdot> - q \<noteq> 0 \<Longrightarrow> p \<noteq> q"
using inf_compl_bot by blast
qed
text \<open>The following result is part of \cite[Proposition 2.4.5(ii)]{SchmidtStroehlein1993}.\<close>
lemma point_not_equal:
assumes "is_point p"
and "is_point q"
shows "p\<noteq>q \<longleftrightarrow> p\<le>-q"
and "p\<le>-q \<longleftrightarrow> p;q\<^sup>T \<le> -1'"
and "p;q\<^sup>T \<le> -1' \<longleftrightarrow> p\<^sup>T;q \<le> 0"
proof -
have "p \<noteq> q \<Longrightarrow> p \<le> - q"
by (metis assms point_not_equal_aux is_point_def vector_compl vector_mult point_singleton
inf.orderI inf.cobounded1)
thus "p\<noteq>q \<longleftrightarrow> p\<le>-q"
by (metis assms(1) galois_aux inf.orderE is_point_def order.refl)
next
show "(p \<le> - q) = (p ; q\<^sup>T \<le> - 1')"
by (simp add: conv_galois_2)
next
show "(p ; q\<^sup>T \<le> - 1') = (p\<^sup>T ; q \<le> 0)"
by (metis assms(2) compl_bot_eq conv_galois_2 galois_aux maddux_141 mult_1_right
point_equations(4))
qed
lemma point_is_point:
"point x \<longleftrightarrow> is_point x"
apply (rule iffI)
apply (simp add: is_point_def point_def surj_one tarski)
using is_point_def is_vector_def mult_assoc point_def sur_def_var1 tarski by fastforce
lemma point_in_vector_or_complement:
assumes "point p"
and "is_vector v"
shows "p \<le> v \<or> p \<le> -v"
proof (cases "p \<le> -v")
assume "p \<le> -v"
thus ?thesis
by simp
next
assume "\<not>(p \<le> -v)"
hence "p\<cdot>v \<noteq> 0"
by (simp add: galois_aux)
hence "1;(p\<cdot>v) = 1"
using assms comp_assoc is_vector_def point_def tarski vector_mult by fastforce
hence "p \<le> p;(p\<cdot>v)\<^sup>T;(p\<cdot>v)"
by (metis inf_top.left_neutral modular_2_var)
also have "... \<le> p;p\<^sup>T;v"
by (simp add: mult_isol_var)
also have "... \<le> v"
using assms(1) comp_assoc point_def ss423conv by fastforce
finally show ?thesis ..
qed
lemma point_in_vector_or_complement_iff:
assumes "point p"
and "is_vector v"
shows "p \<le> v \<longleftrightarrow> \<not>(p \<le> -v)"
by (metis assms annir compl_top_eq galois_aux inf.orderE one_compl point_def ss423conv tarski
top_greatest point_in_vector_or_complement)
lemma different_points_consequences:
assumes "point p"
and "point q"
and "p\<noteq>q"
shows "p\<^sup>T;-q=1"
and "-q\<^sup>T;p=1"
and "-(p\<^sup>T;-q)=0"
and "-(-q\<^sup>T;p)=0"
proof -
have "p \<le> -q"
by (metis assms compl_le_swap1 inf.absorb1 inf.absorb2 point_def point_in_vector_or_complement)
thus 1: "p\<^sup>T;-q=1"
using assms(1) by (metis is_vector_def point_def ss423conv top_le)
thus 2: "-q\<^sup>T;p=1"
using conv_compl conv_one by force
from 1 show "-(p\<^sup>T;-q)=0"
by simp
from 2 show "-(-q\<^sup>T;p)=0"
by simp
qed
text \<open>Some (more) results about singletons\<close>
lemma singleton_pq:
assumes "point p"
and "point q"
shows "singleton (p;q\<^sup>T)"
using assms comp_assoc point_def point_equations(1,3) point_is_point by fastforce
lemma singleton_equal_aux:
assumes "singleton p"
and "singleton q"
and "q\<le>p"
shows "p \<le> q;1"
proof -
have pLp: "p;1;p\<^sup>T \<le>1'"
by (simp add: assms(1) maddux_21 ss423conv)
have "p = 1;(q\<^sup>T;q;1) \<cdot> p"
using tarski
by (metis assms(2) annir singleton_injective inf.commute inf_top.right_neutral inj_triple
mult_assoc surj_one)
also have "... \<le> (1 \<cdot> p;(q\<^sup>T;q;1)\<^sup>T);(q\<^sup>T;q;1 \<cdot> 1;p)"
using dedekind by (metis conv_one)
also have "... \<le> p;1;q\<^sup>T;q;q\<^sup>T;q;1"
by (simp add: comp_assoc mult_isol)
also have "... \<le> p;1;p\<^sup>T;q;q\<^sup>T;q;1"
using assms(3) by (metis comp_assoc conv_iso mult_double_iso)
also have "... \<le> 1';q;q\<^sup>T;q;1"
using pLp using mult_isor by blast
also have "... \<le> q;1"
using assms(2) singleton_singletonT by (simp add: comp_assoc mult_isol)
finally show ?thesis .
qed
lemma singleton_equal:
assumes "singleton p"
and "singleton q"
and "q\<le>p"
shows "q=p"
proof -
have p1: "p \<le> q;1"
using assms by (rule singleton_equal_aux)
have "p\<^sup>T \<le> q\<^sup>T;1"
using assms singleton_equal_aux singleton_transp conv_iso by fastforce
hence p2: "p \<le> 1;q"
using conv_iso by force
have "p \<le> q;1 \<cdot> 1;q"
using p1 p2 inf.boundedI by blast
also have "... \<le> (q \<cdot> 1;q;1);(1 \<cdot> q\<^sup>T;1;q)"
using dedekind by (metis comp_assoc conv_one)
also have "... \<le> q;q\<^sup>T;1;q"
by (simp add: mult_isor comp_assoc)
also have "... \<le> q;1'"
by (metis assms(2) conv_contrav conv_invol conv_one is_inj_def mult_assoc mult_isol
one_idem_mult)
also have "... \<le> q"
by simp
finally have "p \<le> q" .
thus "q=p"
using assms(3) by simp
qed
lemma singleton_nonsplit:
assumes "singleton p"
and "x\<le>p"
shows "x=0 \<or> x=p"
proof (cases "x=0")
assume "x=0"
thus ?thesis ..
next
assume 1: "x\<noteq>0"
have "singleton x"
proof (safe)
show "is_inj (x;1)"
using assms injective_down_closed mult_isor by blast
show "is_inj (x\<^sup>T;1)"
using assms conv_iso injective_down_closed mult_isol_var by blast
show "is_sur (x;1)"
using 1 comp_assoc sur_def_var1 tarski by fastforce
thus "is_sur (x\<^sup>T;1)"
by (metis conv_contrav conv_one mult.semigroup_axioms sur_def_var1 semigroup.assoc)
qed
thus ?thesis
using assms singleton_equal by blast
qed
lemma singleton_nonzero:
assumes "singleton p"
shows "p\<noteq>0"
proof
assume "p = 0"
hence "point 0"
using assms singleton_is_point by fastforce
thus False
by (simp add: is_point_def point_is_point)
qed
lemma singleton_sum:
assumes "singleton p"
shows "p \<le> x+y \<longleftrightarrow> (p\<le>x \<or> p\<le>y)"
proof
show "p \<le> x + y \<Longrightarrow> p \<le> x \<or> p \<le> y"
proof -
assume as: "p \<le> x + y"
show "p \<le> x \<or> p \<le> y"
proof (cases "p\<le>x")
assume "p\<le>x"
thus ?thesis ..
next
assume a:"\<not>(p\<le>x)"
hence "p\<cdot>x \<noteq> p"
using a inf.orderI by fastforce
hence "p \<le> -x"
using assms singleton_nonsplit galois_aux inf_le1 by blast
hence "p\<le>y"
using as by (metis galois_1 inf.orderE)
thus ?thesis
by simp
qed
qed
next
show "p \<le> x \<or> p \<le> y \<Longrightarrow> p \<le> x + y"
using sup.coboundedI1 sup.coboundedI2 by blast
qed
lemma singleton_iff:
"singleton x \<longleftrightarrow> x \<noteq> 0 \<and> x\<^sup>T;1;x + x;1;x\<^sup>T \<le> 1'"
by (smt comp_assoc conv_contrav conv_invol conv_one is_inj_def le_sup_iff one_idem_mult
sur_def_var1 tarski)
lemma singleton_not_atom_in_relation_algebra_tarski:
assumes "p\<noteq>0"
and "\<forall>x . x\<le>p \<longrightarrow> x=0 \<or> x=p"
shows "singleton p"
nitpick [expect=genuine] oops
end (* relation_algebra_tarski *)
subsection \<open>Relation algebras satisfying the point axiom\<close>
class relation_algebra_point = relation_algebra +
assumes point_axiom: "x \<noteq> 0 \<longrightarrow> (\<exists>y z . point y \<and> point z \<and> y;z\<^sup>T \<le> x)"
begin
text \<open>Some (more) results about points\<close>
lemma point_exists:
"\<exists>x . point x"
-by (metis (full_types) eq_iff is_inj_def is_sur_def is_vector_def point_axiom point_def)
+by (metis (full_types) order.eq_iff is_inj_def is_sur_def is_vector_def point_axiom point_def)
lemma point_below_vector:
assumes "is_vector v"
and "v \<noteq> 0"
shows "\<exists>x . point x \<and> x \<le> v"
proof -
from assms(2) obtain y and z where 1: "point y \<and> point z \<and> y;z\<^sup>T \<le> v"
using point_axiom by blast
have "z\<^sup>T;1 = (1;z)\<^sup>T"
using conv_contrav conv_one by simp
hence "y;(1;z)\<^sup>T \<le> v"
using 1 by (metis assms(1) comp_assoc is_vector_def mult_isor)
thus ?thesis
using 1 by (metis conv_one is_vector_def point_def sur_def_var1)
qed
end (* relation_algebra_point *)
class relation_algebra_tarski_point = relation_algebra_tarski + relation_algebra_point
begin
lemma atom_is_singleton:
assumes "p\<noteq>0"
and "\<forall>x . x\<le>p \<longrightarrow> x=0 \<or> x=p"
shows "singleton p"
by (metis assms singleton_nonzero singleton_pq point_axiom)
lemma singleton_iff_atom:
"singleton p \<longleftrightarrow> p\<noteq>0 \<and> (\<forall>x . x\<le>p \<longrightarrow> x=0 \<or> x=p)"
using singleton_nonsplit singleton_nonzero atom_is_singleton by blast
lemma maddux_tarski:
assumes "x\<noteq>0"
shows "\<exists>y . y\<noteq>0 \<and> y\<le>x \<and> is_p_fun y"
proof -
obtain p q where 1: "point p \<and> point q \<and> p;q\<^sup>T \<le> x"
using assms point_axiom by blast
hence 2: "p;q\<^sup>T\<noteq>0"
by (simp add: singleton_nonzero singleton_pq)
have "is_p_fun (p;q\<^sup>T)"
using 1 by (meson singleton_singletonT singleton_pq singleton_transp is_inj_def p_fun_inj)
thus ?thesis
using 1 2 by force
qed
text \<open>Intermediate Point Theorem \cite[Proposition 2.4.8]{SchmidtStroehlein1993}\<close>
lemma intermediate_point_theorem:
assumes "point p"
and "point r"
shows "p \<le> x;y;r \<longleftrightarrow> (\<exists>q . point q \<and> p \<le> x;q \<and> q \<le> y;r)"
proof
assume 1: "p \<le> x;y;r"
let ?v = "x\<^sup>T;p \<cdot> y;r"
have 2: "is_vector ?v"
using assms comp_assoc is_vector_def point_def vector_mult by fastforce
have "?v \<noteq> 0"
using 1 by (metis assms(1) inf.absorb2 is_point_def maddux_141 point_is_point mult.assoc)
hence "\<exists>q . point q \<and> q \<le> ?v"
using 2 point_below_vector by blast
thus "\<exists>q . point q \<and> p \<le> x;q \<and> q \<le> y;r"
using assms(1) point_swap by auto
next
assume "\<exists>q . point q \<and> p \<le> x;q \<and> q \<le> y;r"
thus "p \<le> x;y;r"
using comp_assoc mult_isol order_trans by fastforce
qed
end (* relation_algebra_tarski_point *)
(*
The following shows that rtc can be defined with only 2 axioms.
This should eventually go into AFP/Relation_Algebra_RTC.relation_algebra_rtc.
There the class definition should be replaced with:
class relation_algebra_rtc = relation_algebra + star_op +
assumes rtc_unfoldl: "1' + x ; x\<^sup>\<star> \<le> x\<^sup>\<star>"
and rtc_inductl: "z + x ; y \<le> y \<longrightarrow> x\<^sup>\<star> ; z \<le> y"
and the following lemmas:
*)
context relation_algebra
begin
lemma unfoldl_inductl_implies_unfoldr:
assumes "\<And>x. 1' + x;(rtc x) \<le> rtc x"
and "\<And>x y z. x+y;z \<le> z \<Longrightarrow> rtc(y);x \<le> z"
shows "1' + rtc(x);x \<le> rtc x"
by (metis assms le_sup_iff mult_oner order.trans subdistl_eq sup_absorb2 sup_ge1)
lemma star_transpose_swap:
assumes "\<And>x. 1' + x;(rtc x) \<le> rtc x"
and "\<And>x y z. x+y;z \<le> z \<Longrightarrow> rtc(y);x \<le> z"
shows "rtc(x\<^sup>T) = (rtc x)\<^sup>T"
-apply(simp only: eq_iff; rule conjI)
+apply(simp only: order.eq_iff; rule conjI)
apply (metis assms conv_add conv_contrav conv_e conv_iso mult_1_right
unfoldl_inductl_implies_unfoldr )
by (metis assms conv_add conv_contrav conv_e conv_invol conv_iso mult_1_right
unfoldl_inductl_implies_unfoldr)
lemma unfoldl_inductl_implies_inductr:
assumes "\<And>x. 1' + x;(rtc x) \<le> rtc x"
and "\<And>x y z. x+y;z \<le> z \<Longrightarrow> rtc(y);x \<le> z"
shows "x+z;y \<le> z \<Longrightarrow> x;rtc(y) \<le> z"
by (metis assms conv_add conv_contrav conv_iso star_transpose_swap)
end (* relation_algebra *)
context relation_algebra_rtc
begin
abbreviation tc ("(_\<^sup>+)" [101] 100) where "tc x \<equiv> x;x\<^sup>\<star>"
abbreviation is_acyclic
where "is_acyclic x \<equiv> x\<^sup>+ \<le> -1'"
text \<open>General theorems\<close>
lemma star_denest_10:
assumes "x;y=0"
shows "(x+y)\<^sup>\<star> = y;y\<^sup>\<star>;x\<^sup>\<star>+x\<^sup>\<star>"
using assms bubble_sort sup.commute by auto
lemma star_star_plus:
"x\<^sup>\<star> + y\<^sup>\<star> = x\<^sup>+ + y\<^sup>\<star>"
by (metis (full_types) sup.left_commute star_plus_one star_unfoldl_eq sup.commute)
text \<open>The following two lemmas are from \cite{Guttmann2018b}.\<close>
lemma cancel_separate:
assumes "x ; y \<le> 1'"
shows "x\<^sup>\<star> ; y\<^sup>\<star> \<le> x\<^sup>\<star> + y\<^sup>\<star>"
proof -
have "x ; y\<^sup>\<star> = x + x ; y ; y\<^sup>\<star>"
by (metis comp_assoc conway.dagger_unfoldl_distr distrib_left mult_oner)
also have "... \<le> x + y\<^sup>\<star>"
by (metis assms join_isol star_invol star_plus_one star_subdist_var_2 sup.absorb2 sup.assoc)
also have "... \<le> x\<^sup>\<star> + y\<^sup>\<star>"
using join_iso by fastforce
finally have "x ; (x\<^sup>\<star> + y\<^sup>\<star>) \<le> x\<^sup>\<star> + y\<^sup>\<star>"
by (simp add: distrib_left le_supI1)
thus ?thesis
by (simp add: rtc_inductl)
qed
lemma cancel_separate_inj_converse:
assumes "is_inj x"
shows "x\<^sup>\<star> ; x\<^sup>T\<^sup>\<star> = x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
- apply (rule antisym)
+ apply (rule order.antisym)
using assms cancel_separate is_inj_def apply blast
by (metis conway.dagger_unfoldl_distr le_supI mult_1_right mult_isol sup.cobounded1)
lemma cancel_separate_p_fun_converse:
assumes "is_p_fun x"
shows "x\<^sup>T\<^sup>\<star> ; x\<^sup>\<star> = x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using sup_commute assms cancel_separate_inj_converse p_fun_inj by fastforce
lemma cancel_separate_converse_idempotent:
assumes "is_inj x"
and "is_p_fun x"
shows "(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>) = x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis assms cancel_separate cancel_separate_p_fun_converse church_rosser_equiv is_inj_def
star_denest_var_6)
lemma triple_star:
assumes "is_inj x"
and "is_p_fun x"
shows "x\<^sup>\<star>;x\<^sup>T\<^sup>\<star>;x\<^sup>\<star> = x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (simp add: assms cancel_separate_inj_converse cancel_separate_p_fun_converse)
lemma inj_xxts:
assumes "is_inj x"
shows "x;x\<^sup>T\<^sup>\<star> \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis assms cancel_separate_inj_converse distrib_right less_eq_def star_ext)
lemma plus_top:
"x\<^sup>+;1 = x;1"
by (metis comp_assoc conway.dagger_unfoldr_distr sup_top_left)
lemma top_plus:
"1;x\<^sup>+ = 1;x"
by (metis comp_assoc conway.dagger_unfoldr_distr star_denest_var_2 star_ext star_slide_var
sup_top_left top_unique)
lemma plus_conv:
"(x\<^sup>+)\<^sup>T = x\<^sup>T\<^sup>+"
by (simp add: star_conv star_slide_var)
lemma inj_implies_step_forwards_backwards:
assumes "is_inj x"
shows "x\<^sup>\<star>;(x\<^sup>+\<cdot>1');1 \<le> x\<^sup>T;1"
proof -
have "(x\<^sup>+\<cdot>1');1 \<le> (x\<^sup>\<star>\<cdot>x\<^sup>T);(x\<cdot>(x\<^sup>\<star>)\<^sup>T);1"
by (metis conv_contrav conv_e dedekind mult_1_right mult_isor star_slide_var)
also have "... \<le> (x\<^sup>\<star>\<cdot>x\<^sup>T);1"
by (simp add: comp_assoc mult_isol)
finally have 1: "(x\<^sup>+\<cdot>1');1 \<le> (x\<^sup>\<star>\<cdot>x\<^sup>T);1" .
have "x;(x\<^sup>\<star>\<cdot>x\<^sup>T);1 \<le> (x\<^sup>+\<cdot>x;x\<^sup>T);1"
by (metis inf_idem meet_interchange mult_isor)
also have "... \<le> (x\<^sup>+\<cdot>1');1"
using assms is_inj_def meet_isor mult_isor by fastforce
finally have "x;(x\<^sup>\<star>\<cdot>x\<^sup>T);1 \<le> (x\<^sup>\<star>\<cdot>x\<^sup>T);1"
using 1 by fastforce
hence "x\<^sup>\<star>;(x\<^sup>+\<cdot>1');1 \<le> (x\<^sup>\<star>\<cdot>x\<^sup>T);1"
using 1 by (simp add: comp_assoc rtc_inductl)
thus "x\<^sup>\<star>;(x\<^sup>+\<cdot>1');1 \<le> x\<^sup>T;1"
using inf.cobounded2 mult_isor order_trans by blast
qed
text \<open>Acyclic relations\<close>
text \<open>The following result is from \cite{Guttmann2017c}.\<close>
lemma acyclic_inv:
assumes "is_acyclic t"
and "is_vector v"
and "e \<le> v;-v\<^sup>T"
and "t \<le> v;v\<^sup>T"
shows "is_acyclic (t + e)"
proof -
have "t\<^sup>+;e \<le> t\<^sup>+;v;-v\<^sup>T"
by (simp add: assms(3) mult_assoc mult_isol)
also have "... \<le> v;v\<^sup>T;t\<^sup>\<star>;v;-v\<^sup>T"
by (simp add: assms(4) mult_isor)
also have "... \<le> v;-v\<^sup>T"
by (metis assms(2) mult_double_iso top_greatest is_vector_def mult_assoc)
also have "... \<le> -1'"
by (simp add: conv_galois_1)
finally have 1: "t\<^sup>+;e \<le> -1'" .
have "e \<le> v;-v\<^sup>T"
using assms(3) by simp
also have "... \<le> -1'"
by (simp add: conv_galois_1)
finally have 2: "t\<^sup>+;e + e \<le> -1'"
using 1 by simp
have 3: "e;t\<^sup>\<star> = e"
by (metis assms(2-4) et(1) independence2)
have 4: "e\<^sup>\<star> = 1' + e"
using assms(2-3) ee boffa_var bot_least by blast
have "(t + e)\<^sup>+ = (t + e);t\<^sup>\<star>;(e;t\<^sup>\<star>)\<^sup>\<star>"
by (simp add: comp_assoc)
also have "... = (t + e);t\<^sup>\<star>;(1' + e)"
using 3 4 by simp
also have "... = t\<^sup>+;(1' + e) + e;t\<^sup>\<star>;(1' + e)"
by simp
also have "... = t\<^sup>+;(1' + e) + e;(1' + e)"
using 3 by simp
also have "... = t\<^sup>+;(1' + e) + e"
using 4 assms(2-3) ee independence2 by fastforce
also have "... = t\<^sup>+ + t\<^sup>+;e + e"
by (simp add: distrib_left)
also have "... \<le> -1'"
using assms(1) 2 by simp
finally show ?thesis .
qed
lemma acyclic_single_step:
assumes "is_acyclic x"
shows "x \<le> -1'"
by (metis assms dual_order.trans mult_isol mult_oner star_ref)
lemma acyclic_reachable_points:
assumes "is_point p"
and "is_point q"
and "p \<le> x;q"
and "is_acyclic x"
shows "p\<noteq>q"
proof
assume "p=q"
hence "p \<le> x;q \<cdot> q"
- by (simp add: assms(3) eq_iff inf.absorb2)
+ by (simp add: assms(3) order.eq_iff inf.absorb2)
also have "... = (x \<cdot> 1');q"
using assms(2) inj_distr is_point_def by simp
also have "... \<le> (-1' \<cdot> 1');q"
using acyclic_single_step assms(4) by (metis abel_semigroup.commute inf.abel_semigroup_axioms
meet_isor mult_isor)
also have "... = 0"
by simp
finally have "p \<le> 0" .
thus False
using assms(1) bot_unique is_point_def by blast
qed
lemma acyclic_trans:
assumes "is_acyclic x"
shows "x \<le> -(x\<^sup>T\<^sup>+)"
proof -
have "\<exists>c\<ge>x. c \<le> - (x\<^sup>+)\<^sup>T"
by (metis assms compl_mono conv_galois_2 conv_iso double_compl mult_onel star_1l)
thus ?thesis
by (metis dual_order.trans plus_conv)
qed
lemma acyclic_trans':
assumes "is_acyclic x"
shows "x\<^sup>\<star> \<le> -(x\<^sup>T\<^sup>+)"
proof -
have "x\<^sup>\<star> \<le> - (- (- (x\<^sup>T ; - (- 1'))) ; (x\<^sup>\<star>)\<^sup>T)"
by (metis assms conv_galois_1 conv_galois_2 order_trans star_trans)
then show ?thesis
by (simp add: star_conv)
qed
text \<open>Regressively finite\<close>
lemma regressively_finite_acyclic:
assumes "regressively_finite x"
shows "is_acyclic x"
proof -
have 1: "is_vector ((x\<^sup>+ \<cdot> 1');1)"
by (simp add: is_vector_def mult_assoc)
have "(x\<^sup>+ \<cdot> 1');1 = (x\<^sup>T\<^sup>+ \<cdot> 1');1"
by (metis plus_conv test_converse)
also have "... \<le> x\<^sup>T;(1';x\<^sup>T\<^sup>\<star> \<cdot> x);1"
by (metis conv_invol modular_1_var mult_isor mult_oner mult_onel)
also have "... \<le> x\<^sup>T;(1' \<cdot> x\<^sup>+);x\<^sup>T\<^sup>\<star>;1"
by (metis comp_assoc conv_invol modular_2_var mult_isol mult_isor star_conv)
also have "... = x\<^sup>T;(x\<^sup>+ \<cdot> 1');1"
by (metis comp_assoc conway.dagger_unfoldr_distr inf.commute sup.cobounded1 top_le)
finally have "(x\<^sup>+ \<cdot> 1');1 = 0"
using 1 assms by (simp add: comp_assoc)
thus ?thesis
by (simp add: galois_aux ss_p18)
qed
notation power (infixr "\<up>" 80)
lemma power_suc_below_plus:
"x \<up> Suc n \<le> x\<^sup>+"
apply (induct n)
using mult_isol star_ref apply fastforce
by (simp add: mult_isol_var order_trans)
end (* relation_algebra_rtc *)
class relation_algebra_rtc_tarski = relation_algebra_rtc + relation_algebra_tarski
begin
lemma point_loop_not_acyclic:
assumes "is_point p"
and "p \<le> x \<up> Suc n ; p"
shows "\<not> is_acyclic x"
proof -
have "p \<le> x\<^sup>+ ; p"
by (meson assms dual_order.trans point_def point_is_point ss423bij power_suc_below_plus)
hence "p ; p\<^sup>T \<le> x\<^sup>+"
using assms(1) point_def point_is_point ss423bij by blast
thus ?thesis
using assms(1) order.trans point_not_equal(1) point_not_equal(2) by blast
qed
end
class relation_algebra_rtc_point = relation_algebra_rtc + relation_algebra_point
class relation_algebra_rtc_tarski_point = relation_algebra_rtc_tarski + relation_algebra_rtc_point +
relation_algebra_tarski_point
text \<open>
Finite graphs: the axiom says the algebra has finitely many elements.
This means the relations have a finite base set.
\<close>
class relation_algebra_rtc_tarski_point_finite = relation_algebra_rtc_tarski_point + finite
begin
text \<open>For a finite acyclic relation, the powers eventually vanish.\<close>
lemma acyclic_power_vanishes:
assumes "is_acyclic x"
shows "\<exists>n . x \<up> Suc n = 0"
proof -
let ?n = "card { p . is_point p }"
let ?p = "x \<up> ?n"
have "?p = 0"
proof (rule ccontr)
assume "?p \<noteq> 0"
from this obtain p q where 1: "point p \<and> point q \<and> p;q\<^sup>T \<le> ?p"
using point_axiom by blast
hence 2: "p \<le> ?p;q"
using point_def ss423bij by blast
have "\<forall>n\<le>?n . (\<exists>f. \<forall>i\<le>n . is_point (f i) \<and> (\<forall>j\<le>i . p \<le> x\<up>(?n-i) ; f i \<and> f i \<le> x\<up>(i-j) ; f j))"
proof
fix n
show "n\<le>?n \<longrightarrow> (\<exists>f. \<forall>i\<le>n . is_point (f i) \<and> (\<forall>j\<le>i . p \<le> x\<up>(?n-i) ; f i \<and> f i \<le> x\<up>(i-j) ; f j))"
proof (induct n)
case 0
thus ?case
using 1 2 point_is_point by fastforce
next
case (Suc n)
fix n
assume 3: "n\<le>?n \<longrightarrow> (\<exists>f . \<forall>i\<le>n . is_point (f i) \<and> (\<forall>j\<le>i . p \<le> x \<up> (?n-i) ; f i \<and> f i \<le> x \<up> (i-j) ; f j))"
show "Suc n\<le>?n \<longrightarrow> (\<exists>f . \<forall>i\<le>Suc n . is_point (f i) \<and> (\<forall>j\<le>i . p \<le> x \<up> (?n-i) ; f i \<and> f i \<le> x \<up> (i-j) ; f j))"
proof
assume 4: "Suc n\<le>?n"
from this obtain f where 5: "\<forall>i\<le>n . is_point (f i) \<and> (\<forall>j\<le>i . p \<le> x \<up> (?n-i) ; f i \<and> f i \<le> x \<up> (i-j) ; f j)"
using 3 by auto
have "p \<le> x \<up> (?n-n) ; f n"
using 5 by blast
also have "... = x \<up> (?n-n-one_class.one) ; x ; f n"
using 4 by (metis (no_types) Suc_diff_le diff_Suc_1 diff_Suc_Suc power_Suc2)
finally obtain r where 6: "point r \<and> p \<le> x \<up> (?n-Suc n) ; r \<and> r \<le> x ; f n"
using 1 5 intermediate_point_theorem point_is_point by fastforce
let ?g = "\<lambda>m . if m = Suc n then r else f m"
have "\<forall>i\<le>Suc n . is_point (?g i) \<and> (\<forall>j\<le>i . p \<le> x \<up> (?n-i) ; ?g i \<and> ?g i \<le> x \<up> (i-j) ; ?g j)"
proof
fix i
show "i\<le>Suc n \<longrightarrow> is_point (?g i) \<and> (\<forall>j\<le>i . p \<le> x \<up> (?n-i) ; ?g i \<and> ?g i \<le> x \<up> (i-j) ; ?g j)"
proof (cases "i\<le>n")
case True
thus ?thesis
using 5 by simp
next
case False
have "is_point (?g (Suc n)) \<and> (\<forall>j\<le>Suc n . p \<le> x \<up> (?n-Suc n) ; ?g (Suc n) \<and> ?g (Suc n) \<le> x \<up> (Suc n-j) ; ?g j)"
proof
show "is_point (?g (Suc n))"
using 6 point_is_point by fastforce
next
show "\<forall>j\<le>Suc n . p \<le> x \<up> (?n-Suc n) ; ?g (Suc n) \<and> ?g (Suc n) \<le> x \<up> (Suc n-j) ; ?g j"
proof
fix j
show "j\<le>Suc n \<longrightarrow> p \<le> x \<up> (?n-Suc n) ; ?g (Suc n) \<and> ?g (Suc n) \<le> x \<up> (Suc n-j) ; ?g j"
proof
assume 7: "j\<le>Suc n"
show "p \<le> x \<up> (?n-Suc n) ; ?g (Suc n) \<and> ?g (Suc n) \<le> x \<up> (Suc n-j) ; ?g j"
proof
show "p \<le> x \<up> (?n-Suc n) ; ?g (Suc n)"
using 6 by simp
next
show "?g (Suc n) \<le> x \<up> (Suc n-j) ; ?g j"
proof (cases "j = Suc n")
case True
thus ?thesis
by simp
next
case False
hence "f n \<le> x \<up> (n-j) ; f j"
using 5 7 by fastforce
hence "x ; f n \<le> x \<up> (Suc n-j) ; f j"
using 7 False Suc_diff_le comp_assoc mult_isol by fastforce
thus ?thesis
using 6 False by fastforce
qed
qed
qed
qed
qed
thus ?thesis
by (simp add: False le_Suc_eq)
qed
qed
thus "\<exists>f . \<forall>i\<le>Suc n . is_point (f i) \<and> (\<forall>j\<le>i . p \<le> x \<up> (?n-i) ; f i \<and> f i \<le> x \<up> (i-j) ; f j)"
by auto
qed
qed
qed
from this obtain f where 8: "\<forall>i\<le>?n . is_point (f i) \<and> (\<forall>j\<le>i . p \<le> x \<up> (?n-i) ; f i \<and> f i \<le> x \<up> (i-j) ; f j)"
by fastforce
let ?A = "{ k . k\<le>?n }"
have "f ` ?A \<subseteq> { p . is_point p }"
using 8 by blast
hence "card (f ` ?A) \<le> ?n"
by (simp add: card_mono)
hence "\<not> inj_on f ?A"
by (simp add: pigeonhole)
from this obtain i j where 9: "i \<le> ?n \<and> j \<le> ?n \<and> i \<noteq> j \<and> f i = f j"
by (metis (no_types, lifting) inj_on_def mem_Collect_eq)
show False
apply (cases "i < j")
using 8 9 apply (metis Suc_diff_le Suc_leI assms diff_Suc_Suc order_less_imp_le
point_loop_not_acyclic)
using 8 9 by (metis assms neqE point_loop_not_acyclic Suc_diff_le Suc_leI assms diff_Suc_Suc
order_less_imp_le)
qed
thus ?thesis
by (metis annir power.simps(2))
qed
text \<open>Hence finite acyclic relations are regressively finite.\<close>
lemma acyclic_regressively_finite:
assumes "is_acyclic x"
shows "regressively_finite x"
proof
have "is_acyclic (x\<^sup>T)"
using assms acyclic_trans' compl_le_swap1 order_trans star_ref by blast
from this obtain n where 1: "x\<^sup>T \<up> Suc n = 0"
using acyclic_power_vanishes by fastforce
fix v
show "is_vector v \<and> v \<le> x\<^sup>T;v \<longrightarrow> v = 0"
proof
assume 2: "is_vector v \<and> v \<le> x\<^sup>T;v"
have "v \<le> x\<^sup>T \<up> Suc n ; v"
proof (induct n)
case 0
thus ?case
using 2 by simp
next
case (Suc n)
hence "x\<^sup>T ; v \<le> x\<^sup>T \<up> Suc (Suc n) ; v"
by (simp add: comp_assoc mult_isol)
thus ?case
using 2 dual_order.trans by blast
qed
thus "v = 0"
using 1 by (simp add: le_bot)
qed
qed
lemma acyclic_is_regressively_finite:
"is_acyclic x \<longleftrightarrow> regressively_finite x"
using acyclic_regressively_finite regressively_finite_acyclic by blast
end (* end relation_algebra_rtc_tarski_point_finite *)
end
diff --git a/thys/Relational_Paths/Path_Algorithms.thy b/thys/Relational_Paths/Path_Algorithms.thy
--- a/thys/Relational_Paths/Path_Algorithms.thy
+++ b/thys/Relational_Paths/Path_Algorithms.thy
@@ -1,995 +1,995 @@
(* Title: Correctness of Path Algorithms
Author: Walter Guttmann, Peter Hoefner
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
Peter Hoefner <peter at hoefner-online.de>
*)
section \<open>Correctness of Path Algorithms\<close>
text \<open>
To show that our theory of paths integrates with verification tasks, we verify the correctness of three basic path algorithms.
Algorithms at the presented level are executable and can serve prototyping purposes.
Data refinement can be carried out to move from such algorithms to more efficient programs.
The total-correctness proofs use a library developed in \cite{Guttmann2018c}.
\<close>
theory Path_Algorithms
imports "HOL-Hoare.Hoare_Logic" Rooted_Paths
begin
no_notation
trancl ("(_\<^sup>+)" [1000] 999)
class choose_singleton_point_signature =
fixes choose_singleton :: "'a \<Rightarrow> 'a"
fixes choose_point :: "'a \<Rightarrow> 'a"
class relation_algebra_rtc_tarski_choose_point =
relation_algebra_rtc_tarski + choose_singleton_point_signature +
assumes choose_singleton_singleton: "x \<noteq> 0 \<Longrightarrow> singleton (choose_singleton x)"
assumes choose_singleton_decreasing: "choose_singleton x \<le> x"
assumes choose_point_point: "is_vector x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> point (choose_point x)"
assumes choose_point_decreasing: "choose_point x \<le> x"
begin
no_notation
composition (infixl ";" 75) and
times (infixl "*" 70)
notation
composition (infixl "*" 75)
subsection \<open>Construction of a path\<close>
text \<open>
Our first example is a basic greedy algorithm that constructs a path from a vertex $x$ to a different vertex $y$ of a directed acyclic graph $D$.
\<close>
abbreviation "construct_path_inv q x y D W \<equiv>
is_acyclic D \<and> point x \<and> point y \<and> point q \<and>
D\<^sup>\<star> * q \<le> D\<^sup>T\<^sup>\<star> * x \<and> W \<le> D \<and> terminating_path W \<and>
(W = 0 \<longleftrightarrow> q=y) \<and> (W \<noteq> 0 \<longleftrightarrow> q = start_points W \<and> y = end_points W)"
abbreviation "construct_path_inv_simp q x y D W \<equiv>
is_acyclic D \<and> point x \<and> point y \<and> point q \<and>
D\<^sup>\<star> * q \<le> D\<^sup>T\<^sup>\<star> * x \<and> W \<le> D \<and> terminating_path W \<and>
q = start_points W \<and> y = end_points W"
lemma construct_path_pre:
assumes "is_acyclic D"
and "point y"
and "point x"
and "D\<^sup>\<star> * y \<le> D\<^sup>T\<^sup>\<star> * x"
shows "construct_path_inv y x y D 0"
apply (intro conjI, simp_all add: assms is_inj_def is_p_fun_def path_def)
using assms(2) cycle_iff by fastforce
text \<open>
The following three lemmas are auxiliary lemmas for \<open>construct_path_inv\<close>.
They are pulled out of the main proof to have more structure.
\<close>
lemma path_inv_points:
assumes "construct_path_inv q x y D W \<and> q \<noteq> x"
shows "point q"
and "point (choose_point (D*q))"
using assms apply blast
by (metis assms choose_point_point comp_assoc is_vector_def point_def reachable_implies_predecessor)
lemma path_inv_choose_point_decrease:
assumes "construct_path_inv q x y D W \<and> q \<noteq> x"
shows "W\<noteq>0 \<Longrightarrow> choose_point (D*q) \<le> -((W + choose_point (D*q) * q\<^sup>T)\<^sup>T*1)"
proof -
let ?q = "choose_point (D*q)"
let ?W = "W + ?q * q\<^sup>T"
assume as: "W\<noteq>0"
hence "q*W \<le> W\<^sup>+" (* "connected_root q W" *)
by (metis assms conv_contrav conv_invol conv_iso conv_terminating_path
forward_terminating_path_end_points_1 plus_conv point_def ss423bij
terminating_path_iff)
hence "?q \<cdot> W\<^sup>T*1 \<le> D*q \<cdot> W\<^sup>T\<^sup>+*q"
using choose_point_decreasing meet_iso meet_isor inf_mono assms connected_root_iff2 by simp
also have "... \<le> (D \<cdot> D\<^sup>T\<^sup>+)*q"
by (metis assms inj_distr point_def conv_contrav conv_invol conv_iso meet_isor
mult_isol_var mult_isor star_conv star_slide_var star_subdist sup.commute sup.orderE)
also have "... \<le> 0"
- by (metis acyclic_trans assms conv_zero step_has_target eq_iff galois_aux ss_p18)
+ by (metis acyclic_trans assms conv_zero step_has_target order.eq_iff galois_aux ss_p18)
finally have a: "?q \<le> -(W\<^sup>T*1)"
using galois_aux le_bot by blast
have "point ?q"
using assms by(rule path_inv_points(2))
hence "?q \<le> -(q*?q\<^sup>T*1)"
by (metis assms acyclic_imp_one_step_different_points(2) point_is_point
choose_point_decreasing edge_end end_point_char end_point_no_successor)
with a show ?thesis
by (simp add: inf.boundedI)
qed
lemma end_points:
assumes "construct_path_inv q x y D W \<and> q \<noteq> x"
shows "choose_point (D*q) = start_points (W + choose_point (D*q) * q\<^sup>T)"
and "y = end_points (W + choose_point (D*q) * q\<^sup>T)"
proof -
let ?q = "choose_point (D*q)"
let ?W = "W + ?q * q\<^sup>T"
show 1: "?q = start_points ?W"
- proof (rule antisym)
+ proof (rule order.antisym)
show" start_points ?W \<le> ?q"
by (metis assms(1) path_inv_points(2) acyclic_imp_one_step_different_points(2)
choose_point_decreasing edge_end edge_start sup.commute
- path_concatenation_start_points_approx point_is_point eq_iff sup_bot_left)
+ path_concatenation_start_points_approx point_is_point order.eq_iff sup_bot_left)
show "?q \<le> start_points ?W"
proof -
have a: "?q = ?q*q\<^sup>T*1"
by (metis assms(1) comp_assoc point_equations(1) point_is_point aux4 conv_zero
choose_point_decreasing choose_point_point conv_contrav conv_one point_def
inf.orderE inf_compl_bot inf_compl_bot_right is_vector_def maddux_142
sup_bot_left sur_def_var1)
hence "?q =(q \<cdot> -q) + (?q \<cdot> -q \<cdot> -(?W\<^sup>T*1))"
by (metis assms path_inv_points(2) path_inv_choose_point_decrease
acyclic_imp_one_step_different_points(1) choose_point_decreasing inf.orderE
inf_compl_bot sup_inf_absorb edge_start point_is_point sup_bot_left)
also have "... \<le> (W*1 \<cdot> -(?W\<^sup>T*1) \<cdot> -q) + (?q \<cdot> -q \<cdot> -(?W\<^sup>T*1))"
by simp
also have "... = (W*1 + ?q) \<cdot> -(q + ?W\<^sup>T*1)"
by (metis compl_sup inf_sup_distrib2 meet_assoc sup.commute)
also have "... \<le> ?W*1 \<cdot> -(?W\<^sup>T*1)"
using a by (metis inf.left_commute distrib_right' compl_sup inf.cobounded2)
finally show "?q \<le> start_points ?W" .
qed
qed
show "y = end_points ?W"
proof -
have point_nq: "point ?q"
using assms by(rule path_inv_points(2))
hence yp: "y \<le> -?q"
using 1 assms
by (metis acyclic_imp_one_step_different_points(2) choose_point_decreasing cycle_no_points(1)
finite_iff finite_iff_msc forward_finite_iff_msc path_aux1a path_edge_equals_cycle
point_is_point point_not_equal(1) terminating_iff1)
have "y = y + (W*1 \<cdot> -(W\<^sup>T*1) \<cdot> -(W*1))"
by (simp add: inf.commute)
also have "... = y + (q \<cdot> -(W*1))"
using assms by fastforce
also have "... = y + (q \<cdot> -(W*1) \<cdot> -?q)"
by (metis calculation sup_assoc sup_inf_absorb)
also have "... = (y \<cdot> -?q) + (q \<cdot> -(W*1) \<cdot> -?q)"
using yp by (simp add: inf.absorb1)
also have "... = (W\<^sup>T*1 \<cdot> -(W*1) \<cdot> -?q) + (q \<cdot> -(W*1) \<cdot> -?q)"
using assms by fastforce
also have "... = (W\<^sup>T*1 + q) \<cdot> -(W*1) \<cdot> -?q"
by (simp add: inf_sup_distrib2)
also have "... = (W\<^sup>T*1 + q) \<cdot> -(W*1 + ?q)"
by (simp add: inf.assoc)
also have "... = (W\<^sup>T*1 + q*?q\<^sup>T*1) \<cdot> -(W*1 + ?q*q\<^sup>T*1)"
using point_nq
by(metis assms(1) comp_assoc conv_contrav conv_one is_vector_def point_def sur_def_var1)
also have "... = (?W\<^sup>T)*1 \<cdot> -(?W*1)"
by simp
finally show ?thesis .
qed
qed
lemma construct_path_inv:
assumes "construct_path_inv q x y D W \<and> q \<noteq> x"
shows "construct_path_inv (choose_point (D*q)) x y D (W + choose_point (D*q)*q\<^sup>T)"
proof (intro conjI)
let ?q = "choose_point (D*q)"
let ?W = "W + ?q * q\<^sup>T"
show "is_acyclic D"
using assms by blast
show point_y: "point y"
using assms by blast
show "point x"
using assms by blast
show "?W \<le> D"
using assms choose_point_decreasing le_sup_iff point_def ss423bij inf.boundedE by blast
show "D\<^sup>\<star>*?q \<le> D\<^sup>T\<^sup>\<star>*x"
proof -
have "D\<^sup>+*q \<le> D\<^sup>T\<^sup>\<star>*x"
using assms conv_galois_2 order_trans star_1l by blast
thus ?thesis
by (metis choose_point_decreasing comp_assoc dual_order.trans mult_isol star_slide_var)
qed
show point_nq: "point ?q"
using assms by(rule path_inv_points(2))
show pathW: "path ?W"
proof(cases "W=0")
assume "W=0"
thus ?thesis
using assms edge_is_path point_is_point point_nq by simp
next
assume a: "W\<noteq>0"
have b: "?q*q\<^sup>T \<le> 1*?q*q\<^sup>T*-(?q*q\<^sup>T*1)"
proof -
have "?q*q\<^sup>T \<le>1" by simp
thus ?thesis
using assms point_nq
by(metis different_points_consequences(1) point_def sur_def_var1
acyclic_imp_one_step_different_points(2) choose_point_decreasing comp_assoc
is_vector_def point_def point_equations(3,4) point_is_point)
qed
have c: "W \<le> -(1*W)*W*1"
using assms terminating_path_iff by blast
have d: "(?q*q\<^sup>T)\<^sup>T*1 \<cdot> -((?q*q\<^sup>T)*1) = W*1 \<cdot> -(W\<^sup>T*1)"
using a
by (metis assms path_inv_points(2) acyclic_reachable_points choose_point_decreasing
edge_end point_is_point comp_assoc point_def sur_total total_one)
have e: "?q*q\<^sup>T*1 \<cdot> W\<^sup>T*1 = 0"
proof -
have "?q*q\<^sup>T*1 \<cdot> W\<^sup>T*1 = ?q \<cdot> W\<^sup>T*1"
using assms point_nq
by (metis comp_assoc conv_contrav conv_one is_vector_def point_def sur_def_var1)
also have "... \<le> -(?W\<^sup>T*1) \<cdot> ?W\<^sup>T*1"
using assms path_inv_choose_point_decrease
by (smt a conv_contrav conv_iso conv_one inf_mono less_eq_def subdistl_eq)
also have "... \<le> 0"
using compl_inf_bot eq_refl by blast
finally show ?thesis
using bot_unique by blast
qed
show ?thesis
using b c d e by (metis assms comp_assoc edge_is_path path_concatenation_cycle_free
point_is_point sup.commute point_nq)
qed
show "?W = 0 \<longleftrightarrow> ?q = y"
apply (rule iffI)
apply (metis assms conv_zero dist_alt edge_start inf_compl_bot_right modular_1_aux' modular_2_aux'
point_is_point sup.left_idem sup_bot_left point_nq)
by (smt assms end_points(1) conv_contrav conv_invol cycle_no_points(1) end_point_iff2 has_start_end_points_iff path_aux1b path_edge_equals_cycle point_is_point start_point_iff2 sup_bot_left top_greatest pathW)
show "?W\<noteq>0 \<longleftrightarrow> ?q = start_points ?W \<and> y = end_points ?W"
apply (rule iffI)
using assms end_points apply blast
using assms by force
show "terminating ?W"
by (smt assms end_points end_point_iff2 has_start_end_points_iff point_is_point start_point_iff2
terminating_iff1 pathW point_nq)
qed
theorem construct_path_partial: "VARS p q W
{ is_acyclic D \<and> point y \<and> point x \<and> D\<^sup>\<star>*y \<le> D\<^sup>T\<^sup>\<star>*x }
W := 0;
q := y;
WHILE q \<noteq> x
INV { construct_path_inv q x y D W }
DO p := choose_point (D*q);
W := W + p*q\<^sup>T;
q := p
OD
{ W \<le> D \<and> terminating_path W \<and> (W=0 \<longleftrightarrow> x=y) \<and> (W\<noteq>0 \<longleftrightarrow> x = start_points W \<and> y = end_points W) }"
apply vcg
using construct_path_pre apply blast
using construct_path_inv apply blast
by fastforce
end (* relation_algebra_rtc_tarski_choose_point *)
text \<open>For termination, we additionally need finiteness.\<close>
context finite
begin
lemma decrease_set:
assumes "\<forall>x::'a . Q x \<longrightarrow> P x"
and "P w"
and "\<not> Q w"
shows "card { x . Q x } < card { x . P x }"
by (metis Collect_mono assms card_seteq finite mem_Collect_eq not_le)
end
class relation_algebra_rtc_tarski_choose_point_finite =
relation_algebra_rtc_tarski_choose_point + relation_algebra_rtc_tarski_point_finite
begin
lemma decrease_variant:
assumes "y \<le> z"
and "w \<le> z"
and "\<not> w \<le> y"
shows "card { x . x \<le> y } < card { x . x \<le> z }"
by (metis Collect_mono assms card_seteq linorder_not_le dual_order.trans finite_code mem_Collect_eq)
lemma construct_path_inv_termination:
assumes "construct_path_inv q x y D W \<and> q \<noteq> x"
shows "card { z . z \<le> -(W + choose_point (D*q)*q\<^sup>T) } < card { z . z \<le> -W }"
proof -
let ?q = "choose_point (D*q)"
let ?W = "W + ?q * q\<^sup>T"
show ?thesis
proof (rule decrease_variant)
show "-?W \<le> -W"
by simp
show "?q * q\<^sup>T \<le> -W"
by (metis assms galois_aux inf_compl_bot_right maddux_142 mult_isor order_trans top_greatest)
show "\<not> (?q * q\<^sup>T \<le> -?W)"
using assms end_points(1)
by (smt acyclic_imp_one_step_different_points(2) choose_point_decreasing compl_sup inf.absorb1
inf_compl_bot_right sup.commute sup_bot.left_neutral conv_zero end_points(2))
qed
qed
theorem construct_path_total: "VARS p q W
[ is_acyclic D \<and> point y \<and> point x \<and> D\<^sup>\<star>*y \<le> D\<^sup>T\<^sup>\<star>*x ]
W := 0;
q := y;
WHILE q \<noteq> x
INV { construct_path_inv q x y D W }
VAR { card { z . z \<le> -W } }
DO p := choose_point (D*q);
W := W + p*q\<^sup>T;
q := p
OD
[ W \<le> D \<and> terminating_path W \<and> (W=0 \<longleftrightarrow> x=y) \<and> (W\<noteq>0 \<longleftrightarrow> x = start_points W \<and> y = end_points W) ]"
apply vcg_tc
using construct_path_pre apply blast
apply (rule CollectI, rule conjI)
using construct_path_inv apply blast
using construct_path_inv_termination apply clarsimp
by fastforce
end (* relation_algebra_rtc_tarski_choose_point_finite *)
subsection \<open>Topological sorting\<close>
text \<open>
In our second example we look at topological sorting.
Given a directed acyclic graph, the problem is to construct a linear order of its vertices that contains $x$ before $y$ for each edge $(x,y)$ of the graph.
If the input graph models dependencies between tasks, the output is a linear schedule of the tasks that respects all dependencies.
\<close>
context relation_algebra_rtc_tarski_choose_point
begin
abbreviation topological_sort_inv
where "topological_sort_inv q v R W \<equiv>
regressively_finite R \<and> R \<cdot> v*v\<^sup>T \<le> W\<^sup>+ \<and> terminating_path W \<and> W*1 = v\<cdot>-q \<and>
(W = 0 \<or> q = end_points W) \<and> point q \<and> R*v \<le> v \<and> q \<le> v \<and> is_vector v"
lemma topological_sort_pre:
assumes "regressively_finite R"
shows "topological_sort_inv (choose_point (minimum R 1)) (choose_point (minimum R 1)) R 0"
proof (intro conjI,simp_all add:assms)
let ?q = "choose_point (- (R\<^sup>T * 1))"
show point_q: "point ?q"
using assms by (metis (full_types) annir choose_point_point galois_aux2 is_inj_def is_sur_def
is_vector_def one_idem_mult point_def ss_p18 inf_top_left one_compl)
show "R \<cdot> ?q * ?q\<^sup>T \<le> 0"
- by (metis choose_point_decreasing conv_invol end_point_char eq_iff inf_bot_left schroeder_2)
+ by (metis choose_point_decreasing conv_invol end_point_char order.eq_iff inf_bot_left schroeder_2)
show "path 0"
by (simp add: is_inj_def is_p_fun_def path_def)
show "R*?q \<le> ?q"
by (metis choose_point_decreasing compl_bot_eq conv_galois_1 inf_compl_bot_left2 le_inf_iff)
show "is_vector ?q"
using point_q point_def by blast
qed
lemma topological_sort_inv:
assumes "v \<noteq> 1"
and "topological_sort_inv q v R W"
shows "topological_sort_inv (choose_point (minimum R (- v))) (v +
choose_point (minimum R (- v))) R (W + q * choose_point (minimum R (- v))\<^sup>T)"
proof (intro conjI)
let ?p = "choose_point (minimum R (-v))"
let ?W = "W + q*?p\<^sup>T"
let ?v = "v + ?p"
show point_p: "point ?p"
using assms
by (metis choose_point_point compl_bot_eq double_compl galois_aux2 comp_assoc is_vector_def
vector_compl vector_mult)
hence ep_np: "end_points (q*?p\<^sup>T) = ?p"
using assms(2)
by (metis aux4 choose_point_decreasing edge_end le_supI1 point_in_vector_or_complement_iff
point_is_point)
hence sp_q: "start_points (q*?p\<^sup>T) = q"
using assms(2) point_p
by (metis (no_types, lifting) conv_contrav conv_invol edge_start point_is_point)
hence ep_sp: "W \<noteq> 0 \<Longrightarrow> end_points W = start_points (q*?p\<^sup>T)"
using assms(2) by force
have "W*1 \<cdot> (q*?p\<^sup>T)\<^sup>T*1 = v\<cdot>-q\<cdot>?p"
using assms(2) point_p is_vector_def mult_assoc point_def point_equations(3) point_is_point
by auto
hence 1: "W*1 \<cdot> (q*?p\<^sup>T)\<^sup>T*1 = 0"
by (metis choose_point_decreasing dual_order.trans galois_aux inf.cobounded2 inf.commute)
show "regressively_finite R"
using assms(2) by blast
show "R \<cdot> ?v*?v\<^sup>T \<le> ?W\<^sup>+"
proof -
have a: "R \<cdot> v*v\<^sup>T \<le> ?W\<^sup>+"
using assms(2) by (meson mult_isol_var order.trans order_prop star_subdist)
have b: "R \<cdot> v*?p\<^sup>T \<le> ?W\<^sup>+"
proof -
have "R \<cdot> v*?p\<^sup>T \<le> W*1*?p\<^sup>T + q*?p\<^sup>T"
by (metis inf_le2 assms(2) aux4 double_compl inf_absorb2 distrib_right)
also have "... = W*?p\<^sup>T + q*?p\<^sup>T"
using point_p by (metis conv_contrav conv_one is_vector_def mult_assoc point_def)
also have "... \<le> W\<^sup>+*end_points W*?p\<^sup>T + q*?p\<^sup>T"
using assms(2)
by (meson forward_terminating_path_end_points_1 join_iso mult_isor terminating_path_iff)
also have "... \<le> W\<^sup>+*q*?p\<^sup>T + q*?p\<^sup>T"
using assms(2) by (metis annil eq_refl)
also have "... = W\<^sup>\<star>*q*?p\<^sup>T"
using conway.dagger_unfoldl_distr mult_assoc sup_commute by fastforce
also have "... \<le> ?W\<^sup>+"
by (metis mult_assoc mult_isol_var star_slide_var star_subdist sup_ge2)
finally show ?thesis .
qed
have c: "R \<cdot> ?p*v\<^sup>T \<le> ?W\<^sup>+"
proof -
have "v \<le> -?p"
using choose_point_decreasing compl_le_swap1 inf_le1 order_trans by blast
hence "R*v \<le> -?p"
using assms(2) order.trans by blast
thus ?thesis
by (metis galois_aux inf_le2 schroeder_2)
qed
have d: "R \<cdot> ?p*?p\<^sup>T \<le> ?W\<^sup>+"
proof -
have "R \<cdot> ?p*?p\<^sup>T \<le> R \<cdot> 1'"
using point_p is_inj_def meet_isor point_def by blast
also have "... = 0"
using assms(2) regressively_finite_irreflexive galois_aux by blast
finally show ?thesis
using bot_least inf.absorb_iff2 by simp
qed
have "R \<cdot> ?v*?v\<^sup>T = (R \<cdot> v*v\<^sup>T) + (R \<cdot> v*?p\<^sup>T) + (R \<cdot> ?p*v\<^sup>T) + (R \<cdot> ?p*?p\<^sup>T)"
by (metis conv_add distrib_left distrib_right inf_sup_distrib1 sup.commute sup.left_commute)
also have "... \<le> ?W\<^sup>+"
using a b c d by (simp add: le_sup_iff)
finally show ?thesis .
qed
show pathW: "path ?W"
proof (cases "W = 0")
assume "W = 0"
thus ?thesis
using assms(2) point_p edge_is_path point_is_point sup_bot_left by auto
next
assume a1: "W \<noteq> 0"
have fw_path: "forward_terminating_path W"
using assms(2) terminating_iff by blast
have bw_path: "backward_terminating_path (q*?p\<^sup>T)"
using assms point_p sp_q
by (metis conv_backward_terminating conv_has_start_points conv_path edge_is_path
forward_terminating_iff1 point_is_point start_point_iff2)
show ?thesis
using fw_path bw_path ep_sp 1 a1 path_concatenation_cycle_free by blast
qed
show "terminating ?W"
proof (rule start_end_implies_terminating)
show "has_start_points ?W"
apply (cases "W = 0")
using assms(2) sp_q pathW
apply (metis (no_types, lifting) point_is_point start_point_iff2 sup_bot.left_neutral)
using assms(2) ep_sp 1 pathW
by (metis has_start_end_points_iff path_concatenation_start_points start_point_iff2
terminating_iff1)
show "has_end_points ?W"
apply (cases "W = 0")
using point_p ep_np ep_sp pathW end_point_iff2 point_is_point apply force
using point_p ep_np ep_sp 1 pathW
by (metis end_point_iff2 path_concatenation_end_points point_is_point)
qed
show "?W*1 = ?v\<cdot>-?p"
proof -
have "?W*1 = v"
by (metis assms(2) point_p is_vector_def mult_assoc point_def point_equations(3)
point_is_point aux4 distrib_right' inf_absorb2 sup.commute)
also have "... = v\<cdot>-?p"
by (metis choose_point_decreasing compl_le_swap1 inf.cobounded1 inf.orderE order_trans)
finally show ?thesis
by (simp add: inf_sup_distrib2)
qed
show "?W = 0 \<or> ?p = end_points ?W"
using ep_np ep_sp 1 by (metis path_concatenation_end_points sup_bot_left)
show "R*?v \<le> ?v"
using assms(2)
by (meson choose_point_decreasing conv_galois_1 inf.cobounded2 order.trans sup.coboundedI1
sup_least)
show "?p \<le> ?v"
by simp
show "is_vector ?v"
using assms(2) point_p point_def vector_add by blast
qed
lemma topological_sort_post:
assumes "\<not> v \<noteq> 1"
and "topological_sort_inv q v R W"
shows "R \<le> W\<^sup>+ \<and> terminating_path W \<and> (W + W\<^sup>T)*1 = -1'*1"
proof (intro conjI,simp_all add:assms)
show "R \<le> W\<^sup>+"
using assms by force
show " backward_terminating W \<and> W \<le> 1 * W * (- v + q)"
using assms by force
show "v \<cdot> - q + W\<^sup>T * 1 = - 1' * 1"
proof (cases "W = 0")
assume "W = 0"
thus ?thesis
using assms
by (metis compl_bot_eq conv_one conv_zero double_compl inf_top.left_neutral is_inj_def
le_bot mult_1_right one_idem_mult point_def ss_p18 star_zero sup.absorb2 top_le)
next
assume a1: "W \<noteq> 0"
hence "-1' \<noteq> 0"
using assms backward_terminating_path_irreflexive le_bot by fastforce
hence "1 = 1*-1'*1"
by (simp add: tarski)
also have "... = -1'*1"
by (metis comp_assoc distrib_left mult_1_left sup_top_left distrib_right sup_compl_top)
finally have a: "1 = -1'*1" .
have "W*1 + W\<^sup>T*1 = 1"
using assms a1 by (metis double_compl galois_aux4 inf.absorb_iff2 inf_top.left_neutral)
thus ?thesis
using a by (simp add: assms(2))
qed
qed
theorem topological_sort_partial: "VARS p q v W
{ regressively_finite R }
W := 0;
q := choose_point (minimum R 1);
v := q;
WHILE v \<noteq> 1
INV { topological_sort_inv q v R W }
DO p := choose_point (minimum R (-v));
W := W + q*p\<^sup>T;
q := p;
v := v + p
OD
{ R \<le> W\<^sup>+ \<and> terminating_path W \<and> (W + W\<^sup>T)*1 = -1'*1 }"
apply vcg
using topological_sort_pre apply blast
using topological_sort_inv apply blast
using topological_sort_post by blast
end (* relation_algebra_rtc_tarski_choose_point *)
context relation_algebra_rtc_tarski_choose_point_finite
begin
lemma topological_sort_inv_termination:
assumes "v \<noteq> 1"
and "topological_sort_inv q v R W"
shows "card {z . z \<le> -(v + choose_point (minimum R (-v)))} < card { z . z \<le> -v }"
proof (rule decrease_variant)
let ?p = "choose_point (minimum R (-v))"
let ?v = "v + ?p"
show "-?v \<le> -v"
by simp
show "?p \<le> -v"
using choose_point_decreasing inf.boundedE by blast
have "point ?p"
using assms
by (metis choose_point_point compl_bot_eq double_compl galois_aux2 comp_assoc is_vector_def
vector_compl vector_mult)
thus "\<not> (?p \<le> -?v)"
by (metis annir compl_sup inf.absorb1 inf_compl_bot_right maddux_20 no_end_point_char)
qed
text \<open>
Use precondition \<open>is_acyclic\<close> instead of \<open>regressively_finite\<close>.
They are equivalent for finite graphs.
\<close>
theorem topological_sort_total: "VARS p q v W
[ is_acyclic R ]
W := 0;
q := choose_point (minimum R 1);
v := q;
WHILE v \<noteq> 1
INV { topological_sort_inv q v R W }
VAR { card { z . z \<le> -v } }
DO p := choose_point (minimum R (-v));
W := W + q*p\<^sup>T;
q := p;
v := v + p
OD
[ R \<le> W\<^sup>+ \<and> terminating_path W \<and> (W + W\<^sup>T)*1 = -1'*1 ]"
apply vcg_tc
apply (drule acyclic_regressively_finite)
using topological_sort_pre apply blast
apply (rule CollectI, rule conjI)
using topological_sort_inv apply blast
using topological_sort_inv_termination apply auto[1]
using topological_sort_post by blast
end (* relation_algebra_rtc_tarski_choose_point_finite *)
subsection \<open>Construction of a tree\<close>
text \<open>
Our last application is a correctness proof of an algorithm that constructs a non-empty cycle for a given directed graph.
This works in two steps.
The first step is to construct a directed tree from a given root along the edges of the graph.
\<close>
context relation_algebra_rtc_tarski_choose_point
begin
abbreviation construct_tree_pre
where "construct_tree_pre x y R \<equiv> y \<le> R\<^sup>T\<^sup>\<star>*x \<and> point x"
abbreviation construct_tree_inv
where "construct_tree_inv v x y D R \<equiv> construct_tree_pre x y R \<and> is_acyclic D \<and> is_inj D \<and>
D \<le> R \<and> D*x = 0 \<and> v = x + D\<^sup>T*1 \<and> x*v\<^sup>T \<le> D\<^sup>\<star> \<and> D \<le> v*v\<^sup>T \<and>
is_vector v"
abbreviation construct_tree_post
where "construct_tree_post x y D R \<equiv> is_acyclic D \<and> is_inj D \<and> D \<le> R \<and> D*x = 0 \<and> D\<^sup>T*1 \<le> D\<^sup>T\<^sup>\<star>*x \<and>
D\<^sup>\<star>*y \<le> D\<^sup>T\<^sup>\<star>*x"
lemma construct_tree_pre:
assumes "construct_tree_pre x y R"
shows "construct_tree_inv x x y 0 R"
using assms by (simp add: is_inj_def point_def)
lemma construct_tree_inv_aux:
assumes "\<not> y \<le> v"
and "construct_tree_inv v x y D R"
shows "singleton (choose_singleton (v*-v\<^sup>T \<cdot> R))"
proof (rule choose_singleton_singleton, rule notI)
assume "v*-v\<^sup>T \<cdot> R = 0"
hence "R\<^sup>T\<^sup>\<star>*v \<le> v"
by (metis galois_aux conv_compl conv_galois_1 conv_galois_2 conv_invol double_compl
star_inductl_var)
hence "y = 0"
using assms by (meson mult_isol order_trans sup.cobounded1)
thus False
using assms point_is_point by auto
qed
lemma construct_tree_inv:
assumes "\<not> y \<le> v"
and "construct_tree_inv v x y D R"
shows "construct_tree_inv (v + choose_singleton (v*-v\<^sup>T \<cdot> R)\<^sup>T*1) x y (D +
choose_singleton (v*-v\<^sup>T \<cdot> R)) R"
proof (intro conjI)
let ?e = "choose_singleton (v*-v\<^sup>T \<cdot> R)"
let ?D = "D + ?e"
let ?v = "v + ?e\<^sup>T*1"
have 1: "?e \<le> v*-v\<^sup>T"
using choose_singleton_decreasing inf.boundedE by blast
show "point x"
by (simp add: assms)
show "y \<le> R\<^sup>T\<^sup>\<star>*x"
by (simp add: assms)
show "is_acyclic ?D"
using 1 assms acyclic_inv by fastforce
show "is_inj ?D"
using 1 construct_tree_inv_aux assms injective_inv by blast
show "?D \<le> R"
apply (rule sup.boundedI)
using assms apply blast
using choose_singleton_decreasing inf.boundedE by blast
show "?D*x = 0"
proof -
have "?D*x = ?e*x"
by (simp add: assms)
also have "... \<le> ?e*v"
by (simp add: assms mult_isol)
also have "... \<le> v*-v\<^sup>T*v"
using 1 mult_isor by blast
also have "... = 0"
by (metis assms(2) annir comp_assoc vector_prop1)
finally show ?thesis
using le_bot by blast
qed
show "?v = x + ?D\<^sup>T*1"
by (simp add: assms sup_assoc)
show "x*?v\<^sup>T \<le> ?D\<^sup>\<star>"
proof -
have "x*?v\<^sup>T = x*v\<^sup>T + x*1*?e"
by (simp add: distrib_left mult_assoc)
also have "... \<le> D\<^sup>\<star> + x*1*(?e \<cdot> v*-v\<^sup>T)"
using 1 by (metis assms(2) inf.absorb1 join_iso)
also have "... = D\<^sup>\<star> + x*1*(?e \<cdot> v \<cdot> -v\<^sup>T)"
by (metis assms(2) comp_assoc conv_compl inf.assoc vector_compl vector_meet_comp)
also have "... \<le> D\<^sup>\<star> + x*1*(?e \<cdot> v)"
using join_isol mult_subdistl by fastforce
also have "... = D\<^sup>\<star> + x*(1 \<cdot> v\<^sup>T)*?e"
by (metis assms(2) inf.commute mult_assoc vector_2)
also have "... = D\<^sup>\<star> + x*v\<^sup>T*?e"
by simp
also have "... \<le> D\<^sup>\<star> + D\<^sup>\<star>*?e"
using assms join_isol mult_isor by blast
also have "... \<le> ?D\<^sup>\<star>"
by (meson le_sup_iff prod_star_closure star_ext star_subdist)
finally show ?thesis .
qed
show "?D \<le> ?v*?v\<^sup>T"
proof (rule sup.boundedI)
show "D \<le> ?v*?v\<^sup>T"
using assms
by (meson conv_add distrib_left le_supI1 conv_iso dual_order.trans mult_isol_var order_prop)
have "?e \<le> v*(-v\<^sup>T \<cdot> v\<^sup>T*?e)"
using 1 inf.absorb_iff2 modular_1' by fastforce
also have "... \<le> v*1*?e"
by (simp add: comp_assoc le_infI2 mult_isol_var)
also have "... \<le> ?v*?v\<^sup>T"
by (metis conv_contrav conv_invol conv_iso conv_one mult_assoc mult_isol_var sup.cobounded1
sup_ge2)
finally show "?e \<le> ?v*?v\<^sup>T"
by simp
qed
show "is_vector ?v"
using assms comp_assoc is_vector_def by fastforce
qed
lemma construct_tree_post:
assumes "y \<le> v"
and "construct_tree_inv v x y D R"
shows "construct_tree_post x y D R"
proof -
have "v*x\<^sup>T \<le> D\<^sup>T\<^sup>\<star>"
by (metis (no_types, lifting) assms(2) conv_contrav conv_invol conv_iso star_conv)
hence 1: "v \<le> D\<^sup>T\<^sup>\<star>*x"
using assms point_def ss423bij by blast
hence 2: "D\<^sup>T*1 \<le> D\<^sup>T\<^sup>\<star>*x"
using assms le_supE by blast
have "D\<^sup>\<star>*y \<le> D\<^sup>T\<^sup>\<star>*x"
proof (rule star_inductl, rule sup.boundedI)
show "y \<le> D\<^sup>T\<^sup>\<star>*x"
using 1 assms order.trans by blast
next
have "D*(D\<^sup>T\<^sup>\<star>*x) = D*x + D*D\<^sup>T\<^sup>+*x"
by (metis conway.dagger_unfoldl_distr distrib_left mult_assoc)
also have "... = D*D\<^sup>T\<^sup>+*x"
using assms by simp
also have "... \<le> 1'*D\<^sup>T\<^sup>\<star>*x"
by (metis assms(2) is_inj_def mult_assoc mult_isor)
finally show "D*(D\<^sup>T\<^sup>\<star>*x) \<le> D\<^sup>T\<^sup>\<star>*x"
by simp
qed
thus "construct_tree_post x y D R"
using 2 assms by simp
qed
theorem construct_tree_partial: "VARS e v D
{ construct_tree_pre x y R }
D := 0;
v := x;
WHILE \<not> y \<le> v
INV { construct_tree_inv v x y D R }
DO e := choose_singleton (v*-v\<^sup>T \<cdot> R);
D := D + e;
v := v + e\<^sup>T*1
OD
{ construct_tree_post x y D R }"
apply vcg
using construct_tree_pre apply blast
using construct_tree_inv apply blast
using construct_tree_post by blast
end (* relation_algebra_rtc_tarski_choose_point *)
context relation_algebra_rtc_tarski_choose_point_finite
begin
lemma construct_tree_inv_termination:
assumes " \<not> y \<le> v"
and "construct_tree_inv v x y D R"
shows "card { z . z \<le> -(v + choose_singleton (v*-v\<^sup>T \<cdot> R)\<^sup>T*1) } < card { z . z \<le> -v }"
proof (rule decrease_variant)
let ?e = "choose_singleton (v*-v\<^sup>T \<cdot> R)"
let ?v = "v + ?e\<^sup>T*1"
have 1: "?e \<le> v*-v\<^sup>T"
using choose_singleton_decreasing inf.boundedE by blast
have 2: "singleton ?e"
using construct_tree_inv_aux assms by auto
show "-?v \<le> -v"
by simp
have "?e\<^sup>T \<le> -v*v\<^sup>T"
using 1 conv_compl conv_iso by force
also have "... \<le> -v*1"
by (simp add: mult_isol)
finally show "?e\<^sup>T*1 \<le> -v"
using assms by (metis is_vector_def mult_isor one_compl)
thus "\<not> (?e\<^sup>T*1 \<le> -?v)"
using 2 by (metis annir compl_sup inf.absorb1 inf_compl_bot_right surj_one tarski)
qed
theorem construct_tree_total: "VARS e v D
[ construct_tree_pre x y R ]
D := 0;
v := x;
WHILE \<not> y \<le> v
INV { construct_tree_inv v x y D R }
VAR { card { z . z \<le> -v } }
DO e := choose_singleton (v*-v\<^sup>T \<cdot> R);
D := D + e;
v := v + e\<^sup>T*1
OD
[ construct_tree_post x y D R ]"
apply vcg_tc
using construct_tree_pre apply blast
apply (rule CollectI, rule conjI)
using construct_tree_inv apply blast
using construct_tree_inv_termination apply force
using construct_tree_post by blast
end (* relation_algebra_rtc_tarski_choose_point_finite *)
subsection \<open>Construction of a non-empty cycle\<close>
text \<open>
The second step is to construct a path from the root to a given vertex in the tree.
Adding an edge back to the root gives the cycle.
\<close>
context relation_algebra_rtc_tarski_choose_point
begin
abbreviation comment
where "comment _ \<equiv> SKIP" (* instead of inner comments *)
abbreviation construct_cycle_inv
where "construct_cycle_inv v x y D R \<equiv> construct_tree_inv v x y D R \<and> point y \<and> y*x\<^sup>T \<le> R"
lemma construct_cycle_pre:
assumes " \<not> is_acyclic R"
and "y = choose_point ((R\<^sup>+ \<cdot> 1')*1)"
and "x = choose_point (R\<^sup>\<star>*y \<cdot> R\<^sup>T*y)"
shows "construct_cycle_inv x x y 0 R"
proof(rule conjI, rule_tac [2] conjI)
show point_y: "point y"
using assms by (simp add: choose_point_point is_vector_def mult_assoc galois_aux ss_p18)
have "R\<^sup>\<star>*y \<cdot> R\<^sup>T*y \<noteq> 0"
proof
have "R\<^sup>+ \<cdot> 1' = (R\<^sup>+)\<^sup>T \<cdot> 1'"
by (metis (mono_tags, hide_lams) conv_e conv_times inf.cobounded1 inf.commute
many_strongly_connected_iff_6_eq mult_oner star_subid)
also have "... = R\<^sup>T\<^sup>+ \<cdot> 1'"
using plus_conv by fastforce
also have "... \<le> (R\<^sup>T\<^sup>\<star> \<cdot> R)*R\<^sup>T"
by (metis conv_contrav conv_e conv_invol modular_2_var mult_oner star_slide_var)
also have "... \<le> (R\<^sup>T\<^sup>\<star> \<cdot> R)*1"
by (simp add: mult_isol)
finally have a: "(R\<^sup>+ \<cdot> 1')*1 \<le> (R\<^sup>T\<^sup>\<star> \<cdot> R)*1"
by (metis mult_assoc mult_isor one_idem_mult)
assume "R\<^sup>\<star>*y \<cdot> R\<^sup>T*y = 0"
hence "(R\<^sup>\<star> \<cdot> R\<^sup>T)*y = 0"
using point_y inj_distr point_def by blast
hence "(R\<^sup>\<star> \<cdot> R\<^sup>T)\<^sup>T*1 \<le> -y"
by (simp add: conv_galois_1)
hence "y \<le> -((R\<^sup>\<star> \<cdot> R\<^sup>T)\<^sup>T*1)"
using compl_le_swap1 by blast
also have "... = -((R\<^sup>T\<^sup>\<star> \<cdot> R)*1)"
by (simp add: star_conv)
also have "... \<le> -((R\<^sup>+ \<cdot> 1')*1)"
using a comp_anti by blast
also have "... \<le> -y"
by (simp add: assms galois_aux ss_p18 choose_point_decreasing)
finally have "y = 0"
using inf.absorb2 by fastforce
thus False
using point_y annir point_equations(2) point_is_point tarski by force
qed
hence point_x: "point x"
by (metis point_y assms(3) inj_distr is_vector_def mult_assoc point_def choose_point_point)
hence "y \<le> R\<^sup>T\<^sup>\<star> * x"
by (metis assms(3) point_y choose_point_decreasing inf_le1 order.trans point_swap star_conv)
thus tree_inv: "construct_tree_inv x x y 0 R"
using point_x construct_tree_pre by blast
show "y * x\<^sup>T \<le> R"
proof -
have "x \<le> R\<^sup>\<star>*y \<cdot> R\<^sup>T*y"
using assms(3) choose_point_decreasing by blast
also have "... = (R\<^sup>\<star> \<cdot> R\<^sup>T)*y"
using point_y inj_distr point_def by fastforce
finally have "x*y\<^sup>T \<le> R\<^sup>\<star> \<cdot> R\<^sup>T"
using point_y point_def ss423bij by blast
also have "... \<le> R\<^sup>T"
by simp
finally show ?thesis
using conv_iso by force
qed
qed
lemma construct_cycle_pre2:
assumes "y \<le> v"
and "construct_cycle_inv v x y D R"
shows "construct_path_inv y x y D 0 \<and> D \<le> R \<and> D * x = 0 \<and> y * x\<^sup>T \<le> R"
proof(intro conjI, simp_all add: assms)
show "D\<^sup>\<star> * y \<le> D\<^sup>T\<^sup>\<star> * x"
using assms construct_tree_post by blast
show "path 0"
by (simp add: is_inj_def is_p_fun_def path_def)
show "y \<noteq> 0"
using assms(2) is_point_def point_is_point by blast
qed
lemma construct_cycle_post:
assumes "\<not> q \<noteq> x"
and "(construct_path_inv q x y D W \<and> D \<le> R \<and> D * x = 0 \<and> y * x\<^sup>T \<le> R)"
shows "W + y * x\<^sup>T \<noteq> 0 \<and> W + y * x\<^sup>T \<le> R \<and> cycle (W + y * x\<^sup>T)"
proof(intro conjI)
let ?C = "W + y*x\<^sup>T"
show "?C \<noteq> 0"
by (metis assms acyclic_imp_one_step_different_points(2) no_trivial_inverse point_def ss423bij
sup_bot.monoid_axioms monoid.left_neutral)
show "?C \<le> R"
using assms(2) order_trans sup.boundedI by blast
show "path (W + y * x\<^sup>T)"
by (metis assms construct_tree_pre edge_is_path less_eq_def path_edge_equals_cycle
point_is_point terminating_iff1)
show "many_strongly_connected (W + y * x\<^sup>T)"
by (metis assms construct_tree_pre bot_least conv_zero less_eq_def
path_edge_equals_cycle star_conv star_subid terminating_iff1)
qed
theorem construct_cycle_partial: "VARS e p q v x y C D W
{ \<not> is_acyclic R }
y := choose_point ((R\<^sup>+ \<cdot> 1')*1);
x := choose_point (R\<^sup>\<star>*y \<cdot> R\<^sup>T*y);
D := 0;
v := x;
WHILE \<not> y \<le> v
INV { construct_cycle_inv v x y D R }
DO e := choose_singleton (v*-v\<^sup>T \<cdot> R);
D := D + e;
v := v + e\<^sup>T*1
OD;
comment { is_acyclic D \<and> point y \<and> point x \<and> D\<^sup>\<star>*y \<le> D\<^sup>T\<^sup>\<star>*x };
W := 0;
q := y;
WHILE q \<noteq> x
INV { construct_path_inv q x y D W \<and> D \<le> R \<and> D*x = 0 \<and> y*x\<^sup>T \<le> R }
DO p := choose_point (D*q);
W := W + p*q\<^sup>T;
q := p
OD;
comment { W \<le> D \<and> terminating_path W \<and> (W = 0 \<longleftrightarrow> q=y) \<and> (W \<noteq> 0 \<longleftrightarrow> q = start_points W \<and> y = end_points W) };
C := W + y*x\<^sup>T
{ C \<noteq> 0 \<and> C \<le> R \<and> cycle C }"
apply vcg
using construct_cycle_pre apply blast
using construct_tree_inv apply blast
using construct_cycle_pre2 apply blast
using construct_path_inv apply blast
using construct_cycle_post by blast
end (* relation_algebra_rtc_tarski_choose_point *)
context relation_algebra_rtc_tarski_choose_point_finite
begin
theorem construct_cycle_total: "VARS e p q v x y C D W
[ \<not> is_acyclic R ]
y := choose_point ((R\<^sup>+ \<cdot> 1')*1);
x := choose_point (R\<^sup>\<star>*y \<cdot> R\<^sup>T*y);
D := 0;
v := x;
WHILE \<not> y \<le> v
INV { construct_cycle_inv v x y D R }
VAR { card { z . z \<le> -v } }
DO e := choose_singleton (v*-v\<^sup>T \<cdot> R);
D := D + e;
v := v + e\<^sup>T*1
OD;
comment { is_acyclic D \<and> point y \<and> point x \<and> D\<^sup>\<star>*y \<le> D\<^sup>T\<^sup>\<star>*x };
W := 0;
q := y;
WHILE q \<noteq> x
INV { construct_path_inv q x y D W \<and> D \<le> R \<and> D*x = 0 \<and> y*x\<^sup>T \<le> R }
VAR { card { z . z \<le> -W } }
DO p := choose_point (D*q);
W := W + p*q\<^sup>T;
q := p
OD;
comment { W \<le> D \<and> terminating_path W \<and> (W = 0 \<longleftrightarrow> q=y) \<and> (W \<noteq> 0 \<longleftrightarrow> q = start_points W \<and> y = end_points W)};
C := W + y*x\<^sup>T
[ C \<noteq> 0 \<and> C \<le> R \<and> cycle C ]"
apply vcg_tc
using construct_cycle_pre apply blast
apply (rule CollectI, rule conjI)
using construct_tree_inv apply blast
using construct_tree_inv_termination apply force
using construct_cycle_pre2 apply blast
apply (rule CollectI, rule conjI)
using construct_path_inv apply blast
using construct_path_inv_termination apply clarsimp
using construct_cycle_post by blast
end (* relation_algebra_rtc_tarski_choose_point_finite *)
end
diff --git a/thys/Relational_Paths/Paths.thy b/thys/Relational_Paths/Paths.thy
--- a/thys/Relational_Paths/Paths.thy
+++ b/thys/Relational_Paths/Paths.thy
@@ -1,2307 +1,2307 @@
(* Title: Relational Characterisation of Paths
Author: Walter Guttmann, Peter Hoefner
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
Peter Hoefner <peter at hoefner-online.de>
*)
section \<open>Relational Characterisation of Paths\<close>
text \<open>
This theory provides the relation-algebraic characterisations of paths, as defined in Sections 3--5 of \cite{BerghammerFurusawaGuttmannHoefner2020}.
\<close>
theory Paths
imports More_Relation_Algebra
begin
context relation_algebra_tarski
begin
lemma path_concat_aux_0:
assumes "is_vector v"
and "v \<noteq> 0"
and "w;v\<^sup>T \<le> x"
and "v;z \<le> y"
shows "w;1;z \<le> x;y"
proof -
from tarski assms(1,2) have "1 = 1;v\<^sup>T;v;1"
by (metis conv_contrav conv_one eq_refl inf_absorb1 inf_top_left is_vector_def ra_2)
hence "w;1;z = w;1;v\<^sup>T;v;1;z"
by (simp add: mult_isor mult_isol mult_assoc)
also from assms(1) have "... = w;v\<^sup>T;v;z"
by (metis is_vector_def comp_assoc conv_contrav conv_one)
also from assms(3) have "... \<le> x;v;z"
by (simp add: mult_isor)
also from assms(4) have "... \<le> x;y"
by (simp add: mult_isol mult_assoc)
finally show ?thesis .
qed
end (* context relation_algebra_tarski *)
subsection \<open>Consequences without the Tarski rule\<close>
context relation_algebra_rtc
begin
text \<open>Definitions for path classifications\<close>
abbreviation connected
where "connected x \<equiv> x;1;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
abbreviation many_strongly_connected
where "many_strongly_connected x \<equiv> x\<^sup>\<star> = x\<^sup>T\<^sup>\<star>"
abbreviation one_strongly_connected
where "one_strongly_connected x \<equiv> x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>\<star>"
definition path
where "path x \<equiv> connected x \<and> is_p_fun x \<and> is_inj x"
abbreviation cycle
where "cycle x \<equiv> path x \<and> many_strongly_connected x"
abbreviation start_points
where "start_points x \<equiv> x;1 \<cdot> -(x\<^sup>T;1)"
abbreviation end_points
where "end_points x \<equiv> x\<^sup>T;1 \<cdot> -(x;1)"
abbreviation no_start_points
where "no_start_points x \<equiv> x;1 \<le> x\<^sup>T;1"
abbreviation no_end_points
where "no_end_points x \<equiv> x\<^sup>T;1 \<le> x;1"
abbreviation no_start_end_points
where "no_start_end_points x \<equiv> x;1 = x\<^sup>T;1"
abbreviation has_start_points
where "has_start_points x \<equiv> 1 = -(1;x);x;1"
abbreviation has_end_points
where "has_end_points x \<equiv> 1 = 1;x;-(x;1)"
abbreviation has_start_end_points
where "has_start_end_points x \<equiv> 1 = -(1;x);x;1 \<cdot> 1;x;-(x;1)"
abbreviation backward_terminating
where "backward_terminating x \<equiv> x \<le> -(1;x);x;1"
abbreviation forward_terminating
where "forward_terminating x \<equiv> x \<le> 1;x;-(x;1)"
abbreviation terminating
where "terminating x \<equiv> x \<le> -(1;x);x;1 \<cdot> 1;x;-(x;1)"
abbreviation backward_finite
where "backward_finite x \<equiv> x \<le> x\<^sup>T\<^sup>\<star> + -(1;x);x;1"
abbreviation forward_finite
where "forward_finite x \<equiv> x \<le> x\<^sup>T\<^sup>\<star> + 1;x;-(x;1)"
abbreviation finite
where "finite x \<equiv> x \<le> x\<^sup>T\<^sup>\<star> + (-(1;x);x;1 \<cdot> 1;x;-(x;1))"
abbreviation no_start_points_path
where "no_start_points_path x \<equiv> path x \<and> no_start_points x"
abbreviation no_end_points_path
where "no_end_points_path x \<equiv> path x \<and> no_end_points x"
abbreviation no_start_end_points_path
where "no_start_end_points_path x \<equiv> path x \<and> no_start_end_points x"
abbreviation has_start_points_path
where "has_start_points_path x \<equiv> path x \<and> has_start_points x"
abbreviation has_end_points_path
where "has_end_points_path x \<equiv> path x \<and> has_end_points x"
abbreviation has_start_end_points_path
where "has_start_end_points_path x \<equiv> path x \<and> has_start_end_points x"
abbreviation backward_terminating_path
where "backward_terminating_path x \<equiv> path x \<and> backward_terminating x"
abbreviation forward_terminating_path
where "forward_terminating_path x \<equiv> path x \<and> forward_terminating x"
abbreviation terminating_path
where "terminating_path x \<equiv> path x \<and> terminating x"
abbreviation backward_finite_path
where "backward_finite_path x \<equiv> path x \<and> backward_finite x"
abbreviation forward_finite_path
where "forward_finite_path x \<equiv> path x \<and> forward_finite x"
abbreviation finite_path
where "finite_path x \<equiv> path x \<and> finite x"
text \<open>General properties\<close>
lemma reachability_from_z_in_y:
assumes "x \<le> y\<^sup>\<star>;z"
and "x \<cdot> z = 0"
shows "x \<le> y\<^sup>+;z"
by (metis assms conway.dagger_unfoldl_distr galois_1 galois_aux inf.orderE)
lemma reachable_imp:
assumes "point p"
and "point q"
and "p\<^sup>\<star>;q \<le> p\<^sup>T\<^sup>\<star>;p"
shows "p \<le> p\<^sup>\<star>;q"
by (metis assms conway.dagger_unfoldr_distr le_supE point_swap star_conv)
text \<open>Basic equivalences\<close>
lemma no_start_end_points_iff:
"no_start_end_points x \<longleftrightarrow> no_start_points x \<and> no_end_points x"
by fastforce
lemma has_start_end_points_iff:
"has_start_end_points x \<longleftrightarrow> has_start_points x \<and> has_end_points x"
by (metis inf_eq_top_iff)
lemma terminating_iff:
"terminating x \<longleftrightarrow> backward_terminating x \<and> forward_terminating x"
by simp
lemma finite_iff:
"finite x \<longleftrightarrow> backward_finite x \<and> forward_finite x"
by (simp add: sup_inf_distrib1 inf.boundedI)
lemma no_start_end_points_path_iff:
"no_start_end_points_path x \<longleftrightarrow> no_start_points_path x \<and> no_end_points_path x"
by fastforce
lemma has_start_end_points_path_iff:
"has_start_end_points_path x \<longleftrightarrow> has_start_points_path x \<and> has_end_points_path x"
using has_start_end_points_iff by blast
lemma terminating_path_iff:
"terminating_path x \<longleftrightarrow> backward_terminating_path x \<and> forward_terminating_path x"
by fastforce
lemma finite_path_iff:
"finite_path x \<longleftrightarrow> backward_finite_path x \<and> forward_finite_path x"
using finite_iff by fastforce
text \<open>Closure under converse\<close>
lemma connected_conv:
"connected x \<longleftrightarrow> connected (x\<^sup>T)"
by (metis comp_assoc conv_add conv_contrav conv_iso conv_one star_conv)
lemma conv_many_strongly_connected:
"many_strongly_connected x \<longleftrightarrow> many_strongly_connected (x\<^sup>T)"
by fastforce
lemma conv_one_strongly_connected:
"one_strongly_connected x \<longleftrightarrow> one_strongly_connected (x\<^sup>T)"
by (metis comp_assoc conv_contrav conv_iso conv_one star_conv)
lemma conv_path:
"path x \<longleftrightarrow> path (x\<^sup>T)"
using connected_conv inj_p_fun path_def by fastforce
lemma conv_cycle:
"cycle x \<longleftrightarrow> cycle (x\<^sup>T)"
using conv_path by fastforce
lemma conv_no_start_points:
"no_start_points x \<longleftrightarrow> no_end_points (x\<^sup>T)"
by simp
lemma conv_no_start_end_points:
"no_start_end_points x \<longleftrightarrow> no_start_end_points (x\<^sup>T)"
by fastforce
lemma conv_has_start_points:
"has_start_points x \<longleftrightarrow> has_end_points (x\<^sup>T)"
by (metis comp_assoc conv_compl conv_contrav conv_invol conv_one)
lemma conv_has_start_end_points:
"has_start_end_points x \<longleftrightarrow> has_start_end_points (x\<^sup>T)"
by (metis comp_assoc conv_compl conv_contrav conv_invol conv_one inf_eq_top_iff)
lemma conv_backward_terminating:
"backward_terminating x \<longleftrightarrow> forward_terminating (x\<^sup>T)"
by (metis comp_assoc conv_compl conv_contrav conv_iso conv_one)
lemma conv_terminating:
"terminating x \<longleftrightarrow> terminating (x\<^sup>T)"
apply (rule iffI)
apply (metis conv_compl conv_contrav conv_one conv_times inf.commute le_iff_inf mult_assoc)
by (metis conv_compl conv_contrav conv_invol conv_one conv_times inf.commute le_iff_inf mult_assoc)
lemma conv_backward_finite:
"backward_finite x \<longleftrightarrow> forward_finite (x\<^sup>T)"
by (metis comp_assoc conv_add conv_compl conv_contrav conv_iso conv_one star_conv)
lemma conv_finite:
"finite x \<longleftrightarrow> finite (x\<^sup>T)"
by (metis finite_iff conv_backward_finite conv_invol)
lemma conv_no_start_points_path:
"no_start_points_path x \<longleftrightarrow> no_end_points_path (x\<^sup>T)"
using conv_path by fastforce
lemma conv_no_start_end_points_path:
"no_start_end_points_path x \<longleftrightarrow> no_start_end_points_path (x\<^sup>T)"
using conv_path by fastforce
lemma conv_has_start_points_path:
"has_start_points_path x \<longleftrightarrow> has_end_points_path (x\<^sup>T)"
using conv_has_start_points conv_path by fastforce
lemma conv_has_start_end_points_path:
"has_start_end_points_path x \<longleftrightarrow> has_start_end_points_path (x\<^sup>T)"
using conv_has_start_end_points conv_path by fastforce
lemma conv_backward_terminating_path:
"backward_terminating_path x \<longleftrightarrow> forward_terminating_path (x\<^sup>T)"
using conv_backward_terminating conv_path by fastforce
lemma conv_terminating_path:
"terminating_path x \<longleftrightarrow> terminating_path (x\<^sup>T)"
using conv_path conv_terminating by fastforce
lemma conv_backward_finite_path:
"backward_finite_path x \<longleftrightarrow> forward_finite_path (x\<^sup>T)"
using conv_backward_finite conv_path by fastforce
lemma conv_finite_path:
"finite_path x \<longleftrightarrow> finite_path (x\<^sup>T)"
using conv_finite conv_path by blast
text \<open>Equivalences for \<open>connected\<close>\<close>
lemma connected_iff2:
assumes "is_inj x"
and "is_p_fun x"
shows "connected x \<longleftrightarrow> x;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
proof
assume 1: "connected x"
have "x;1;x\<^sup>T \<le> x;1;x;x\<^sup>T"
by (metis conv_invol modular_var_3 vector_meet_comp_x')
also have "... \<le> (x\<^sup>+ + x\<^sup>T\<^sup>\<star>);x\<^sup>T"
using 1 mult_isor star_star_plus by fastforce
also have "... \<le> x\<^sup>\<star>;x;x\<^sup>T + x\<^sup>T\<^sup>\<star>"
using join_isol star_slide_var by simp
also from assms(1) have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis is_inj_def comp_assoc join_iso mult_1_right mult_isol)
finally show "x;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>" .
next
assume 2: "x;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
have "x;1;x \<le> x;1;x\<^sup>T;x"
by (simp add: modular_var_3 vector_meet_comp_x)
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>+);x"
using 2 by (metis mult_isor star_star_plus sup_commute)
also have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>;x\<^sup>T;x"
using join_iso star_slide_var by simp
also from assms(2) have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis comp_assoc is_p_fun_def join_isol mult_1_right mult_isol)
finally show "connected x" .
qed
lemma connected_iff3:
assumes "is_inj x"
and "is_p_fun x"
shows "connected x \<longleftrightarrow> x\<^sup>T;1;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis assms connected_conv connected_iff2 inj_p_fun p_fun_inj conv_invol add_commute)
lemma connected_iff4:
"connected x \<longleftrightarrow> x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis connected_conv conv_invol add_commute)
lemma connected_iff5:
"connected x \<longleftrightarrow> x\<^sup>+;1;x\<^sup>+ \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using comp_assoc plus_top top_plus by fastforce
lemma connected_iff6:
assumes "is_inj x"
and "is_p_fun x"
shows "connected x \<longleftrightarrow> x\<^sup>+;1;(x\<^sup>+)\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using assms connected_iff2 comp_assoc plus_conv plus_top top_plus by fastforce
lemma connected_iff7:
assumes "is_inj x"
and "is_p_fun x"
shows "connected x \<longleftrightarrow> (x\<^sup>+)\<^sup>T;1;x\<^sup>+ \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis assms connected_iff3 conv_contrav conv_invol conv_one top_plus vector_meet_comp_x)
lemma connected_iff8:
"connected x \<longleftrightarrow> (x\<^sup>+)\<^sup>T;1;(x\<^sup>+)\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
by (metis connected_iff4 comp_assoc conv_contrav conv_invol conv_one plus_conv star_conv top_plus)
text \<open>Equivalences and implications for \<open>many_strongly_connected\<close>\<close>
lemma many_strongly_connected_iff_1:
"many_strongly_connected x \<longleftrightarrow> x\<^sup>T \<le> x\<^sup>\<star>"
apply (rule iffI,simp)
-by (metis conv_invol conv_iso eq_iff star_conv star_invol star_iso)
+by (metis conv_invol conv_iso order.eq_iff star_conv star_invol star_iso)
lemma many_strongly_connected_iff_2:
"many_strongly_connected x \<longleftrightarrow> x\<^sup>T \<le> x\<^sup>+"
proof
assume as: "many_strongly_connected x"
hence "x\<^sup>T \<le> x\<^sup>\<star> \<cdot> (-(1') + x)"
by (metis many_strongly_connected_iff_1 loop_backward_forward inf_greatest)
also have "... \<le> (x\<^sup>\<star> \<cdot> -(1')) + (x\<^sup>\<star> \<cdot> x)"
by (simp add: inf_sup_distrib1)
also have "... \<le> x\<^sup>+"
- by (metis as eq_iff mult_1_right mult_isol star_ref sup.absorb1 conv_invol eq_refl galois_1
+ by (metis as order.eq_iff mult_1_right mult_isol star_ref sup.absorb1 conv_invol eq_refl galois_1
inf.absorb_iff1 inf.commute star_unfoldl_eq sup_mono many_strongly_connected_iff_1)
finally show "x\<^sup>T \<le> x\<^sup>+" .
next
show "x\<^sup>T \<le> x\<^sup>+ \<Longrightarrow> many_strongly_connected x"
using order_trans star_1l many_strongly_connected_iff_1 by blast
qed
lemma many_strongly_connected_iff_3:
"many_strongly_connected x \<longleftrightarrow> x \<le> x\<^sup>T\<^sup>\<star>"
by (metis conv_invol many_strongly_connected_iff_1)
lemma many_strongly_connected_iff_4:
"many_strongly_connected x \<longleftrightarrow> x \<le> x\<^sup>T\<^sup>+"
by (metis conv_invol many_strongly_connected_iff_2)
lemma many_strongly_connected_iff_5:
"many_strongly_connected x \<longleftrightarrow> x\<^sup>\<star>;x\<^sup>T \<le> x\<^sup>+"
by (metis comp_assoc conv_contrav conway.dagger_unfoldr_distr star_conv star_denest_var_2
star_invol star_trans_eq star_unfoldl_eq sup.boundedE many_strongly_connected_iff_2)
lemma many_strongly_connected_iff_6:
"many_strongly_connected x \<longleftrightarrow> x\<^sup>T;x\<^sup>\<star> \<le> x\<^sup>+"
by (metis dual_order.trans star_1l star_conv star_inductl_star star_invol star_slide_var
many_strongly_connected_iff_1 many_strongly_connected_iff_5)
lemma many_strongly_connected_iff_7:
"many_strongly_connected x \<longleftrightarrow> x\<^sup>T\<^sup>+ = x\<^sup>+"
-by (metis antisym conv_invol star_slide_var star_unfoldl_eq many_strongly_connected_iff_5)
+by (metis order.antisym conv_invol star_slide_var star_unfoldl_eq many_strongly_connected_iff_5)
lemma many_strongly_connected_iff_5_eq:
"many_strongly_connected x \<longleftrightarrow> x\<^sup>\<star>;x\<^sup>T = x\<^sup>+"
by (metis order.refl star_slide_var many_strongly_connected_iff_5 many_strongly_connected_iff_7)
lemma many_strongly_connected_iff_6_eq:
"many_strongly_connected x \<longleftrightarrow> x\<^sup>T;x\<^sup>\<star> = x\<^sup>+"
using many_strongly_connected_iff_6 many_strongly_connected_iff_7 by force
lemma many_strongly_connected_implies_no_start_end_points:
assumes "many_strongly_connected x"
shows "no_start_end_points x"
by (metis assms conway.dagger_unfoldl_distr mult_assoc sup_top_left conv_invol
many_strongly_connected_iff_7)
lemma many_strongly_connected_implies_8:
assumes "many_strongly_connected x"
shows "x;x\<^sup>T \<le> x\<^sup>+"
by (simp add: assms mult_isol)
lemma many_strongly_connected_implies_9:
assumes "many_strongly_connected x"
shows "x\<^sup>T;x \<le> x\<^sup>+"
by (metis assms eq_refl phl_cons1 star_ext star_slide_var)
lemma many_strongly_connected_implies_10:
assumes "many_strongly_connected x"
shows "x;x\<^sup>T;x\<^sup>\<star> \<le> x\<^sup>+"
by (simp add: assms comp_assoc mult_isol)
lemma many_strongly_connected_implies_10_eq:
assumes "many_strongly_connected x"
shows "x;x\<^sup>T;x\<^sup>\<star> = x\<^sup>+"
-proof (rule antisym)
+proof (rule order.antisym)
show "x;x\<^sup>T;x\<^sup>\<star> \<le> x\<^sup>+"
by (simp add: assms comp_assoc mult_isol)
next
have "x\<^sup>+ \<le> x;x\<^sup>T;x;x\<^sup>\<star>"
using mult_isor x_leq_triple_x by blast
thus "x\<^sup>+ \<le> x;x\<^sup>T;x\<^sup>\<star>"
by (simp add: comp_assoc mult_isol order_trans)
qed
lemma many_strongly_connected_implies_11:
assumes "many_strongly_connected x"
shows "x\<^sup>\<star>;x\<^sup>T;x \<le> x\<^sup>+"
by (metis assms conv_contrav conv_iso mult_isol star_1l star_slide_var)
lemma many_strongly_connected_implies_11_eq:
assumes "many_strongly_connected x"
shows "x\<^sup>\<star>;x\<^sup>T;x = x\<^sup>+"
by (metis assms comp_assoc conv_invol many_strongly_connected_iff_5_eq
many_strongly_connected_implies_10_eq)
lemma many_strongly_connected_implies_12:
assumes "many_strongly_connected x"
shows "x\<^sup>\<star>;x;x\<^sup>T \<le> x\<^sup>+"
by (metis assms comp_assoc mult_isol star_1l star_slide_var)
lemma many_strongly_connected_implies_12_eq:
assumes "many_strongly_connected x"
shows "x\<^sup>\<star>;x;x\<^sup>T = x\<^sup>+"
by (metis assms comp_assoc star_slide_var many_strongly_connected_implies_10_eq)
lemma many_strongly_connected_implies_13:
assumes "many_strongly_connected x"
shows "x\<^sup>T;x;x\<^sup>\<star> \<le> x\<^sup>+"
by (metis assms star_slide_var many_strongly_connected_implies_11 mult.assoc)
lemma many_strongly_connected_implies_13_eq:
assumes "many_strongly_connected x"
shows "x\<^sup>T;x;x\<^sup>\<star> = x\<^sup>+"
by (metis assms conv_invol many_strongly_connected_iff_7 many_strongly_connected_implies_10_eq)
lemma many_strongly_connected_iff_8:
assumes "is_p_fun x"
shows "many_strongly_connected x \<longleftrightarrow> x;x\<^sup>T \<le> x\<^sup>+"
apply (rule iffI)
apply (simp add: mult_isol)
apply (simp add: many_strongly_connected_iff_1)
by (metis comp_assoc conv_invol dual_order.trans mult_isol x_leq_triple_x assms comp_assoc
dual_order.trans is_p_fun_def order.refl prod_star_closure star_ref)
lemma many_strongly_connected_iff_9:
assumes "is_inj x"
shows "many_strongly_connected x \<longleftrightarrow> x\<^sup>T;x \<le> x\<^sup>+"
by (metis assms conv_contrav conv_iso inj_p_fun star_conv star_slide_var
many_strongly_connected_iff_1 many_strongly_connected_iff_8)
lemma many_strongly_connected_iff_10:
assumes "is_p_fun x"
shows "many_strongly_connected x \<longleftrightarrow> x;x\<^sup>T;x\<^sup>\<star> \<le> x\<^sup>+"
apply (rule iffI)
apply (simp add: comp_assoc mult_isol)
by (metis assms mult_isol mult_oner order_trans star_ref many_strongly_connected_iff_8)
lemma many_strongly_connected_iff_10_eq:
assumes "is_p_fun x"
shows "many_strongly_connected x \<longleftrightarrow> x;x\<^sup>T;x\<^sup>\<star> = x\<^sup>+"
using assms many_strongly_connected_iff_10 many_strongly_connected_implies_10_eq by fastforce
lemma many_strongly_connected_iff_11:
assumes "is_inj x"
shows "many_strongly_connected x \<longleftrightarrow> x\<^sup>\<star>;x\<^sup>T;x \<le> x\<^sup>+"
by (metis assms comp_assoc conv_contrav conv_iso inj_p_fun plus_conv star_conv
many_strongly_connected_iff_10 many_strongly_connected_iff_2)
lemma many_strongly_connected_iff_11_eq:
assumes "is_inj x"
shows "many_strongly_connected x \<longleftrightarrow> x\<^sup>\<star>;x\<^sup>T;x = x\<^sup>+"
using assms many_strongly_connected_iff_11 many_strongly_connected_implies_11_eq by fastforce
lemma many_strongly_connected_iff_12:
assumes "is_p_fun x"
shows "many_strongly_connected x \<longleftrightarrow> x\<^sup>\<star>;x;x\<^sup>T \<le> x\<^sup>+"
by (metis assms dual_order.trans mult_double_iso mult_oner star_ref star_slide_var
many_strongly_connected_iff_8 many_strongly_connected_implies_12)
lemma many_strongly_connected_iff_12_eq:
assumes "is_p_fun x"
shows "many_strongly_connected x \<longleftrightarrow> x\<^sup>\<star>;x;x\<^sup>T = x\<^sup>+"
using assms many_strongly_connected_iff_12 many_strongly_connected_implies_12_eq by fastforce
lemma many_strongly_connected_iff_13:
assumes "is_inj x"
shows "many_strongly_connected x \<longleftrightarrow> x\<^sup>T;x;x\<^sup>\<star> \<le> x\<^sup>+"
by (metis assms comp_assoc conv_contrav conv_iso inj_p_fun star_conv star_slide_var
many_strongly_connected_iff_1 many_strongly_connected_iff_12)
lemma many_strongly_connected_iff_13_eq:
assumes "is_inj x"
shows "many_strongly_connected x \<longleftrightarrow> x\<^sup>T;x;x\<^sup>\<star> = x\<^sup>+"
using assms many_strongly_connected_iff_13 many_strongly_connected_implies_13_eq by fastforce
text \<open>Equivalences and implications for \<open>one_strongly_connected\<close>\<close>
lemma one_strongly_connected_iff:
"one_strongly_connected x \<longleftrightarrow> connected x \<and> many_strongly_connected x"
apply (rule iffI)
apply (metis top_greatest x_leq_triple_x mult_double_iso top_greatest dual_order.trans
many_strongly_connected_iff_1 comp_assoc conv_contrav conv_invol conv_iso le_supI2
star_conv)
by (metis comp_assoc conv_contrav conv_iso conv_one conway.dagger_denest star_conv star_invol
star_sum_unfold star_trans_eq)
lemma one_strongly_connected_iff_1:
"one_strongly_connected x \<longleftrightarrow> x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>+"
proof
assume 1: "one_strongly_connected x"
have "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>T;x;x\<^sup>T;1;x\<^sup>T"
by (metis conv_invol mult_isor x_leq_triple_x)
also from 1 have "... \<le> x\<^sup>T;x;x\<^sup>\<star>"
by (metis distrib_left mult_assoc sup.absorb_iff1)
also from 1 have "... \<le> x\<^sup>+"
using many_strongly_connected_implies_13 one_strongly_connected_iff by blast
finally show "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>+"
.
next
assume "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>+"
thus "one_strongly_connected x"
using dual_order.trans star_1l by blast
qed
lemma one_strongly_connected_iff_1_eq:
"one_strongly_connected x \<longleftrightarrow> x\<^sup>T;1;x\<^sup>T = x\<^sup>+"
apply (rule iffI, simp_all)
by (metis comp_assoc conv_contrav conv_invol mult_double_iso plus_conv star_slide_var top_greatest
- top_plus many_strongly_connected_implies_10_eq one_strongly_connected_iff eq_iff
+ top_plus many_strongly_connected_implies_10_eq one_strongly_connected_iff order.eq_iff
one_strongly_connected_iff_1)
lemma one_strongly_connected_iff_2:
"one_strongly_connected x \<longleftrightarrow> x;1;x \<le> x\<^sup>T\<^sup>\<star>"
by (metis conv_invol eq_refl less_eq_def one_strongly_connected_iff)
lemma one_strongly_connected_iff_3:
"one_strongly_connected x \<longleftrightarrow> x;1;x \<le> x\<^sup>T\<^sup>+"
by (metis comp_assoc conv_contrav conv_invol conv_iso conv_one star_conv
one_strongly_connected_iff_1)
lemma one_strongly_connected_iff_3_eq:
"one_strongly_connected x \<longleftrightarrow> x;1;x = x\<^sup>T\<^sup>+"
by (metis conv_invol one_strongly_connected_iff_1_eq one_strongly_connected_iff_2)
lemma one_strongly_connected_iff_4_eq:
"one_strongly_connected x \<longleftrightarrow> x\<^sup>T;1;x = x\<^sup>+"
apply (rule iffI)
apply (metis comp_assoc top_plus many_strongly_connected_iff_7 one_strongly_connected_iff
one_strongly_connected_iff_1_eq)
by (metis comp_assoc conv_contrav conv_invol conv_one plus_conv top_plus
one_strongly_connected_iff_1_eq)
lemma one_strongly_connected_iff_5_eq:
"one_strongly_connected x \<longleftrightarrow> x;1;x\<^sup>T = x\<^sup>+"
using comp_assoc conv_contrav conv_invol conv_one plus_conv top_plus many_strongly_connected_iff_7
one_strongly_connected_iff one_strongly_connected_iff_3_eq by metis
lemma one_strongly_connected_iff_6_aux:
"x;x\<^sup>+ \<le> x;1;x"
by (metis comp_assoc maddux_21 mult_isol top_plus)
lemma one_strongly_connected_implies_6_eq:
assumes "one_strongly_connected x"
shows "x;1;x = x;x\<^sup>+"
by (metis assms comp_assoc many_strongly_connected_iff_7 many_strongly_connected_implies_10_eq
one_strongly_connected_iff one_strongly_connected_iff_3_eq)
lemma one_strongly_connected_iff_7_aux:
"x\<^sup>+ \<le> x;1;x"
by (metis le_infI maddux_20 maddux_21 plus_top top_plus vector_meet_comp_x')
lemma one_strongly_connected_implies_7_eq:
assumes "one_strongly_connected x"
shows "x;1;x = x\<^sup>+"
using assms many_strongly_connected_iff_7 one_strongly_connected_iff one_strongly_connected_iff_3_eq
by force
lemma one_strongly_connected_implies_8:
assumes "one_strongly_connected x"
shows "x;1;x \<le> x\<^sup>\<star>"
using assms one_strongly_connected_iff by fastforce
lemma one_strongly_connected_iff_4:
assumes "is_inj x"
shows "one_strongly_connected x \<longleftrightarrow> x\<^sup>T;1;x \<le> x\<^sup>+"
proof
assume "one_strongly_connected x"
thus "x\<^sup>T;1;x \<le> x\<^sup>+"
by (simp add: one_strongly_connected_iff_4_eq)
next
assume 1: "x\<^sup>T;1;x \<le> x\<^sup>+"
hence "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>\<star>;x;x\<^sup>T"
by (metis mult_isor star_slide_var comp_assoc conv_invol modular_var_3 vector_meet_comp_x
order.trans)
also from assms have "... \<le> x\<^sup>\<star>"
using comp_assoc is_inj_def mult_isol mult_oner by fastforce
finally show "one_strongly_connected x"
using dual_order.trans star_1l by fastforce
qed
lemma one_strongly_connected_iff_5:
assumes "is_p_fun x"
shows "one_strongly_connected x \<longleftrightarrow> x;1;x\<^sup>T \<le> x\<^sup>+"
apply (rule iffI)
using one_strongly_connected_iff_5_eq apply simp
by (metis assms comp_assoc mult_double_iso order.trans star_slide_var top_greatest top_plus
many_strongly_connected_iff_12 many_strongly_connected_iff_7 one_strongly_connected_iff_3)
lemma one_strongly_connected_iff_6:
assumes "is_p_fun x"
and "is_inj x"
shows "one_strongly_connected x \<longleftrightarrow> x;1;x \<le> x;x\<^sup>+"
proof
assume "one_strongly_connected x"
thus "x;1;x \<le> x;x\<^sup>+"
by (simp add: one_strongly_connected_implies_6_eq)
next
assume 1: "x;1;x \<le> x;x\<^sup>+"
have "x\<^sup>T;1;x \<le> x\<^sup>T;x;x\<^sup>T;1;x"
by (metis conv_invol mult_isor x_leq_triple_x)
also have "... \<le> x\<^sup>T;x;1;x"
by (metis comp_assoc mult_double_iso top_greatest)
also from 1 have "... \<le> x\<^sup>T;x;x\<^sup>+"
by (simp add: comp_assoc mult_isol)
also from assms(1) have "... \<le> x\<^sup>+"
by (metis comp_assoc is_p_fun_def mult_isor mult_onel)
finally show "one_strongly_connected x"
using assms(2) one_strongly_connected_iff_4 by blast
qed
lemma one_strongly_connected_iff_6_eq:
assumes "is_p_fun x"
and "is_inj x"
shows "one_strongly_connected x \<longleftrightarrow> x;1;x = x;x\<^sup>+"
apply (rule iffI)
using one_strongly_connected_implies_6_eq apply blast
by (simp add: assms one_strongly_connected_iff_6)
text \<open>Start points and end points\<close>
lemma start_end_implies_terminating:
assumes "has_start_points x"
and "has_end_points x"
shows "terminating x"
using assms by simp
lemma start_points_end_points_conv:
"start_points x = end_points (x\<^sup>T)"
by simp
lemma start_point_at_most_one:
assumes "path x"
shows "is_inj (start_points x)"
proof -
have isvec: "is_vector (x;1 \<cdot> -(x\<^sup>T;1))"
by (simp add: comp_assoc is_vector_def one_compl vector_1)
have "x;1 \<cdot> 1;x\<^sup>T \<le> x;1;x;x\<^sup>T"
by (metis comp_assoc conv_contrav conv_one inf.cobounded2 mult_1_right mult_isol one_conv ra_2)
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);x\<^sup>T"
using \<open>path x\<close> by (metis path_def mult_isor)
also have "... = x\<^sup>T + x\<^sup>+;x\<^sup>T + x\<^sup>T\<^sup>+"
by (simp add: star_slide_var)
also have "... \<le> x\<^sup>T\<^sup>+ + x\<^sup>+;x\<^sup>T + x\<^sup>T\<^sup>+"
by (metis add_iso mult_1_right star_unfoldl_eq subdistl)
also have "... \<le> x\<^sup>\<star>;x;x\<^sup>T + x\<^sup>T\<^sup>+"
by (simp add: star_slide_var add_comm)
also have "... \<le> x\<^sup>\<star>;1' + x\<^sup>T\<^sup>+"
using \<open>path x\<close> by (metis path_def is_inj_def comp_assoc distrib_left join_iso less_eq_def)
also have "... = 1' + x\<^sup>\<star>;x + x\<^sup>T;x\<^sup>T\<^sup>\<star>"
by simp
also have "... \<le> 1' + 1;x + x\<^sup>T;1"
by (metis join_isol mult_isol mult_isor sup.mono top_greatest)
finally have aux: "x;1 \<cdot> 1;x\<^sup>T \<le> 1' + 1;x + x\<^sup>T;1" .
from aux have "x;1 \<cdot> 1;x\<^sup>T \<cdot> -(x\<^sup>T;1) \<cdot> -(1;x) \<le> 1'"
by (simp add: galois_1 sup_commute)
hence "(x;1 \<cdot> -(x\<^sup>T;1)) \<cdot> (x;1 \<cdot> -(x\<^sup>T;1))\<^sup>T \<le> 1'"
by (simp add: conv_compl inf.assoc inf.left_commute)
with isvec have "(x;1 \<cdot> -(x\<^sup>T;1)) ; (x;1 \<cdot> -(x\<^sup>T;1))\<^sup>T \<le> 1'"
by (metis vector_meet_comp')
thus "is_inj (start_points x)"
by (simp add: conv_compl is_inj_def)
qed
lemma start_point_zero_point:
assumes "path x"
shows "start_points x = 0 \<or> is_point (start_points x)"
using assms start_point_at_most_one comp_assoc is_point_def is_vector_def vector_compl vector_mult
by simp
lemma start_point_iff1:
assumes "path x"
shows "is_point (start_points x) \<longleftrightarrow> \<not>(no_start_points x)"
using assms start_point_zero_point galois_aux2 is_point_def by blast
lemma end_point_at_most_one:
assumes "path x"
shows "is_inj (end_points x)"
by (metis assms conv_path compl_bot_eq conv_invol inj_def_var1 is_point_def top_greatest
start_point_zero_point)
lemma end_point_zero_point:
assumes "path x"
shows "end_points x = 0 \<or> is_point (end_points x)"
using assms conv_path start_point_zero_point by fastforce
lemma end_point_iff1:
assumes "path x"
shows "is_point (end_points x) \<longleftrightarrow> \<not>(no_end_points x)"
using assms end_point_zero_point galois_aux2 is_point_def by blast
lemma predecessor_point':
assumes "path x"
and "point s"
and "point e"
and "e;s\<^sup>T \<le> x"
shows "x;s = e"
-proof (rule antisym)
+proof (rule order.antisym)
show 1: "e \<le> x ; s"
using assms(2,4) point_def ss423bij by blast
show "x ; s \<le> e"
proof -
have "e\<^sup>T ; (x ; s) = 1"
- using 1 by (metis assms(3) eq_iff is_vector_def point_def ss423conv top_greatest)
+ using 1 by (metis assms(3) order.eq_iff is_vector_def point_def ss423conv top_greatest)
thus ?thesis
- by (metis assms(1-3) comp_assoc conv_contrav conv_invol eq_iff inj_compose is_vector_def
+ by (metis assms(1-3) comp_assoc conv_contrav conv_invol order.eq_iff inj_compose is_vector_def
mult_isol path_def point_def ss423conv sur_def_var1 top_greatest)
qed
qed
lemma predecessor_point:
assumes "path x"
and "point s"
and "point e"
and "e;s\<^sup>T \<le> x"
shows "point(x;s)"
using predecessor_point' assms by blast
lemma points_of_path_iff:
shows "(x + x\<^sup>T);1 = x\<^sup>T;1 + start_points(x)"
and "(x + x\<^sup>T);1 = x;1 + end_points(x)"
using aux9 inf.commute sup.commute by auto
text \<open>Path concatenation preliminaries\<close>
lemma path_concat_aux_1:
assumes "x;1 \<cdot> y;1 \<cdot> y\<^sup>T;1 = 0"
and "end_points x = start_points y"
shows "x;1 \<cdot> y;1 = 0"
proof -
have "x;1 \<cdot> y;1 = (x;1 \<cdot> y;1 \<cdot> y\<^sup>T;1) + (x;1 \<cdot> y;1 \<cdot> -(y\<^sup>T;1))"
by simp
also from assms(1) have "... = x;1 \<cdot> y;1 \<cdot> -(y\<^sup>T;1)"
by (metis aux6_var de_morgan_3 inf.left_commute inf_compl_bot inf_sup_absorb)
also from assms(2) have "... = x;1 \<cdot> x\<^sup>T;1 \<cdot> -(x;1)"
by (simp add: inf.assoc)
also have "... = 0"
by (simp add: inf.commute inf.assoc)
finally show ?thesis .
qed
lemma path_concat_aux_2:
assumes "x;1 \<cdot> x\<^sup>T;1 \<cdot> y\<^sup>T;1 = 0"
and "end_points x = start_points y"
shows "x\<^sup>T;1 \<cdot> y\<^sup>T;1 = 0"
proof -
have "y\<^sup>T;1 \<cdot> x\<^sup>T;1 \<cdot> (x\<^sup>T)\<^sup>T;1 = 0"
using assms(1) inf.assoc inf.commute by force
thus ?thesis
by (metis assms(2) conv_invol inf.commute path_concat_aux_1)
qed
lemma path_concat_aux3_1:
assumes "path x"
shows "x;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
proof -
have "x;1;x\<^sup>T \<le> x;1;x\<^sup>T;x;x\<^sup>T"
by (metis comp_assoc conv_invol mult_isol x_leq_triple_x)
also have "... \<le> x;1;x;x\<^sup>T"
by (metis mult_isol mult_isor mult_assoc top_greatest)
also from assms have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);x\<^sup>T"
using path_def comp_assoc mult_isor by blast
also have "... = x\<^sup>\<star>;x;x\<^sup>T + x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (simp add: star_slide_var star_star_plus)
also have "... \<le> x\<^sup>\<star>;1' + x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (metis assms path_def is_inj_def join_iso mult_isol mult_assoc)
also have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using join_isol by simp
finally show ?thesis .
qed
lemma path_concat_aux3_2:
assumes "path x"
shows "x\<^sup>T;1;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
proof -
have "x\<^sup>T;1;x \<le> x\<^sup>T;x;x\<^sup>T;1;x"
by (metis comp_assoc conv_invol mult_isor x_leq_triple_x)
also have "... \<le> x\<^sup>T;x;1;x"
by (metis mult_isol mult_isor mult_assoc top_greatest)
also from assms have "... \<le> x\<^sup>T;(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>)"
by (simp add: comp_assoc mult_isol path_def)
also have "... = x\<^sup>T;x;x\<^sup>\<star> + x\<^sup>T;x\<^sup>T\<^sup>\<star>"
by (simp add: comp_assoc distrib_left star_star_plus)
also have "... \<le> 1';x\<^sup>\<star> + x\<^sup>T;x\<^sup>T\<^sup>\<star>"
by (metis assms path_def is_p_fun_def join_iso mult_isor mult_assoc)
also have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using join_isol by simp
finally show ?thesis .
qed
lemma path_concat_aux3_3:
assumes "path x"
shows "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
proof -
have "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>T;x;x\<^sup>T;1;x\<^sup>T"
by (metis comp_assoc conv_invol mult_isor x_leq_triple_x)
also have "... \<le> x\<^sup>T;x;1;x\<^sup>T"
by (metis mult_isol mult_isor mult_assoc top_greatest)
also from assms have "... \<le> x\<^sup>T;(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>)"
using path_concat_aux3_1 by (simp add: mult_assoc mult_isol)
also have "... = x\<^sup>T;x;x\<^sup>\<star> + x\<^sup>T;x\<^sup>T\<^sup>\<star>"
by (simp add: comp_assoc distrib_left star_star_plus)
also have "... \<le> 1';x\<^sup>\<star> + x\<^sup>T;x\<^sup>T\<^sup>\<star>"
by (metis assms path_def is_p_fun_def join_iso mult_isor mult_assoc)
also have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using join_isol by simp
finally show ?thesis .
qed
lemma path_concat_aux_3:
assumes "path x"
and "y \<le> x\<^sup>+ + x\<^sup>T\<^sup>+"
and "z \<le> x\<^sup>+ + x\<^sup>T\<^sup>+"
shows "y;1;z \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
proof -
from assms(2,3) have "y;1;z \<le> (x\<^sup>+ + x\<^sup>T\<^sup>+);1;(x\<^sup>+ + x\<^sup>T\<^sup>+)"
using mult_isol_var mult_isor by blast
also have "... = x\<^sup>+;1;x\<^sup>+ + x\<^sup>+;1;x\<^sup>T\<^sup>+ + x\<^sup>T\<^sup>+;1;x\<^sup>+ + x\<^sup>T\<^sup>+;1;x\<^sup>T\<^sup>+"
by (simp add: distrib_left sup_commute sup_left_commute)
also have "... = x;x\<^sup>\<star>;1;x\<^sup>\<star>;x + x;x\<^sup>\<star>;1;x\<^sup>T\<^sup>\<star>;x\<^sup>T + x\<^sup>T;x\<^sup>T\<^sup>\<star>;1;x\<^sup>\<star>;x + x\<^sup>T;x\<^sup>T\<^sup>\<star>;1;x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (simp add: comp_assoc star_slide_var)
also have "... \<le> x;1;x + x;x\<^sup>\<star>;1;x\<^sup>T\<^sup>\<star>;x\<^sup>T + x\<^sup>T;x\<^sup>T\<^sup>\<star>;1;x\<^sup>\<star>;x + x\<^sup>T;x\<^sup>T\<^sup>\<star>;1;x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (metis comp_assoc mult_double_iso top_greatest join_iso)
also have "... \<le> x;1;x + x;1;x\<^sup>T + x\<^sup>T;x\<^sup>T\<^sup>\<star>;1;x\<^sup>\<star>;x + x\<^sup>T;x\<^sup>T\<^sup>\<star>;1;x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (metis comp_assoc mult_double_iso top_greatest join_iso join_isol)
also have "... \<le> x;1;x + x;1;x\<^sup>T + x\<^sup>T;1;x + x\<^sup>T;x\<^sup>T\<^sup>\<star>;1;x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (metis comp_assoc mult_double_iso top_greatest join_iso join_isol)
also have "... \<le> x;1;x + x;1;x\<^sup>T + x\<^sup>T;1;x + x\<^sup>T;1;x\<^sup>T"
by (metis comp_assoc mult_double_iso top_greatest join_isol)
also have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using assms(1) path_def path_concat_aux3_1 path_concat_aux3_2 path_concat_aux3_3 join_iso join_isol
by simp
finally show ?thesis .
qed
lemma path_concat_aux_4:
"x\<^sup>\<star> + x\<^sup>T\<^sup>\<star> \<le> x\<^sup>\<star> + x\<^sup>T;1"
by (metis star_star_plus add_comm join_isol mult_isol top_greatest)
lemma path_concat_aux_5:
assumes "path x"
and "y \<le> start_points x"
and "z \<le> x + x\<^sup>T"
shows "y;1;z \<le> x\<^sup>\<star>"
proof -
from assms(1) have "x;1;x \<le> x\<^sup>\<star> + x\<^sup>T;1"
using path_def path_concat_aux_4 dual_order.trans by blast
hence aux1: "x;1;x \<cdot> -(x\<^sup>T;1) \<le> x\<^sup>\<star>"
by (simp add: galois_1 sup_commute)
from assms(1) have "x;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T;1"
using dual_order.trans path_concat_aux3_1 path_concat_aux_4 by blast
hence aux2: "x;1;x\<^sup>T \<cdot> -(x\<^sup>T;1) \<le> x\<^sup>\<star>"
by (simp add: galois_1 sup_commute)
from assms(2,3) have "y;1;z \<le> (x;1 \<cdot> -(x\<^sup>T;1));1;(x + x\<^sup>T)"
by (simp add: mult_isol_var mult_isor)
also have "... = (x;1 \<cdot> -(x\<^sup>T;1));1;x + (x;1 \<cdot> -(x\<^sup>T;1));1;x\<^sup>T"
using distrib_left by blast
also have "... = (x;1 \<cdot> -(x\<^sup>T;1) \<cdot> 1;x) + (x;1 \<cdot> -(x\<^sup>T;1));1;x\<^sup>T"
by (metis comp_assoc inf_top_right is_vector_def one_idem_mult vector_1 vector_compl)
also have "... = (x;1 \<cdot> -(x\<^sup>T;1) \<cdot> 1;x) + (x;1 \<cdot> -(x\<^sup>T;1) \<cdot> 1;x\<^sup>T)"
by (metis comp_assoc inf_top_right is_vector_def one_idem_mult vector_1 vector_compl)
also have "... = (x;1;x \<cdot> -(x\<^sup>T;1)) + (x;1;x\<^sup>T -(x\<^sup>T;1))"
using vector_meet_comp_x vector_meet_comp_x' diff_eq inf.assoc inf.commute by simp
also from aux1 aux2 have "... \<le> x\<^sup>\<star>"
by (simp add: diff_eq join_iso)
finally show ?thesis .
qed
lemma path_conditions_disjoint_points_iff:
"x;1 \<cdot> (x\<^sup>T;1 + y;1) \<cdot> y\<^sup>T;1 = 0 \<and> start_points x \<cdot> end_points y = 0 \<longleftrightarrow> x;1 \<cdot> y\<^sup>T;1 = 0"
proof
assume 1: "x ; 1 \<cdot> y\<^sup>T ; 1 = 0"
hence g1: "x ; 1 \<cdot> (x\<^sup>T ; 1 + y ; 1) \<cdot> y\<^sup>T ; 1 = 0"
by (metis inf.left_commute inf_bot_right inf_commute)
have g2: "start_points x \<cdot> end_points y = 0"
using 1 by (metis compl_inf_bot inf.assoc inf.commute inf.left_idem)
show "x;1 \<cdot> (x\<^sup>T;1 + y;1) \<cdot> y\<^sup>T;1 = 0 \<and> start_points x \<cdot> end_points y = 0"
using g1 and g2 by simp
next
assume a: "x;1 \<cdot> (x\<^sup>T;1 + y;1) \<cdot> y\<^sup>T;1 = 0 \<and> start_points x \<cdot> end_points y = 0"
from a have a1: "x;1 \<cdot> x\<^sup>T;1 \<cdot> y\<^sup>T;1 = 0"
by (simp add: inf.commute inf_sup_distrib1)
from a have a2: "x;1 \<cdot> y;1 \<cdot> y\<^sup>T;1 = 0"
by (simp add: inf.commute inf_sup_distrib1)
from a have a3: "start_points x \<cdot> end_points y = 0"
by blast
have "x;1 \<cdot> y\<^sup>T;1 = x;1 \<cdot> x\<^sup>T;1 \<cdot> y\<^sup>T;1 + x;1 \<cdot> -(x\<^sup>T;1) \<cdot> y\<^sup>T;1"
by (metis aux4 inf_sup_distrib2)
also from a1 have "... = x;1 \<cdot> -(x\<^sup>T;1) \<cdot> y\<^sup>T;1"
using sup_bot_left by blast
also have "... = x;1 \<cdot> -(x\<^sup>T;1) \<cdot> y;1 \<cdot> y\<^sup>T;1 + x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y;1) \<cdot> y\<^sup>T;1"
by (metis aux4 inf_sup_distrib2)
also have "... \<le> x;1 \<cdot> y;1 \<cdot> y\<^sup>T;1 + x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y;1) \<cdot> y\<^sup>T;1"
using join_iso meet_iso by simp
also from a2 have "... = start_points x \<cdot> end_points y"
using sup_bot_left inf.commute inf.left_commute by simp
also from a3 have "... = 0"
by blast
finally show "x;1 \<cdot> y\<^sup>T;1 = 0"
using le_bot by blast
qed
end (* end relation_algebra_rtc *)
subsection \<open>Consequences with the Tarski rule\<close>
context relation_algebra_rtc_tarski
begin
text \<open>General theorems\<close>
lemma reachable_implies_predecessor:
assumes "p \<noteq> q"
and "point p"
and "point q"
and "x\<^sup>\<star>;q \<le> x\<^sup>T\<^sup>\<star>;p"
shows "x;q \<noteq> 0"
proof
assume contra: "x;q=0"
with assms(4) have "q \<le> x\<^sup>T\<^sup>\<star>;p"
by (simp add: independence1)
hence "p \<le> x\<^sup>\<star>;q"
by (metis assms(2,3) point_swap star_conv)
with contra assms(2,3) have "p=q"
by (simp add: independence1 is_point_def point_singleton point_is_point)
with assms(1) show False
by simp
qed
lemma acyclic_imp_one_step_different_points:
assumes "is_acyclic x"
and "point p"
and "point q"
and "p \<le> x;q"
shows "p \<le> -q" and "p \<noteq> q"
using acyclic_reachable_points assms point_is_point point_not_equal(1) by auto
text \<open>Start points and end points\<close>
lemma start_point_iff2:
assumes "path x"
shows "is_point (start_points x) \<longleftrightarrow> has_start_points x"
proof -
have "has_start_points x \<longleftrightarrow> 1 \<le> -(1;x);x;1"
- by (simp add: eq_iff)
+ by (simp add: order.eq_iff)
also have "... \<longleftrightarrow> 1 \<le> 1;x\<^sup>T;-(x\<^sup>T;1)"
by (metis comp_assoc conv_compl conv_contrav conv_iso conv_one)
also have "... \<longleftrightarrow> 1 \<le> 1;(x;1 \<cdot> -(x\<^sup>T;1))"
by (metis (no_types) conv_contrav conv_one inf.commute is_vector_def one_idem_mult ra_2 vector_1
vector_meet_comp_x)
also have "... \<longleftrightarrow> 1 = 1;(x;1 \<cdot> -(x\<^sup>T;1))"
- by (simp add: eq_iff)
+ by (simp add: order.eq_iff)
also have "... \<longleftrightarrow> x;1 \<cdot> -(x\<^sup>T;1) \<noteq> 0"
by (metis tarski comp_assoc one_compl ra_1 ss_p18)
also have "... \<longleftrightarrow> is_point (start_points x)"
using assms is_point_def start_point_zero_point by blast
finally show ?thesis ..
qed
lemma end_point_iff2:
assumes "path x"
shows "is_point (end_points x) \<longleftrightarrow> has_end_points x"
by (metis assms conv_invol conv_has_start_points conv_path start_point_iff2)
lemma edge_is_path:
assumes "is_point p"
and "is_point q"
shows "path (p;q\<^sup>T)"
apply (unfold path_def; intro conjI)
apply (metis assms comp_assoc is_point_def le_supI1 star_ext vector_rectangle point_equations(3))
apply (metis is_p_fun_def assms comp_assoc conv_contrav conv_invol is_inj_def is_point_def
vector_2_var vector_meet_comp_x' point_equations)
by (metis is_inj_def assms conv_invol conv_times is_point_def p_fun_mult_var vector_meet_comp)
lemma edge_start:
assumes "is_point p"
and "is_point q"
and "p \<noteq> q"
shows "start_points (p;q\<^sup>T) = p"
using assms by (simp add: comp_assoc point_equations(1,3) point_not_equal inf.absorb1)
lemma edge_end:
assumes "is_point p"
and "is_point q"
and "p \<noteq> q"
shows "end_points (p;q\<^sup>T) = q"
using assms edge_start by simp
lemma loop_no_start:
assumes "is_point p"
shows "start_points (p;p\<^sup>T) = 0"
by simp
lemma loop_no_end:
assumes "is_point p"
shows "end_points (p;p\<^sup>T) = 0"
by simp
lemma start_point_no_predecessor:
"x;start_points(x) = 0"
by (metis inf_top.right_neutral modular_1_aux')
lemma end_point_no_successor:
"x\<^sup>T;end_points(x) = 0"
by (metis conv_invol start_point_no_predecessor)
lemma start_to_end:
assumes "path x"
shows "start_points(x);end_points(x)\<^sup>T \<le> x\<^sup>\<star>"
proof (cases "end_points(x) = 0")
assume "end_points(x) = 0"
thus ?thesis
by simp
next
assume ass: "end_points(x) \<noteq> 0"
hence nz: "x;end_points(x) \<noteq> 0"
by (metis comp_res_aux compl_bot_eq inf.left_idem)
have a: "x;end_points(x);end_points(x)\<^sup>T \<le> x + x\<^sup>T"
by (metis end_point_at_most_one assms(1) is_inj_def comp_assoc mult_isol mult_oner le_supI1)
have "start_points(x);end_points(x)\<^sup>T = start_points(x);1;end_points(x)\<^sup>T"
using ass by (simp add: comp_assoc is_vector_def one_compl vector_1)
also have "... = start_points(x);1;x;end_points(x);1;end_points(x)\<^sup>T"
using nz tarski by (simp add: comp_assoc)
also have "... = start_points(x);1;x;end_points(x);end_points(x)\<^sup>T"
using ass by (simp add: comp_assoc is_vector_def one_compl vector_1)
also with a assms(1) have "... \<le> x\<^sup>\<star>"
using path_concat_aux_5 comp_assoc eq_refl by simp
finally show ?thesis .
qed
lemma path_acyclic:
assumes "has_start_points_path x"
shows "is_acyclic x"
proof -
let ?r = "start_points(x)"
have pt: "point(?r)"
using assms point_is_point start_point_iff2 by blast
have "x\<^sup>+\<cdot>1' = (x\<^sup>+)\<^sup>T\<cdot>x\<^sup>+\<cdot>1'"
by (metis conv_e conv_times inf.assoc inf.left_idem inf_le2 many_strongly_connected_iff_7
mult_oner star_subid)
also have "... \<le> x\<^sup>T;1\<cdot>x\<^sup>+\<cdot>1'"
by (metis conv_contrav inf.commute maddux_20 meet_double_iso plus_top star_conv star_slide_var)
finally have "?r;(x\<^sup>+\<cdot>1') \<le> ?r;(x\<^sup>T;1\<cdot>x\<^sup>+\<cdot>1')"
using mult_isol by blast
also have "... = (?r\<cdot>1;x);(x\<^sup>+\<cdot>1')"
by (metis (no_types, lifting) comp_assoc conv_contrav conv_invol conv_one inf.assoc
is_vector_def one_idem_mult vector_2)
also have "... = ?r;x;(x\<^sup>+\<cdot>1')"
by (metis comp_assoc inf_top.right_neutral is_vector_def one_compl one_idem_mult vector_1)
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);(x\<^sup>+\<cdot>1')"
using assms(1) mult_isor
by (meson connected_iff4 dual_order.trans mult_subdistr path_concat_aux3_3)
also have "... = x\<^sup>\<star>;(x\<^sup>+\<cdot>1') + x\<^sup>T\<^sup>+;(x\<^sup>+\<cdot>1')"
by (metis distrib_right star_star_plus sup.commute)
also have "... \<le> x\<^sup>\<star>;(x\<^sup>+\<cdot>1') + x\<^sup>T;1"
by (metis join_isol mult_isol plus_top top_greatest)
finally have "?r;(x\<^sup>+\<cdot>1');1 \<le> x\<^sup>\<star>;(x\<^sup>+\<cdot>1');1 + x\<^sup>T;1"
by (metis distrib_right inf_absorb2 mult_assoc mult_subdistr one_idem_mult)
hence 1: "?r;(x\<^sup>+\<cdot>1');1 \<le> x\<^sup>T;1"
using assms(1) path_def inj_implies_step_forwards_backwards sup_absorb2 by simp
have "x\<^sup>+\<cdot>1' \<le> (x\<^sup>+\<cdot>1');1"
by (simp add: maddux_20)
also have "... \<le> ?r\<^sup>T;?r;(x\<^sup>+\<cdot>1');1"
using pt comp_assoc point_def ss423conv by fastforce
also have "... \<le> ?r\<^sup>T;x\<^sup>T;1"
using 1 by (simp add: comp_assoc mult_isol)
also have "... = 0"
by (metis start_point_no_predecessor annil conv_contrav conv_zero)
finally show ?thesis
using galois_aux le_bot by blast
qed
text \<open>Equivalences for \<open>terminating\<close>\<close>
lemma backward_terminating_iff1:
assumes "path x"
shows "backward_terminating x \<longleftrightarrow> has_start_points x \<or> x = 0"
proof
assume "backward_terminating x"
hence "1;x;1 \<le> 1;-(1;x);x;1;1"
by (metis mult_isor mult_isol comp_assoc)
also have "... = -(1;x);x;1"
by (metis conv_compl conv_contrav conv_invol conv_one mult_assoc one_compl one_idem_mult)
finally have "1;x;1 \<le> -(1;x);x;1" .
with tarski show "has_start_points x \<or> x = 0"
by (metis top_le)
next
show "has_start_points x \<or> x = 0 \<Longrightarrow> backward_terminating x"
by fastforce
qed
lemma backward_terminating_iff2_aux:
assumes "path x"
shows "x;1 \<cdot> 1;x\<^sup>T \<cdot> -(1;x) \<le> x\<^sup>T\<^sup>\<star>"
proof -
have "x;1 \<cdot> 1;x\<^sup>T \<le> x;1;x;x\<^sup>T"
by (metis conv_invol modular_var_3 vector_meet_comp_x vector_meet_comp_x')
also from assms have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);x\<^sup>T"
using path_def mult_isor by blast
also have "... \<le> x\<^sup>\<star>;x;x\<^sup>T + x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (simp add: star_star_plus star_slide_var add_comm)
also from assms have "... \<le> x\<^sup>\<star>;1' + x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (metis path_def is_inj_def join_iso mult_assoc mult_isol)
also have "... = x\<^sup>+ + x\<^sup>T\<^sup>\<star>"
by (metis mult_1_right star_slide_var star_star_plus sup.commute)
also have "... \<le> x\<^sup>T\<^sup>\<star> + 1;x"
by (metis join_iso mult_isor star_slide_var top_greatest add_comm)
finally have "x;1 \<cdot> 1;x\<^sup>T \<le> x\<^sup>T\<^sup>\<star> + 1;x" .
thus ?thesis
by (simp add: galois_1 sup.commute)
qed
lemma backward_terminating_iff2:
assumes "path x"
shows "backward_terminating x \<longleftrightarrow> x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
proof
assume "backward_terminating x"
with assms have "has_start_points x \<or> x = 0"
by (simp add: backward_terminating_iff1)
thus "x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
proof
assume "x = 0"
thus ?thesis
by simp
next
assume "has_start_points x"
hence aux1: "1 = 1;x\<^sup>T;-(x\<^sup>T;1)"
by (metis comp_assoc conv_compl conv_contrav conv_one)
have "x = x \<cdot> 1"
by simp
also have "... \<le> (x;-(1;x) \<cdot> 1;x\<^sup>T);-(x\<^sup>T;1)"
by (metis inf.commute aux1 conv_compl conv_contrav conv_invol conv_one modular_2_var)
also have "... = (x;1 \<cdot> -(1;x) \<cdot> 1;x\<^sup>T);-(x\<^sup>T;1)"
by (metis comp_assoc conv_compl conv_contrav conv_invol conv_one inf.commute inf_top_left
one_compl ra_1)
also from assms have "... \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
using backward_terminating_iff2_aux inf.commute inf.assoc mult_isor by fastforce
finally show "x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)" .
qed
next
assume "x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
hence"x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1) \<cdot> x"
by simp
also have "... = (x\<^sup>T\<^sup>\<star> \<cdot> -(1;x));1 \<cdot> x"
by (metis one_compl conv_compl conv_contrav conv_invol conv_one inf_top_left ra_2)
also have "... \<le> (x\<^sup>T\<^sup>\<star> \<cdot> -(1;x)) ; (1 \<cdot> (x\<^sup>\<star> \<cdot> -(1;x)\<^sup>T);x)"
by (metis (mono_tags) conv_compl conv_invol conv_times modular_1_var star_conv)
also have "... \<le> -(1;x);x\<^sup>\<star>;x"
by (simp add: mult_assoc mult_isol_var)
also have "... \<le> -(1;x);x;1"
by (simp add: mult_assoc mult_isol star_slide_var)
finally show "backward_terminating x" .
qed
lemma backward_terminating_iff3_aux:
assumes "path x"
shows "x\<^sup>T;1 \<cdot> 1;x\<^sup>T \<cdot> -(1;x) \<le> x\<^sup>T\<^sup>\<star>"
proof -
have "x\<^sup>T;1 \<cdot> 1;x\<^sup>T \<le> x\<^sup>T;1;x;x\<^sup>T"
by (metis conv_invol modular_var_3 vector_meet_comp_x vector_meet_comp_x')
also from assms have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);x\<^sup>T"
using mult_isor path_concat_aux3_2 by blast
also have "... \<le> x\<^sup>\<star>;x;x\<^sup>T + x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (simp add: star_star_plus star_slide_var add_comm)
also from assms have "... \<le> x\<^sup>\<star>;1' + x\<^sup>T\<^sup>\<star>;x\<^sup>T"
by (metis path_def is_inj_def join_iso mult_assoc mult_isol)
also have "... = x\<^sup>+ + x\<^sup>T\<^sup>\<star>"
by (metis mult_1_right star_slide_var star_star_plus sup.commute)
also have "... \<le> x\<^sup>T\<^sup>\<star> + 1;x"
by (metis join_iso mult_isor star_slide_var top_greatest add_comm)
finally have "x\<^sup>T;1 \<cdot> 1;x\<^sup>T \<le> x\<^sup>T\<^sup>\<star> + 1;x" .
thus ?thesis
by (simp add: galois_1 sup.commute)
qed
lemma backward_terminating_iff3:
assumes "path x"
shows "backward_terminating x \<longleftrightarrow> x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
proof
assume "backward_terminating x"
with assms have "has_start_points x \<or> x = 0"
by (simp add: backward_terminating_iff1)
thus "x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
proof
assume "x = 0"
thus ?thesis
by simp
next
assume "has_start_points x"
hence aux1: "1 = 1;x\<^sup>T;-(x\<^sup>T;1)"
by (metis comp_assoc conv_compl conv_contrav conv_one)
have "x\<^sup>T = x\<^sup>T \<cdot> 1"
by simp
also have "... \<le> (x\<^sup>T;-(1;x) \<cdot> 1;x\<^sup>T);-(x\<^sup>T;1)"
by (metis inf.commute aux1 conv_compl conv_contrav conv_invol conv_one modular_2_var)
also have "... = (x\<^sup>T;1 \<cdot> -(1;x) \<cdot> 1;x\<^sup>T);-(x\<^sup>T;1)"
by (metis comp_assoc conv_compl conv_contrav conv_invol conv_one inf.commute inf_top_left one_compl ra_1)
also from assms have "... \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
using backward_terminating_iff3_aux inf.commute inf.assoc mult_isor by fastforce
finally show "x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)" .
qed
next
have 1: "-(1;x) \<cdot> x = 0"
by (simp add: galois_aux2 inf.commute maddux_21)
assume "x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
hence "x = -(1;x);x\<^sup>\<star> \<cdot> x"
by (metis (mono_tags, lifting) conv_compl conv_contrav conv_iso conv_one inf.absorb2 star_conv)
also have "... = (-(1;x);x\<^sup>+ + -(1;x);1') \<cdot> x"
by (metis distrib_left star_unfoldl_eq sup_commute)
also have "... = -(1;x);x\<^sup>+ \<cdot> x + -(1;x) \<cdot> x"
by (simp add: inf_sup_distrib2)
also have "... \<le> -(1;x);x\<^sup>+"
using 1 by simp
also have "... \<le> -(1;x);x;1"
by (simp add: mult_assoc mult_isol star_slide_var)
finally show "backward_terminating x" .
qed
lemma backward_terminating_iff4:
assumes "path x"
shows "backward_terminating x \<longleftrightarrow> x \<le> -(1;x);x\<^sup>\<star>"
apply (subst backward_terminating_iff3)
apply (rule assms)
by (metis (mono_tags, lifting) conv_compl conv_iso star_conv conv_contrav conv_one)
lemma forward_terminating_iff1:
assumes "path x"
shows "forward_terminating x \<longleftrightarrow> has_end_points x \<or> x = 0"
by (metis comp_assoc eq_refl le_bot one_compl tarski top_greatest)
lemma forward_terminating_iff2:
assumes "path x"
shows "forward_terminating x \<longleftrightarrow> x\<^sup>T \<le> x\<^sup>\<star>;-(x;1)"
by (metis assms backward_terminating_iff1 backward_terminating_iff2 end_point_iff2
forward_terminating_iff1 compl_bot_eq conv_compl conv_invol conv_one conv_path
double_compl start_point_iff2)
lemma forward_terminating_iff3:
assumes "path x"
shows "forward_terminating x \<longleftrightarrow> x \<le> x\<^sup>\<star>;-(x;1)"
by (metis assms backward_terminating_iff1 backward_terminating_iff3 end_point_iff2
forward_terminating_iff1 compl_bot_eq conv_compl conv_invol conv_one conv_path
double_compl start_point_iff2)
lemma forward_terminating_iff4:
assumes "path x"
shows "forward_terminating x \<longleftrightarrow> x \<le> -(1;x\<^sup>T);x\<^sup>T\<^sup>\<star>"
using forward_terminating_iff2 conv_contrav conv_iso star_conv assms conv_compl by force
lemma terminating_iff1:
assumes "path x"
shows "terminating x \<longleftrightarrow> has_start_end_points x \<or> x = 0"
using assms backward_terminating_iff1 forward_terminating_iff1 by fastforce
lemma terminating_iff2:
assumes "path x"
shows "terminating x \<longleftrightarrow> x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1) \<cdot> -(1;x\<^sup>T);x\<^sup>T\<^sup>\<star>"
using assms backward_terminating_iff2 forward_terminating_iff2 conv_compl conv_iso star_conv
by force
lemma terminating_iff3:
assumes "path x"
shows "terminating x \<longleftrightarrow> x \<le> x\<^sup>\<star>;-(x;1) \<cdot> -(1;x);x\<^sup>\<star>"
using assms backward_terminating_iff4 forward_terminating_iff3 by fastforce
lemma backward_terminating_path_irreflexive:
assumes "backward_terminating_path x"
shows "x \<le> -1'"
proof -
have 1: "x;x\<^sup>T \<le> 1'"
using assms is_inj_def path_def by blast
have "x;(x\<^sup>T \<cdot> 1') \<le> x;x\<^sup>T \<cdot> x"
by (metis inf.bounded_iff inf.commute mult_1_right mult_subdistl)
also have "... \<le> 1' \<cdot> x"
using 1 meet_iso by blast
also have "... = 1' \<cdot> x\<^sup>T"
by (metis conv_e conv_times inf.cobounded1 is_test_def test_eq_conv)
finally have 2: "x\<^sup>T;-(x\<^sup>T \<cdot> 1') \<le> -(x\<^sup>T \<cdot> 1')"
by (metis compl_le_swap1 conv_galois_1 inf.commute)
have "x\<^sup>T \<cdot> 1' \<le> x\<^sup>T;1"
by (simp add: le_infI1 maddux_20)
hence "-(x\<^sup>T;1) \<le> -(x\<^sup>T \<cdot> 1')"
using compl_mono by blast
hence "x\<^sup>T;-(x\<^sup>T \<cdot> 1') + -(x\<^sup>T;1) \<le> -(x\<^sup>T \<cdot> 1')"
using 2 by (simp add: le_supI)
hence "x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1) \<le> -(x\<^sup>T \<cdot> 1')"
by (simp add: rtc_inductl)
hence "x\<^sup>T \<cdot> 1' \<cdot> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1) = 0"
by (simp add: compl_le_swap1 galois_aux)
hence "x\<^sup>T \<cdot> 1' = 0"
using assms backward_terminating_iff3 inf.order_iff le_infI1 by blast
hence "x \<cdot> 1' = 0"
by (simp add: conv_self_conjugate)
thus ?thesis
by (simp add: galois_aux)
qed
lemma forward_terminating_path_end_points_1:
assumes "forward_terminating_path x"
shows "x \<le> x\<^sup>+;end_points x"
proof -
have 1: "-(x;1) \<cdot> x = 0"
by (simp add: galois_aux maddux_20)
have "x = x\<^sup>\<star>;-(x;1) \<cdot> x"
using assms forward_terminating_iff3 inf.absorb2 by fastforce
also have "... = (x\<^sup>+;-(x;1) + 1';-(x;1)) \<cdot> x"
by (simp add: sup.commute)
also have "... = x\<^sup>+;-(x;1) \<cdot> x + -(x;1) \<cdot> x"
using inf_sup_distrib2 by fastforce
also have "... = x\<^sup>+;-(x;1) \<cdot> x"
using 1 by simp
also have "... \<le> x\<^sup>+;(-(x;1) \<cdot> (x\<^sup>+)\<^sup>T;x)"
using modular_1_var by blast
also have "... = x\<^sup>+;(-(x;1) \<cdot> x\<^sup>T\<^sup>+;x)"
using plus_conv by fastforce
also have "... \<le> x\<^sup>+;end_points x"
by (metis inf_commute inf_top_right modular_1' mult_subdistl plus_conv plus_top)
finally show ?thesis .
qed
lemma forward_terminating_path_end_points_2:
assumes "forward_terminating_path x"
shows "x\<^sup>T \<le> x\<^sup>\<star>;end_points x"
proof -
have "x\<^sup>T \<le> x\<^sup>T;x;x\<^sup>T"
by (metis conv_invol x_leq_triple_x)
also have "... \<le> x\<^sup>T;x;1"
using mult_isol top_greatest by blast
also have "... \<le> x\<^sup>T;x\<^sup>+;end_points x;1"
by (metis assms forward_terminating_path_end_points_1 comp_assoc mult_isol mult_isor)
also have "... = x\<^sup>T;x\<^sup>+;end_points x"
by (metis inf_commute mult_assoc one_compl ra_1)
also have "... \<le> x\<^sup>\<star>;end_points x"
by (metis assms comp_assoc compl_le_swap1 conv_galois_1 conv_invol p_fun_compl path_def)
finally show ?thesis .
qed
lemma forward_terminating_path_end_points_3:
assumes "forward_terminating_path x"
shows "start_points x \<le> x\<^sup>+;end_points x"
proof -
have "start_points x \<le> x\<^sup>+;end_points x;1"
using assms forward_terminating_path_end_points_1 comp_assoc mult_isor inf.coboundedI1
by blast
also have "... = x\<^sup>+;end_points x"
by (metis inf_commute mult_assoc one_compl ra_1 )
finally show ?thesis .
qed
lemma backward_terminating_path_start_points_1:
assumes "backward_terminating_path x"
shows "x\<^sup>T \<le> x\<^sup>T\<^sup>+;start_points x"
using assms forward_terminating_path_end_points_1 conv_backward_terminating_path by fastforce
lemma backward_terminating_path_start_points_2:
assumes "backward_terminating_path x"
shows "x \<le> x\<^sup>T\<^sup>\<star>;start_points x"
using assms forward_terminating_path_end_points_2 conv_backward_terminating_path by fastforce
lemma backward_terminating_path_start_points_3:
assumes "backward_terminating_path x"
shows "end_points x \<le> x\<^sup>T\<^sup>+;start_points x"
using assms forward_terminating_path_end_points_3 conv_backward_terminating_path by fastforce
(* lemma not shown in the paper; not necessary for other theorems *)
lemma path_aux1a:
assumes "forward_terminating_path x"
shows "x \<noteq> 0 \<longleftrightarrow> end_points x \<noteq> 0"
using assms end_point_iff2 forward_terminating_iff1 end_point_iff1 galois_aux2 by force
(* lemma not shown in the paper; not necessary for other theorems *)
lemma path_aux1b:
assumes "backward_terminating_path y"
shows "y \<noteq> 0 \<longleftrightarrow> start_points y \<noteq> 0"
using assms start_point_iff2 backward_terminating_iff1 start_point_iff1 galois_aux2 by force
(* lemma not shown in the paper; not necessary for other theorems *)
lemma path_aux1:
assumes "forward_terminating_path x"
and "backward_terminating_path y"
shows "x \<noteq> 0 \<or> y \<noteq> 0 \<longleftrightarrow> end_points x \<noteq> 0 \<or> start_points y \<noteq> 0"
using assms path_aux1a path_aux1b by blast
text \<open>Equivalences for \<open>finite\<close>\<close>
lemma backward_finite_iff_msc:
"backward_finite x \<longleftrightarrow> many_strongly_connected x \<or> backward_terminating x"
proof
assume 1: "backward_finite x"
thus "many_strongly_connected x \<or> backward_terminating x"
proof (cases "-(1;x);x;1 = 0")
assume "-(1;x);x;1 = 0"
thus "many_strongly_connected x \<or> backward_terminating x"
using 1 by (metis conv_invol many_strongly_connected_iff_1 sup_bot_right)
next
assume "-(1;x);x;1 \<noteq> 0"
hence "1;-(1;x);x;1 = 1"
by (simp add: comp_assoc tarski)
hence "-(1;x);x;1 = 1"
by (metis comp_assoc conv_compl conv_contrav conv_invol conv_one one_compl)
thus "many_strongly_connected x \<or> backward_terminating x"
using 1 by simp
qed
next
assume "many_strongly_connected x \<or> backward_terminating x"
thus "backward_finite x"
by (metis star_ext sup.coboundedI1 sup.coboundedI2)
qed
lemma forward_finite_iff_msc:
"forward_finite x \<longleftrightarrow> many_strongly_connected x \<or> forward_terminating x"
by (metis backward_finite_iff_msc conv_backward_finite conv_backward_terminating conv_invol)
lemma finite_iff_msc:
"finite x \<longleftrightarrow> many_strongly_connected x \<or> terminating x"
using backward_finite_iff_msc forward_finite_iff_msc finite_iff by fastforce
text \<open>Path concatenation\<close>
lemma path_concatenation:
assumes "forward_terminating_path x"
and "backward_terminating_path y"
and "end_points x = start_points y"
and "x;1 \<cdot> (x\<^sup>T;1 + y;1) \<cdot> y\<^sup>T;1 = 0"
shows "path (x+y)"
proof (cases "y = 0")
assume "y = 0"
thus ?thesis
using assms(1) by fastforce
next
assume as: "y \<noteq> 0"
show ?thesis
proof (unfold path_def; intro conjI)
from assms(4) have a: "x;1 \<cdot> x\<^sup>T;1 \<cdot> y\<^sup>T;1 + x;1 \<cdot> y;1 \<cdot> y\<^sup>T;1= 0"
by (simp add: inf_sup_distrib1 inf_sup_distrib2)
hence aux1: "x;1 \<cdot> x\<^sup>T;1 \<cdot> y\<^sup>T;1 = 0"
using sup_eq_bot_iff by blast
from a have aux2: "x;1 \<cdot> y;1 \<cdot> y\<^sup>T;1= 0"
using sup_eq_bot_iff by blast
show "is_inj (x + y)"
proof (unfold is_inj_def; auto simp add: distrib_left)
show "x;x\<^sup>T \<le> 1'"
using assms(1) path_def is_inj_def by blast
show "y;y\<^sup>T \<le> 1'"
using assms(2) path_def is_inj_def by blast
have "y;x\<^sup>T = 0"
by (metis assms(3) aux1 annir comp_assoc conv_one le_bot modular_var_2 one_idem_mult
path_concat_aux_2 schroeder_2)
thus "y;x\<^sup>T \<le> 1'"
using bot_least le_bot by blast
thus "x;y\<^sup>T \<le> 1'"
using conv_iso by force
qed
show "is_p_fun (x + y)"
proof (unfold is_p_fun_def; auto simp add: distrib_left)
show "x\<^sup>T;x \<le> 1'"
using assms(1) path_def is_p_fun_def by blast
show "y\<^sup>T;y \<le> 1'"
using assms(2) path_def is_p_fun_def by blast
have "y\<^sup>T;x \<le> y\<^sup>T;(y;1 \<cdot> x;1)"
by (metis conjugation_prop2 inf.commute inf_top.left_neutral maddux_20 mult_isol order_trans
schroeder_1_var)
also have "... = 0"
using assms(3) aux2 annir inf_commute path_concat_aux_1 by fastforce
finally show "y\<^sup>T;x \<le> 1'"
using bot_least le_bot by blast
thus "x\<^sup>T;y \<le> 1'"
using conv_iso by force
qed
show "connected (x + y)"
proof (auto simp add: distrib_left)
have "x;1;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using assms(1) path_def by simp
also have "... \<le> (x\<^sup>\<star>;y\<^sup>\<star>)\<^sup>\<star> + (x\<^sup>T\<^sup>\<star>;y\<^sup>T\<^sup>\<star>)\<^sup>\<star>"
using join_iso join_isol star_subdist by simp
finally show "x;1;x \<le> (x\<^sup>\<star>;y\<^sup>\<star>)\<^sup>\<star> + (x\<^sup>T\<^sup>\<star>;y\<^sup>T\<^sup>\<star>)\<^sup>\<star>" .
have "y;1;y \<le> y\<^sup>\<star> + y\<^sup>T\<^sup>\<star>"
using assms(2) path_def by simp
also have "... \<le> (x\<^sup>\<star>;y\<^sup>\<star>)\<^sup>\<star> + (x\<^sup>T\<^sup>\<star>;y\<^sup>T\<^sup>\<star>)\<^sup>\<star>"
by (metis star_denest star_subdist sup.mono sup_commute)
finally show "y;1;y \<le> (x\<^sup>\<star>;y\<^sup>\<star>)\<^sup>\<star> + (x\<^sup>T\<^sup>\<star>;y\<^sup>T\<^sup>\<star>)\<^sup>\<star>" .
show "y;1;x \<le> (x\<^sup>\<star>;y\<^sup>\<star>)\<^sup>\<star> + (x\<^sup>T\<^sup>\<star>;y\<^sup>T\<^sup>\<star>)\<^sup>\<star>"
proof -
have "(y;1);1;(1;x) \<le> y\<^sup>T\<^sup>\<star>;x\<^sup>T\<^sup>\<star>"
proof (rule_tac v="start_points y" in path_concat_aux_0)
show "is_vector (start_points y)"
by (metis is_vector_def comp_assoc one_compl one_idem_mult ra_1)
show "start_points y \<noteq> 0"
using as
by (metis assms(2) conv_compl conv_contrav conv_one inf.orderE inf_bot_right
inf_top.right_neutral maddux_141)
have "(start_points y);1;y\<^sup>T \<le> y\<^sup>\<star>"
by (rule path_concat_aux_5) (simp_all add: assms(2))
thus "y;1;(start_points y)\<^sup>T \<le> y\<^sup>T\<^sup>\<star>"
by (metis (mono_tags, lifting) conv_iso comp_assoc conv_contrav conv_invol conv_one
star_conv)
have "end_points x;1;x \<le> x\<^sup>T\<^sup>\<star>"
apply (rule path_concat_aux_5)
using assms(1) conv_path by simp_all
thus "start_points y;(1;x) \<le> x\<^sup>T\<^sup>\<star>"
by (metis assms(3) mult_assoc)
qed
thus ?thesis
by (metis comp_assoc le_supI2 less_eq_def one_idem_mult star_denest star_subdist_var_1
sup.commute)
qed
show "x;1;y \<le> (x\<^sup>\<star>;y\<^sup>\<star>)\<^sup>\<star> + (x\<^sup>T\<^sup>\<star>;y\<^sup>T\<^sup>\<star>)\<^sup>\<star>"
proof -
have "(x;1);1;(1;y) \<le> x\<^sup>\<star>;y\<^sup>\<star>"
proof (rule_tac v="start_points y" in path_concat_aux_0)
show "is_vector (start_points y)"
by (simp add: comp_assoc is_vector_def one_compl vector_1_comm)
show "start_points y \<noteq> 0"
using as assms(2,4) backward_terminating_iff1 galois_aux2 start_point_iff1 start_point_iff2
by blast
have "end_points x;1;x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>"
apply (rule path_concat_aux_5)
using assms(1) conv_path by simp_all
hence "(end_points x;1;x\<^sup>T)\<^sup>T \<le> (x\<^sup>T\<^sup>\<star>)\<^sup>T"
using conv_iso by blast
thus "x;1;(start_points y)\<^sup>T \<le> x\<^sup>\<star>"
by (simp add: assms(3) comp_assoc star_conv)
have "start_points y;1;y \<le> y\<^sup>\<star>"
by (rule path_concat_aux_5) (simp_all add: assms(2))
thus "start_points y;(1;y) \<le> y\<^sup>\<star>"
by (simp add: mult_assoc)
qed
thus ?thesis
by (metis comp_assoc dual_order.trans le_supI1 one_idem_mult star_ext)
qed
qed
qed
qed
lemma path_concatenation_with_edge:
assumes "x\<noteq>0"
and "forward_terminating_path x"
and "is_point q"
and "q \<le> -(1;x)"
shows "path (x+(end_points x);q\<^sup>T)"
proof (rule path_concatenation)
from assms(1,2) have 1: "is_point(end_points x)"
using end_point_zero_point path_aux1a by blast
show 2: "backward_terminating_path ((end_points x);q\<^sup>T)"
apply (intro conjI)
apply (metis edge_is_path 1 assms(3))
by (metis assms(2-4) 1 bot_least comp_assoc compl_le_swap1 conv_galois_2 double_compl
end_point_iff1 le_supE point_equations(1) tarski top_le)
thus "end_points x = start_points ((end_points x);q\<^sup>T)"
by (metis assms(3) 1 edge_start comp_assoc compl_top_eq double_compl inf.absorb_iff2 inf.commute
inf_top_right modular_2_aux' point_equations(2))
show "x;1 \<cdot> (x\<^sup>T;1 + ((end_points x);q\<^sup>T);1) \<cdot> ((end_points x);q\<^sup>T)\<^sup>T;1 = 0"
using 2 by (metis assms(3,4) annir compl_le_swap1 compl_top_eq conv_galois_2 double_compl
inf.absorb_iff2 inf.commute modular_1' modular_2_aux' point_equations(2))
show "forward_terminating_path x"
by (simp add: assms(2))
qed
lemma path_concatenation_cycle_free:
assumes "forward_terminating_path x"
and "backward_terminating_path y"
and "end_points x = start_points y"
and "x;1 \<cdot> y\<^sup>T;1 = 0"
shows "path (x+y)"
apply (rule path_concatenation,simp_all add: assms)
by (metis assms(4) inf.left_commute inf_bot_right inf_commute)
lemma path_concatenation_start_points_approx:
assumes "end_points x = start_points y"
shows "start_points (x+y) \<le> start_points x"
proof -
have "start_points (x+y) = x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1) + y;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1)"
by (simp add: inf.assoc inf_sup_distrib2)
also with assms(1) have "... = x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1) + x\<^sup>T;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(x;1)"
by (metis inf.assoc inf.left_commute)
also have "... = x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1)"
by simp
also have "... \<le> start_points x"
using inf_le1 by blast
finally show ?thesis .
qed
lemma path_concatenation_end_points_approx:
assumes "end_points x = start_points y"
shows "end_points (x+y) \<le> end_points y"
proof -
have "end_points (x+y) = x\<^sup>T;1 \<cdot> -(x;1) \<cdot> -(y;1) + y\<^sup>T;1 \<cdot> -(x;1) \<cdot> -(y;1)"
by (simp add: inf.assoc inf_sup_distrib2)
also from assms(1) have "... = y;1 \<cdot> -(y\<^sup>T;1) \<cdot> -(y;1) + y\<^sup>T;1 \<cdot> -(x;1) \<cdot> -(y;1)"
by simp
also have "... = y\<^sup>T;1 \<cdot> -(x;1) \<cdot> -(y;1)"
by (simp add: inf.commute)
also have "... \<le> end_points y"
using inf_le1 meet_iso by blast
finally show ?thesis .
qed
lemma path_concatenation_start_points:
assumes "end_points x = start_points y"
and "x;1 \<cdot> y\<^sup>T;1 = 0"
shows "start_points (x+y) = start_points x"
proof -
from assms(2) have aux: "x;1 \<cdot> -(y\<^sup>T;1) = x;1"
by (simp add: galois_aux inf.absorb1)
have "start_points (x+y) = (x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1)) + (y;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1))"
by (simp add: inf_sup_distrib2 inf.assoc)
also from assms(1) have "... = (x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1)) + (x\<^sup>T;1 \<cdot> -(x;1) \<cdot> -(x\<^sup>T;1))"
using inf.assoc inf.commute by simp
also have "... = (x;1 \<cdot> -(x\<^sup>T;1) \<cdot> -(y\<^sup>T;1))"
by (simp add: inf.assoc)
also from aux have "... = x;1 \<cdot> -(x\<^sup>T;1)"
by (metis inf.assoc inf.commute)
finally show ?thesis .
qed
lemma path_concatenation_end_points:
assumes "end_points x = start_points y"
and "x;1 \<cdot> y\<^sup>T;1 = 0"
shows "end_points (x+y) = end_points y"
proof -
from assms(2) have aux: "y\<^sup>T;1 \<cdot> -(x;1) = y\<^sup>T;1"
using galois_aux inf.absorb1 inf_commute by blast
have "end_points (x+y) = (x\<^sup>T;1 + y\<^sup>T;1) \<cdot> -(x;1) \<cdot> -(y;1)"
using inf.assoc by simp
also from assms(1) have "... = (y;1 \<cdot> -(y\<^sup>T;1) \<cdot> -(y;1)) + (y\<^sup>T;1 \<cdot> -(x;1) \<cdot> -(y;1))"
by (simp add: inf_sup_distrib2)
also have "... = y\<^sup>T;1 \<cdot> -(x;1) \<cdot> -(y;1)"
by (simp add: inf.assoc)
also from aux have "... = y\<^sup>T;1 \<cdot> -(y;1)"
by (metis inf.assoc inf.commute)
finally show ?thesis .
qed
lemma path_concatenation_cycle_free_complete:
assumes "forward_terminating_path x"
and "backward_terminating_path y"
and "end_points x = start_points y"
and "x;1 \<cdot> y\<^sup>T;1 = 0"
shows "path (x+y) \<and> start_points (x+y) = start_points x \<and> end_points (x+y) = end_points y"
using assms path_concatenation_cycle_free path_concatenation_end_points path_concatenation_start_points
by blast
text \<open>Path restriction (path from a given point)\<close>
lemma reachable_points_iff:
assumes "point p"
shows "(x\<^sup>T\<^sup>\<star>;p \<cdot> x) = (x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x"
-proof (rule antisym)
+proof (rule order.antisym)
show "(x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x \<le> x\<^sup>T\<^sup>\<star>;p \<cdot> x"
proof (rule le_infI)
show "(x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x \<le> x\<^sup>T\<^sup>\<star>;p"
proof -
have "(x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x \<le> x\<^sup>T\<^sup>\<star>;p;1"
by (simp add: mult_isol_var)
also have "... \<le> x\<^sup>T\<^sup>\<star>;p"
- using assms by (simp add: comp_assoc eq_iff point_equations(1) point_is_point)
+ using assms by (simp add: comp_assoc order.eq_iff point_equations(1) point_is_point)
finally show ?thesis .
qed
show "(x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x \<le> x"
by (metis inf_le2 mult_isor mult_onel)
qed
show "x\<^sup>T\<^sup>\<star>;p \<cdot> x \<le> (x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x"
proof -
have "(x\<^sup>T\<^sup>\<star>;p);x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>;p + -1'"
by (metis assms comp_assoc is_vector_def mult_isol point_def sup.coboundedI1 top_greatest)
hence aux: "(-(x\<^sup>T\<^sup>\<star>;p) \<cdot> 1');x \<le> -(x\<^sup>T\<^sup>\<star>;p)"
using compl_mono conv_galois_2 by fastforce
have "x = (x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x + (-(x\<^sup>T\<^sup>\<star>;p) \<cdot> 1');x"
by (metis aux4 distrib_right inf_commute mult_1_left)
also with aux have "... \<le> (x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x + -(x\<^sup>T\<^sup>\<star>;p)"
using join_isol by blast
finally have "x \<le> (x\<^sup>T\<^sup>\<star>;p \<cdot> 1');x + -(x\<^sup>T\<^sup>\<star>;p)" .
thus ?thesis
using galois_2 inf.commute by fastforce
qed
qed
lemma path_from_given_point:
assumes "path x"
and "point p"
shows "path(x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
and "start_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) \<le> p"
and "end_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) \<le> end_points(x)"
proof (unfold path_def; intro conjI)
show uni: "is_p_fun (x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
by (metis assms(1) inf_commute is_p_fun_def p_fun_mult_var path_def)
show inj: "is_inj (x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
by (metis abel_semigroup.commute assms(1) conv_times inf.abel_semigroup_axioms inj_p_fun
is_p_fun_def p_fun_mult_var path_def)
show "connected (x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
proof -
let ?t="x\<^sup>T\<^sup>\<star>;p \<cdot> 1'"
let ?u="-(x\<^sup>T\<^sup>\<star>;p) \<cdot> 1'"
(* some aux statements about ?t and ?u *)
have t_plus_u: "?t + ?u = 1'"
by (simp add: inf.commute)
have t_times_u: "?t ; ?u \<le> 0"
by (simp add: inf.left_commute is_test_def test_comp_eq_mult)
have t_conv: "?t\<^sup>T=?t"
using inf.cobounded2 is_test_def test_eq_conv by blast
have txu_zero: "?t;x;?u \<le> 0"
proof -
have "x\<^sup>T;?t;1 \<le> -?u"
proof -
have "x\<^sup>T;?t;1 \<le> x\<^sup>T;x\<^sup>T\<^sup>\<star>;p"
using assms(2)
by (simp add: is_vector_def mult.semigroup_axioms mult_isol_var mult_subdistr order.refl
point_def semigroup.assoc)
also have "... \<le> -?u"
by (simp add: le_supI1 mult_isor)
finally show ?thesis .
qed
thus ?thesis
by (metis compl_bot_eq compl_le_swap1 conv_contrav conv_galois_1 t_conv)
qed
hence txux_zero: "?t;x;?u;x \<le> 0"
using annil le_bot by fastforce
(* end some aux statements about ?t and ?u *)
have tx_leq: "?t;x\<^sup>\<star> \<le> (?t;x)\<^sup>\<star>"
proof -
have "?t;x\<^sup>\<star> = ?t;(?t;x + ?u;x)\<^sup>\<star>"
using t_plus_u by (metis distrib_right' mult_onel)
also have "... = ?t;(?u;x;(?u;x)\<^sup>\<star>;(?t;x)\<^sup>\<star>+(?t;x)\<^sup>\<star>)"
using txux_zero star_denest_10 by (simp add: comp_assoc le_bot)
also have "... = ?t;?u;x;(?u;x)\<^sup>\<star>;(?t;x)\<^sup>\<star>+?t;(?t;x)\<^sup>\<star>"
by (simp add: comp_assoc distrib_left)
also have "... \<le> 0;x;(?u;x)\<^sup>\<star>;(?t;x)\<^sup>\<star>+?t;(?t;x)\<^sup>\<star>"
using le_bot t_times_u by blast
also have "... \<le>(?t;x)\<^sup>\<star>"
by (metis annil inf.commute inf_bot_right le_supI mult_onel mult_subdistr)
finally show ?thesis .
qed
hence aux: "?t;x\<^sup>\<star>;?t \<le> (?t;x)\<^sup>\<star>"
using inf.cobounded2 order.trans prod_star_closure star_ref by blast
with t_conv have aux_trans: "?t;x\<^sup>T\<^sup>\<star>;?t \<le> (?t;x)\<^sup>T\<^sup>\<star>"
by (metis comp_assoc conv_contrav conv_self_conjugate_var g_iso star_conv)
from aux aux_trans have "?t;(x\<^sup>\<star>+x\<^sup>T\<^sup>\<star>);?t \<le> (?t;x)\<^sup>\<star> + (?t;x)\<^sup>T\<^sup>\<star>"
by (metis sup_mono distrib_right' distrib_left)
with assms(1) path_concat_aux3_1 have "?t;(x;1;x\<^sup>T);?t \<le> (?t;x)\<^sup>\<star> + (?t;x)\<^sup>T\<^sup>\<star>"
using dual_order.trans mult_double_iso by blast
with t_conv have "(?t;x);1;(?t;x)\<^sup>T \<le> (?t;x)\<^sup>\<star> + (?t;x)\<^sup>T\<^sup>\<star>"
using comp_assoc conv_contrav by fastforce
with connected_iff2 show ?thesis
using assms(2) inj reachable_points_iff uni by fastforce
qed
next
show "start_points (x\<^sup>T\<^sup>\<star>;p \<cdot> x) \<le> p"
proof -
have 1: "is_vector (x\<^sup>T\<^sup>\<star>;p)"
using assms(2) by (simp add: is_vector_def mult_assoc point_def)
hence "(x\<^sup>T\<^sup>\<star>;p \<cdot> x);1 \<le> x\<^sup>T\<^sup>\<star>;p"
by (simp add: inf.commute vector_1_comm)
also have "... = x\<^sup>T\<^sup>+;p + p"
by (simp add: sup.commute)
finally have 2: "(x\<^sup>T\<^sup>\<star>;p \<cdot> x);1 \<cdot> -(x\<^sup>T\<^sup>+;p) \<le> p"
using galois_1 by blast
have "(x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 = (x\<^sup>T \<cdot> (x\<^sup>T\<^sup>\<star>;p)\<^sup>T);1"
by (simp add: inf.commute)
also have "... = x\<^sup>T;(x\<^sup>T\<^sup>\<star>;p \<cdot> 1)"
using 1 vector_2 by blast
also have "... = x\<^sup>T\<^sup>+;p"
by (simp add: comp_assoc)
finally show "start_points (x\<^sup>T\<^sup>\<star>;p \<cdot> x) \<le> p"
using 2 by simp
qed
next
show "end_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) \<le> end_points(x)"
proof -
have 1: "is_vector (x\<^sup>T\<^sup>\<star>;p)"
using assms(2) by (simp add: is_vector_def mult_assoc point_def)
have "(x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 = ((x\<^sup>T\<^sup>\<star>;p)\<^sup>T \<cdot> x\<^sup>T);1"
by (simp add: star_conv)
also have "... = x\<^sup>T;(x\<^sup>T\<^sup>\<star>;p \<cdot> 1)"
using 1 vector_2 inf.commute by fastforce
also have "... \<le> x\<^sup>T\<^sup>\<star>;p"
using comp_assoc mult_isor by fastforce
finally have 2: "(x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 \<cdot> -(x\<^sup>T\<^sup>\<star>;p) = 0"
using galois_aux2 by blast
have "(x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 \<cdot> -((x\<^sup>T\<^sup>\<star>;p \<cdot> x);1) = (x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 \<cdot> (-(x\<^sup>T\<^sup>\<star>;p) + -(x;1))"
using 1 vector_1 by fastforce
also have "... = (x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 \<cdot> -(x\<^sup>T\<^sup>\<star>;p) + (x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 \<cdot> -(x;1)"
using inf_sup_distrib1 by blast
also have "... = (x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 \<cdot> -(x;1)"
using 2 by simp
also have "... \<le> x\<^sup>T;1 \<cdot> -(x;1)"
using meet_iso mult_subdistr_var by fastforce
finally show ?thesis .
qed
qed
lemma path_from_given_point':
assumes "has_start_points_path x"
and "point p"
and "p \<le> x;1" (* p has a successor hence path not empty *)
shows "path(x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
and "start_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) = p"
and "end_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) = end_points(x)"
proof -
show "path(x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
using assms path_from_given_point(1) by blast
next
show "start_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) = p"
- proof (simp only: eq_iff; rule conjI)
+ proof (simp only: order.eq_iff; rule conjI)
show "start_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) \<le> p"
using assms path_from_given_point(2) by blast
show "p \<le> start_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
proof -
have 1: "is_vector(x\<^sup>T\<^sup>\<star>;p)"
using assms(2) comp_assoc is_vector_def point_equations(1) point_is_point by fastforce
hence a: "p \<le> (x\<^sup>T\<^sup>\<star>;p \<cdot> x);1"
by (metis vector_1 assms(3) conway.dagger_unfoldl_distr inf.orderI inf_greatest
inf_sup_absorb)
have "x\<^sup>T\<^sup>+;p \<cdot> p \<le> (x\<^sup>T\<^sup>+ \<cdot> 1'); p"
using assms(2) inj_distr point_def by fastforce
also have "... \<le> (-1'\<^sup>T \<cdot> 1'); p"
using assms(1) path_acyclic
by (metis conv_contrav conv_e meet_iso mult_isor star_conv star_slide_var test_converse)
also have "... \<le> 0"
by simp
finally have 2: "x\<^sup>T\<^sup>+;p \<cdot> p \<le> 0" .
have b: "p \<le> -((x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1)"
proof -
have "(x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1 = ((x\<^sup>T\<^sup>\<star>;p)\<^sup>T \<cdot> x\<^sup>T);1"
by (simp add: star_conv)
also have "... = x\<^sup>T;(x\<^sup>T\<^sup>\<star>;p \<cdot> 1)"
using 1 vector_2 inf.commute by fastforce
also have "... = x\<^sup>T;x\<^sup>T\<^sup>\<star>;p"
by (simp add: comp_assoc)
also have "... \<le> -p"
using 2 galois_aux le_bot by blast
finally show ?thesis
using compl_le_swap1 by blast
qed
with a show ?thesis
by simp
qed
qed
next
show "end_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) = end_points(x)"
- proof (simp only: eq_iff; rule conjI)
+ proof (simp only: order.eq_iff; rule conjI)
show "end_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x) \<le> end_points(x)"
using assms path_from_given_point(3) by blast
show "end_points(x) \<le> end_points(x\<^sup>T\<^sup>\<star>;p \<cdot> x)"
proof -
have 1: "is_vector(x\<^sup>T\<^sup>\<star>;p)"
using assms(2) comp_assoc is_vector_def point_equations(1) point_is_point by fastforce
have 2: "is_vector(end_points(x))"
by (simp add: comp_assoc is_vector_def one_compl vector_1_comm)
have a: "end_points(x) \<le> (x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1"
proof -
have "x\<^sup>T;1 \<cdot> 1;x\<^sup>T = x\<^sup>T;1;x\<^sup>T"
by (simp add: vector_meet_comp_x')
also have "... \<le> x\<^sup>T\<^sup>\<star> + x\<^sup>\<star>"
using assms(1) path_concat_aux3_3 sup.commute by fastforce
also have "... = x\<^sup>T\<^sup>\<star> + x\<^sup>+"
by (simp add: star_star_plus sup.commute)
also have "... \<le> x\<^sup>T\<^sup>\<star> + x;1"
using join_isol mult_isol by fastforce
finally have "end_points(x) \<cdot> 1;x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>"
by (metis galois_1 inf.assoc inf.commute sup_commute)
hence "end_points(x) \<cdot> p\<^sup>T \<le> x\<^sup>T\<^sup>\<star>"
using assms(3)
by (metis conv_contrav conv_iso conv_one dual_order.trans inf.cobounded1 inf.right_idem
inf_mono)
hence "end_points(x) ; p\<^sup>T \<le> x\<^sup>T\<^sup>\<star>"
using assms(2) 2 by (simp add: point_def vector_meet_comp)
hence "end_points(x) \<le> x\<^sup>T\<^sup>\<star>;p"
using assms(2) point_def ss423bij by blast
hence "x\<^sup>T;1 \<le> x\<^sup>T\<^sup>\<star>;p + x;1"
by (simp add: galois_1 sup_commute)
hence "x\<^sup>T;1 \<le> x\<^sup>T\<^sup>+;p + p + x;1"
by (metis conway.dagger_unfoldl_distr sup_commute)
hence "x\<^sup>T;1 \<le> x\<^sup>T\<^sup>+;p + x;1"
by (simp add: assms(3) sup.absorb2 sup.assoc)
hence "end_points(x) \<le> x\<^sup>T\<^sup>+;p"
by (simp add: galois_1 sup_commute)
also have "... = (x\<^sup>T\<^sup>\<star>;p \<cdot> x)\<^sup>T;1"
using 1 inf_commute mult_assoc vector_2 by fastforce
finally show ?thesis .
qed
have "x\<^sup>T;1 \<cdot> (x\<^sup>T\<^sup>\<star>;p \<cdot> x);1 \<le> x;1"
by (simp add: le_infI2 mult_isor)
hence b: "end_points(x) \<le> -((x\<^sup>T\<^sup>\<star>;p \<cdot> x);1)"
using galois_1 galois_2 by blast
with a show ?thesis
by simp
qed
qed
qed
text \<open>Cycles\<close>
lemma selfloop_is_cycle:
assumes "is_point x"
shows "cycle (x;x\<^sup>T)"
by (simp add: assms edge_is_path)
lemma start_point_no_cycle:
assumes "has_start_points_path x"
shows "\<not> cycle x"
using assms many_strongly_connected_implies_no_start_end_points no_start_end_points_iff
start_point_iff1 start_point_iff2 by blast
lemma end_point_no_cycle:
assumes "has_end_points_path x"
shows "\<not> cycle x"
using assms end_point_iff2 end_point_iff1 many_strongly_connected_implies_no_start_end_points
no_start_end_points_iff by blast
lemma cycle_no_points:
assumes "cycle x"
shows "start_points x = 0"
and "end_points x = 0"
by (metis assms inf_compl_bot many_strongly_connected_implies_no_start_end_points)+
text \<open>Path concatenation to cycle\<close>
lemma path_path_equals_cycle_aux:
assumes "has_start_end_points_path x"
and "has_start_end_points_path y"
and "start_points x = end_points y"
and "end_points x = start_points y"
shows "x \<le> (x+y)\<^sup>T\<^sup>\<star>"
proof-
let ?e = "end_points(x)"
let ?s = "start_points(x)"
have sp: "is_point(?s)"
using assms(1) start_point_iff2 has_start_end_points_path_iff by blast
have ep: "is_point(?e)"
using assms(1) end_point_iff2 has_start_end_points_path_iff by blast
have "x \<le> x\<^sup>T\<^sup>\<star>;?s;1 \<cdot> 1;?e\<^sup>T;x\<^sup>T\<^sup>\<star>"
by (metis assms(1) backward_terminating_path_start_points_2 end_point_iff2 ep
forward_terminating_iff1 forward_terminating_path_end_points_2 comp_assoc
conv_contrav conv_invol conv_iso inf.boundedI point_equations(1) point_equations(4)
star_conv sp start_point_iff2)
also have "... = x\<^sup>T\<^sup>\<star>;?s;1;?e\<^sup>T;x\<^sup>T\<^sup>\<star>"
by (metis inf_commute inf_top_right ra_1)
also have "... = x\<^sup>T\<^sup>\<star>;?s;?e\<^sup>T;x\<^sup>T\<^sup>\<star>"
by (metis ep comp_assoc point_equations(4))
also have "... \<le> x\<^sup>T\<^sup>\<star>;y\<^sup>T\<^sup>\<star>;x\<^sup>T\<^sup>\<star>"
by (metis (mono_tags, lifting) assms(2-4) start_to_end comp_assoc conv_contrav conv_invol
conv_iso mult_double_iso star_conv)
also have "... = (x\<^sup>\<star>;y\<^sup>\<star>;x\<^sup>\<star>)\<^sup>T"
by (simp add: comp_assoc star_conv)
also have "... \<le> ((x+y)\<^sup>\<star>;(x+y)\<^sup>\<star>;(x+y)\<^sup>\<star>)\<^sup>T"
by (metis conv_invol conv_iso prod_star_closure star_conv star_denest star_ext star_iso
star_trans_eq sup_ge1)
also have "... = (x+y)\<^sup>T\<^sup>\<star>"
by (metis star_conv star_trans_eq)
finally show x: "x \<le> (x+y)\<^sup>T\<^sup>\<star>" .
qed
lemma path_path_equals_cycle:
assumes "has_start_end_points_path x"
and "has_start_end_points_path y"
and "start_points x = end_points y"
and "end_points x = start_points y"
and "x;1 \<cdot> (x\<^sup>T;1 + y;1) \<cdot> y\<^sup>T;1 = 0"
shows "cycle(x + y)"
proof (intro conjI)
show "path (x + y)"
apply (rule path_concatenation)
using assms by(simp_all add:has_start_end_points_iff)
show "many_strongly_connected (x + y)"
by (metis path_path_equals_cycle_aux assms(1-4) sup.commute le_supI many_strongly_connected_iff_3)
qed
lemma path_edge_equals_cycle:
assumes "has_start_end_points_path x"
shows "cycle(x + end_points(x);(start_points x)\<^sup>T)"
proof (rule path_path_equals_cycle)
let ?s = "start_points x"
let ?e = "end_points x"
let ?y = "(?e;?s\<^sup>T)"
have sp: "is_point(?s)"
using start_point_iff2 assms has_start_end_points_path_iff by blast
have ep: "is_point(?e)"
using end_point_iff2 assms has_start_end_points_path_iff by blast
show "has_start_end_points_path x"
using assms by blast
show "has_start_end_points_path ?y"
using edge_is_path
by (metis assms edge_end edge_start end_point_iff2 end_point_iff1 galois_aux2
has_start_end_points_iff inf.left_idem inf_compl_bot_right start_point_iff2)
show "?s = end_points ?y"
by (metis sp ep edge_end annil conv_zero inf.left_idem inf_compl_bot_right)
thus "?e = start_points ?y"
by (metis edge_start ep conv_contrav conv_invol sp)
show "x;1 \<cdot> (x\<^sup>T;1 + ?e;?s\<^sup>T;1) \<cdot> (?e;?s\<^sup>T)\<^sup>T;1 = 0"
proof -
have "x;1 \<cdot> (x\<^sup>T;1 + ?e;?s\<^sup>T;1) \<cdot> (?e;?s\<^sup>T)\<^sup>T;1 = x;1 \<cdot> (x\<^sup>T;1 + ?e;1;?s\<^sup>T;1) \<cdot> (?s;?e\<^sup>T);1"
using sp comp_assoc point_equations(3) by fastforce
also have "... = x;1 \<cdot> (x\<^sup>T;1 + ?e;1) \<cdot> ?s;1"
by (metis sp ep comp_assoc point_equations(1,3))
also have "... \<le> 0"
by (simp add: sp ep inf.assoc point_equations(1))
finally show ?thesis
using bot_unique by blast
qed
qed
text \<open>Break cycles\<close>
lemma cycle_remove_edge:
assumes "cycle x"
and "point s"
and "point e"
and "e;s\<^sup>T \<le> x"
shows "path(x \<cdot> -(e;s\<^sup>T))"
and "start_points (x \<cdot> -(e;s\<^sup>T)) \<le> s"
and "end_points (x \<cdot> -(e;s\<^sup>T)) \<le> e"
proof -
show "path(x \<cdot> -(e;s\<^sup>T))"
proof (unfold path_def; intro conjI)
show 1: "is_p_fun(x \<cdot> -(e;s\<^sup>T))"
using assms(1) path_def is_p_fun_def p_fun_mult_var by blast
show 2: "is_inj(x \<cdot> -(e;s\<^sup>T))"
using assms(1) path_def inf.cobounded1 injective_down_closed by blast
show "connected (x \<cdot> -(e;s\<^sup>T))"
proof -
have "x\<^sup>\<star> = ((x \<cdot> -(e;s\<^sup>T)) + e;s\<^sup>T)\<^sup>\<star>"
by (metis assms(4) aux4_comm inf.absorb2)
also have "... = (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; (e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>)\<^sup>\<star>"
by simp
also have "... = (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; (1' + e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>;(e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>)\<^sup>\<star>)"
by fastforce
also have "... = (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>;(e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>)\<^sup>\<star>"
by (simp add: distrib_left mult_assoc)
also have "... = (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; e;(s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>;e)\<^sup>\<star>;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>"
by (simp add: comp_assoc star_slide)
also have "... \<le> (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; e;1;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>"
using top_greatest join_isol mult_double_iso by (metis mult_assoc)
also have "... = (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>"
using assms(3) by (simp add: comp_assoc is_vector_def point_def)
finally have 3: "x\<^sup>\<star> \<le> (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>" .
from assms(4) have "e;s\<^sup>T \<le> e;e\<^sup>T;x"
using assms(3) comp_assoc mult_isol point_def ss423conv by fastforce
also have "... \<le> e;e\<^sup>T;(x\<^sup>\<star>)\<^sup>T"
using assms(1) many_strongly_connected_iff_3 mult_isol star_conv by fastforce
also have "... \<le> e;e\<^sup>T;((x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>)\<^sup>T"
using 3 conv_iso mult_isol by blast
also have "... \<le> e;e\<^sup>T;((x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star> ; s;e\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>)"
by (simp add: star_conv comp_assoc)
also have "... \<le> e;e\<^sup>T;(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star> + e;e\<^sup>T;(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star> ; s;e\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>"
by (simp add: comp_assoc distrib_left)
also have "... \<le> e;e\<^sup>T;(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star> + e;1;e\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>"
by (metis comp_assoc join_isol mult_isol mult_isor top_greatest)
also have "... \<le> e;e\<^sup>T;(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star> + e;e\<^sup>T;(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>"
using assms(3) by (simp add: point_equations(1) point_is_point)
also have "... = e;e\<^sup>T;(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>"
by simp
also have "... \<le> 1';(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>"
using assms(3) is_inj_def point_def join_iso mult_isor by blast
finally have 4: "e;s\<^sup>T \<le>(x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>"
by simp
have "(x \<cdot> -(e;s\<^sup>T));1;(x \<cdot> -(e;s\<^sup>T)) \<le> x;1;x"
by (simp add: mult_isol_var)
also have "...\<le> x\<^sup>\<star>"
using assms(1) connected_iff4 one_strongly_connected_iff one_strongly_connected_implies_8
path_concat_aux3_3 by blast
also have "... \<le> (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; e;s\<^sup>T ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>"
by (rule 3)
also have "... \<le> (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> ; (x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star> ; (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star>"
using 4 by (metis comp_assoc join_isol mult_isol mult_isor)
also have "... \<le> (x \<cdot> -(e;s\<^sup>T))\<^sup>\<star> + (x \<cdot> -(e;s\<^sup>T))\<^sup>T\<^sup>\<star>"
using 1 2 triple_star by force
finally show ?thesis .
qed
qed
next
show "start_points (x \<cdot> -(e;s\<^sup>T)) \<le> s"
proof -
have 1: "is_vector(-s)"
using assms(2) by (simp add: point_def vector_compl)
have "(x \<cdot> -(e;s\<^sup>T));1 \<cdot> -s \<le> x;1 \<cdot> -s"
using meet_iso mult_subdistr by blast
also have "... \<le> x\<^sup>T;1 \<cdot> -s"
using assms(1) many_strongly_connected_implies_no_start_end_points meet_iso
no_start_end_points_path_iff by blast
also have "... \<le> (x\<^sup>T \<cdot> -s);1"
using 1 by (simp add: vector_1_comm)
also have "... \<le> (x\<^sup>T \<cdot> -(s;e\<^sup>T));1"
by (metis 1 galois_aux inf.boundedI inf.cobounded1 inf.commute mult_isor schroeder_2
vector_1_comm)
also have "... = (x \<cdot> -(e;s\<^sup>T))\<^sup>T;1"
by (simp add: conv_compl)
finally show ?thesis
by (simp add: galois_1 sup_commute)
qed
next
show "end_points (x \<cdot> -(e;s\<^sup>T)) \<le> e"
proof -
have 1: "is_vector(-e)"
using assms(3) by (simp add: point_def vector_compl)
have "(x \<cdot> -(e;s\<^sup>T))\<^sup>T;1 \<cdot> -e \<le> x\<^sup>T;1 \<cdot> -e"
using meet_iso mult_subdistr by simp
also have "... \<le> x;1 \<cdot> -e"
using assms(1) many_strongly_connected_implies_no_start_end_points meet_iso
no_start_end_points_path_iff by blast
also have "... \<le> (x \<cdot> -e);1"
using 1 by (simp add: vector_1_comm)
also have "... \<le> (x \<cdot> -(e;s\<^sup>T));1"
by (metis 1 galois_aux inf.boundedI inf.cobounded1 inf.commute mult_isor schroeder_2
vector_1_comm)
finally show ?thesis
by (simp add: galois_1 sup_commute)
qed
qed
lemma cycle_remove_edge':
assumes "cycle x"
and "point s"
and "point e"
and "s\<noteq>e"
and "e;s\<^sup>T \<le> x"
shows "path(x \<cdot> -(e;s\<^sup>T))"
and "s = start_points (x \<cdot> -(e;s\<^sup>T))"
and "e = end_points (x \<cdot> -(e;s\<^sup>T))"
proof -
show "path (x \<cdot> - (e ; s\<^sup>T))"
using assms(1,2,3,5) cycle_remove_edge(1) by blast
next
show "s = start_points (x \<cdot> - (e ; s\<^sup>T))"
- proof (simp only: eq_iff; rule conjI)
+ proof (simp only: order.eq_iff; rule conjI)
show "s \<le> start_points (x \<cdot> - (e ; s\<^sup>T))"
proof -
have a: "s \<le> (x \<cdot> - (e ; s\<^sup>T));1"
proof -
have 1: "is_vector(-e)"
using assms(3) point_def vector_compl by blast
from assms(2-4) have "s = s \<cdot> -e"
using comp_assoc edge_end point_equations(1) point_equations(3) point_is_point by fastforce
also have "... \<le> x\<^sup>T;e \<cdot> -e"
using assms(3,5) conv_iso meet_iso point_def ss423conv by fastforce
also have "... \<le> x;1 \<cdot> -e"
by (metis assms(1) many_strongly_connected_implies_no_start_end_points meet_iso mult_isol
top_greatest)
also have "... \<le> (x \<cdot> -e);1"
using 1 by (simp add: vector_1_comm)
also have "... \<le> (x \<cdot> - (e ; s\<^sup>T));1"
by (metis assms(3) comp_anti is_vector_def meet_isor mult_isol mult_isor point_def
top_greatest)
finally show ?thesis .
qed
have b: "s \<le> -((x \<cdot> - (e ; s\<^sup>T))\<^sup>T;1)"
proof -
have 1: "x;s =e"
using assms predecessor_point' by blast
have "s \<cdot> x\<^sup>T = s;(e\<^sup>T+-(e\<^sup>T)) \<cdot> x\<^sup>T"
using assms(2) point_equations(1) point_is_point by fastforce
also have "... = s;e\<^sup>T \<cdot> x\<^sup>T"
by (metis 1 conv_contrav inf.commute inf_sup_absorb modular_1')
also have "... \<le> e\<^sup>T"
by (metis assms(3) inf.coboundedI1 mult_isor point_equations(4) point_is_point
top_greatest)
finally have "s \<cdot> x\<^sup>T \<le> s \<cdot> e\<^sup>T"
by simp
also have "... \<le> s ; e\<^sup>T"
using assms(2,3) by (simp add: point_def vector_meet_comp)
finally have 2: "s \<cdot> x\<^sup>T \<cdot> -(s ; e\<^sup>T) = 0"
using galois_aux2 by blast
thus ?thesis
proof -
have "s ; e\<^sup>T = e\<^sup>T \<cdot> s"
using assms(2,3) inf_commute point_def vector_meet_comp by force
thus ?thesis
using 2
by (metis assms(2,3) conv_compl conv_invol conv_one conv_times galois_aux
inf.assoc point_def point_equations(1) point_is_point schroeder_2
vector_meet_comp)
qed
qed
with a show ?thesis
by simp
qed
show "start_points (x \<cdot> - (e ; s\<^sup>T)) \<le> s"
using assms(1,2,3,5) cycle_remove_edge(2) by blast
qed
next
show "e = end_points (x \<cdot> - (e ; s\<^sup>T))"
- proof (simp only: eq_iff; rule conjI)
+ proof (simp only: order.eq_iff; rule conjI)
show "e \<le> end_points (x \<cdot> - (e ; s\<^sup>T))"
(* just copied and adapted the proof of the previous case (start_point) *)
proof -
have a: "e \<le> (x \<cdot> - (e ; s\<^sup>T))\<^sup>T;1"
proof -
have 1: "is_vector(-s)"
using assms(2) point_def vector_compl by blast
from assms(2-4) have "e = e \<cdot> -s"
using comp_assoc edge_end point_equations(1) point_equations(3) point_is_point by fastforce
also have "... \<le> x;s \<cdot> -s"
using assms(2,5) meet_iso point_def ss423bij by fastforce
also have "... \<le> x\<^sup>T;1 \<cdot> -s"
by (metis assms(1) many_strongly_connected_implies_no_start_end_points meet_iso mult_isol
top_greatest)
also have "... \<le> (x\<^sup>T \<cdot> -s);1"
using 1 by (simp add: vector_1_comm)
also have "... \<le> (x\<^sup>T \<cdot> - (s ; e\<^sup>T));1"
by (metis assms(2) comp_anti is_vector_def meet_isor mult_isol mult_isor point_def
top_greatest)
finally show ?thesis
by (simp add: conv_compl)
qed
have b: "e \<le> -((x \<cdot> - (e ; s\<^sup>T));1)"
proof -
have 1: "x\<^sup>T;e =s"
using assms predecessor_point' by (metis conv_contrav conv_invol conv_iso conv_path)
have "e \<cdot> x = e;(s\<^sup>T+-(s\<^sup>T)) \<cdot> x"
using assms(3) point_equations(1) point_is_point by fastforce
also have "... = e;s\<^sup>T \<cdot> x"
by (metis 1 conv_contrav conv_invol inf.commute inf_sup_absorb modular_1')
also have "... \<le> s\<^sup>T"
by (metis assms(2) inf.coboundedI1 mult_isor point_equations(4) point_is_point top_greatest)
finally have "e \<cdot> x \<le> e \<cdot> s\<^sup>T"
by simp
also have "... \<le> e ; s\<^sup>T"
using assms(2,3) by (simp add: point_def vector_meet_comp)
finally have 2: "e \<cdot> x \<cdot> -(e ; s\<^sup>T) = 0"
using galois_aux2 by blast
thus ?thesis
proof -
have "e ; s\<^sup>T = s\<^sup>T \<cdot> e"
using assms(2,3) inf_commute point_def vector_meet_comp by force
thus ?thesis
using 2
by (metis assms(2,3) conv_one galois_aux inf.assoc point_def point_equations(1)
point_is_point schroeder_2 vector_meet_comp)
qed
qed
with a show ?thesis
by simp
qed
show "end_points (x \<cdot> - (e ; s\<^sup>T)) \<le> e"
using assms(1,2,3,5) cycle_remove_edge(3) by blast
qed
qed
end (* context relation_algebra_rtc_tarski *)
end
diff --git a/thys/Relational_Paths/Rooted_Paths.thy b/thys/Relational_Paths/Rooted_Paths.thy
--- a/thys/Relational_Paths/Rooted_Paths.thy
+++ b/thys/Relational_Paths/Rooted_Paths.thy
@@ -1,1272 +1,1272 @@
(* Title: Relational Characterisation of Rooted Paths
Author: Walter Guttmann, Peter Hoefner
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
Peter Hoefner <peter at hoefner-online.de>
*)
section \<open>Relational Characterisation of Rooted Paths\<close>
text \<open>
We characterise paths together with a designated root.
This is important as often algorithms start with a single vertex, and then build up a path, a tree or another structure.
An example is Dijkstra's shortest path algorithm.
\<close>
theory Rooted_Paths
imports Paths
begin
context relation_algebra
begin
text \<open>General theorems\<close>
lemma step_has_target:
assumes "x;r \<noteq> 0"
shows "x\<^sup>T;1 \<noteq> 0"
using assms inf.commute inf_bot_right schroeder_1 by fastforce
lemma end_point_char:
"x\<^sup>T;p = 0 \<longleftrightarrow> p \<le> -(x;1)"
-using antisym bot_least compl_bot_eq conv_galois_1 by fastforce
+using order.antisym bot_least compl_bot_eq conv_galois_1 by fastforce
end (* relation_algebra *)
context relation_algebra_tarski
begin
text \<open>General theorems concerning points\<close>
lemma successor_point:
assumes "is_inj x"
and "point r"
and "x;r \<noteq> 0"
shows "point (x;r)"
using assms
by (simp add: inj_compose is_point_def is_vector_def mult_assoc point_is_point)
lemma no_end_point_char:
assumes "point p"
shows "x\<^sup>T;p \<noteq> 0 \<longleftrightarrow> p \<le> x;1"
by (simp add: assms comp_assoc end_point_char is_vector_def point_in_vector_or_complement_iff)
lemma no_end_point_char_converse:
assumes "point p"
shows "x;p \<noteq> 0 \<longleftrightarrow> p \<le> x\<^sup>T;1"
using assms no_end_point_char by force
end (* relation_algebra_tarski *)
subsection \<open>Consequences without the Tarski rule\<close>
context relation_algebra_rtc
begin
text \<open>Definitions for path classifications\<close>
definition path_root
where "path_root r x \<equiv> r;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star> \<and> is_inj x \<and> is_p_fun x \<and> point r"
abbreviation connected_root
where "connected_root r x \<equiv> r;x \<le> x\<^sup>+"
definition backward_finite_path_root
where "backward_finite_path_root r x \<equiv> connected_root r x \<and> is_inj x \<and> is_p_fun x \<and> point r"
abbreviation backward_terminating_path_root
where "backward_terminating_path_root r x \<equiv> backward_finite_path_root r x \<and> x;r = 0"
abbreviation cycle_root
where "cycle_root r x \<equiv> r;x \<le> x\<^sup>+ \<cdot> x\<^sup>T;1 \<and> is_inj x \<and> is_p_fun x \<and> point r"
abbreviation non_empty_cycle_root
where "non_empty_cycle_root r x \<equiv> backward_finite_path_root r x \<and> r \<le> x\<^sup>T;1"
abbreviation finite_path_root_end
where "finite_path_root_end r x e \<equiv> backward_finite_path_root r x \<and> point e \<and> r \<le> x\<^sup>\<star>;e"
abbreviation terminating_path_root_end
where "terminating_path_root_end r x e \<equiv> finite_path_root_end r x e \<and> x\<^sup>T;e = 0"
text \<open>Equivalent formulations of \<open>connected_root\<close>\<close>
lemma connected_root_iff1:
assumes "point r"
shows "connected_root r x \<longleftrightarrow> 1;x \<le> r\<^sup>T;x\<^sup>+"
by (metis assms comp_assoc is_vector_def point_def ss423conv)
lemma connected_root_iff2:
assumes "point r"
shows "connected_root r x \<longleftrightarrow> x\<^sup>T;1 \<le> x\<^sup>T\<^sup>+;r"
by (metis assms conv_contrav conv_invol conv_iso conv_one star_conv star_slide_var
connected_root_iff1)
lemma connected_root_aux:
"x\<^sup>T\<^sup>+;r \<le> x\<^sup>T;1"
by (simp add: comp_assoc mult_isol)
lemma connected_root_iff3:
assumes "point r"
shows "connected_root r x \<longleftrightarrow> x\<^sup>T;1 = x\<^sup>T\<^sup>+;r"
-using assms antisym connected_root_aux connected_root_iff2 by fastforce
+using assms order.antisym connected_root_aux connected_root_iff2 by fastforce
lemma connected_root_iff4:
assumes "point r"
shows "connected_root r x \<longleftrightarrow> 1;x = r\<^sup>T;x\<^sup>+"
by (metis assms conv_contrav conv_invol conv_one star_conv star_slide_var connected_root_iff3)
text \<open>Consequences of \<open>connected_root\<close>\<close>
lemma has_root_contra:
assumes "connected_root r x"
and "point r"
and "x\<^sup>T;r = 0"
shows "x = 0"
using assms comp_assoc independence1 conv_zero ss_p18 connected_root_iff3
by force
lemma has_root:
assumes "connected_root r x"
and "point r"
and "x \<noteq> 0"
shows "x\<^sup>T;r \<noteq> 0"
using has_root_contra assms by blast
lemma connected_root_move_root:
assumes "connected_root r x"
and "q \<le> x\<^sup>\<star>;r"
shows "connected_root q x"
by (metis assms comp_assoc mult_isol phl_cons1 star_slide_var star_trans_eq)
lemma root_cycle_converse:
assumes "connected_root r x"
and "point r"
and "x;r \<noteq> 0"
shows "x\<^sup>T;r \<noteq> 0"
using assms conv_zero has_root by fastforce
text \<open>Rooted paths\<close>
lemma path_iff_aux_1:
assumes "bijective r"
shows "r;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star> \<longleftrightarrow> x \<le> r\<^sup>T;(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>)"
by (simp add: assms ss423conv)
lemma path_iff_aux_2:
assumes "bijective r"
shows "r;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star> \<longleftrightarrow> x\<^sup>T \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);r"
proof -
have "((x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);r)\<^sup>T = r\<^sup>T;(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>)"
by (metis conv_add conv_contrav conv_invol star_conv sup.commute)
thus ?thesis
by (metis assms conv_invol conv_iso path_iff_aux_1)
qed
lemma path_iff_backward:
assumes "is_inj x"
and "is_p_fun x"
and "point r"
and "r;x \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
shows "connected x"
proof -
have "x\<^sup>T;1;x\<^sup>T \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);r;1;x\<^sup>T"
using assms(3,4) path_iff_aux_2 mult_isor point_def by blast
also have "... = (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);r;1;x\<^sup>T;x;x\<^sup>T"
using assms(1) comp_assoc inj_p_fun p_fun_triple by fastforce
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);r;x;x\<^sup>T"
by (metis assms(3) mult_double_iso top_greatest point_def is_vector_def comp_assoc)
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);x\<^sup>T"
by (metis assms(4) comp_assoc mult_double_iso)
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>)"
using le_supI2 mult_isol star_ext by blast
also have "... = x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using assms(1,2) cancel_separate_converse_idempotent by fastforce
finally show ?thesis
by (metis conv_add conv_contrav conv_invol conv_one mult_assoc star_conv sup.orderE sup.orderI
sup_commute)
qed
lemma empty_path_root_end:
assumes "terminating_path_root_end r x e"
shows "e = r \<longleftrightarrow> x = 0"
apply(standard)
using assms has_root backward_finite_path_root_def apply blast
-by (metis assms antisym conv_e conv_zero independence1 is_inj_def mult_oner point_swap
+by (metis assms order.antisym conv_e conv_zero independence1 is_inj_def mult_oner point_swap
backward_finite_path_root_def ss423conv sur_def_var1 x_leq_triple_x)
lemma path_root_acyclic:
assumes "path_root r x"
and "x;r = 0"
shows "is_acyclic x"
proof -
have "x\<^sup>+\<cdot>1' = (x\<^sup>+)\<^sup>T\<cdot>x\<^sup>+\<cdot>1'"
by (metis conv_e conv_times inf.assoc inf.left_idem inf_le2 many_strongly_connected_iff_7 mult_oner star_subid)
also have "... \<le> x\<^sup>T;1\<cdot>x\<^sup>+\<cdot>1'"
by (metis conv_contrav inf.commute maddux_20 meet_double_iso plus_top star_conv star_slide_var)
finally have "r;(x\<^sup>+\<cdot>1') \<le> r;(x\<^sup>T;1\<cdot>x\<^sup>+\<cdot>1')"
using mult_isol by blast
also have "... = (r\<cdot>1;x);(x\<^sup>+\<cdot>1')"
by (metis (no_types, lifting) comp_assoc conv_contrav conv_invol conv_one inf.assoc is_vector_def one_idem_mult vector_2)
also have "... = r;x;(x\<^sup>+\<cdot>1')"
by (metis assms(1) path_root_def point_def inf_top_right vector_1)
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);(x\<^sup>+\<cdot>1')"
using assms(1) mult_isor path_root_def by blast
also have "... = x\<^sup>\<star>;(x\<^sup>+\<cdot>1') + x\<^sup>T\<^sup>+;(x\<^sup>+\<cdot>1')"
by (metis distrib_right star_star_plus sup.commute)
also have "... \<le> x\<^sup>\<star>;(x\<^sup>+\<cdot>1') + x\<^sup>T;1"
by (metis join_isol mult_isol plus_top top_greatest)
finally have "r;(x\<^sup>+\<cdot>1');1 \<le> x\<^sup>\<star>;(x\<^sup>+\<cdot>1');1 + x\<^sup>T;1"
by (metis distrib_right inf_absorb2 mult_assoc mult_subdistr one_idem_mult)
hence 1: "r;(x\<^sup>+\<cdot>1');1 \<le> x\<^sup>T;1"
by (metis assms(1) inj_implies_step_forwards_backwards sup_absorb2 path_root_def)
have "x\<^sup>+\<cdot>1' \<le> (x\<^sup>+\<cdot>1');1"
by (simp add: maddux_20)
also have "... \<le> r\<^sup>T;r;(x\<^sup>+\<cdot>1');1"
by (metis assms(1) comp_assoc order.refl point_def ss423conv path_root_def)
also have "... \<le> r\<^sup>T;x\<^sup>T;1"
using 1 by (simp add: comp_assoc mult_isol)
also have "... = 0"
using assms(2) annil conv_contrav conv_zero by force
finally show ?thesis
using galois_aux le_bot by blast
qed
text \<open>Start points and end points\<close>
lemma start_points_in_root_aux:
assumes "backward_finite_path_root r x"
shows "x;1 \<le> x\<^sup>T\<^sup>\<star>;r"
proof -
have "x;1 \<le> x;x\<^sup>T\<^sup>+;r"
by (metis assms inf_top.left_neutral modular_var_2 mult_assoc connected_root_iff3
backward_finite_path_root_def)
also have "... \<le> 1';x\<^sup>T\<^sup>\<star>;r"
by (metis assms is_inj_def mult_assoc mult_isor backward_finite_path_root_def)
finally show ?thesis
by simp
qed
lemma start_points_in_root:
assumes "backward_finite_path_root r x"
shows "start_points x \<le> r"
using assms galois_1 sup_commute connected_root_iff3 backward_finite_path_root_def
start_points_in_root_aux by fastforce
lemma start_points_not_zero_contra:
assumes "connected_root r x"
and "point r"
and "start_points x = 0"
and "x;r = 0"
shows "x = 0"
proof -
have "x;1 \<le> x\<^sup>T;1"
using assms(3) galois_aux by force
also have "... \<le> -r"
using assms(4) comp_res compl_bot_eq by blast
finally show ?thesis
using assms(1,2) has_root_contra galois_aux schroeder_1 by force
qed
lemma start_points_not_zero:
assumes "connected_root r x"
and "point r"
and "x \<noteq> 0"
and "x;r = 0"
shows "start_points x \<noteq> 0"
using assms start_points_not_zero_contra by blast
text \<open>Backwards terminating and backwards finite\<close>
lemma backward_terminating_path_root_aux:
assumes "backward_terminating_path_root r x"
shows "x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
proof -
have "x\<^sup>T\<^sup>\<star>;r \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
using assms comp_res compl_bot_eq compl_le_swap1 mult_isol by blast
thus ?thesis
using assms dual_order.trans maddux_20 start_points_in_root_aux by blast
qed
lemma backward_finite_path_connected_aux:
assumes "backward_finite_path_root r x"
shows "x\<^sup>T;r;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
proof -
have "x\<^sup>T;r;x\<^sup>T \<cdot> r\<^sup>T = x\<^sup>T;r;(x\<^sup>T \<cdot> r\<^sup>T)"
by (metis conv_invol conv_times vector_1_comm comp_assoc conv_contrav assms
backward_finite_path_root_def point_def)
also have "... \<le> x\<^sup>T;r;r\<^sup>T"
by (simp add: mult_isol)
also have 1: "... \<le> x\<^sup>T"
by (metis assms comp_assoc is_inj_def mult_1_right mult_isol point_def
backward_finite_path_root_def)
also have "... \<le> x\<^sup>T\<^sup>\<star>"
by simp
finally have 2: "x\<^sup>T;r;x\<^sup>T \<cdot> r\<^sup>T \<le> x\<^sup>T\<^sup>\<star>" .
let ?v = "x;1 \<cdot> -r"
have "?v \<le> x\<^sup>T\<^sup>+;r"
by (simp add: assms galois_1 start_points_in_root_aux)
hence "r\<^sup>T;x \<cdot> ?v \<le> r\<^sup>T;x \<cdot> x\<^sup>T\<^sup>+;r"
using meet_isor by blast
also have 3: "... = x\<^sup>T\<^sup>+;r \<cdot> 1;r\<^sup>T;x"
by (metis assms conv_contrav conv_one inf_commute is_vector_def point_def
backward_finite_path_root_def)
also have "... = (x\<^sup>T\<^sup>+;r \<cdot> 1);r\<^sup>T;x"
using 3 by (metis comp_assoc inf_commute is_vector_def star_conv vector_1 assms
backward_finite_path_root_def point_def)
also have "... = x\<^sup>T\<^sup>+;r;r\<^sup>T;x"
by simp
also have "... \<le> x\<^sup>T\<^sup>+;x"
using 1 by (metis mult_assoc mult_isol mult_isor star_slide_var)
also have "... = x\<^sup>T\<^sup>\<star>;x\<^sup>T;x"
by (simp add: star_slide_var)
also have "... \<le> x\<^sup>T\<^sup>\<star>"
by (metis assms backward_finite_path_root_def is_p_fun_def mult_1_right mult_assoc mult_isol_var
star_1l star_inductl_star)
finally have 4: "x\<^sup>T;r \<cdot> ?v\<^sup>T \<le> x\<^sup>\<star>"
using conv_iso star_conv by force
have "x\<^sup>T;r;x\<^sup>T \<cdot> -r\<^sup>T = (x\<^sup>T;r \<cdot> 1);x\<^sup>T \<cdot> -r\<^sup>T"
by simp
also have "... = x\<^sup>T;r \<cdot> 1;x\<^sup>T \<cdot> -r\<^sup>T"
by (metis inf.commute is_vector_def comp_assoc vector_1 assms backward_finite_path_root_def
point_def)
also have "... \<le> x\<^sup>\<star>"
using 4 by (simp add: conv_compl inf.assoc)
finally have "(x\<^sup>T;r;x\<^sup>T \<cdot> -r\<^sup>T) + (x\<^sup>T;r;x\<^sup>T \<cdot> r\<^sup>T) \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using 2 sup.mono by blast
thus ?thesis
by fastforce
qed
lemma backward_finite_path_connected:
assumes "backward_finite_path_root r x"
shows "connected x"
proof -
from assms obtain r where 1: "backward_finite_path_root r x" ..
have "x\<^sup>T;(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>) = x\<^sup>T;(1' + x\<^sup>+) + x\<^sup>T\<^sup>+"
by (simp add: distrib_left)
also have "... = x\<^sup>T;x\<^sup>+ + x\<^sup>T\<^sup>+"
using calculation distrib_left star_star_plus by fastforce
also have "... \<le> 1';x\<^sup>\<star> + x\<^sup>T\<^sup>+"
using 1 by (metis add_iso comp_assoc is_p_fun_def mult_isor backward_finite_path_root_def)
also have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using join_isol by fastforce
finally have "x\<^sup>T;r;x\<^sup>T + x\<^sup>T;(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>) \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using 1 backward_finite_path_connected_aux by simp
hence "x\<^sup>T\<^sup>\<star>;x\<^sup>T;r;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using star_inductl comp_assoc by simp
hence "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using 1 backward_finite_path_root_def connected_root_iff3 star_slide_var by fastforce
thus ?thesis
by (metis (mono_tags, lifting) sup.commute comp_assoc conv_add conv_contrav conv_invol conv_iso
conv_one star_conv)
qed
lemma backward_finite_path_root_path:
assumes "backward_finite_path_root r x"
shows "path x"
using assms path_def backward_finite_path_connected backward_finite_path_root_def by blast
lemma backward_finite_path_root_path_root:
assumes "backward_finite_path_root r x"
shows "path_root r x"
using assms backward_finite_path_root_def le_supI1 star_star_plus path_root_def by fastforce
lemma zero_backward_terminating_path_root:
assumes "point r"
shows "backward_terminating_path_root r 0"
by (simp add: assms is_inj_def is_p_fun_def backward_finite_path_root_def)
lemma backward_finite_path_root_move_root:
assumes "backward_finite_path_root r x"
and "point q"
and "q \<le> x\<^sup>\<star>;r"
shows "backward_finite_path_root q x"
using assms connected_root_move_root backward_finite_path_root_def by blast
text \<open>Cycle\<close>
lemma non_empty_cycle_root_var_axioms_1:
"non_empty_cycle_root r x \<longleftrightarrow> x\<^sup>T;1 \<le> x\<^sup>T\<^sup>+;r \<and> is_inj x \<and> is_p_fun x \<and> point r \<and> r \<le> x\<^sup>T;1"
using connected_root_iff2 backward_finite_path_root_def by blast
lemma non_empty_cycle_root_loop:
assumes "non_empty_cycle_root r x"
shows "r \<le> x\<^sup>T\<^sup>+;r"
using assms connected_root_iff3 backward_finite_path_root_def by fastforce
lemma cycle_root_end_empty:
assumes "terminating_path_root_end r x e"
and "many_strongly_connected x"
shows "x = 0"
by (metis assms has_root_contra point_swap backward_finite_path_root_def
backward_finite_path_root_move_root star_conv)
lemma cycle_root_end_empty_var:
assumes "terminating_path_root_end r x e"
and "x \<noteq> 0"
shows "\<not> many_strongly_connected x"
using assms cycle_root_end_empty by blast
text \<open>Terminating path\<close>
lemma terminating_path_root_end_connected:
assumes "terminating_path_root_end r x e"
shows "x;1 \<le> x\<^sup>+;e"
proof -
have "x;1 \<le> x;x\<^sup>T;1"
by (metis comp_assoc inf_top.left_neutral modular_var_2)
also have "... = x;x\<^sup>T\<^sup>+;r"
using assms backward_finite_path_root_def connected_root_iff3 comp_assoc by fastforce
also have "... \<le> x;x\<^sup>T\<^sup>+;x\<^sup>\<star>;e"
by (simp add: assms comp_assoc mult_isol)
also have "... = x;x\<^sup>T;(x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);e"
using assms cancel_separate_p_fun_converse comp_assoc backward_finite_path_root_def by fastforce
also have "... = x;x\<^sup>T;(x\<^sup>+ + x\<^sup>T\<^sup>\<star>);e"
by (simp add: star_star_plus)
also have "... = x;x\<^sup>T;x\<^sup>+;e + x;x\<^sup>T\<^sup>+;e"
by (simp add: comp_assoc distrib_left)
also have "... = x;x\<^sup>T;x\<^sup>+;e"
by (simp add: assms comp_assoc independence1)
also have "... \<le> x\<^sup>+;e"
by (metis assms annil independence1 is_inj_def mult_isor mult_oner backward_finite_path_root_def)
finally show ?thesis .
qed
lemma terminating_path_root_end_forward_finite:
assumes "terminating_path_root_end r x e"
shows "backward_finite_path_root e (x\<^sup>T)"
using assms terminating_path_root_end_connected inj_p_fun connected_root_iff2
backward_finite_path_root_def by force
end (* relation_algebra_rtc *)
subsection \<open>Consequences with the Tarski rule\<close>
context relation_algebra_rtc_tarski
begin
text \<open>Some (more) results about points\<close>
lemma point_reachable_converse:
assumes "is_vector v"
and "v \<noteq> 0"
and "point r"
and "v \<le> x\<^sup>T\<^sup>+;r"
shows "r \<le> x\<^sup>+;v"
proof -
have "v\<^sup>T;v \<noteq> 0"
by (metis assms(2) inf.idem inf_bot_right mult_1_right schroeder_1)
hence "1;v\<^sup>T;v = 1"
using assms(1) is_vector_def mult_assoc tarski by force
hence 1: "r = r;v\<^sup>T;v"
by (metis assms(3) is_vector_def mult_assoc point_def)
have "v;r\<^sup>T \<le> x\<^sup>T\<^sup>+"
using assms(3,4) point_def ss423bij by simp
hence "r;v\<^sup>T \<le> x\<^sup>+"
by (metis conv_contrav conv_invol conv_iso star_conv star_slide_var)
thus ?thesis
using 1 by (metis mult_isor)
qed
text \<open>Roots\<close>
lemma root_in_start_points:
assumes "connected_root r x"
and "is_vector r"
and "x \<noteq> 0"
and "x;r = 0"
shows "r \<le> start_points x"
proof -
have "r = r;x;1"
by (metis assms(2,3) comp_assoc is_vector_def tarski)
also have "... \<le> x;1"
by (metis assms(1) comp_assoc one_idem_mult phl_seq top_greatest)
finally show ?thesis
using assms(4) comp_res compl_bot_eq compl_le_swap1 inf.boundedI by blast
qed
lemma root_equals_start_points:
assumes "backward_terminating_path_root r x"
and "x \<noteq> 0"
shows "r = start_points x"
-using assms antisym point_def backward_finite_path_root_def start_points_in_root root_in_start_points
+using assms order.antisym point_def backward_finite_path_root_def start_points_in_root root_in_start_points
by fastforce
lemma root_equals_end_points:
assumes "backward_terminating_path_root r (x\<^sup>T)"
and "x \<noteq> 0"
shows "r = end_points x"
by (metis assms conv_invol step_has_target ss_p18 root_equals_start_points)
lemma root_in_edge_sources:
assumes "connected_root r x"
and "x \<noteq> 0"
and "is_vector r"
shows "r \<le> x;1"
proof -
have "r;1;x;1 \<le> x\<^sup>+;1"
using assms(1,3) is_vector_def mult_isor by fastforce
thus ?thesis
by (metis assms(2) comp_assoc conway.dagger_unfoldl_distr dual_order.trans maddux_20 sup.commute
sup_absorb2 tarski top_greatest)
qed
text \<open>Rooted Paths\<close>
lemma non_empty_path_root_iff_aux:
assumes "path_root r x"
and "x \<noteq> 0"
shows "r \<le> (x + x\<^sup>T);1"
proof -
have "(r;x \<cdot> 1');1 = (x\<^sup>T;r\<^sup>T \<cdot> 1');1"
by (metis conv_contrav conv_e conv_times inf.cobounded2 is_test_def test_eq_conv)
also have "... \<le> x\<^sup>T;r\<^sup>T;1"
using mult_subdistr by blast
also have "... \<le> x\<^sup>T;1"
by (metis mult_assoc mult_double_iso one_idem_mult top_greatest)
finally have 1: "(r;x \<cdot> 1');1 \<le> x\<^sup>T;1" .
have "r \<le> r;1;x;1"
using assms(2) comp_assoc maddux_20 tarski by fastforce
also have "... = r;x;1"
using assms(1) path_root_def point_def is_vector_def by simp
also have "... = (r;x \<cdot> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>));1"
using assms(1) path_root_def by (simp add: inf.absorb_iff1)
also have "... = (r;x \<cdot> (x\<^sup>+ + x\<^sup>T\<^sup>+ + 1'));1"
by (metis star_star_plus star_unfoldl_eq sup_commute sup_left_commute)
also have "... \<le> (x\<^sup>+ + x\<^sup>T\<^sup>+ + (r;x \<cdot> 1'));1"
by (metis inf_le2 inf_sup_distrib1 mult_isor order_refl sup_mono)
also have "... \<le> x;1 + x\<^sup>T;1 + (r;x \<cdot> 1');1"
by (simp add: plus_top)
also have "... = x;1 + x\<^sup>T;1"
using 1 sup.coboundedI2 sup.order_iff by fastforce
finally show ?thesis
by simp
qed
text \<open>Backwards terminating and backwards finite\<close>
lemma backward_terminating_path_root_2:
assumes "backward_terminating_path_root r x"
shows "backward_terminating x"
using assms backward_terminating_iff2 path_def backward_terminating_path_root_aux
backward_finite_path_connected backward_finite_path_root_def by blast
lemma backward_terminating_path_root:
assumes "backward_terminating_path_root r x"
shows "backward_terminating_path x"
using assms backward_finite_path_root_path backward_terminating_path_root_2 by fastforce
text \<open>(Non-empty) Cycle\<close>
lemma cycle_iff:
assumes "point r"
shows "x;r \<noteq> 0 \<longleftrightarrow> r \<le> x\<^sup>T;1"
by (simp add: assms no_end_point_char_converse)
lemma non_empty_cycle_root_iff:
assumes "connected_root r x"
and "point r"
shows "x;r \<noteq> 0 \<longleftrightarrow> r \<le> x\<^sup>T\<^sup>+;r"
using assms connected_root_iff3 cycle_iff by simp
lemma backward_finite_path_root_terminating_or_cycle:
"backward_finite_path_root r x \<longleftrightarrow> backward_terminating_path_root r x \<or> non_empty_cycle_root r x"
using cycle_iff backward_finite_path_root_def by blast
lemma non_empty_cycle_root_msc:
assumes "non_empty_cycle_root r x"
shows "many_strongly_connected x"
proof -
let ?p = "x\<^sup>T;r"
have 1: "is_point ?p"
unfolding is_point_def
using conjI assms is_vector_def mult_assoc point_def inj_compose p_fun_inj
cycle_iff backward_finite_path_root_def root_cycle_converse by fastforce
have "?p \<le> x\<^sup>T\<^sup>+;?p"
by (metis assms comp_assoc mult_isol star_slide_var non_empty_cycle_root_loop)
hence "?p \<le> x\<^sup>+;?p"
using 1 bot_least point_def point_is_point point_reachable_converse by blast
also have "... = x\<^sup>\<star>;(x;x\<^sup>T);r"
by (metis comp_assoc star_slide_var)
also have "... \<le> x\<^sup>\<star>;1';r"
using assms is_inj_def mult_double_iso backward_finite_path_root_def by blast
finally have 2: "?p \<le> x\<^sup>\<star>;r"
by simp
have "x\<^sup>T;x\<^sup>\<star>;r = ?p + x\<^sup>T;x\<^sup>+;r"
by (metis conway.dagger_unfoldl_distr distrib_left mult_assoc)
also have "... \<le> ?p + 1';x\<^sup>\<star>;r"
by (metis assms is_p_fun_def join_isol mult_assoc mult_isor backward_finite_path_root_def)
also have "... = x\<^sup>\<star>;r"
using 2 by (simp add: sup_absorb2)
finally have 3: "x\<^sup>T\<^sup>\<star>;r \<le> x\<^sup>\<star>;r"
by (metis star_inductl comp_assoc conway.dagger_unfoldl_distr le_supI order_prop)
have "x\<^sup>T \<le> x\<^sup>T\<^sup>+;r"
by (metis assms maddux_20 connected_root_iff3 backward_finite_path_root_def)
also have "... \<le> x\<^sup>\<star>;r"
using 3 by (metis assms conway.dagger_unfoldl_distr sup_absorb2 non_empty_cycle_root_loop)
finally have 4: "x\<^sup>T \<le> x\<^sup>\<star>;r" .
have "x\<^sup>T \<le> x\<^sup>T;x;x\<^sup>T"
by (metis conv_invol x_leq_triple_x)
also have "... \<le> 1;x;x\<^sup>T"
by (simp add: mult_isor)
also have "... = r\<^sup>T;x\<^sup>+;x\<^sup>T"
using assms connected_root_iff4 backward_finite_path_root_def by fastforce
also have "... \<le> r\<^sup>T;x\<^sup>\<star>"
by (metis assms is_inj_def mult_1_right mult_assoc mult_isol backward_finite_path_root_def
star_slide_var)
finally have "x\<^sup>T \<le> x\<^sup>\<star>;r \<cdot> r\<^sup>T;x\<^sup>\<star>"
using 4 by simp
also have "... = x\<^sup>\<star>;r \<cdot> 1;r\<^sup>T;x\<^sup>\<star>"
by (metis assms conv_contrav conv_one is_vector_def point_def backward_finite_path_root_def)
also have "... = (x\<^sup>\<star>;r \<cdot> 1);r\<^sup>T;x\<^sup>\<star>"
by (metis (no_types, lifting) assms is_vector_def mult_assoc point_def
backward_finite_path_root_def vector_1)
also have "... = x\<^sup>\<star>;r;r\<^sup>T;x\<^sup>\<star>"
by simp
also have "... \<le> x\<^sup>\<star>;x\<^sup>\<star>"
by (metis assms is_inj_def mult_1_right mult_assoc mult_isol mult_isor point_def
backward_finite_path_root_def)
also have "... \<le> x\<^sup>\<star>"
by simp
finally show ?thesis
by (simp add: many_strongly_connected_iff_1)
qed
lemma non_empty_cycle_root_msc_cycle:
assumes "non_empty_cycle_root r x"
shows "cycle x"
using assms backward_finite_path_root_path non_empty_cycle_root_msc by fastforce
lemma non_empty_cycle_root_non_empty:
assumes "non_empty_cycle_root r x"
shows "x \<noteq> 0"
using assms cycle_iff annil backward_finite_path_root_def by blast
lemma non_empty_cycle_root_rtc_symmetric:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>\<star>;r = x\<^sup>T\<^sup>\<star>;r"
using assms non_empty_cycle_root_msc by fastforce
lemma non_empty_cycle_root_point_exchange:
assumes "non_empty_cycle_root r x"
and "point p"
shows "r \<le> x\<^sup>\<star>;p \<longleftrightarrow> p \<le> x\<^sup>\<star>;r"
by (metis assms(1,2) inj_sur_semi_swap point_def non_empty_cycle_root_msc
backward_finite_path_root_def star_conv)
lemma non_empty_cycle_root_rtc_tc:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>\<star>;r = x\<^sup>+;r"
-proof (rule antisym)
+proof (rule order.antisym)
have "r \<le> x\<^sup>+;r"
using assms many_strongly_connected_iff_7 non_empty_cycle_root_loop non_empty_cycle_root_msc
by simp
thus "x\<^sup>\<star>;r \<le> x\<^sup>+;r"
using sup_absorb2 by fastforce
next
show "x\<^sup>+;r \<le> x\<^sup>\<star>;r"
by (simp add: mult_isor)
qed
lemma non_empty_cycle_root_no_start_end_points:
assumes "non_empty_cycle_root r x"
shows "x;1 = x\<^sup>T;1"
using assms many_strongly_connected_implies_no_start_end_points non_empty_cycle_root_msc by blast
lemma non_empty_cycle_root_move_root:
assumes "non_empty_cycle_root r x"
and "point q"
and "q \<le> x\<^sup>\<star>;r"
shows "non_empty_cycle_root q x"
by (metis assms cycle_iff dual_order.trans backward_finite_path_root_move_root start_points_in_root
root_equals_start_points non_empty_cycle_root_non_empty)
lemma non_empty_cycle_root_loop_converse:
assumes "non_empty_cycle_root r x"
shows "r \<le> x\<^sup>+;r"
using assms less_eq_def non_empty_cycle_root_rtc_tc by fastforce
lemma non_empty_cycle_root_move_root_same_reachable:
assumes "non_empty_cycle_root r x"
and "point q"
and "q \<le> x\<^sup>\<star>;r"
shows "x\<^sup>\<star>;r = x\<^sup>\<star>;q"
by (metis assms many_strongly_connected_iff_7 connected_root_iff3 connected_root_move_root
backward_finite_path_root_def non_empty_cycle_root_msc non_empty_cycle_root_rtc_tc)
lemma non_empty_cycle_root_move_root_same_reachable_2:
assumes "non_empty_cycle_root r x"
and "point q"
and "q \<le> x\<^sup>\<star>;r"
shows "x\<^sup>\<star>;r = x\<^sup>T\<^sup>\<star>;q"
using assms non_empty_cycle_root_move_root_same_reachable non_empty_cycle_root_msc by simp
lemma non_empty_cycle_root_move_root_msc:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>T\<^sup>\<star>;q = x\<^sup>\<star>;q"
using assms non_empty_cycle_root_msc by simp
lemma non_empty_cycle_root_move_root_rtc_tc:
assumes "non_empty_cycle_root r x"
and "point q"
and "q \<le> x\<^sup>\<star>;r"
shows "x\<^sup>\<star>;q = x\<^sup>+;q"
using assms non_empty_cycle_root_move_root non_empty_cycle_root_rtc_tc by blast
lemma non_empty_cycle_root_move_root_loop_converse:
assumes "non_empty_cycle_root r x"
and "point q"
and "q \<le> x\<^sup>\<star>;r"
shows "q \<le> x\<^sup>T\<^sup>+;q"
using assms non_empty_cycle_root_loop non_empty_cycle_root_move_root by blast
lemma non_empty_cycle_root_move_root_loop:
assumes "non_empty_cycle_root r x"
and "point q"
and "q \<le> x\<^sup>\<star>;r"
shows "q \<le> x\<^sup>+;q"
using assms non_empty_cycle_root_loop_converse non_empty_cycle_root_move_root by blast
lemma non_empty_cycle_root_msc_plus:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>+;r = x\<^sup>T\<^sup>+;r"
using assms many_strongly_connected_iff_7 non_empty_cycle_root_msc by fastforce
lemma non_empty_cycle_root_tc_start_points:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>+;r = x;1"
by (metis assms connected_root_iff3 backward_finite_path_root_def non_empty_cycle_root_msc_plus
non_empty_cycle_root_no_start_end_points)
lemma non_empty_cycle_root_rtc_start_points:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>\<star>;r = x;1"
by (simp add: assms non_empty_cycle_root_rtc_tc non_empty_cycle_root_tc_start_points)
lemma non_empty_cycle_root_converse_start_end_points:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>T \<le> x;1;x"
by (metis assms conv_contrav conv_invol conv_one inf.boundedI maddux_20 maddux_21 vector_meet_comp_x
non_empty_cycle_root_no_start_end_points)
lemma non_empty_cycle_root_start_end_points_plus:
assumes "non_empty_cycle_root r x"
shows "x;1;x \<le> x\<^sup>+"
-using assms eq_iff one_strongly_connected_iff one_strongly_connected_implies_7_eq
+using assms order.eq_iff one_strongly_connected_iff one_strongly_connected_implies_7_eq
backward_finite_path_connected non_empty_cycle_root_msc by blast
lemma non_empty_cycle_root_converse_plus:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>T \<le> x\<^sup>+"
using assms many_strongly_connected_iff_2 non_empty_cycle_root_msc by blast
lemma non_empty_cycle_root_plus_converse:
assumes "non_empty_cycle_root r x"
shows "x\<^sup>+ = x\<^sup>T\<^sup>+"
using assms many_strongly_connected_iff_7 non_empty_cycle_root_msc by fastforce
lemma non_empty_cycle_root_converse:
assumes "non_empty_cycle_root r x"
shows "non_empty_cycle_root r (x\<^sup>T)"
by (metis assms conv_invol inj_p_fun connected_root_iff3 backward_finite_path_root_def
non_empty_cycle_root_msc_plus non_empty_cycle_root_tc_start_points)
lemma non_empty_cycle_root_move_root_forward:
assumes "non_empty_cycle_root r x"
and "point q"
and "r \<le> x\<^sup>\<star>;q"
shows "non_empty_cycle_root q x"
by (metis assms backward_finite_path_root_move_root non_empty_cycle_root_no_start_end_points
non_empty_cycle_root_point_exchange non_empty_cycle_root_rtc_start_points)
lemma non_empty_cycle_root_move_root_forward_cycle:
assumes "non_empty_cycle_root r x"
and "point q"
and "r \<le> x\<^sup>\<star>;q"
shows "x;q \<noteq> 0 \<and> x\<^sup>T;q \<noteq> 0"
by (metis assms comp_assoc independence1 ss_p18 non_empty_cycle_root_move_root_forward
non_empty_cycle_root_msc_plus non_empty_cycle_root_non_empty
non_empty_cycle_root_tc_start_points)
lemma non_empty_cycle_root_equivalences:
assumes "non_empty_cycle_root r x"
and "point q"
shows "(r \<le> x\<^sup>\<star>;q \<longleftrightarrow> q \<le> x\<^sup>\<star>;r)"
and "(r \<le> x\<^sup>\<star>;q \<longleftrightarrow> x;q \<noteq> 0)"
and "(r \<le> x\<^sup>\<star>;q \<longleftrightarrow> x\<^sup>T;q \<noteq> 0)"
and "(r \<le> x\<^sup>\<star>;q \<longleftrightarrow> q \<le> x;1)"
and "(r \<le> x\<^sup>\<star>;q \<longleftrightarrow> q \<le> x\<^sup>T;1)"
using assms cycle_iff no_end_point_char non_empty_cycle_root_no_start_end_points
non_empty_cycle_root_point_exchange non_empty_cycle_root_rtc_start_points
by metis+
lemma non_empty_cycle_root_chord:
assumes "non_empty_cycle_root r x"
and "point p"
and "point q"
and "r \<le> x\<^sup>\<star>;p"
and "r \<le> x\<^sup>\<star>;q"
shows "p \<le> x\<^sup>\<star>;q"
using assms non_empty_cycle_root_move_root_same_reachable non_empty_cycle_root_point_exchange
by fastforce
lemma non_empty_cycle_root_var_axioms_2:
"non_empty_cycle_root r x \<longleftrightarrow> x;1 \<le> x\<^sup>+;r \<and> is_inj x \<and> is_p_fun x \<and> point r \<and> r \<le> x;1"
apply (rule iffI)
- apply (metis eq_iff backward_finite_path_root_def non_empty_cycle_root_no_start_end_points
+ apply (metis order.eq_iff backward_finite_path_root_def non_empty_cycle_root_no_start_end_points
non_empty_cycle_root_tc_start_points)
by (metis conv_invol p_fun_inj connected_root_iff2 connected_root_iff3
non_empty_cycle_root_var_axioms_1 non_empty_cycle_root_msc_plus
non_empty_cycle_root_rtc_start_points non_empty_cycle_root_rtc_tc)
lemma non_empty_cycle_root_var_axioms_3:
"non_empty_cycle_root r x \<longleftrightarrow> x;1 \<le> x\<^sup>+;r \<and> is_inj x \<and> is_p_fun x \<and> point r \<and> r \<le> x\<^sup>+;x;1"
apply (rule iffI)
apply (metis comp_assoc eq_refl backward_finite_path_root_def star_inductl_var_eq2
non_empty_cycle_root_no_start_end_points non_empty_cycle_root_rtc_start_points
non_empty_cycle_root_tc_start_points)
by (metis annir comp_assoc conv_contrav no_end_point_char non_empty_cycle_root_var_axioms_2)
lemma non_empty_cycle_root_subset_equals:
assumes "non_empty_cycle_root r x"
and "non_empty_cycle_root r y"
and "x \<le> y"
shows "x = y"
proof -
have "y;x\<^sup>T\<^sup>\<star>;r = y;x\<^sup>T\<^sup>+;r"
using assms(1) comp_assoc non_empty_cycle_root_msc non_empty_cycle_root_msc_plus
non_empty_cycle_root_rtc_tc by fastforce
also have "... \<le> y;y\<^sup>T;x\<^sup>T\<^sup>\<star>;r"
using assms(3) comp_assoc conv_iso mult_double_iso by fastforce
also have "... \<le> x\<^sup>T\<^sup>\<star>;r"
using assms(2) backward_finite_path_root_def is_inj_def
by (meson dual_order.trans mult_isor order.refl prod_star_closure star_ref)
finally have "r + y;x\<^sup>T\<^sup>\<star>;r \<le> x\<^sup>T\<^sup>\<star>;r"
by (metis conway.dagger_unfoldl_distr le_supI sup.cobounded1)
hence "y\<^sup>\<star>;r \<le> x\<^sup>T\<^sup>\<star>;r"
by (simp add: comp_assoc rtc_inductl)
hence "y;1 \<le> x;1"
using assms(1,2) non_empty_cycle_root_msc non_empty_cycle_root_rtc_start_points by fastforce
thus ?thesis
using assms(2,3) backward_finite_path_root_def ss422iv by blast
qed
lemma non_empty_cycle_root_subset_equals_change_root:
assumes "non_empty_cycle_root r x"
and "non_empty_cycle_root q y"
and "x \<le> y"
shows "x = y"
proof -
have "r \<le> y;1"
by (metis assms(1,3) dual_order.trans mult_isor non_empty_cycle_root_no_start_end_points)
hence "non_empty_cycle_root r y"
by (metis assms(1,2) connected_root_move_root backward_finite_path_root_def
non_empty_cycle_root_no_start_end_points non_empty_cycle_root_rtc_start_points)
thus ?thesis
using assms(1,3) non_empty_cycle_root_subset_equals by blast
qed
lemma non_empty_cycle_root_equivalences_2:
assumes "non_empty_cycle_root r x"
shows "(v \<le> x\<^sup>\<star>;r \<longleftrightarrow> v \<le> x\<^sup>T;1)"
and "(v \<le> x\<^sup>\<star>;r \<longleftrightarrow> v \<le> x;1)"
using assms non_empty_cycle_root_no_start_end_points non_empty_cycle_root_rtc_start_points
by metis+
lemma cycle_root_non_empty:
assumes "x \<noteq> 0"
shows "cycle_root r x \<longleftrightarrow> non_empty_cycle_root r x"
proof
assume 1: "cycle_root r x"
have "r \<le> r;1;x;1"
using assms comp_assoc maddux_20 tarski by fastforce
also have "... \<le> (x\<^sup>+ \<cdot> x\<^sup>T;1);1"
using 1 by (simp add: is_vector_def mult_isor point_def)
also have "... \<le> x\<^sup>T;1"
by (simp add: ra_1)
finally show "non_empty_cycle_root r x"
using 1 backward_finite_path_root_def inf.boundedE by blast
next
assume "non_empty_cycle_root r x"
thus "cycle_root r x"
by (metis backward_finite_path_root_def inf.orderE maddux_20 non_empty_cycle_root_plus_converse
ra_1)
qed
text \<open>Start points and end points\<close>
lemma start_points_path_aux:
assumes "backward_finite_path_root r x"
and "start_points x \<noteq> 0"
shows "x;r = 0"
by (metis assms compl_inf_bot inf.commute non_empty_cycle_root_no_start_end_points
backward_finite_path_root_terminating_or_cycle)
lemma start_points_path:
assumes "backward_finite_path_root r x"
and "start_points x \<noteq> 0"
shows "backward_terminating_path_root r x"
by (simp add: assms start_points_path_aux)
lemma root_in_start_points_2:
assumes "backward_finite_path_root r x"
and "start_points x \<noteq> 0"
shows "r \<le> start_points x"
by (metis assms conv_zero eq_refl galois_aux2 root_equals_start_points start_points_path_aux)
lemma root_equals_start_points_2:
assumes "backward_finite_path_root r x"
and "start_points x \<noteq> 0"
shows "r = start_points x"
by (metis assms inf_bot_left ss_p18 root_equals_start_points start_points_path)
lemma start_points_injective:
assumes "backward_finite_path_root r x"
shows "is_inj (start_points x)"
by (metis assms compl_bot_eq inj_def_var1 point_def backward_finite_path_root_def top_greatest
root_equals_start_points_2)
lemma backward_terminating_path_root_aux_2:
assumes "backward_finite_path_root r x"
and "start_points x \<noteq> 0 \<or> x = 0"
shows "x \<le> x\<^sup>T\<^sup>\<star>;-(x\<^sup>T;1)"
using assms bot_least backward_terminating_path_root_aux start_points_path by blast
lemma start_points_not_zero_iff:
assumes "backward_finite_path_root r x"
shows "x;r = 0 \<and> x \<noteq> 0 \<longleftrightarrow> start_points x \<noteq> 0"
by (metis assms conv_zero inf_compl_bot backward_finite_path_root_def start_points_not_zero_contra
start_points_path_aux)
text \<open>Backwards terminating and backwards finite: Part II\<close>
lemma backward_finite_path_root_acyclic_terminating_aux:
assumes "backward_finite_path_root r x"
and "is_acyclic x"
shows "x;r = 0"
proof (cases "x = 0")
assume "x = 0"
thus ?thesis
by simp
next
assume "x \<noteq> 0"
hence 1: "r \<le> x;1"
using assms(1) has_root_contra no_end_point_char backward_finite_path_root_def by blast
have "r\<cdot>(x\<^sup>T;1) = r\<cdot>(x\<^sup>T\<^sup>+;r)"
using assms(1) connected_root_iff3 backward_finite_path_root_def by fastforce
also have "... \<le> r\<cdot>(-1';r)"
by (metis assms(2) conv_compl conv_contrav conv_e conv_iso meet_isor mult_isor star_conv
star_slide_var)
also have "... = 0"
by (metis (no_types) assms(1) inj_distr annil inf_compl_bot mult_1_left point_def
backward_finite_path_root_def)
finally have "r \<le> start_points x"
using 1 galois_aux inf.boundedI le_bot by blast
thus ?thesis
using assms(1) annir le_bot start_points_path by blast
qed
lemma backward_finite_path_root_acyclic_terminating_iff:
assumes "backward_finite_path_root r x"
shows "is_acyclic x \<longleftrightarrow> x;r = 0"
apply (rule iffI)
apply (simp add: assms backward_finite_path_root_acyclic_terminating_aux)
using assms backward_finite_path_root_path_root path_root_acyclic by blast
lemma backward_finite_path_root_acyclic_terminating:
assumes "backward_finite_path_root r x"
and "is_acyclic x"
shows "backward_terminating_path_root r x"
by (simp add: assms backward_finite_path_root_acyclic_terminating_aux)
lemma non_empty_cycle_root_one_strongly_connected:
assumes "non_empty_cycle_root r x"
shows "one_strongly_connected x"
by (metis assms one_strongly_connected_iff order_trans star_1l star_star_plus sup.absorb2
non_empty_cycle_root_msc non_empty_cycle_root_start_end_points_plus)
lemma backward_finite_path_root_nodes_reachable:
assumes "backward_finite_path_root r x"
and "v \<le> x;1 + x\<^sup>T;1"
and "is_sur v"
shows "r \<le> x\<^sup>\<star>;v"
proof -
have "v \<le> x;1 + x\<^sup>T\<^sup>+;r"
using assms connected_root_iff3 backward_finite_path_root_def by fastforce
also have "... \<le> x\<^sup>T\<^sup>\<star>;r + x\<^sup>T\<^sup>+;r"
using assms(1) join_iso start_points_in_root_aux by blast
also have "... = x\<^sup>T\<^sup>\<star>;r"
using mult_isor sup.absorb1 by fastforce
finally show ?thesis
using assms(1,3)
by (simp add: inj_sur_semi_swap point_def backward_finite_path_root_def star_conv
inj_sur_semi_swap_short)
qed
lemma terminating_path_root_end_backward_terminating:
assumes "terminating_path_root_end r x e"
shows "backward_terminating_path_root r x"
using assms non_empty_cycle_root_move_root_forward_cycle
backward_finite_path_root_terminating_or_cycle by blast
lemma terminating_path_root_end_converse:
assumes "terminating_path_root_end r x e"
shows "terminating_path_root_end e (x\<^sup>T) r"
by (metis assms terminating_path_root_end_backward_terminating backward_finite_path_root_def
conv_invol terminating_path_root_end_forward_finite point_swap star_conv)
lemma terminating_path_root_end_forward_terminating:
assumes "terminating_path_root_end r x e"
shows "backward_terminating_path_root e (x\<^sup>T)"
using assms terminating_path_root_end_converse by blast
end (* relation_algebra_rtc_tarski *)
subsection \<open>Consequences with the Tarski rule and the point axiom\<close>
context relation_algebra_rtc_tarski_point
begin
text \<open>Rooted paths\<close>
lemma path_root_iff:
"(\<exists>r . path_root r x) \<longleftrightarrow> path x"
proof
assume "\<exists>r . path_root r x"
thus "path x"
using path_def path_iff_backward point_def path_root_def by blast
next
assume 1: "path x"
show "\<exists>r . path_root r x"
proof (cases "x = 0")
assume "x = 0"
thus ?thesis
by (simp add: is_inj_def is_p_fun_def point_exists path_root_def)
next
assume "\<not>(x = 0)"
hence "x;1 \<noteq> 0"
by (simp add: ss_p18)
from this obtain r where 2: "point r \<and> r \<le> x;1"
using comp_assoc is_vector_def one_idem_mult point_below_vector by fastforce
hence "r;x \<le> x;1;x"
by (simp add: mult_isor)
also have "... \<le> x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>"
using 1 path_def by blast
finally show ?thesis
using 1 2 path_def path_root_def by blast
qed
qed
lemma non_empty_path_root_iff:
"(\<exists>r . path_root r x \<and> r \<le> (x + x\<^sup>T);1) \<longleftrightarrow> path x \<and> x \<noteq> 0"
apply (rule iffI)
using non_empty_cycle_root_non_empty path_root_def zero_backward_terminating_path_root path_root_iff
apply fastforce
using path_root_iff non_empty_path_root_iff_aux by blast
text \<open>(Non-empty) Cycle\<close>
lemma non_empty_cycle_root_iff:
"(\<exists>r . non_empty_cycle_root r x) \<longleftrightarrow> cycle x \<and> x \<noteq> 0"
proof
assume "\<exists>r . non_empty_cycle_root r x"
thus "cycle x \<and> x \<noteq> 0"
using non_empty_cycle_root_msc_cycle non_empty_cycle_root_non_empty by fastforce
next
assume 1: "cycle x \<and> x \<noteq> 0"
hence "x\<^sup>T;1 \<noteq> 0"
using many_strongly_connected_implies_no_start_end_points ss_p18 by blast
from this obtain r where 2: "point r \<and> r \<le> x\<^sup>T;1"
using comp_assoc is_vector_def one_idem_mult point_below_vector by fastforce
have 3: "x\<^sup>T;1;x\<^sup>T \<le> x\<^sup>\<star>"
using 1 one_strongly_connected_iff path_def by blast
have "r;x \<le> x\<^sup>T;1;x"
using 2 by (simp add: is_vector_def mult_isor point_def)
also have "... \<le> x\<^sup>T;1;x;x\<^sup>T;x"
using comp_assoc mult_isol x_leq_triple_x by fastforce
also have "... \<le> x\<^sup>T;1;x\<^sup>T;x"
by (metis mult_assoc mult_double_iso top_greatest)
also have "... \<le> x\<^sup>\<star>;x"
using 3 mult_isor by blast
finally have "connected_root r x"
by (simp add: star_slide_var)
hence "non_empty_cycle_root r x"
using 1 2 path_def backward_finite_path_root_def by fastforce
thus "\<exists>r . non_empty_cycle_root r x" ..
qed
lemma non_empty_cycle_subset_equals:
assumes "cycle x"
and "cycle y"
and "x \<le> y"
and "x \<noteq> 0"
shows "x = y"
by (metis assms le_bot non_empty_cycle_root_subset_equals_change_root non_empty_cycle_root_iff)
lemma cycle_root_iff:
"(\<exists>r . cycle_root r x) \<longleftrightarrow> cycle x"
proof (cases "x = 0")
assume "x = 0"
thus ?thesis
using path_def point_exists by fastforce
next
assume "x \<noteq> 0"
thus ?thesis
using cycle_root_non_empty non_empty_cycle_root_iff by simp
qed
text \<open>Backwards terminating and backwards finite\<close>
lemma backward_terminating_path_root_iff:
"(\<exists>r . backward_terminating_path_root r x) \<longleftrightarrow> backward_terminating_path x"
proof
assume "\<exists>r . backward_terminating_path_root r x"
thus "backward_terminating_path x"
using backward_terminating_path_root by fastforce
next
assume 1: "backward_terminating_path x"
show "\<exists>r . backward_terminating_path_root r x"
proof (cases "x = 0")
assume "x = 0"
thus ?thesis
using point_exists zero_backward_terminating_path_root by blast
next
let ?r = "start_points x"
assume "x \<noteq> 0"
hence 2: "is_point ?r"
using 1 start_point_iff2 backward_terminating_iff1 by fastforce
have 3: "x;?r = 0"
by (metis inf_top.right_neutral modular_1_aux')
have "x;1;x \<le> x;1;x;x\<^sup>T;x"
using comp_assoc mult_isol x_leq_triple_x by fastforce
also have "... \<le> (x\<^sup>\<star> + x\<^sup>T\<^sup>\<star>);x\<^sup>T;x"
using 1 mult_isor path_def by blast
also have "... = (1' + x\<^sup>+ + x\<^sup>T\<^sup>+);x\<^sup>T;x"
by (metis star_star_plus star_unfoldl_eq sup.commute)
also have "... = x\<^sup>T;x + x\<^sup>+;x\<^sup>T;x + x\<^sup>T\<^sup>+;x\<^sup>T;x"
by (metis distrib_right' mult_onel)
also have "... = x\<^sup>T;(x + x\<^sup>T\<^sup>\<star>;x\<^sup>T;x) + x\<^sup>+;x\<^sup>T;x"
using comp_assoc distrib_left sup.commute sup.assoc by simp
also have "... \<le> x\<^sup>T;1 + x\<^sup>+;x\<^sup>T;x"
using join_iso mult_isol by fastforce
also have "... \<le> x\<^sup>T;1 + x\<^sup>+;1'"
using 1 by (metis comp_assoc join_isol mult_isol path_def is_p_fun_def)
finally have "-(x\<^sup>T;1) \<cdot> x;1;x \<le> x\<^sup>+"
by (simp add: galois_1 inf.commute)
hence "?r;x \<le> x\<^sup>+"
by (metis inf_commute one_compl ra_1)
hence "backward_terminating_path_root ?r x"
using 1 2 3 by (simp add: point_is_point backward_finite_path_root_def path_def)
thus ?thesis ..
qed
qed
lemma non_empty_backward_terminating_path_root_iff:
"backward_terminating_path_root (start_points x) x \<longleftrightarrow> backward_terminating_path x \<and> x \<noteq> 0"
apply (rule iffI)
apply (metis backward_finite_path_root_path backward_terminating_path_root_2 conv_zero
inf.cobounded1 non_empty_cycle_root_non_empty)
using backward_terminating_path_root_iff root_equals_start_points by blast
lemma non_empty_backward_terminating_path_root_iff':
"backward_finite_path_root (start_points x) x \<longleftrightarrow> backward_terminating_path x \<and> x \<noteq> 0"
using start_point_no_predecessor non_empty_backward_terminating_path_root_iff by simp
lemma backward_finite_path_root_iff:
"(\<exists>r . backward_finite_path_root r x) \<longleftrightarrow> backward_finite_path x"
proof
assume "\<exists>r . backward_finite_path_root r x"
thus "backward_finite_path x"
by (meson backward_finite_iff_msc non_empty_cycle_root_msc backward_finite_path_root_path
backward_finite_path_root_terminating_or_cycle backward_terminating_path_root)
next
assume "backward_finite_path x"
thus "\<exists>r . backward_finite_path_root r x"
by (metis backward_finite_iff_msc point_exists non_empty_cycle_root_iff
zero_backward_terminating_path_root backward_terminating_path_root_iff)
qed
lemma non_empty_backward_finite_path_root_iff:
"(\<exists>r . backward_finite_path_root r x \<and> r \<le> x;1) \<longleftrightarrow> backward_finite_path x \<and> x \<noteq> 0"
apply (rule iffI)
apply (metis backward_finite_path_root_iff annir backward_finite_path_root_def le_bot
no_end_point_char ss_p18)
using backward_finite_path_root_iff backward_finite_path_root_def point_def root_in_edge_sources by blast
text \<open>Terminating\<close>
lemma terminating_path_root_end_aux:
assumes "terminating_path x"
shows "\<exists>r e . terminating_path_root_end r x e"
proof (cases "x = 0")
assume "x = 0"
thus ?thesis
using point_exists zero_backward_terminating_path_root by fastforce
next
assume 1: "\<not>(x = 0)"
have 2: "backward_terminating_path x"
using assms by simp
from this obtain r where 3: "backward_terminating_path_root r x"
using backward_terminating_path_root_iff by blast
have "backward_terminating_path (x\<^sup>T)"
using 2 by (metis assms backward_terminating_iff1 conv_backward_terminating_path conv_invol
conv_zero inf_top.left_neutral)
from this obtain e where 4: "backward_terminating_path_root e (x\<^sup>T)"
using backward_terminating_path_root_iff by blast
have "r \<le> x;1"
using 1 3 root_in_edge_sources backward_finite_path_root_def point_def by fastforce
also have "... = x\<^sup>+;e"
using 4 connected_root_iff3 backward_finite_path_root_def by fastforce
also have "... \<le> x\<^sup>\<star>;e"
by (simp add: mult_isor)
finally show ?thesis
using 3 4 backward_finite_path_root_def by blast
qed
lemma terminating_path_root_end_iff:
"(\<exists>r e . terminating_path_root_end r x e) \<longleftrightarrow> terminating_path x"
proof
assume 1: "\<exists>r e . terminating_path_root_end r x e"
show "terminating_path x"
proof (cases "x = 0")
assume "x = 0"
thus ?thesis
by (simp add: is_inj_def is_p_fun_def path_def)
next
assume "\<not>(x = 0)"
hence 2: "\<not> many_strongly_connected x"
using 1 cycle_root_end_empty by blast
hence 3: "backward_terminating_path x"
using 1 backward_terminating_path_root terminating_path_root_end_backward_terminating by blast
have "\<exists>e . backward_finite_path_root e (x\<^sup>T)"
using 1 terminating_path_root_end_converse by blast
hence "backward_terminating_path (x\<^sup>T)"
using 1 backward_terminating_path_root terminating_path_root_end_converse by blast
hence "forward_terminating_path x"
by (simp add: conv_backward_terminating_path)
thus ?thesis
using 3 by (simp add: inf.boundedI)
qed
next
assume "terminating_path x"
thus "\<exists>r e . terminating_path_root_end r x e"
using terminating_path_root_end_aux by blast
qed
lemma non_empty_terminating_path_root_end_iff:
"terminating_path_root_end (start_points x) x (end_points x) \<longleftrightarrow> terminating_path x \<and> x \<noteq> 0"
apply (rule iffI)
apply (metis conv_zero non_empty_backward_terminating_path_root_iff terminating_path_root_end_iff)
using terminating_path_root_end_iff terminating_path_root_end_forward_terminating
root_equals_end_points terminating_path_root_end_backward_terminating root_equals_start_points
by blast
lemma non_empty_finite_path_root_end_iff:
"finite_path_root_end (start_points x) x (end_points x) \<longleftrightarrow> terminating_path x \<and> x \<noteq> 0"
using non_empty_terminating_path_root_end_iff end_point_no_successor by simp
end (* relation_algebra_rtc_tarski_point *)
end
diff --git a/thys/Residuated_Lattices/Involutive_Residuated.thy b/thys/Residuated_Lattices/Involutive_Residuated.thy
--- a/thys/Residuated_Lattices/Involutive_Residuated.thy
+++ b/thys/Residuated_Lattices/Involutive_Residuated.thy
@@ -1,138 +1,138 @@
(* Title: Involutive Residuated Structures
Author: Victor Gomes
Maintainer: Victor Gomes <vborgesferreiragomes1 at sheffield.ac.uk>
*)
section \<open>Involutive Residuated Structures\<close>
theory Involutive_Residuated
imports Residuated_Lattices
begin
class uminus' =
fixes uminus' :: "'a \<Rightarrow> 'a" ("-'' _" [81] 80)
text \<open>
Involutive posets is a structure where the double negation property holds for the
negation operations, and a Galois connection for negations exists.
\<close>
class involutive_order = order + uminus + uminus' +
assumes gn: "x \<le> -'y \<longleftrightarrow> y \<le> -x"
and dn1[simp]: "-'(-x) = x"
and dn2[simp]: "-(-'x) = x"
(* The involutive pair (-', -) is compatible with multiplication *)
class involutive_pogroupoid = order + times + involutive_order +
assumes ipg1: "x\<cdot>y \<le> z \<longleftrightarrow> (-z)\<cdot>x \<le> -y"
and ipg2: "x\<cdot>y \<le> z \<longleftrightarrow> y\<cdot>(-'z) \<le> -'x"
begin
lemma neg_antitone: "x \<le> y \<Longrightarrow> -y \<le> -x"
by (metis local.dn1 local.gn)
lemma neg'_antitone: "x \<le> y \<Longrightarrow> -'y \<le> -'x"
by (metis local.dn2 local.gn)
subclass pogroupoid
proof
fix x y z assume assm: "x \<le> y"
show "x \<cdot> z \<le> y \<cdot> z"
by (metis assm local.ipg2 local.order_refl local.order_trans neg'_antitone)
show "z \<cdot> x \<le> z \<cdot> y"
by (metis assm local.dual_order.trans local.ipg1 local.order_refl neg_antitone)
qed
abbreviation inv_resl :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"inv_resl y x \<equiv> -(x\<cdot>(-'y))"
abbreviation inv_resr :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"inv_resr x y \<equiv> -'((-y)\<cdot>x)"
sublocale residuated_pogroupoid _ _ _ inv_resl inv_resr
proof
fix x y z
show "(x \<le> - (y \<cdot> -' z)) = (x \<cdot> y \<le> z)"
by (metis local.gn local.ipg2)
show "(x \<cdot> y \<le> z) = (y \<le> -' (- z \<cdot> x))"
by (metis local.gn local.ipg1)
qed
end
class division_order = order + residual_l_op + residual_r_op +
assumes div_galois: "x \<le> z \<leftarrow> y \<longleftrightarrow> y \<le> x \<rightarrow> z"
class involutive_division_order = division_order + involutive_order +
assumes contraposition: "y \<rightarrow> -x = -'y \<leftarrow> x"
context involutive_pogroupoid begin
sublocale involutive_division_order _ _ inv_resl inv_resr
proof
fix x y z
show "(x \<le> - (y \<cdot> -' z)) = (y \<le> -' (- z \<cdot> x))"
by (metis local.gn local.ipg1 local.ipg2)
show "-' (- (- x) \<cdot> y) = - (x \<cdot> -' (-' y))"
- by (metis local.dn1 local.dn2 local.eq_iff local.gn local.jipsen1l local.jipsen1r local.resl_galois local.resr_galois)
+ by (metis local.dn1 local.dn2 order.eq_iff local.gn local.jipsen1l local.jipsen1r local.resl_galois local.resr_galois)
qed
lemma inv_resr_neg [simp]: "inv_resr (-x) (-y) = inv_resl x y"
by (metis local.contraposition local.dn1)
lemma inv_resl_neg' [simp]: "inv_resl (-'x) (-'y) = inv_resr x y"
by (metis local.contraposition local.dn2)
lemma neg'_mult_resl: "-'((-y)\<cdot>(-x)) = inv_resl x (-'y)"
by (metis inv_resr_neg local.dn2)
lemma neg_mult_resr: "-((-'y)\<cdot>(-'x)) = inv_resr (-x) y"
by (metis neg'_mult_resl)
lemma resr_de_morgan1: "-'(inv_resr (-y) (-x)) = -'(inv_resl y x)"
by (metis local.dn1 neg_mult_resr)
lemma resr_de_morgan2: "-(inv_resl (-'x) (-'y)) = -(inv_resr x y)"
by (metis inv_resl_neg')
end
text \<open>
We prove that an involutive division poset is equivalent to an involutive po-groupoid
by a lemma to avoid cyclic definitions
\<close>
lemma (in involutive_division_order) inv_pogroupoid:
"class.involutive_pogroupoid (\<lambda>x y. -(y \<rightarrow> -'x)) uminus uminus' (\<le>) (<)"
proof
fix x y z
have "(- (y \<rightarrow> -' x) \<le> z) = (-z \<le> -y \<leftarrow> x)"
by (metis local.contraposition local.dn1 local.dn2 local.gn local.div_galois)
thus "(- (y \<rightarrow> -' x) \<le> z) = (- (x \<rightarrow> -' (- z)) \<le> - y)"
by (metis local.contraposition local.div_galois local.dn1 local.dn2 local.gn)
moreover have "(- (x \<rightarrow> -' (- z)) \<le> - y) = (- (-' z \<rightarrow> -' y) \<le> -' x)"
apply (auto, metis local.contraposition local.div_galois local.dn1 local.dn2 local.gn)
by (metis local.contraposition local.div_galois local.dn1 local.dn2 local.gn)
ultimately show "(- (y \<rightarrow> -' x) \<le> z) = (- (-' z \<rightarrow> -' y) \<le> -' x)"
by metis
qed
context involutive_pogroupoid begin
definition negation_constant :: "'a \<Rightarrow> bool" where
"negation_constant a \<equiv> \<forall>x. -'x = inv_resr x a \<and> -x = inv_resl a x"
definition division_unit :: "'a \<Rightarrow> bool" where
"division_unit a \<equiv> \<forall>x. x = inv_resr a x \<and> x = inv_resl x a"
lemma neg_iff_div_unit: "(\<exists>a. negation_constant a) \<longleftrightarrow> (\<exists>b. division_unit b)"
unfolding negation_constant_def division_unit_def
apply safe
apply (rule_tac x="-a" in exI, auto)
apply (metis local.dn1 local.dn2)
apply (metis local.dn2)
apply (rule_tac x="-b" in exI, auto)
apply (metis local.contraposition)
apply (metis local.dn2)
done
end
end
diff --git a/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy b/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy
--- a/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy
+++ b/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy
@@ -1,665 +1,665 @@
(* Title: Residuated Boolean Algebras
Author: Victor Gomes <vborgesferreiragomes1 at sheffield.ac.uk>
Maintainer: Georg Struth <g.struth@sheffield.ac.uk>
*)
section \<open>Residuated Boolean Algebras\<close>
theory Residuated_Boolean_Algebras
imports Residuated_Lattices
begin
subsection \<open>Conjugation on Boolean Algebras\<close>
text \<open>
Similarly, as in the previous section, we define the conjugation for
arbitrary residuated functions on boolean algebras.
\<close>
context boolean_algebra
begin
lemma inf_bot_iff_le: "x \<sqinter> y = \<bottom> \<longleftrightarrow> x \<le> -y"
by (metis le_iff_inf inf_sup_distrib1 inf_top_right sup_bot.left_neutral sup_compl_top compl_inf_bot inf.assoc inf_bot_right)
lemma le_iff_inf_bot: "x \<le> y \<longleftrightarrow> x \<sqinter> -y = \<bottom>"
by (metis inf_bot_iff_le compl_le_compl_iff inf_commute)
lemma indirect_eq: "(\<And>z. x \<le> z \<longleftrightarrow> y \<le> z) \<Longrightarrow> x = y"
- by (metis eq_iff)
+ by (metis order.eq_iff)
text \<open>
Let $B$ be a boolean algebra. The maps $f$ and $g$ on $B$ are
a pair of conjugates if and only if for all $x, y \in B$,
$f(x) \sqcap y = \bot \Leftrightarrow x \sqcap g(t) = \bot$.
\<close>
definition conjugation_pair :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where
"conjugation_pair f g \<equiv> \<forall>x y. f(x) \<sqinter> y = \<bottom> \<longleftrightarrow> x \<sqinter> g(y) = \<bottom>"
lemma conjugation_pair_commute: "conjugation_pair f g \<Longrightarrow> conjugation_pair g f"
by (auto simp: conjugation_pair_def inf_commute)
lemma conjugate_iff_residuated: "conjugation_pair f g = residuated_pair f (\<lambda>x. -g(-x))"
apply (clarsimp simp: conjugation_pair_def residuated_pair_def inf_bot_iff_le)
by (metis double_compl)
lemma conjugate_residuated: "conjugation_pair f g \<Longrightarrow> residuated_pair f (\<lambda>x. -g(-x))"
by (metis conjugate_iff_residuated)
lemma residuated_iff_conjugate: "residuated_pair f g = conjugation_pair f (\<lambda>x. -g(-x))"
apply (clarsimp simp: conjugation_pair_def residuated_pair_def inf_bot_iff_le)
by (metis double_compl)
text \<open>
A map $f$ has a conjugate pair if and only if it is residuated.
\<close>
lemma conj_residuatedI1: "\<exists>g. conjugation_pair f g \<Longrightarrow> residuated f"
by (metis conjugate_iff_residuated residuated_def)
lemma conj_residuatedI2: "\<exists>g. conjugation_pair g f \<Longrightarrow> residuated f"
by (metis conj_residuatedI1 conjugation_pair_commute)
lemma exist_conjugateI[intro]: "residuated f \<Longrightarrow> \<exists>g. conjugation_pair f g"
by (metis residuated_def residuated_iff_conjugate)
lemma exist_conjugateI2[intro]: "residuated f \<Longrightarrow> \<exists>g. conjugation_pair g f"
by (metis exist_conjugateI conjugation_pair_commute)
text \<open>
The conjugate of a residuated function $f$ is unique.
\<close>
lemma unique_conjugate[intro]: "residuated f \<Longrightarrow> \<exists>!g. conjugation_pair f g"
proof -
{
fix g h x assume "conjugation_pair f g" and "conjugation_pair f h"
hence "g = h"
apply (unfold conjugation_pair_def)
apply (rule ext)
- apply (rule antisym)
+ apply (rule order.antisym)
by (metis le_iff_inf_bot inf_commute inf_compl_bot)+
}
moreover assume "residuated f"
ultimately show ?thesis by force
qed
lemma unique_conjugate2[intro]: "residuated f \<Longrightarrow> \<exists>!g. conjugation_pair g f"
by (metis unique_conjugate conjugation_pair_commute)
text \<open>
Since the conjugate of a residuated map is unique, we define a
conjugate operation.
\<close>
definition conjugate :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)" where
"conjugate f \<equiv> THE g. conjugation_pair g f"
lemma conjugate_iff_def: "residuated f \<Longrightarrow> f(x) \<sqinter> y = \<bottom> \<longleftrightarrow> x \<sqinter> conjugate f y = \<bottom>"
apply (clarsimp simp: conjugate_def dest!: unique_conjugate)
apply (subgoal_tac "(THE g. conjugation_pair g f) = g")
apply (clarsimp simp add: conjugation_pair_def)
apply (rule the1_equality)
by (auto intro: conjugation_pair_commute)
lemma conjugateI1: "residuated f \<Longrightarrow> f(x) \<sqinter> y = \<bottom> \<Longrightarrow> x \<sqinter> conjugate f y = \<bottom>"
by (metis conjugate_iff_def)
lemma conjugateI2: "residuated f \<Longrightarrow> x \<sqinter> conjugate f y = \<bottom> \<Longrightarrow> f(x) \<sqinter> y = \<bottom>"
by (metis conjugate_iff_def)
text \<open>
Few more lemmas about conjugation follow.
\<close>
lemma residuated_conj1: "residuated f \<Longrightarrow> conjugation_pair f (conjugate f)"
using conjugateI1 conjugateI2 conjugation_pair_def by auto
lemma residuated_conj2: "residuated f \<Longrightarrow> conjugation_pair (conjugate f) f"
using conjugateI1 conjugateI2 conjugation_pair_def inf_commute by auto
lemma conj_residuated: "residuated f \<Longrightarrow> residuated (conjugate f)"
by (force dest!: residuated_conj2 intro: conj_residuatedI1)
lemma conj_involution: "residuated f \<Longrightarrow> conjugate (conjugate f) = f"
by (metis conj_residuated residuated_conj1 residuated_conj2 unique_conjugate)
lemma residual_conj_eq: "residuated f \<Longrightarrow> residual (conjugate f) = (\<lambda>x. -f(-x))"
apply (unfold residual_def)
apply (rule the1_equality)
apply (rule residual_unique)
apply (auto intro: conj_residuated conjugate_residuated residuated_conj2)
done
lemma residual_conj_eq_ext: "residuated f \<Longrightarrow> residual (conjugate f) x = -f(-x)"
by (metis residual_conj_eq)
lemma conj_iso: "residuated f \<Longrightarrow> x \<le> y \<Longrightarrow> conjugate f x \<le> conjugate f y"
by (metis conj_residuated res_iso)
lemma conjugate_strict: "residuated f \<Longrightarrow> conjugate f \<bottom> = \<bottom>"
by (metis conj_residuated residuated_strict)
lemma conjugate_sup: "residuated f \<Longrightarrow> conjugate f (x \<squnion> y) = conjugate f x \<squnion> conjugate f y"
by (metis conj_residuated residuated_sup)
lemma conjugate_subinf: "residuated f \<Longrightarrow> conjugate f (x \<sqinter> y) \<le> conjugate f x \<sqinter> conjugate f y"
by (auto simp: conj_iso)
text \<open>
Next we prove some lemmas from Maddux's article. Similar lemmas have been proved in AFP entry
for relation algebras. They should be consolidated in the future.
\<close>
lemma maddux1: "residuated f \<Longrightarrow> f(x \<sqinter> - conjugate f(y)) \<le> f(x) \<sqinter> -y"
proof -
assume assm: "residuated f"
hence "f(x \<sqinter> - conjugate f(y)) \<le> f x"
by (metis inf_le1 res_iso)
moreover have "f(x \<sqinter> - conjugate f (y)) \<sqinter> y = \<bottom>"
by (metis assm conjugateI2 inf_bot_iff_le inf_le2)
ultimately show ?thesis
by (metis inf_bot_iff_le le_inf_iff)
qed
lemma maddux1': "residuated f \<Longrightarrow> conjugate f(x \<sqinter> -f(y)) \<le> conjugate f(x) \<sqinter> -y"
by (metis conj_involution conj_residuated maddux1)
lemma maddux2: "residuated f \<Longrightarrow> f(x) \<sqinter> y \<le> f(x \<sqinter> conjugate f y)"
proof -
assume resf: "residuated f"
obtain z where z_def: "z = f(x \<sqinter> conjugate f y)" by auto
hence "f(x \<sqinter> conjugate f y) \<sqinter> -z = \<bottom>"
by (metis inf_compl_bot)
hence "x \<sqinter> conjugate f y \<sqinter> conjugate f (-z) = \<bottom>"
by (metis conjugate_iff_def resf)
hence "x \<sqinter> conjugate f (y \<sqinter> -z) = \<bottom>"
apply (subgoal_tac "conjugate f (y \<sqinter> -z) \<le> conjugate f y \<sqinter> conjugate f (-z)")
apply (metis (no_types, hide_lams) dual_order.trans inf.commute inf_bot_iff_le inf_left_commute)
by (metis conj_iso inf_le2 inf_top.left_neutral le_inf_iff resf)
hence "f(x) \<sqinter> y \<sqinter> -z = \<bottom>"
by (metis conjugateI2 inf.assoc resf)
thus ?thesis
by (metis double_compl inf_bot_iff_le z_def)
qed
lemma maddux2': "residuated f \<Longrightarrow> conjugate f(x) \<sqinter> y \<le> conjugate f(x \<sqinter> f y)"
by (metis conj_involution conj_residuated maddux2)
lemma residuated_conjugate_ineq: "residuated f \<Longrightarrow> conjugate f x \<le> y \<longleftrightarrow> x \<le> -f(-y)"
by (metis conj_residuated residual_galois residual_conj_eq)
lemma residuated_comp_closed: "residuated f \<Longrightarrow> residuated g \<Longrightarrow> residuated (f o g)"
by (auto simp add: residuated_def residuated_pair_def)
lemma conjugate_comp: "residuated f \<Longrightarrow> residuated g \<Longrightarrow> conjugate (f o g) = conjugate g o conjugate f"
proof (rule ext, rule indirect_eq)
fix x y
assume assms: "residuated f" "residuated g"
have "conjugate (f o g) x \<le> y \<longleftrightarrow> x \<le> -f(g(-y))"
apply (subst residuated_conjugate_ineq)
using assms by (auto intro!: residuated_comp_closed)
also have "... \<longleftrightarrow> conjugate g (conjugate f x) \<le> y"
using assms by (simp add: residuated_conjugate_ineq)
finally show "(conjugate (f \<circ> g) x \<le> y) = ((conjugate g \<circ> conjugate f) x \<le> y)"
by auto
qed
lemma conjugate_comp_ext: "residuated f \<Longrightarrow> residuated g \<Longrightarrow> conjugate (\<lambda>x. f (g x)) x = conjugate g (conjugate f x)"
using conjugate_comp by (simp add: comp_def)
end (* boolean_algebra *)
context complete_boolean_algebra begin
text \<open>
On a complete boolean algebra, it is possible to give an explicit
definition of conjugation.
\<close>
lemma conjugate_eq: "residuated f \<Longrightarrow> conjugate f y = \<Sqinter>{x. y \<le> -f(-x)}"
proof -
assume assm: "residuated f" obtain g where g_def: "g = conjugate f" by auto
have "g y = \<Sqinter>{x. x \<ge> g y}"
- by (auto intro!: antisym Inf_lower Inf_greatest)
+ by (auto intro!: order.antisym Inf_lower Inf_greatest)
also have "... = \<Sqinter>{x. -x \<sqinter> g y = \<bottom>}"
by (simp add: inf_bot_iff_le)
also have "... = \<Sqinter>{x. f(-x) \<sqinter> y = \<bottom>}"
by (metis conjugate_iff_def assm g_def)
finally show ?thesis
by (simp add: g_def le_iff_inf_bot inf_commute)
qed
end (* complete_boolean_algebra *)
subsection \<open>Residuated Boolean Structures\<close>
text \<open>
In this section, we present various residuated structures based on
boolean algebras.
The left and right conjugation of the multiplicative operation is
defined, and a number of facts is derived.
\<close>
class residuated_boolean_algebra = boolean_algebra + residuated_pogroupoid
begin
subclass residuated_lgroupoid ..
definition conjugate_l :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<lhd>" 60) where
"x \<lhd> y \<equiv> -(-x \<leftarrow> y)"
definition conjugate_r :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<rhd>" 60) where
"x \<rhd> y \<equiv> -(x \<rightarrow> -y)"
lemma residual_conjugate_r: "x \<rightarrow> y = -(x \<rhd> -y)"
by (metis conjugate_r_def double_compl)
lemma residual_conjugate_l: "x \<leftarrow> y = -(-x \<lhd> y)"
by (metis conjugate_l_def double_compl)
lemma conjugation_multl: "x\<cdot>y \<sqinter> z = \<bottom> \<longleftrightarrow> x \<sqinter> (z \<lhd> y) = \<bottom>"
by (metis conjugate_l_def double_compl le_iff_inf_bot resl_galois)
lemma conjugation_multr: "x\<cdot>y \<sqinter> z = \<bottom> \<longleftrightarrow> y \<sqinter> (x \<rhd> z) = \<bottom>"
by (metis conjugate_r_def inf_bot_iff_le le_iff_inf_bot resr_galois)
lemma conjugation_conj: "(x \<lhd> y) \<sqinter> z = \<bottom> \<longleftrightarrow> y \<sqinter> (z \<rhd> x) = \<bottom>"
by (metis inf_commute conjugation_multr conjugation_multl)
lemma conjugation_pair_multl [simp]: "conjugation_pair (\<lambda>x. x\<cdot>y) (\<lambda>x. x \<lhd> y)"
by (simp add: conjugation_pair_def conjugation_multl)
lemma conjugation_pair_multr [simp]: "conjugation_pair (\<lambda>x. y\<cdot>x) (\<lambda>x. y \<rhd> x)"
by (simp add: conjugation_pair_def conjugation_multr)
lemma conjugation_pair_conj [simp]: "conjugation_pair (\<lambda>x. y \<lhd> x) (\<lambda>x. x \<rhd> y)"
by (simp add: conjugation_pair_def conjugation_conj)
lemma residuated_conjl1 [simp]: "residuated (\<lambda>x. x \<lhd> y)"
by (metis conj_residuatedI2 conjugation_pair_multl)
lemma residuated_conjl2 [simp]: "residuated (\<lambda>x. y \<lhd> x)"
by (metis conj_residuatedI1 conjugation_pair_conj)
lemma residuated_conjr1 [simp]: "residuated (\<lambda>x. y \<rhd> x)"
by (metis conj_residuatedI2 conjugation_pair_multr)
lemma residuated_conjr2 [simp]: "residuated (\<lambda>x. x \<rhd> y)"
by (metis conj_residuatedI2 conjugation_pair_conj)
lemma conjugate_multr [simp]: "conjugate (\<lambda>x. y\<cdot>x) = (\<lambda>x. y \<rhd> x)"
by (metis conjugation_pair_multr residuated_conj1 residuated_multr unique_conjugate)
lemma conjugate_conjr1 [simp]: "conjugate (\<lambda>x. y \<rhd> x) = (\<lambda>x. y\<cdot>x)"
by (metis conjugate_multr conj_involution residuated_multr)
lemma conjugate_multl [simp]: "conjugate (\<lambda>x. x\<cdot>y) = (\<lambda>x. x \<lhd> y)"
by (metis conjugation_pair_multl residuated_conj1 residuated_multl unique_conjugate)
lemma conjugate_conjl1 [simp]: "conjugate (\<lambda>x. x \<lhd> y) = (\<lambda>x. x\<cdot>y)"
proof -
have "conjugate (conjugate (\<lambda>x. x\<cdot>y)) = conjugate (\<lambda>x. x \<lhd> y)" by simp
thus ?thesis
by (metis conj_involution[OF residuated_multl])
qed
lemma conjugate_conjl2[simp]: "conjugate (\<lambda>x. y \<lhd> x) = (\<lambda>x. x \<rhd> y)"
by (metis conjugation_pair_conj unique_conjugate residuated_conj1 residuated_conjl2)
lemma conjugate_conjr2[simp]: "conjugate (\<lambda>x. x \<rhd> y) = (\<lambda>x. y \<lhd> x)"
proof -
have "conjugate (conjugate (\<lambda>x. y \<lhd> x)) = conjugate (\<lambda>x. x \<rhd> y)" by simp
thus ?thesis
by (metis conj_involution[OF residuated_conjl2])
qed
lemma conjl1_iso: "x \<le> y \<Longrightarrow> x \<lhd> z \<le> y \<lhd> z"
by (metis conjugate_l_def compl_mono resl_iso)
lemma conjl2_iso: "x \<le> y \<Longrightarrow> z \<lhd> x \<le> z \<lhd> y"
by (metis res_iso residuated_conjl2)
lemma conjr1_iso: "x \<le> y \<Longrightarrow> z \<rhd> x \<le> z \<rhd> y"
by (metis res_iso residuated_conjr1)
lemma conjr2_iso: "x \<le> y \<Longrightarrow> x \<rhd> z \<le> y \<rhd> z"
by (metis conjugate_r_def compl_mono resr_antitonel)
lemma conjl1_sup: "z \<lhd> (x \<squnion> y) = (z \<lhd> x) \<squnion> (z \<lhd> y)"
by (metis conjugate_l_def compl_inf resl_distr)
lemma conjl2_sup: "(x \<squnion> y) \<lhd> z = (x \<lhd> z) \<squnion> (y \<lhd> z)"
by (metis (poly_guards_query) residuated_sup residuated_conjl1)
lemma conjr1_sup: "z \<rhd> (x \<squnion> y) = (z \<rhd> x) \<squnion> (z \<rhd> y)"
by (metis residuated_sup residuated_conjr1)
lemma conjr2_sup: "(x \<squnion> y) \<rhd> z = (x \<rhd> z) \<squnion> (y \<rhd> z)"
by (metis conjugate_r_def compl_inf resr_distl)
lemma conjl1_strict: "\<bottom> \<lhd> x = \<bottom>"
by (metis residuated_strict residuated_conjl1)
lemma conjl2_strict: "x \<lhd> \<bottom> = \<bottom>"
by (metis residuated_strict residuated_conjl2)
lemma conjr1_strict: "\<bottom> \<rhd> x = \<bottom>"
by (metis residuated_strict residuated_conjr2)
lemma conjr2_strict: "x \<rhd> \<bottom> = \<bottom>"
by (metis residuated_strict residuated_conjr1)
lemma conjl1_iff: "x \<lhd> y \<le> z \<longleftrightarrow> x \<le> -(-z\<cdot>y)"
by (metis conjugate_l_def compl_le_swap1 compl_le_swap2 resl_galois)
lemma conjl2_iff: "x \<lhd> y \<le> z \<longleftrightarrow> y \<le> -(-z \<rhd> x)"
by (metis conjl1_iff conjugate_r_def compl_le_swap2 double_compl resr_galois)
lemma conjr1_iff: "x \<rhd> y \<le> z \<longleftrightarrow> y \<le> -(x\<cdot>-z)"
by (metis conjugate_r_def compl_le_swap1 double_compl resr_galois)
lemma conjr2_iff: "x \<rhd> y \<le> z \<longleftrightarrow> x \<le> -(y \<lhd> -z)"
by (metis conjugation_conj double_compl inf.commute le_iff_inf_bot)
text \<open>
We apply Maddux's lemmas regarding conjugation of an arbitrary residuated function
for each of the 6 functions.
\<close>
lemma maddux1a: "a\<cdot>(x \<sqinter> -(a \<rhd> y)) \<le> a\<cdot>x"
by (insert maddux1 [of "\<lambda>x. a\<cdot>x"]) simp
lemma maddux1a': "a\<cdot>(x \<sqinter> -(a \<rhd> y)) \<le> -y"
by (insert maddux1 [of "\<lambda>x. a\<cdot>x"]) simp
lemma maddux1b: "(x \<sqinter> -(y \<lhd> a))\<cdot>a \<le> x\<cdot>a"
by (insert maddux1 [of "\<lambda>x. x\<cdot>a"]) simp
lemma maddux1b': "(x \<sqinter> -(y \<lhd> a))\<cdot>a \<le> -y"
by (insert maddux1 [of "\<lambda>x. x\<cdot>a"]) simp
lemma maddux1c: " a \<lhd> x \<sqinter> -(y \<rhd> a) \<le> a \<lhd> x"
by (insert maddux1 [of "\<lambda>x. a \<lhd> x"]) simp
lemma maddux1c': "a \<lhd> x \<sqinter> -(y \<rhd> a) \<le> -y"
by (insert maddux1 [of "\<lambda>x. a \<lhd> x"]) simp
lemma maddux1d: "a \<rhd> x \<sqinter> -(a\<cdot>y) \<le> a \<rhd> x"
by (insert maddux1 [of "\<lambda>x. a \<rhd> x"]) simp
lemma maddux1d': "a \<rhd> x \<sqinter> -(a\<cdot>y) \<le> -y"
by (insert maddux1 [of "\<lambda>x. a \<rhd> x"]) simp
lemma maddux1e: "x \<sqinter> -(y\<cdot>a) \<lhd> a \<le> x \<lhd> a"
by (insert maddux1 [of "\<lambda>x. x \<lhd> a"]) simp
lemma maddux1e': "x \<sqinter> -(y\<cdot>a) \<lhd> a \<le> -y"
by (insert maddux1 [of "\<lambda>x. x \<lhd> a"]) simp
lemma maddux1f: "x \<sqinter> -(a \<lhd> y) \<rhd> a \<le> x \<rhd> a"
by (insert maddux1 [of "\<lambda>x. x \<rhd> a"]) simp
lemma maddux1f': "x \<sqinter> -(a \<lhd> y) \<rhd> a \<le> -y"
by (insert maddux1 [of "\<lambda>x. x \<rhd> a"]) simp
lemma maddux2a: "a\<cdot>x \<sqinter> y \<le> a\<cdot>(x \<sqinter> (a \<rhd> y))"
by (insert maddux2 [of "\<lambda>x. a\<cdot>x"]) simp
lemma maddux2b: "x\<cdot>a \<sqinter> y \<le> (x \<sqinter> (y \<lhd> a))\<cdot>a"
by (insert maddux2 [of "\<lambda>x. x\<cdot>a"]) simp
lemma maddux2c: "(a \<lhd> x) \<sqinter> y \<le> a \<lhd> (x \<sqinter> (y \<rhd> a))"
by (insert maddux2 [of "\<lambda>x. a \<lhd> x"]) simp
lemma maddux2d: "(a \<rhd> x) \<sqinter> y \<le> a \<rhd> (x \<sqinter> a\<cdot>y)"
by (insert maddux2 [of "\<lambda>x. a \<rhd> x"]) simp
lemma maddux2e: "(x \<lhd> a) \<sqinter> y \<le> (x \<sqinter> y\<cdot>a) \<lhd> a"
by (insert maddux2 [of "\<lambda>x. x \<lhd> a"]) simp
lemma maddux2f: "(x \<rhd> a) \<sqinter> y \<le> (x \<sqinter> (a \<lhd> y)) \<rhd> a"
by (insert maddux2 [of "\<lambda>x. x \<rhd> a"]) simp
text \<open>
The multiplicative operation $\cdot$ on a residuated boolean algebra is generally not
associative. We prove some equivalences related to associativity.
\<close>
lemma res_assoc_iff1: "(\<forall>x y z. x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z) \<longleftrightarrow> (\<forall>x y z. x \<rhd> (y \<rhd> z) = y\<cdot>x \<rhd> z)"
proof safe
fix x y z assume "\<forall>x y z. x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z"
thus "x \<rhd> (y \<rhd> z) = y \<cdot> x \<rhd> z"
using conjugate_comp_ext[of "\<lambda>z. y\<cdot>z" "\<lambda>z. x\<cdot>z"] by auto
next
fix x y z assume "\<forall>x y z. x \<rhd> (y \<rhd> z) = y\<cdot>x \<rhd> z"
thus "x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z"
using conjugate_comp_ext[of "\<lambda>z. y \<rhd> z" "\<lambda>z. x \<rhd> z"] by auto
qed
lemma res_assoc_iff2: "(\<forall>x y z. x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z) \<longleftrightarrow> (\<forall>x y z. x \<lhd> (y \<cdot> z) = (x \<lhd> z) \<lhd> y)"
proof safe
fix x y z assume "\<forall>x y z. x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z"
hence "\<forall>x y z. (x\<cdot>y)\<cdot>z = x\<cdot>(y\<cdot>z)" by simp
thus "x \<lhd> (y \<cdot> z) = (x \<lhd> z) \<lhd> y"
using conjugate_comp_ext[of "\<lambda>x. x\<cdot>z" "\<lambda>x. x\<cdot>y"] by auto
next
fix x y z assume "\<forall>x y z. x \<lhd> (y \<cdot> z) = (x \<lhd> z) \<lhd> y"
hence "\<forall>x y z. (x \<lhd> z) \<lhd> y = x \<lhd> (y \<cdot> z)" by simp
thus "x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z"
using conjugate_comp_ext[of "\<lambda>z. z \<lhd> y" "\<lambda>x. x \<lhd> z"] by auto
qed
lemma res_assoc_iff3: "(\<forall>x y z. x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z) \<longleftrightarrow> (\<forall>x y z. (x \<rhd> y) \<lhd> z = x \<rhd> (y \<lhd> z))"
proof safe
fix x y z assume "\<forall>x y z. x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z"
thus "(x \<rhd> y) \<lhd> z = x \<rhd> (y \<lhd> z)"
using conjugate_comp_ext[of "\<lambda>u. x\<cdot>u" "\<lambda>u. u\<cdot>z"] and
conjugate_comp_ext[of "\<lambda>u. u\<cdot>z" "\<lambda>u. x\<cdot>u", symmetric]
by auto
next
fix x y z assume "\<forall>x y z. (x \<rhd> y) \<lhd> z = x \<rhd> (y \<lhd> z)"
thus "x\<cdot>(y\<cdot>z) = (x\<cdot>y)\<cdot>z"
using conjugate_comp_ext[of "\<lambda>u. x \<rhd> u" "\<lambda>u. u \<lhd> z"] and
conjugate_comp_ext[of "\<lambda>u. u \<lhd> z" "\<lambda>u. x \<rhd> u", symmetric]
by auto
qed
end (* residuated_boolean_algebra *)
class unital_residuated_boolean = residuated_boolean_algebra + one +
assumes mult_onel [simp]: "x\<cdot>1 = x"
and mult_oner [simp]: "1\<cdot>x = x"
begin
text \<open>
The following equivalences are taken from J{\'o}sson and Tsinakis.
\<close>
lemma jonsson1a: "(\<exists>f. \<forall>x y. x \<rhd> y = f(x)\<cdot>y) \<longleftrightarrow> (\<forall>x y. x \<rhd> y = (x \<rhd> 1)\<cdot>y)"
apply standard
apply force
apply (rule_tac x="\<lambda>x. x \<rhd> 1" in exI)
apply force
done
lemma jonsson1b: "(\<forall>x y. x \<rhd> y = (x \<rhd> 1)\<cdot>y) \<longleftrightarrow> (\<forall>x y. x\<cdot>y = (x \<rhd> 1) \<rhd> y)"
proof safe
fix x y
assume "\<forall>x y. x \<rhd> y = (x \<rhd> 1)\<cdot>y"
hence "conjugate (\<lambda>y. x \<rhd> y) = conjugate (\<lambda>y. (x \<rhd> 1)\<cdot>y)" by metis
thus "x\<cdot>y = (x \<rhd> 1) \<rhd> y" by simp
next
fix x y
assume "\<forall>x y. x \<cdot> y = x \<rhd> 1 \<rhd> y"
thus "x \<rhd> y = (x \<rhd> 1) \<cdot> y"
by (metis mult_onel)
qed
lemma jonsson1c: "(\<forall>x y. x \<rhd> y = (x \<rhd> 1)\<cdot>y) \<longleftrightarrow> (\<forall>x y. y \<lhd> x = 1 \<lhd> (x \<lhd> y))"
proof safe
fix x y
assume "\<forall>x y. x \<rhd> y = (x \<rhd> 1)\<cdot>y"
hence "(\<lambda>x. x \<rhd> y) = (\<lambda>x. (x \<rhd> 1)\<cdot>y)" by metis
hence "(\<lambda>x. x \<rhd> y) = (\<lambda>x. x\<cdot>y) o (\<lambda>x. x \<rhd> 1)" by force
hence "conjugate (\<lambda>x. y \<lhd> x) = (\<lambda>x. x\<cdot>y) o conjugate (\<lambda>x. 1 \<lhd> x)" by simp
hence "conjugate (conjugate (\<lambda>x. y \<lhd> x)) = conjugate ((\<lambda>x. x\<cdot>y) o conjugate (\<lambda>x. 1 \<lhd> x))" by simp
hence "(\<lambda>x. y \<lhd> x) = conjugate ((\<lambda>x. x\<cdot>y) o conjugate (\<lambda>x. 1 \<lhd> x))" by simp
also have "... = conjugate (conjugate (\<lambda>x. 1 \<lhd> x)) o conjugate (\<lambda>x. x\<cdot>y)"
by (subst conjugate_comp[symmetric]) simp_all
finally show "y \<lhd> x = 1 \<lhd> (x \<lhd> y)" by simp
next
fix x y
assume "\<forall>x y. y \<lhd> x = 1 \<lhd> (x \<lhd> y)"
hence "(\<lambda>x. y \<lhd> x) = (\<lambda>x. 1 \<lhd> (x \<lhd> y))" by metis
hence "(\<lambda>x. y \<lhd> x) = (\<lambda>x. 1 \<lhd> x) o conjugate (\<lambda>x. x\<cdot>y)" by force
hence "conjugate (\<lambda>x. y \<lhd> x) = conjugate ((\<lambda>x. 1 \<lhd> x) o conjugate (\<lambda>x. x\<cdot>y))" by metis
also have "... = conjugate (conjugate (\<lambda>x. x\<cdot>y)) o conjugate (\<lambda>x. 1 \<lhd> x)"
by (subst conjugate_comp[symmetric]) simp_all
finally have "(\<lambda>x. x \<rhd> y) = (\<lambda>x. x\<cdot>y) o (\<lambda>x. x \<rhd> 1)" by simp
hence "(\<lambda>x. x \<rhd> y) = (\<lambda>x. (x \<rhd> 1) \<cdot> y)" by (simp add: comp_def)
thus "x \<rhd> y = (x \<rhd> 1) \<cdot> y" by metis
qed
lemma jonsson2a: "(\<exists>g. \<forall>x y. x \<lhd> y = x\<cdot>g(y)) \<longleftrightarrow> (\<forall>x y. x \<lhd> y = x\<cdot>(1 \<lhd> y))"
apply standard
apply force
apply (rule_tac x="\<lambda>x. 1 \<lhd> x" in exI)
apply force
done
lemma jonsson2b: "(\<forall>x y. x \<lhd> y = x\<cdot>(1 \<lhd> y)) \<longleftrightarrow> (\<forall>x y. x\<cdot>y = x \<lhd> (1 \<lhd> y))"
proof safe
fix x y
assume "\<forall>x y. x \<lhd> y = x\<cdot>(1 \<lhd> y)"
hence "conjugate (\<lambda>x. x \<lhd> y) = conjugate (\<lambda>x. x\<cdot>(1 \<lhd> y))" by metis
thus "x\<cdot>y = x \<lhd> (1 \<lhd> y)" by simp metis
next
fix x y
assume "\<forall>x y. x\<cdot>y = x \<lhd> (1 \<lhd> y)"
hence "(\<lambda>x. x\<cdot>y) = (\<lambda>x. x \<lhd> (1 \<lhd> y))" by metis
hence "conjugate (\<lambda>x. x\<cdot>y) = conjugate (\<lambda>x. x \<lhd> (1 \<lhd> y))" by metis
thus "x \<lhd> y = x \<cdot> (1 \<lhd> y)" by simp metis
qed
lemma jonsson2c: "(\<forall>x y. x \<lhd> y = x\<cdot>(1 \<lhd> y)) \<longleftrightarrow> (\<forall>x y. y \<rhd> x = (x \<rhd> y) \<rhd> 1)"
proof safe
fix x y
assume "\<forall>x y. x \<lhd> y = x\<cdot>(1 \<lhd> y)"
hence "(\<lambda>y. x \<lhd> y) = (\<lambda>y. x\<cdot>(1 \<lhd> y))" by metis
hence "(\<lambda>y. x \<lhd> y) = (\<lambda>y. x\<cdot>y) o (\<lambda>y. 1 \<lhd> y)" by force
hence "conjugate (\<lambda>y. y \<rhd> x) = (\<lambda>y. x\<cdot>y) o conjugate (\<lambda>y. y \<rhd> 1)" by force
hence "conjugate (conjugate (\<lambda>y. y \<rhd> x)) = conjugate ((\<lambda>y. x\<cdot>y) o conjugate (\<lambda>y. y \<rhd> 1))" by metis
hence "(\<lambda>y. y \<rhd> x) = conjugate ((\<lambda>y. x\<cdot>y) o conjugate (\<lambda>y. y \<rhd> 1))" by simp
also have "... = conjugate (conjugate (\<lambda>y. y \<rhd> 1)) o conjugate (\<lambda>y. x\<cdot>y)"
by (subst conjugate_comp[symmetric]) simp_all
finally have "(\<lambda>y. y \<rhd> x) = (\<lambda>y. x \<rhd> y \<rhd> 1)" by (simp add: comp_def)
thus "y \<rhd> x = x \<rhd> y \<rhd> 1" by metis
next
fix x y
assume "\<forall>x y. y \<rhd> x = x \<rhd> y \<rhd> 1"
hence "(\<lambda>y. y \<rhd> x) = (\<lambda>y. x \<rhd> y \<rhd> 1)" by force
hence "(\<lambda>y. y \<rhd> x) = (\<lambda>y. y \<rhd> 1) o conjugate (\<lambda>y. x\<cdot>y)" by force
hence "conjugate (\<lambda>y. y \<rhd> x) = conjugate ((\<lambda>y. y \<rhd> 1) o conjugate (\<lambda>y. x\<cdot>y))" by metis
also have "... = conjugate (conjugate (\<lambda>y. x\<cdot>y)) o conjugate (\<lambda>y. y \<rhd> 1)"
by (subst conjugate_comp[symmetric]) simp_all
finally have "(\<lambda>y. x \<lhd> y) = (\<lambda>y. x\<cdot>y) o (\<lambda>y. 1 \<lhd> y)"
by (metis conjugate_conjr1 conjugate_conjr2 conjugate_multr)
thus "x \<lhd> y = x \<cdot> (1 \<lhd> y)" by (simp add: comp_def)
qed
lemma jonsson3a: "(\<forall>x. (x \<rhd> 1) \<rhd> 1 = x) \<longleftrightarrow> (\<forall>x. 1 \<lhd> (1 \<lhd> x) = x)"
proof safe
fix x assume "\<forall>x. x \<rhd> 1 \<rhd> 1 = x"
thus "1 \<lhd> (1 \<lhd> x) = x"
- by (metis compl_le_swap1 compl_le_swap2 conjr2_iff eq_iff)
+ by (metis compl_le_swap1 compl_le_swap2 conjr2_iff order.eq_iff)
next
fix x assume "\<forall>x. 1 \<lhd> (1 \<lhd> x) = x"
thus "x \<rhd> 1 \<rhd> 1 = x"
by (metis conjugate_l_def conjugate_r_def double_compl jipsen2r)
qed
lemma jonsson3b: "(\<forall>x. (x \<rhd> 1) \<rhd> 1 = x) \<Longrightarrow> (x \<sqinter> y) \<rhd> 1 = (x \<rhd> 1) \<sqinter> (y \<rhd> 1)"
-proof (rule antisym, auto simp: conjr2_iso)
+proof (rule order.antisym, auto simp: conjr2_iso)
assume assm: "\<forall>x. (x \<rhd> 1) \<rhd> 1 = x"
hence "(x \<rhd> 1) \<sqinter> (y \<rhd> 1) \<rhd> 1 = x \<sqinter> (((x \<rhd> 1) \<sqinter> (y \<rhd> 1) \<rhd> 1) \<sqinter> y)"
by (metis (no_types) conjr2_iso inf.cobounded2 inf.commute inf.orderE)
hence "(x \<rhd> 1) \<sqinter> (y \<rhd> 1) \<rhd> 1 \<le> x \<sqinter> y"
using inf.orderI inf_left_commute by presburger
thus "(x \<rhd> 1) \<sqinter> (y \<rhd> 1) \<le> x \<sqinter> y \<rhd> 1"
using assm by (metis (no_types) conjr2_iso)
qed
lemma jonsson3c: "\<forall>x. (x \<rhd> 1) \<rhd> 1 = x \<Longrightarrow> x \<rhd> 1 = 1 \<lhd> x"
proof (rule indirect_eq)
fix z
assume assms: "\<forall>x. (x \<rhd> 1) \<rhd> 1 = x"
hence "(x \<rhd> 1) \<sqinter> -z = \<bottom> \<longleftrightarrow> ((x \<rhd> 1) \<sqinter> -z) \<rhd> 1 = \<bottom>"
by (metis compl_sup conjugation_conj double_compl inf_bot_right sup_bot.left_neutral)
also have "... \<longleftrightarrow> -z\<cdot>x \<sqinter> 1 = \<bottom>"
by (metis assms jonsson3b conjugation_multr)
finally have "(x \<rhd> 1) \<sqinter> -z = \<bottom> \<longleftrightarrow> (1 \<lhd> x) \<sqinter> -z = \<bottom>"
by (metis conjugation_multl inf.commute)
thus "(x \<rhd> 1 \<le> z) \<longleftrightarrow> (1 \<lhd> x \<le> z)"
by (metis le_iff_inf_bot)
qed
end (* unital_residuated_boolean *)
class residuated_boolean_semigroup = residuated_boolean_algebra + semigroup_mult
begin
subclass residuated_boolean_algebra ..
text \<open>
The following lemmas hold trivially, since they are equivalent to associativity.
\<close>
lemma res_assoc1: "x \<rhd> (y \<rhd> z) = y\<cdot>x \<rhd> z"
by (metis res_assoc_iff1 mult_assoc)
lemma res_assoc2: "x \<lhd> (y \<cdot> z) = (x \<lhd> z) \<lhd> y"
by (metis res_assoc_iff2 mult_assoc)
lemma res_assoc3: "(x \<rhd> y) \<lhd> z = x \<rhd> (y \<lhd> z)"
by (metis res_assoc_iff3 mult_assoc)
end (*residuated_boolean_semigroup *)
class residuated_boolean_monoid = residuated_boolean_algebra + monoid_mult
begin
subclass unital_residuated_boolean
by standard auto
subclass residuated_lmonoid ..
lemma jonsson4: "(\<forall>x y. x \<lhd> y = x\<cdot>(1 \<lhd> y)) \<longleftrightarrow> (\<forall>x y. x \<rhd> y = (x \<rhd> 1)\<cdot>y)"
proof safe
fix x y assume assms: "\<forall>x y. x \<lhd> y = x\<cdot>(1 \<lhd> y)"
have "x \<rhd> y = (y \<rhd> x) \<rhd> 1"
by (metis assms jonsson2c)
also have "... = (y \<rhd> ((x \<rhd> 1) \<rhd> 1)) \<rhd> 1"
by (metis assms jonsson2b jonsson3a mult_oner)
also have "... = (((x \<rhd> 1)\<cdot>y) \<rhd> 1) \<rhd> 1"
by (metis conjugate_r_def double_compl resr3)
also have "... = (x \<rhd> 1)\<cdot>y"
by (metis assms jonsson2b jonsson3a mult_oner)
finally show "x \<rhd> y = (x \<rhd> 1)\<cdot>y" .
next
fix x y assume assms: "\<forall>x y. x \<rhd> y = (x \<rhd> 1)\<cdot>y"
have "y \<lhd> x = 1 \<lhd> (x \<lhd> y)"
by (metis assms jonsson1c)
also have "... = 1 \<lhd> ((1 \<lhd> (1 \<lhd> x)) \<lhd> y)"
by (metis assms conjugate_l_def double_compl jonsson1c mult_1_right resl3)
also have "... = 1 \<lhd> (1 \<lhd> (y\<cdot>(1 \<lhd> x)))"
by (metis conjugate_l_def double_compl resl3)
also have "... = y\<cdot>(1 \<lhd> x)"
by (metis assms jonsson1b jonsson1c jonsson3c mult_onel)
finally show "y \<lhd> x = y\<cdot>(1 \<lhd> x)".
qed
end (* residuated_boolean_monoid *)
end
diff --git a/thys/Residuated_Lattices/Residuated_Lattices.thy b/thys/Residuated_Lattices/Residuated_Lattices.thy
--- a/thys/Residuated_Lattices/Residuated_Lattices.thy
+++ b/thys/Residuated_Lattices/Residuated_Lattices.thy
@@ -1,730 +1,730 @@
(* Title: Residuated Lattices
Author: Victor Gomes <vborgesferreiragomes1 at sheffield.ac.uk>
Maintainer: Georg Struth <g.struth@sheffield.ac.uk>
*)
theory Residuated_Lattices
imports Kleene_Algebra.Signatures
begin
notation
times (infixl "\<cdot>" 70)
notation
bot ("\<bottom>") and
top ("\<top>") and
inf (infixl "\<sqinter>" 65) and
sup (infixl "\<squnion>" 65) and
Inf ("\<Sqinter>_" [900] 900) and
Sup ("\<Squnion>_" [900] 900)
syntax
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10)
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10)
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
section \<open>Residuated Lattices\<close>
subsection \<open>Residuated Functions on a Partial Order\<close>
text \<open>
We follow Galatos and \emph{al.} to define residuated funtions on
partial orders. Material from articles by Maddux, and J{\'o}sson and Tsinakis are also considered.
This development should in the future be expanded to functions or categories where the sources and targets have
different type.
Let $P$ be a partial order, or a poset.
A map $f : P \to P$ is residuated if there exists a map
$g: P \to P$ such that $f(x) \le y \Leftrightarrow x \le g(y)$
for all $x, y \in P$.
That is, they are adjoints of a Galois connection.
\<close>
context order begin
definition residuated_pair :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where
"residuated_pair f g \<equiv> \<forall>x y. f(x) \<le> y \<longleftrightarrow> x \<le> g(y)"
theorem residuated_pairI [intro]:
assumes "\<forall>x y. x \<le> y \<longrightarrow> f x \<le> f y"
and "\<forall>x y. x \<le> y \<longrightarrow> g x \<le> g y"
and "\<forall>x. x \<le> (g o f) x"
and "\<forall>x. (f o g) x \<le> x"
shows "residuated_pair f g"
by (metis assms comp_apply local.order_trans residuated_pair_def antisym)
definition residuated :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where
"residuated f \<equiv> \<exists>g. residuated_pair f g"
text \<open>
If a map $f$ is residuated, then its residual $g$ is unique.
\<close>
lemma residual_unique: "residuated f \<Longrightarrow> \<exists>!g. residuated_pair f g"
unfolding residuated_def residuated_pair_def
by (metis ext eq_refl order.antisym)
text \<open>
Since the residual of a map $f$ is unique, it makes sense to
define a residual operator.
\<close>
definition residual :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)" where
"residual f \<equiv> THE g. residuated_pair f g"
lemma residual_galois: "residuated f \<Longrightarrow> f(x) \<le> y \<longleftrightarrow> x \<le> residual f y"
apply (clarsimp simp: residuated_def residuated_pair_def)
apply (subgoal_tac "residual f = g")
apply (simp_all add: residual_def)
apply (rule the1_equality)
apply (metis residuated_def residuated_pair_def residual_unique)
by (simp add: residuated_pair_def)
lemma residual_galoisI1: "residuated f \<Longrightarrow> f(x) \<le> y \<Longrightarrow> x \<le> residual f y"
by (metis residual_galois)
lemma residual_galoisI2: "residuated f \<Longrightarrow> x \<le> residual f y \<Longrightarrow> f(x) \<le> y"
by (metis residual_galois)
text \<open>
A closure operator on a poset is a map that is expansive, isotone and
idempotent.
The composition of the residual of a function $f$ with $f$ is a closure operator.
\<close>
definition closure :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where
"closure f \<equiv> (\<forall>x. x \<le> f x) \<and> (\<forall>x y. x \<le> y \<longrightarrow> f x \<le> f y) \<and> (\<forall>x. f(f x) = f x)"
lemma res_c1: "residuated f \<Longrightarrow> x \<le> residual f (f x)"
by (metis local.order.refl residual_galois)
lemma res_c2: "residuated f \<Longrightarrow> x \<le> y \<Longrightarrow> residual f (f x) \<le> residual f (f y)"
by (metis local.order.refl local.order_trans residual_galois)
lemma res_c3: "residuated f \<Longrightarrow> residual f (f (residual f (f x))) = residual f (f x)"
- by (metis local.eq_iff local.order_trans res_c1 residual_galois)
+ by (metis order.eq_iff local.order_trans res_c1 residual_galois)
lemma res_closure: "residuated f \<Longrightarrow> closure (residual f o f)"
by (auto simp: closure_def intro: res_c1 res_c2 res_c3)
text \<open>
Dually, an interior operator on a poset is a map that is contractive, isotone and
idempotent.
The composition of $f$ with its residual is an interior operator.
\<close>
definition interior :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where
"interior f \<equiv> (\<forall>x. f x \<le> x) \<and> (\<forall>x y. x \<le> y \<longrightarrow> f x \<le> f y) \<and> (\<forall>x. f(f x) = f x)"
lemma res_i1: "residuated f \<Longrightarrow> f (residual f x) \<le> x"
by (metis local.order.refl residual_galois)
lemma res_i2: "residuated f \<Longrightarrow> x \<le> y \<Longrightarrow> f (residual f x) \<le> f (residual f y)"
by (metis local.order.refl local.order_trans residual_galois)
lemma res_i3: "residuated f \<Longrightarrow> f (residual f (f (residual f x))) = f (residual f x)"
by (metis local.antisym_conv res_c1 res_c3 res_i1 residual_galois)
lemma res_interior: "residuated f \<Longrightarrow> interior (f o residual f)"
by (auto simp: interior_def intro: res_i1 res_i2 res_i3)
text \<open>Next we show a few basic lemmas about residuated maps.\<close>
lemma res_iso: "residuated f \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
by (metis local.order.trans res_c1 residual_galois)
lemma res_residual_iso: "residuated f \<Longrightarrow> x \<le> y \<Longrightarrow> residual f x \<le> residual f y"
by (metis local.order.trans res_i1 residual_galois)
lemma res_comp1 [simp]: "residuated f \<Longrightarrow> f o residual f o f = f"
proof -
assume resf: "residuated f"
{
fix x
have "(f o residual f o f) x = f x"
- by (metis resf comp_apply local.eq_iff res_c1 res_i1 res_iso)
+ by (metis resf comp_apply order.eq_iff res_c1 res_i1 res_iso)
}
thus ?thesis
by (metis ext)
qed
lemma res_comp2 [simp]: "residuated f \<Longrightarrow> residual f o f o residual f = residual f"
proof -
assume resf: "residuated f"
{
fix x
have "(residual f o f o residual f) x = residual f x"
- by (metis resf comp_apply local.eq_iff res_c1 res_i1 res_residual_iso)
+ by (metis resf comp_apply order.eq_iff res_c1 res_i1 res_residual_iso)
}
thus ?thesis
by (metis ext)
qed
end (* order *)
text \<open>
A residuated function $f$ preserves all existing joins.
Dually, its residual preserves all existing meets.
We restrict our attention to semilattices, where binary joins or meets exist, and to
complete lattices, where arbitrary joins and meets exist.
\<close>
lemma (in semilattice_sup) residuated_sup: "residuated f \<Longrightarrow> f (x \<squnion> y) = f x \<squnion> f y"
-proof (rule antisym)
+proof (rule order.antisym)
assume assm: "residuated f"
thus "f (x \<squnion> y) \<le> f x \<squnion> f y"
by (metis local.residual_galoisI1 local.residual_galoisI2 local.sup.bounded_iff local.sup_ge1 local.sup_ge2)
show "f x \<squnion> f y \<le> f (x \<squnion> y)"
by (metis assm local.res_iso local.sup.bounded_iff local.sup_ge1 local.sup_ge2)
qed
lemma (in semilattice_inf) residuated_inf: "residuated f \<Longrightarrow> residual f (x \<sqinter> y) = residual f x \<sqinter> residual f y"
-proof (rule antisym)
+proof (rule order.antisym)
assume assm: "residuated f"
thus "residual f (x \<sqinter> y) \<le> residual f x \<sqinter> residual f y"
by (metis local.inf.boundedI local.inf.cobounded1 local.inf.cobounded2 local.res_residual_iso)
show "residual f x \<sqinter> residual f y \<le> residual f (x \<sqinter> y)"
by (metis assm local.inf.bounded_iff local.inf.cobounded1 local.inf.cobounded2 local.residual_galoisI1 local.residual_galoisI2)
qed
context bounded_lattice begin
lemma residuated_strict: "residuated f \<Longrightarrow> f \<bottom> = \<bottom>"
by (metis local.bot_least local.bot_unique local.res_i1 local.res_iso)
lemma res_top: "residuated f \<Longrightarrow> residual f \<top> = \<top>"
by (metis local.residual_galoisI1 local.top_greatest local.top_unique)
end
context complete_lattice begin
text \<open>
On a complete lattice, a function $f$ is residuated if and only if
it preserves arbitrary (possibly infinite) joins.
Dually, a function $g$ is a residual of a residuated funtion $f$
if and only if $g$ preserver arbitrary meets.
\<close>
lemma residual_eq1: "residuated f \<Longrightarrow> residual f y = \<Squnion> {x. f x \<le> y}"
-proof (rule antisym)
+proof (rule order.antisym)
assume assm: "residuated f"
thus "residual f y \<le> \<Squnion>{x. f x \<le> y}"
by (auto simp: res_i1 intro!: Sup_upper)
show "\<Squnion>{x. f x \<le> y} \<le> residual f y"
by (auto simp: assm intro!: Sup_least residual_galoisI1)
qed
lemma residual_eq2: "residuated f \<Longrightarrow> f x = \<Sqinter> {y. x \<le> residual f y}"
-proof (rule antisym)
+proof (rule order.antisym)
assume assm: "residuated f"
thus "f x \<le> \<Sqinter>{y. x \<le> residual f y}"
by (auto intro: Inf_greatest residual_galoisI2)
show "\<Sqinter>{y. x \<le> residual f y} \<le> f x"
using assm by (auto simp: res_c1 intro!: Inf_lower)
qed
lemma residuated_Sup: "residuated f \<Longrightarrow> f(\<Squnion> X) = \<Squnion>{f x | x. x \<in> X}"
-proof (rule antisym)
+proof (rule order.antisym)
assume assm: "residuated f"
obtain y where y_def: "y = \<Squnion>{f x |x. x \<in> X}"
by auto
hence "\<forall>x \<in> X. f x \<le> y"
by (auto simp: y_def intro: Sup_upper)
hence "\<forall>x \<in> X. x \<le> residual f y"
by (auto simp: assm intro!: residual_galoisI1)
hence "\<Squnion>X \<le> residual f y"
by (auto intro: Sup_least)
thus "f(\<Squnion> X) \<le> \<Squnion>{f x |x. x \<in> X}"
by (metis y_def assm residual_galoisI2)
qed (clarsimp intro!: Sup_least res_iso Sup_upper)
lemma residuated_Inf: "residuated f \<Longrightarrow> residual f(\<Sqinter> X) = \<Sqinter>{residual f x | x. x \<in> X}"
-proof (rule antisym, clarsimp intro!: Inf_greatest res_residual_iso Inf_lower)
+proof (rule order.antisym, clarsimp intro!: Inf_greatest res_residual_iso Inf_lower)
assume assm: "residuated f"
obtain y where y_def: "y = \<Sqinter>{residual f x |x. x \<in> X}"
by auto
hence "\<forall>x \<in> X. y \<le> residual f x"
by (auto simp: y_def intro: Inf_lower)
hence "\<forall>x \<in> X. f y \<le> x"
by (metis assm residual_galoisI2)
hence "f y \<le> \<Sqinter> X"
by (auto intro: Inf_greatest)
thus "\<Sqinter>{residual f x |x. x \<in> X} \<le> residual f (\<Sqinter>X)"
by (auto simp: assm y_def intro!: residual_galoisI1)
qed
lemma Sup_sup: "\<forall>X. f(\<Squnion> X) = \<Squnion>{f x | x. x \<in> X} \<Longrightarrow> f (x \<squnion> y) = f x \<squnion> f y"
apply (erule_tac x="{x, y}" in allE)
by (force intro: Sup_eqI)
lemma Sup_residuatedI: "\<forall>X. f(\<Squnion> X) = \<Squnion>{f x | x. x \<in> X} \<Longrightarrow> residuated f"
proof (unfold residuated_def residuated_pair_def, standard+)
fix x y
assume "f x \<le> y"
thus "x \<le> \<Squnion>{x. f x \<le> y}"
by (clarsimp intro!: Sup_upper)
next
fix x y
assume assm: "\<forall>X. f (\<Squnion>X) = \<Squnion>{f x |x. x \<in> X}"
hence f_iso: "\<forall>x y. x \<le> y \<longrightarrow> f x \<le> f y"
using Sup_sup by (auto simp: le_iff_sup)
assume "x \<le> \<Squnion>{x. f x \<le> y}"
hence "f x \<le> f(\<Squnion>{x. f x \<le> y})"
by (metis f_iso)
also have "... = \<Squnion>{f x | x . f x \<le> y}"
using assm by auto
finally show "f x \<le> y"
apply (rule order_trans)
by (auto intro!: Sup_least)
qed
lemma Inf_inf: "\<forall>X. f(\<Sqinter> X) = \<Sqinter>{f x | x. x \<in> X} \<Longrightarrow> f (x \<sqinter> y) = f x \<sqinter> f y"
apply (erule_tac x="{x, y}" in allE)
by (force intro: Inf_eqI)
lemma Inf_residuatedI: "\<forall>X. \<Sqinter>{g x | x. x \<in> X} = g (\<Sqinter> X) \<Longrightarrow> \<exists>f. residuated_pair f g"
proof (unfold residuated_pair_def, standard+)
fix x y
assume "x \<le> g y"
thus "\<Sqinter>{y. x \<le> g y} \<le> y"
by (clarsimp intro!: Inf_lower)
next
fix x y
assume assm: "\<forall>X. \<Sqinter>{g x | x. x \<in> X} = g (\<Sqinter> X)"
hence g_iso: "\<forall>x y. x \<le> y \<longrightarrow> g x \<le> g y"
using Inf_inf by (auto simp: le_iff_inf)
assume "\<Sqinter>{y. x \<le> g y} \<le> y"
hence "g (\<Sqinter>{y. x \<le> g y}) \<le> g y"
by (metis g_iso)
hence "(\<Sqinter>{g y | y. x \<le> g y}) \<le> g y"
using assm apply (erule_tac x="{y. x \<le> g y}" in allE)
by (auto intro!: Inf_lower)
thus "x \<le> g y"
apply (rule local.dual_order.trans)
by (auto intro: Inf_greatest)
qed
end (* complete lattices *)
subsection \<open>Residuated Structures\<close>
text \<open>
In this section, we define residuated algebraic structures, starting from the simplest of all,
a \emph{residuated partial ordered groupoid}, to
\emph{residuated l-monoids}, which are residuated
lattices where the multiplicative operator forms a monoid.
\<close>
class pogroupoid = order + times +
assumes mult_isor: "x \<le> y \<Longrightarrow> x \<cdot> z \<le> y \<cdot> z"
and mult_isol: "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
text \<open>
A residuated partial ordered groupoid is simply a partial order and a multiplicative groupoid with
two extra operators satisfying the residuation laws.
It is straighforward to prove that multiplication is compatible with the order,
that is, multiplication is isotone.
Most of the lemmas below come in pairs; they are related by opposition duality.
Formalising this duality is left for future work.
\<close>
class residual_r_op =
fixes residual_r :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<rightarrow>" 60)
class residual_l_op =
fixes residual_l :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<leftarrow>" 60)
class residuated_pogroupoid = order + times + residual_l_op + residual_r_op +
assumes resl_galois: "x \<le> z \<leftarrow> y \<longleftrightarrow> x \<cdot> y \<le> z"
and resr_galois: "x \<cdot> y \<le> z \<longleftrightarrow> y \<le> x \<rightarrow> z"
begin
lemma reslI [intro]: "x \<cdot> y \<le> z \<Longrightarrow> x \<le> z \<leftarrow> y"
by (metis resl_galois)
lemma resrI [intro]: "x \<cdot> y \<le> z \<Longrightarrow> y \<le> x \<rightarrow> z"
by (metis resr_galois)
lemma residuated_pair_multl [simp]: "residuated_pair (\<lambda>x. x \<cdot> y) (\<lambda>x. x \<leftarrow> y)"
by (auto simp: residuated_pair_def resl_galois)
lemma residuated_pair_multr [simp]: "residuated_pair (\<lambda>y. x \<cdot> y) (\<lambda>y. x \<rightarrow> y)"
by (auto simp: residuated_pair_def resr_galois)
text \<open>
Multiplication is then obviously residuated.
\<close>
lemma residuated_multl [simp]: "residuated (\<lambda>x. x \<cdot> y)"
by (metis residuated_def residuated_pair_multl)
lemma residuated_multr [simp]: "residuated (\<lambda>y. x \<cdot> y)"
by (metis residuated_def residuated_pair_multr)
lemma resl_eq [simp]: "residual (\<lambda>x. x \<cdot> y) = (\<lambda>x. x \<leftarrow> y)"
unfolding residual_def apply (rule the1_equality)
by (auto simp: intro!: residual_unique)
lemma resr_eq [simp]: "residual (\<lambda>y. x \<cdot> y) = (\<lambda>y. x \<rightarrow> y)"
unfolding residual_def apply (rule the1_equality)
by (auto simp: intro!: residual_unique)
text \<open>
Next we prove a few lemmas, all of which are instantiation of more general facts about residuated functions.
\<close>
lemma res_lc1: "x \<le> x \<cdot> y \<leftarrow> y"
by auto
lemma res_lc2: "x \<le> y \<Longrightarrow> x \<cdot> z \<leftarrow> z \<le> y \<cdot> z \<leftarrow> z"
by (metis local.res_c2 resl_eq residuated_multl)
lemma res_lc3 [simp]: "(x \<cdot> y \<leftarrow> y) \<cdot> y \<leftarrow> y = x \<cdot> y \<leftarrow> y"
by (metis local.res_c3 resl_eq residuated_multl)
lemma res_rc1: "x \<le> y \<rightarrow> y \<cdot> x"
by auto
lemma res_rc2: "x \<le> y \<Longrightarrow> z \<rightarrow> z \<cdot> x \<le> z \<rightarrow> z \<cdot> y"
by (metis local.res_c2 resr_eq residuated_multr)
lemma res_rc3 [simp]: "y \<rightarrow> y \<cdot> (y \<rightarrow> y \<cdot> x) = y \<rightarrow> y \<cdot> x"
by (metis local.res_c3 resr_eq residuated_multr)
lemma res_li1: "(x \<leftarrow> y) \<cdot> y \<le> x"
by (metis local.res_i1 resl_eq residuated_multl)
lemma res_li2: "x \<le> y \<Longrightarrow> (x \<leftarrow> z) \<cdot> z \<le> (y \<leftarrow> z) \<cdot> z"
by (metis local.res_i2 resl_eq residuated_multl)
lemma res_li3 [simp]: "((x \<leftarrow> y) \<cdot> y \<leftarrow> y) \<cdot> y = (x \<leftarrow> y) \<cdot> y"
by (metis local.res_i3 resl_eq residuated_multl)
lemma res_ri1: "y \<cdot> (y \<rightarrow> x) \<le> x"
by (metis local.res_i1 resr_eq residuated_multr)
lemma res_ri2: "x \<le> y \<Longrightarrow> z \<cdot> (z \<rightarrow> x) \<le> z \<cdot> (z \<rightarrow> y)"
by (metis local.res_i2 resr_eq residuated_multr)
lemma res_ri3 [simp]: "y \<cdot> (y \<rightarrow> y \<cdot> (y \<rightarrow> x)) = y \<cdot> (y \<rightarrow> x)"
by (metis local.res_i3 resr_eq residuated_multr)
subclass pogroupoid
proof
fix x y z
show "x \<le> y \<Longrightarrow> x \<cdot> z \<le> y \<cdot> z"
by (metis local.res_iso residuated_multl)
show "x \<le> y \<Longrightarrow> z \<cdot> x \<le> z \<cdot> y"
by (metis local.res_iso residuated_multr)
qed
lemma resl_iso: "x \<le> y \<Longrightarrow> x \<leftarrow> z \<le> y \<leftarrow> z"
by (metis res_residual_iso resl_eq residuated_multl)
lemma resr_iso: "x \<le> y \<Longrightarrow> z \<rightarrow> x \<le> z \<rightarrow> y"
by (metis res_residual_iso resr_eq residuated_multr)
lemma resl_comp1 [simp]: "(x \<cdot> y \<leftarrow> y) \<cdot> y = x \<cdot> y"
- by (metis local.antisym local.mult_isor res_lc1 res_li1)
+ by (metis order.antisym local.mult_isor res_lc1 res_li1)
lemma resl_comp2 [simp]: "(x \<leftarrow> y) \<cdot> y \<leftarrow> y = x \<leftarrow> y"
- by (metis local.eq_iff res_lc1 res_li1 resl_iso)
+ by (metis order.eq_iff res_lc1 res_li1 resl_iso)
lemma resr_comp1 [simp]: "y \<cdot> (y \<rightarrow> y \<cdot> x) = y \<cdot> x"
- by (metis local.antisym local.mult_isol res_rc1 res_ri1)
+ by (metis order.antisym local.mult_isol res_rc1 res_ri1)
lemma resr_comp2 [simp]: "y \<rightarrow> y \<cdot> (y \<rightarrow> x) = y \<rightarrow> x"
- by (metis local.eq_iff res_rc1 res_ri1 resr_iso)
+ by (metis order.eq_iff res_rc1 res_ri1 resr_iso)
lemma resl_antitoner: "x \<le> y \<longrightarrow> z \<leftarrow> y \<le> z \<leftarrow> x"
by (metis local.dual_order.trans local.mult_isol res_li1 reslI)
lemma resr_antitonel: "x \<le> y \<longrightarrow> y \<rightarrow> z \<le> x \<rightarrow> z"
by (metis local.dual_order.trans local.resl_galois res_ri1 resrI)
text \<open>The following lemmas are taken from Galatos and \emph{al.}\<close>
lemma jipsen1l: "x \<le> y \<leftarrow> (x \<rightarrow> y)"
by (metis res_ri1 reslI)
lemma jipsen1r: "x \<le> (y \<leftarrow> x) \<rightarrow> y"
by (metis res_li1 resrI)
lemma jipsen2l: "(y \<leftarrow> (x \<rightarrow> y)) \<rightarrow> y = x \<rightarrow> y"
- by (metis jipsen1l jipsen1r local.eq_iff local.resr_antitonel)
+ by (metis jipsen1l jipsen1r order.eq_iff local.resr_antitonel)
lemma jipsen2r: "y \<leftarrow> ((y \<leftarrow> x) \<rightarrow> y) = y \<leftarrow> x"
- by (metis jipsen1l jipsen1r local.eq_iff local.resl_antitoner)
+ by (metis jipsen1l jipsen1r order.eq_iff local.resl_antitoner)
end (* residuated_pogroupoid *)
text \<open>
In a residuated partial ordered semigroup, the multiplicative operator is associative.
\<close>
class residuated_posemigroup = semigroup_mult + residuated_pogroupoid
begin
lemma resl_trans: "(x \<leftarrow> y) \<cdot> (y \<leftarrow> z) \<le> x \<leftarrow> z"
by (metis local.res_li1 local.resl_antitoner local.resl_galois mult_assoc)
lemma resr_trans: "(x \<rightarrow> y) \<cdot> (y \<rightarrow> z) \<le> x \<rightarrow> z"
by (metis local.res_ri1 local.resr_antitonel local.resr_galois mult_assoc)
lemma resr1: "(x \<rightarrow> y) \<cdot> z \<le> (x \<rightarrow> y \<cdot> z)"
by (metis local.mult_isor local.res_ri1 local.resrI mult_assoc)
lemma resr2: "x \<rightarrow> y \<le> z \<cdot> x \<rightarrow> z \<cdot> y"
by (metis local.mult_isol local.res_ri1 local.resrI mult_assoc)
lemma resr3: "x \<cdot> y \<rightarrow> z = y \<rightarrow> (x \<rightarrow> z)"
- by (metis local.eq_iff local.resr_galois mult_assoc)
+ by (metis order.eq_iff local.resr_galois mult_assoc)
lemma resl1: "z \<cdot> (x \<leftarrow> y) \<le> (z \<cdot> x \<leftarrow> y)"
by (metis local.mult_isol local.res_li1 local.reslI mult_assoc)
lemma resl2: "x \<leftarrow> y \<le> x \<cdot> z \<leftarrow> y \<cdot> z"
by (metis local.mult_isor local.res_li1 local.reslI mult_assoc)
lemma resl3: "x \<leftarrow> y \<cdot> z = (x \<leftarrow> z) \<leftarrow> y"
- by (metis local.eq_iff local.resl_galois mult_assoc)
+ by (metis order.eq_iff local.resl_galois mult_assoc)
lemma residual_assoc: "x \<rightarrow> (y \<leftarrow> z) = (x \<rightarrow> y) \<leftarrow> z"
-proof (rule antisym)
+proof (rule order.antisym)
show "x \<rightarrow> (y \<leftarrow> z) \<le> (x \<rightarrow> y) \<leftarrow> z"
by (metis local.res_ri1 local.resl_galois local.resr_galois mult_assoc)
show "(x \<rightarrow> y) \<leftarrow> z \<le> x \<rightarrow> (y \<leftarrow> z)"
by (metis local.res_li1 local.resl_galois local.resr_galois mult_assoc)
qed
end (* residuated_posemigroup *)
text \<open>
In a residuated partially ordered monoid, the multiplicative operator has a unit $1$; that is,
its reduct forms a monoid.
\<close>
class residuated_pomonoid = monoid_mult + residuated_pogroupoid
begin
subclass residuated_posemigroup ..
lemma resl_unit: "x \<leftarrow> 1 = x"
by (metis local.mult_1_right local.resl_comp1)
lemma resr_unit: "1 \<rightarrow> x = x"
by (metis local.mult_1_left local.resr_comp1)
lemma mult_one_resl: "x \<cdot> (1 \<leftarrow> y) \<le> x \<leftarrow> y"
by (metis local.mult_1_right local.resl1)
lemma mult_one_resr: "(x \<rightarrow> 1) \<cdot> y \<le> x \<rightarrow> y"
by (metis local.mult_1_left local.resr1)
text \<open>More lemmas from Galatos and \emph{al.} follow.\<close>
lemma jipsen3l: "1 \<le> x \<leftarrow> x"
by (metis local.jipsen1l resr_unit)
lemma jipsen3r: "1 \<le> x \<rightarrow> x"
by (metis local.jipsen1r resl_unit)
lemma jipsen4l [simp]: "(x \<leftarrow> x) \<cdot> x = x"
by (metis local.mult_1_left local.resl_comp1)
lemma jipsen4r [simp]: "x \<cdot> (x \<rightarrow> x) = x"
by (metis local.mult_1_right local.resr_comp1)
lemma jipsen5l [simp]: "(x \<leftarrow> x) \<cdot> (x \<leftarrow> x) = x \<leftarrow> x"
by (metis jipsen4l local.resl3)
lemma jipsen5r [simp]: "(x \<rightarrow> x) \<cdot> (x \<rightarrow> x) = x \<rightarrow> x"
by (metis jipsen4r local.resr3)
lemma jipsen6l: "y \<leftarrow> x \<le> (y \<leftarrow> z) \<leftarrow> (x \<leftarrow> z)"
by (metis local.resl_galois local.resl_trans)
lemma jipsen6r: "x \<rightarrow> y \<le> (z \<rightarrow> x) \<rightarrow> (z \<rightarrow> y)"
by (metis local.resr_galois local.resr_trans)
lemma jipsen7l: "y \<leftarrow> x \<le> (z \<leftarrow> y) \<rightarrow> (z \<leftarrow> x)"
by (metis local.resr_galois local.resl_trans)
lemma jipsen7r: "x \<rightarrow> y \<le> (x \<rightarrow> z) \<leftarrow> (y \<rightarrow> z)"
by (metis local.resl_galois local.resr_trans)
end (* residuated_pomonoid *)
text \<open>
Residuated partially ordered groupoids can be expanded residuated join semilattice ordered groupoids.
They are used as a base for action algebras, which are expansions
of Kleene algebras by operations of residuation.
Action algebras can be found in the AFP entry for Kleene algebras.
\<close>
class residuated_sup_lgroupoid = semilattice_sup + residuated_pogroupoid
begin
lemma distl: "x \<cdot> (y \<squnion> z) = x \<cdot> y \<squnion> x \<cdot> z"
by (metis local.residuated_multr local.residuated_sup)
lemma distr: "(x \<squnion> y) \<cdot> z = x \<cdot> z \<squnion> y \<cdot> z"
by (metis local.residuated_multl local.residuated_sup)
lemma resl_subdist_var: "x \<leftarrow> z \<le> (x \<squnion> y) \<leftarrow> z"
by (metis local.resl_iso local.sup_ge1)
lemma resl_subdist: "(x \<leftarrow> z) \<squnion> (y \<leftarrow> z) \<le> (x \<squnion> y) \<leftarrow> z"
by (metis local.le_sup_iff local.sup_commute resl_subdist_var)
lemma resr_subdist_var: "(x \<rightarrow> y) \<le> x \<rightarrow> (y \<squnion> z)"
by (metis local.resr_iso local.sup_ge1)
lemma resr_subdist: "(x \<rightarrow> y) \<squnion> (x \<rightarrow> z) \<le> x \<rightarrow> (y \<squnion> z)"
by (metis sup_commute sup_least resr_subdist_var)
lemma resl_superdist_var: "x \<leftarrow> (y \<squnion> z) \<le> x \<leftarrow> y"
by (metis local.le_sup_iff local.res_li1 local.reslI local.resr_galois)
lemma resr_superdist_var: "(x \<squnion> y) \<rightarrow> z \<le> x \<rightarrow> z"
by (metis local.le_sup_iff local.res_ri1 local.resl_galois local.resrI)
end (* residuated_sup_lgroupoid *)
text \<open>
Similarly, semigroup and monoid variants can be defined.
\<close>
class residuated_sup_lsemigroup = semilattice_sup + residuated_posemigroup
subclass (in residuated_sup_lsemigroup) residuated_sup_lgroupoid ..
class residuated_sup_lmonoid = semilattice_sup + residuated_posemigroup
subclass (in residuated_sup_lmonoid) residuated_sup_lsemigroup ..
text \<open>
By lattice duality, we define residuated meet semillatice ordered groupoid.
\<close>
class residuated_inf_lgroupoid = semilattice_inf + residuated_pogroupoid
begin
lemma resl_dist: "(x \<sqinter> y) \<leftarrow> z = (x \<leftarrow> z) \<sqinter> (y \<leftarrow> z)"
by (metis local.resl_eq local.residuated_inf local.residuated_multl)
lemma resr_dist: "x \<rightarrow> (y \<sqinter> z) = (x \<rightarrow> y) \<sqinter> (x \<rightarrow> z)"
by (metis local.resr_eq local.residuated_inf local.residuated_multr)
lemma resl_inf_superdist: "x \<leftarrow> y \<le> x \<leftarrow> (y \<sqinter> z)"
by (metis local.inf_le1 local.resl_antitoner)
lemma resr_inf_superdist_var: "x \<rightarrow> y \<le> (x \<sqinter> z) \<rightarrow> y"
by (metis local.inf_le1 local.resr_antitonel)
end (* residuated_inf_lgroupoid *)
class residuated_inf_lsemigroup = semilattice_inf + residuated_posemigroup
subclass (in residuated_inf_lsemigroup) residuated_inf_lgroupoid ..
class residuated_inf_lmonoid = semilattice_inf + residuated_posemigroup
subclass (in residuated_inf_lmonoid) residuated_inf_lsemigroup ..
text \<open>
Finally, we obtain residuated lattice groupoids.
\<close>
class residuated_lgroupoid = lattice + residuated_pogroupoid
begin
subclass residuated_sup_lgroupoid ..
lemma resl_distr: "z \<leftarrow> (x \<squnion> y) = (z \<leftarrow> x) \<sqinter> (z \<leftarrow> y)"
-proof (rule antisym)
+proof (rule order.antisym)
show "z \<leftarrow> x \<squnion> y \<le> (z \<leftarrow> x) \<sqinter> (z \<leftarrow> y)"
by (metis local.inf.bounded_iff local.resl_superdist_var local.sup_commute)
show "(z \<leftarrow> x) \<sqinter> (z \<leftarrow> y) \<le> z \<leftarrow> x \<squnion> y"
by (metis local.inf.cobounded1 local.inf.cobounded2 local.resl_galois local.resr_galois local.sup.bounded_iff)
qed
lemma resr_distl: "(x \<squnion> y) \<rightarrow> z = (x \<rightarrow> z) \<sqinter> (y \<rightarrow> z)"
-proof (rule antisym)
+proof (rule order.antisym)
show "x \<squnion> y \<rightarrow> z \<le> (x \<rightarrow> z) \<sqinter> (y \<rightarrow> z)"
by (metis local.inf.bounded_iff local.resr_antitonel local.resr_superdist_var local.sup_ge2)
show "(x \<rightarrow> z) \<sqinter> (y \<rightarrow> z) \<le> x \<squnion> y \<rightarrow> z"
by (metis local.inf.cobounded1 local.inf.cobounded2 local.resl_galois local.resr_galois local.sup_least)
qed
end (* residuated_lgroupoid *)
class residuated_lsemigroup = lattice + residuated_sup_lsemigroup
subclass (in residuated_lsemigroup) residuated_lgroupoid ..
class residuated_lmonoid = lattice + residuated_sup_lmonoid
subclass (in residuated_lmonoid) residuated_lsemigroup ..
class residuated_blgroupoid = bounded_lattice + residuated_pogroupoid
begin
lemma multl_strict [simp]: "x \<cdot> \<bottom> = \<bottom>"
by (metis local.residuated_multr local.residuated_strict)
lemma multr_strict [simp]: "\<bottom> \<cdot> x = \<bottom>"
by (metis local.residuated_multl local.residuated_strict)
lemma resl_top [simp]: "\<top> \<leftarrow> x = \<top>"
by (metis local.res_top local.residuated_multl local.resl_eq)
lemma resl_impl [simp]: "x \<leftarrow> \<bottom> = \<top>"
by (metis local.resl_comp2 multl_strict resl_top)
lemma resr_top [simp]: "x \<rightarrow> \<top> = \<top>"
by (metis local.resrI local.top_greatest local.top_unique)
lemma resr_impl [simp]: "\<bottom> \<rightarrow> x = \<top>"
by (metis local.resr_comp2 multr_strict resr_top)
end (* residuated_blgroupoid *)
class residuated_blsemigroup = bounded_lattice + residuated_sup_lsemigroup
subclass (in residuated_blsemigroup) residuated_blgroupoid ..
class residuated_blmonoid = bounded_lattice + residuated_sup_lmonoid
subclass (in residuated_blmonoid) residuated_blsemigroup ..
class residuated_clgroupoid = complete_lattice + residuated_pogroupoid
begin
lemma resl_eq_def: "y \<leftarrow> x = \<Squnion> {z. z \<cdot> x \<le> y}"
by (metis local.residual_eq1 local.residuated_multl local.resl_eq)
lemma resr_eq_def: "x \<rightarrow> y = \<Squnion> {z. x \<cdot> z \<le> y}"
by (metis local.residual_eq1 local.residuated_multr local.resr_eq)
lemma multl_def: "x \<cdot> y = \<Sqinter> {z. x \<le> z \<leftarrow> y}"
proof -
have "\<Sqinter>{ya. x \<le> residual (\<lambda>a. a \<cdot> y) ya} = \<Sqinter>{z. x \<le> z \<leftarrow> y}"
by simp
thus ?thesis
by (metis residuated_multl residual_eq2)
qed
lemma multr_def: "x \<cdot> y = \<Sqinter> {z. y \<le> x \<rightarrow> z}"
proof -
have "\<Sqinter>{ya. y \<le> residual ((\<cdot>) x) ya} = \<Sqinter>{z. y \<le> x \<rightarrow> z}"
by simp
thus ?thesis
by (metis residuated_multr residual_eq2)
qed
end (* residuated_clgroupoid *)
class residuated_clsemigroup = complete_lattice + residuated_sup_lsemigroup
subclass (in residuated_clsemigroup) residuated_clgroupoid ..
class residuated_clmonoid = complete_lattice + residuated_sup_lmonoid
subclass (in residuated_clmonoid) residuated_clsemigroup ..
end
diff --git a/thys/Signature_Groebner/More_MPoly.thy b/thys/Signature_Groebner/More_MPoly.thy
--- a/thys/Signature_Groebner/More_MPoly.thy
+++ b/thys/Signature_Groebner/More_MPoly.thy
@@ -1,304 +1,305 @@
(* Author: Alexander Maletzky *)
section \<open>More Properties of Power-Products and Multivariate Polynomials\<close>
theory More_MPoly
imports Prelims Polynomials.MPoly_Type_Class_Ordered
begin
subsection \<open>Power-Products\<close>
lemma (in comm_powerprod) minus_plus': "s adds t \<Longrightarrow> u + (t - s) = (u + t) - s"
using add_commute minus_plus by auto
context ulcs_powerprod
begin
lemma lcs_alt_2:
assumes "a + x = b + y"
shows "lcs x y = (b + y) - gcs a b"
proof -
have "a + (lcs x y + gcs a b) = lcs (a + x) (a + y) + gcs a b" by (simp only: lcs_plus_left ac_simps)
also have "... = lcs (b + y) (a + y) + gcs a b" by (simp only: assms)
also have "... = (lcs a b + y) + gcs a b" by (simp only: lcs_plus_right lcs_comm)
also have "... = (gcs a b + lcs a b) + y" by (simp only: ac_simps)
also have "... = a + (b + y)" by (simp only: gcs_plus_lcs, simp add: ac_simps)
finally have "(lcs x y + gcs a b) - gcs a b = (b + y) - gcs a b" by simp
thus ?thesis by simp
qed
corollary lcs_alt_1:
assumes "a + x = b + y"
shows "lcs x y = (a + x) - gcs a b"
proof -
have "lcs x y = lcs y x" by (simp only: lcs_comm)
also from assms[symmetric] have "... = (a + x) - gcs b a" by (rule lcs_alt_2)
also have "... = (a + x) - gcs a b" by (simp only: gcs_comm)
finally show ?thesis .
qed
corollary lcs_minus_1:
assumes "a + x = b + y"
shows "lcs x y - x = a - gcs a b"
by (simp add: lcs_alt_1[OF assms] diff_right_commute)
corollary lcs_minus_2:
assumes "a + x = b + y"
shows "lcs x y - y = b - gcs a b"
by (simp add: lcs_alt_2[OF assms] diff_right_commute)
lemma gcs_minus:
assumes "u adds s" and "u adds t"
shows "gcs (s - u) (t - u) = gcs s t - u"
proof -
from assms have "gcs s t = gcs ((s - u) + u) ((t - u) + u)" by (simp add: minus_plus)
also have "... = gcs (s - u) (t - u) + u" by (simp only: gcs_plus_right)
finally show ?thesis by simp
qed
corollary gcs_minus_gcs: "gcs (s - (gcs s t)) (t - (gcs s t)) = 0"
by (simp add: gcs_minus gcs_adds gcs_adds_2)
end (* ulcs_powerprod *)
subsection \<open>Miscellaneous\<close>
lemma poly_mapping_rangeE:
assumes "c \<in> Poly_Mapping.range p"
obtains k where "k \<in> keys p" and "c = lookup p k"
using assms by (transfer, auto)
lemma poly_mapping_range_nonzero: "0 \<notin> Poly_Mapping.range p"
by (transfer, auto)
lemma (in term_powerprod) Keys_range_vectorize_poly: "Keys (Poly_Mapping.range (vectorize_poly p)) = pp_of_term ` keys p"
proof
show "Keys (Poly_Mapping.range (vectorize_poly p)) \<subseteq> pp_of_term ` keys p"
proof
fix t
assume "t \<in> Keys (Poly_Mapping.range (vectorize_poly p))"
then obtain q where "q \<in> Poly_Mapping.range (vectorize_poly p)" and "t \<in> keys q" by (rule in_KeysE)
from this(1) obtain k where q: "q = lookup (vectorize_poly p) k" by (metis DiffE imageE range.rep_eq)
with \<open>t \<in> keys q\<close> have "term_of_pair (t, k) \<in> keys p"
by (metis in_keys_iff lookup_proj_poly lookup_vectorize_poly)
hence "pp_of_term (term_of_pair (t, k)) \<in> pp_of_term ` keys p" by (rule imageI)
thus "t \<in> pp_of_term ` keys p" by (simp only: pp_of_term_of_pair)
qed
next
show "pp_of_term ` keys p \<subseteq> Keys (Poly_Mapping.range (vectorize_poly p))"
proof
fix t
assume "t \<in> pp_of_term ` keys p"
then obtain x where "x \<in> keys p" and "t = pp_of_term x" ..
from this(2) have "term_of_pair (t, component_of_term x) = x" by (simp add: term_of_pair_pair)
with \<open>x \<in> keys p\<close> have "lookup p (term_of_pair (t, component_of_term x)) \<noteq> 0"
by (simp add: in_keys_iff)
hence "lookup (proj_poly (component_of_term x) p) t \<noteq> 0" by (simp add: lookup_proj_poly)
hence t: "t \<in> keys (proj_poly (component_of_term x) p)"
by (simp add: in_keys_iff)
from \<open>x \<in> keys p\<close> have "component_of_term x \<in> keys (vectorize_poly p)"
by (simp add: keys_vectorize_poly)
from t show "t \<in> Keys (Poly_Mapping.range (vectorize_poly p))"
proof (rule in_KeysI)
have "proj_poly (component_of_term x) p = lookup (vectorize_poly p) (component_of_term x)"
by (simp only: lookup_vectorize_poly)
also from \<open>component_of_term x \<in> keys (vectorize_poly p)\<close>
have "... \<in> Poly_Mapping.range (vectorize_poly p)" by (rule in_keys_lookup_in_range)
finally show "proj_poly (component_of_term x) p \<in> Poly_Mapping.range (vectorize_poly p)" .
qed
qed
qed
subsection \<open>@{const ordered_term.lt} and @{const ordered_term.higher}\<close>
context ordered_term
begin
lemma lt_lookup_vectorize: "punit.lt (lookup (vectorize_poly p) (component_of_term (lt p))) = lp p"
proof (cases "p = 0")
case True
thus ?thesis by (simp add: vectorize_zero min_term_def pp_of_term_of_pair)
next
case False
show ?thesis
proof (rule punit.lt_eqI_keys)
from False have "lt p \<in> keys p" by (rule lt_in_keys)
thus "lp p \<in> keys (lookup (vectorize_poly p) (component_of_term (lt p)))"
by (simp add: lookup_vectorize_poly keys_proj_poly)
next
fix t
assume "t \<in> keys (lookup (vectorize_poly p) (component_of_term (lt p)))"
also have "... = pp_of_term ` {x\<in>keys p. component_of_term x = component_of_term (lt p)}"
by (simp only: lookup_vectorize_poly keys_proj_poly)
finally obtain v where "v \<in> keys p" and 1: "component_of_term v = component_of_term (lt p)"
and t: "t = pp_of_term v" by auto
from this(1) have "v \<preceq>\<^sub>t lt p" by (rule lt_max_keys)
show "t \<preceq> lp p"
proof (rule ccontr)
assume "\<not> t \<preceq> lp p"
hence "lp p \<prec> pp_of_term v" by (simp add: t)
hence "lp p \<noteq> pp_of_term v" and "lp p \<preceq> pp_of_term v" by simp_all
note this(2)
moreover from 1 have "component_of_term (lt p) \<le> component_of_term v" by simp
ultimately have "lt p \<preceq>\<^sub>t v" by (rule ord_termI)
- with \<open>v \<preceq>\<^sub>t lt p\<close> have "v = lt p" by (rule ord_term_lin.antisym)
+ with \<open>v \<preceq>\<^sub>t lt p\<close> have "v = lt p"
+ by simp
with \<open>lp p \<noteq> pp_of_term v\<close> show False by simp
qed
qed
qed
lemma lower_higher_zeroI: "u \<preceq>\<^sub>t v \<Longrightarrow> lower (higher p v) u = 0"
by (simp add: lower_eq_zero_iff lookup_higher_when)
lemma lookup_minus_higher: "lookup (p - higher p v) u = (lookup p u when u \<preceq>\<^sub>t v)"
by (auto simp: lookup_minus lookup_higher_when when_def)
lemma keys_minus_higher: "keys (p - higher p v) = {u \<in> keys p. u \<preceq>\<^sub>t v}"
by (rule set_eqI, simp add: lookup_minus_higher conj_commute flip: lookup_not_eq_zero_eq_in_keys)
lemma lt_minus_higher: "v \<in> keys p \<Longrightarrow> lt (p - higher p v) = v"
by (rule lt_eqI_keys, simp_all add: keys_minus_higher)
lemma lc_minus_higher: "v \<in> keys p \<Longrightarrow> lc (p - higher p v) = lookup p v"
by (simp add: lc_def lt_minus_higher lookup_minus_higher)
lemma tail_minus_higher: "v \<in> keys p \<Longrightarrow> tail (p - higher p v) = lower p v"
by (rule poly_mapping_eqI, simp add: lookup_tail_when lt_minus_higher lookup_lower_when lookup_minus_higher cong: when_cong)
end (* ordered_term *)
subsection \<open>@{const gd_term.dgrad_p_set}\<close>
lemma (in gd_term) dgrad_p_set_closed_mult_scalar:
assumes "dickson_grading d" and "p \<in> punit.dgrad_p_set d m" and "r \<in> dgrad_p_set d m"
shows "p \<odot> r \<in> dgrad_p_set d m"
proof (rule dgrad_p_setI)
fix v
assume "v \<in> keys (p \<odot> r)"
then obtain t u where "t \<in> keys p" and "u \<in> keys r" and v: "v = t \<oplus> u"
by (rule in_keys_mult_scalarE)
from assms(2) \<open>t \<in> keys p\<close> have "d t \<le> m" by (rule punit.dgrad_p_setD[simplified])
moreover from assms(3) \<open>u \<in> keys r\<close> have "d (pp_of_term u) \<le> m" by (rule dgrad_p_setD)
ultimately have "d (t + pp_of_term u) \<le> m" using assms(1) by (simp add: dickson_gradingD1)
thus "d (pp_of_term v) \<le> m" by (simp only: v pp_of_term_splus)
qed
subsection \<open>Regular Sequences\<close>
definition is_regular_sequence :: "('a::comm_powerprod \<Rightarrow>\<^sub>0 'b::comm_ring_1) list \<Rightarrow> bool"
where "is_regular_sequence fs \<longleftrightarrow> (\<forall>j<length fs. \<forall>q. q * fs ! j \<in> ideal (set (take j fs)) \<longrightarrow>
q \<in> ideal (set (take j fs)))"
lemma is_regular_sequenceD:
"is_regular_sequence fs \<Longrightarrow> j < length fs \<Longrightarrow> q * fs ! j \<in> ideal (set (take j fs)) \<Longrightarrow>
q \<in> ideal (set (take j fs))"
by (simp add: is_regular_sequence_def)
lemma is_regular_sequence_Nil: "is_regular_sequence []"
by (simp add: is_regular_sequence_def)
lemma is_regular_sequence_snocI:
assumes "\<And>q. q * f \<in> ideal (set fs) \<Longrightarrow> q \<in> ideal (set fs)" and "is_regular_sequence fs"
shows "is_regular_sequence (fs @ [f])"
proof (simp add: is_regular_sequence_def, intro impI allI)
fix j q
assume 1: "j < Suc (length fs)" and 2: "q * (fs @ [f]) ! j \<in> ideal (set (take j fs))"
show "q \<in> ideal (set (take j fs))"
proof (cases "j = length fs")
case True
from 2 have "q * f \<in> ideal (set fs)" by (simp add: True)
hence "q \<in> ideal (set fs)" by (rule assms(1))
thus ?thesis by (simp add: True)
next
case False
with 1 have "j < length fs" by simp
with 2 have "q * fs ! j \<in> ideal (set (take j fs))" by (simp add: nth_append)
with assms(2) \<open>j < length fs\<close> show "q \<in> ideal (set (take j fs))" by (rule is_regular_sequenceD)
qed
qed
lemma is_regular_sequence_snocD:
assumes "is_regular_sequence (fs @ [f])"
shows "\<And>q. q * f \<in> ideal (set fs) \<Longrightarrow> q \<in> ideal (set fs)"
and "is_regular_sequence fs"
proof -
fix q
assume 1: "q * f \<in> ideal (set fs)"
note assms
moreover have "length fs < length (fs @ [f])" by simp
moreover from 1 have "q * (fs @ [f]) ! (length fs) \<in> ideal (set (take (length fs) (fs @ [f])))"
by simp
ultimately have "q \<in> ideal (set (take (length fs) (fs @ [f])))" by (rule is_regular_sequenceD)
thus "q \<in> ideal (set fs)" by simp
next
show "is_regular_sequence fs" unfolding is_regular_sequence_def
proof (intro impI allI)
fix j q
assume 1: "j < length fs" and 2: "q * fs ! j \<in> ideal (set (take j fs))"
note assms
moreover from 1 have "j < length (fs @ [f])" by simp
moreover from 1 2 have "q * (fs @ [f]) ! j \<in> ideal (set (take j (fs @ [f])))"
by (simp add: nth_append)
ultimately have "q \<in> ideal (set (take j (fs @ [f])))" by (rule is_regular_sequenceD)
with 1 show "q \<in> ideal (set (take j fs))" by simp
qed
qed
lemma is_regular_sequence_removeAll_zero:
assumes "is_regular_sequence fs"
shows "is_regular_sequence (removeAll 0 fs)"
using assms
proof (induct fs rule: rev_induct)
case Nil
show ?case by (simp add: is_regular_sequence_Nil)
next
case (snoc f fs)
have "set (removeAll 0 fs) = set fs - {0}" by simp
also have "ideal ... = ideal (set fs)" by (fact ideal.span_Diff_zero)
finally have eq: "ideal (set (removeAll 0 fs)) = ideal (set fs)" .
from snoc(2) have *: "is_regular_sequence fs" by (rule is_regular_sequence_snocD)
show ?case
proof (simp, intro conjI impI)
show "is_regular_sequence (removeAll 0 fs @ [f])"
proof (rule is_regular_sequence_snocI)
fix q
assume "q * f \<in> ideal (set (removeAll 0 fs))"
hence "q * f \<in> ideal (set fs)" by (simp only: eq)
with snoc(2) have "q \<in> ideal (set fs)" by (rule is_regular_sequence_snocD)
thus "q \<in> ideal (set (removeAll 0 fs))" by (simp only: eq)
next
from * show "is_regular_sequence (removeAll 0 fs)" by (rule snoc.hyps)
qed
next
from * show "is_regular_sequence (removeAll 0 fs)" by (rule snoc.hyps)
qed
qed
lemma is_regular_sequence_remdups:
assumes "is_regular_sequence fs"
shows "is_regular_sequence (rev (remdups (rev fs)))"
using assms
proof (induct fs rule: rev_induct)
case Nil
show ?case by (simp add: is_regular_sequence_Nil)
next
case (snoc f fs)
from snoc(2) have *: "is_regular_sequence fs" by (rule is_regular_sequence_snocD)
show ?case
proof (simp, intro conjI impI)
show "is_regular_sequence (rev (remdups (rev fs)) @ [f])"
proof (rule is_regular_sequence_snocI)
fix q
assume "q * f \<in> ideal (set (rev (remdups (rev fs))))"
hence "q * f \<in> ideal (set fs)" by simp
with snoc(2) have "q \<in> ideal (set fs)" by (rule is_regular_sequence_snocD)
thus "q \<in> ideal (set (rev (remdups (rev fs))))" by simp
next
from * show "is_regular_sequence (rev (remdups (rev fs)))" by (rule snoc.hyps)
qed
next
from * show "is_regular_sequence (rev (remdups (rev fs)))" by (rule snoc.hyps)
qed
qed
end (* theory *)
diff --git a/thys/Signature_Groebner/Signature_Examples.thy b/thys/Signature_Groebner/Signature_Examples.thy
--- a/thys/Signature_Groebner/Signature_Examples.thy
+++ b/thys/Signature_Groebner/Signature_Examples.thy
@@ -1,378 +1,378 @@
(* Author: Alexander Maletzky *)
section \<open>Sample Computations with Signature-Based Algorithms\<close>
theory Signature_Examples
imports Signature_Groebner Groebner_Bases.Benchmarks Groebner_Bases.Code_Target_Rat
begin
subsection \<open>Setup\<close>
lift_definition except_pp :: "('a, 'b) pp \<Rightarrow> 'a set \<Rightarrow> ('a, 'b::zero) pp" is except .
lemma hom_grading_varnum_pp: "hom_grading (varnum_pp::('a::countable, 'b::add_wellorder) pp \<Rightarrow> nat)"
proof -
define f where "f = (\<lambda>n t. (except_pp t (- {x. elem_index x < n}))::('a, 'b) pp)"
show ?thesis unfolding hom_grading_def hom_grading_fun_def
proof (intro exI allI conjI impI)
fix n s t
show "f n (s + t) = f n s + f n t" unfolding f_def by transfer (rule except_plus)
next
fix n t
show "varnum_pp (f n t) \<le> n" unfolding f_def
by transfer (simp add: varnum_le_iff keys_except)
next
fix n t
show "varnum_pp t \<le> n \<Longrightarrow> f n t = t" unfolding f_def
by transfer (auto simp: except_id_iff varnum_le_iff)
qed
qed
instance pp :: (countable, add_wellorder) quasi_pm_powerprod
by (standard, intro exI conjI, fact dickson_grading_varnum_pp, fact hom_grading_varnum_pp)
subsubsection \<open>Projections of Term Orders to Orders on Power-Products\<close>
definition proj_comp :: "(('a::nat, 'b::nat) pp \<times> nat) nat_term_order \<Rightarrow> ('a, 'b) pp \<Rightarrow> ('a, 'b) pp \<Rightarrow> order"
where "proj_comp cmp = (\<lambda>x y. nat_term_compare cmp (x, 0) (y, 0))"
definition proj_ord :: "(('a::nat, 'b::nat) pp \<times> nat) nat_term_order \<Rightarrow> ('a, 'b) pp nat_term_order"
where "proj_ord cmp = Abs_nat_term_order (proj_comp cmp)"
text \<open>In principle, @{const proj_comp} and @{const proj_ord} could be defined more generally on type
@{typ "'a \<times> nat"}, but then @{typ 'a} would have to belong to some new type-class which is the
intersection of @{class nat_pp_term} and @{class nat_pp_compare} and additionally requires
@{prop "rep_nat_term x = (rep_nat_pp x, 0)"}.\<close>
lemma comparator_proj_comp: "comparator (proj_comp cmp)"
proof -
interpret cmp: comparator "nat_term_compare cmp" by (rule comparator_nat_term_compare)
show ?thesis unfolding proj_comp_def
proof
fix x y :: "('a, 'b) pp"
show "invert_order (nat_term_compare cmp (x, 0) (y, 0)) = nat_term_compare cmp (y, 0) (x, 0)"
by (simp only: cmp.sym)
next
fix x y :: "('a, 'b) pp"
assume "nat_term_compare cmp (x, 0) (y, 0) = Eq"
hence "(x, 0) = (y, 0::nat)" by (rule cmp.weak_eq)
thus "x = y" by simp
next
fix x y z :: "('a, 'b) pp"
assume "nat_term_compare cmp (x, 0) (y, 0) = Lt" and "nat_term_compare cmp (y, 0) (z, 0) = Lt"
- thus "nat_term_compare cmp (x, 0) (z, 0) = Lt" by (rule cmp.trans)
+ thus "nat_term_compare cmp (x, 0) (z, 0) = Lt" by (rule cmp.comp_trans)
qed
qed
lemma nat_term_comp_proj_comp: "nat_term_comp (proj_comp cmp)"
proof -
have 1: "fst (rep_nat_term (u, i)) = rep_nat_pp u" for u::"('a, 'b) pp" and i::nat
by (simp add: rep_nat_term_prod_def)
have 2: "snd (rep_nat_term (u, i)) = i" for u::"('a, 'b) pp" and i::nat
by (simp add: rep_nat_term_prod_def rep_nat_nat_def)
show ?thesis
proof (rule nat_term_compI)
fix u v :: "('a, 'b) pp"
assume a: "fst (rep_nat_term u) = 0"
note nat_term_comp_nat_term_compare
moreover have "snd (rep_nat_term (u, 0::nat)) = snd (rep_nat_term (v, 0::nat))" by (simp only: 2)
moreover from a have "fst (rep_nat_term (u, 0::nat)) = 0" by (simp add: 1 rep_nat_term_pp_def)
ultimately have "nat_term_compare cmp (u, 0) (v, 0) \<noteq> Gt" by (rule nat_term_compD1)
thus "proj_comp cmp u v \<noteq> Gt" by (simp add: proj_comp_def)
next
fix u v :: "('a, 'b) pp"
assume "snd (rep_nat_term u) < snd (rep_nat_term v)"
thus "proj_comp cmp u v = Lt" by (simp add: rep_nat_term_pp_def)
next
fix t u v :: "('a, 'b) pp"
assume "proj_comp cmp u v = Lt"
hence "nat_term_compare cmp (u, 0) (v, 0) = Lt" by (simp add: proj_comp_def)
with nat_term_comp_nat_term_compare have "nat_term_compare cmp (splus (t, 0) (u, 0)) (splus (t, 0) (v, 0)) = Lt"
by (rule nat_term_compD3)
thus "proj_comp cmp (splus t u) (splus t v) = Lt"
by (simp add: proj_comp_def splus_prod_def pprod.splus_def splus_pp_term)
next
fix u v a b :: "('a, 'b) pp"
assume u: "fst (rep_nat_term u) = fst (rep_nat_term a)" and v: "fst (rep_nat_term v) = fst (rep_nat_term b)"
and a: "proj_comp cmp a b = Lt"
note nat_term_comp_nat_term_compare
moreover from u have "fst (rep_nat_term (u, 0::nat)) = fst (rep_nat_term (a, 0::nat))"
by (simp add: 1 rep_nat_term_pp_def)
moreover from v have "fst (rep_nat_term (v, 0::nat)) = fst (rep_nat_term (b, 0::nat))"
by (simp add: 1 rep_nat_term_pp_def)
moreover have "snd (rep_nat_term (u, 0::nat)) = snd (rep_nat_term (v, 0::nat))"
and "snd (rep_nat_term (a, 0::nat)) = snd (rep_nat_term (b, 0::nat))" by (simp_all only: 2)
moreover from a have "nat_term_compare cmp (a, 0) (b, 0) = Lt" by (simp add: proj_comp_def)
ultimately have "nat_term_compare cmp (u, 0) (v, 0) = Lt" by (rule nat_term_compD4)
thus "proj_comp cmp u v = Lt" by (simp add: proj_comp_def)
qed
qed
corollary nat_term_compare_proj_ord: "nat_term_compare (proj_ord cmp) = proj_comp cmp"
unfolding proj_ord_def using comparator_proj_comp nat_term_comp_proj_comp
by (rule nat_term_compare_Abs_nat_term_order_id)
lemma proj_ord_LEX [code]: "proj_ord LEX = LEX"
proof -
have "nat_term_compare (proj_ord LEX) = nat_term_compare LEX"
by (auto simp: nat_term_compare_proj_ord nat_term_compare_LEX proj_comp_def lex_comp
lex_comp_aux_def rep_nat_term_prod_def rep_nat_term_pp_def intro!: ext split: order.split)
thus ?thesis by (simp only: nat_term_compare_inject)
qed
lemma proj_ord_DRLEX [code]: "proj_ord DRLEX = DRLEX"
proof -
have "nat_term_compare (proj_ord DRLEX) = nat_term_compare DRLEX"
by (auto simp: nat_term_compare_proj_ord nat_term_compare_DRLEX proj_comp_def deg_comp pot_comp
lex_comp lex_comp_aux_def rep_nat_term_prod_def rep_nat_term_pp_def intro!: ext split: order.split)
thus ?thesis by (simp only: nat_term_compare_inject)
qed
lemma proj_ord_DEG [code]: "proj_ord (DEG to) = DEG (proj_ord to)"
proof -
have "nat_term_compare (proj_ord (DEG to)) = nat_term_compare (DEG (proj_ord to))"
by (simp add: nat_term_compare_proj_ord nat_term_compare_DEG proj_comp_def deg_comp
rep_nat_term_prod_def rep_nat_term_pp_def)
thus ?thesis by (simp only: nat_term_compare_inject)
qed
lemma proj_ord_POT [code]: "proj_ord (POT to) = proj_ord to"
proof -
have "nat_term_compare (proj_ord (POT to)) = nat_term_compare (proj_ord to)"
by (simp add: nat_term_compare_proj_ord nat_term_compare_POT proj_comp_def pot_comp
rep_nat_term_prod_def rep_nat_term_pp_def)
thus ?thesis by (simp only: nat_term_compare_inject)
qed
subsubsection \<open>Locale Interpretation\<close>
locale qpm_nat_inf_term = gd_nat_term "\<lambda>x. x" "\<lambda>x. x" to
for to::"(('a::nat, 'b::nat) pp \<times> nat) nat_term_order"
begin
sublocale aux: qpm_inf_term "\<lambda>x. x" "\<lambda>x. x"
"le_of_nat_term_order (proj_ord to)"
"lt_of_nat_term_order (proj_ord to)"
"le_of_nat_term_order to"
"lt_of_nat_term_order to"
proof intro_locales
(*
show "class.preorder (le_of_nat_term_order (proj_ord to)) (lt_of_nat_term_order (proj_ord to))"
and "class.order_axioms (le_of_nat_term_order (proj_ord to))"
and "class.linorder_axioms (le_of_nat_term_order (proj_ord to))"
using linorder_le_of_nat_term_order[of "proj_ord to"] by (simp_all add: class.linorder_def class.order_def)
next*)
show "ordered_powerprod_axioms (le_of_nat_term_order (proj_ord to))"
by (unfold_locales, fact le_of_nat_term_order_zero_min, auto dest: le_of_nat_term_order_plus_monotone simp: ac_simps)
next
show "ordered_term_axioms (\<lambda>x. x) (\<lambda>x. x) (le_of_nat_term_order (proj_ord to)) (le_of_nat_term_order to)"
proof
fix v w t
assume "le_of_nat_term_order to v w"
thus "le_of_nat_term_order to (local.splus t v) (local.splus t w)"
by (simp add: le_of_nat_term_order nat_term_compare_splus splus_eq_splus)
next
fix v w
assume "le_of_nat_term_order (proj_ord to) (pp_of_term v) (pp_of_term w)"
and "component_of_term v \<le> component_of_term w"
hence "nat_term_compare to (fst v, 0) (fst w, 0) \<noteq> Gt" and "snd v \<le> snd w"
by (simp_all add: le_of_nat_term_order nat_term_compare_proj_ord proj_comp_def)
from comparator_nat_term_compare nat_term_comp_nat_term_compare _ _ _ _ this(1)
have "nat_term_compare to v w \<noteq> Gt"
by (rule nat_term_compD4'') (simp_all add: rep_nat_term_prod_def ord_iff[symmetric] \<open>snd v \<le> snd w\<close>)
thus "le_of_nat_term_order to v w" by (simp add: le_of_nat_term_order)
qed
qed
end
text \<open>We must define the following two constants outside the global interpretation, since otherwise
their types are too general.\<close>
definition splus_pprod :: "('a::nat, 'b::nat) pp \<Rightarrow> _"
where "splus_pprod = pprod.splus"
definition adds_term_pprod :: "(('a::nat, 'b::nat) pp \<times> _) \<Rightarrow> _"
where "adds_term_pprod = pprod.adds_term"
global_interpretation pprod': qpm_nat_inf_term to
rewrites "pprod.pp_of_term = fst"
and "pprod.component_of_term = snd"
and "pprod.splus = splus_pprod"
and "pprod.adds_term = adds_term_pprod"
and "punit.monom_mult = monom_mult_punit"
and "pprod'.aux.punit.lt = lt_punit (proj_ord to)"
and "pprod'.aux.punit.lc = lc_punit (proj_ord to)"
and "pprod'.aux.punit.tail = tail_punit (proj_ord to)"
for to :: "(('a::nat, 'b::nat) pp \<times> nat) nat_term_order"
defines max_pprod = pprod'.ord_term_lin.max
and Koszul_syz_sigs_aux_pprod = pprod'.aux.Koszul_syz_sigs_aux
and Koszul_syz_sigs_pprod = pprod'.aux.Koszul_syz_sigs
and find_sig_reducer_pprod = pprod'.aux.find_sig_reducer
and sig_trd_spp_body_pprod = pprod'.aux.sig_trd_spp_body
and sig_trd_spp_aux_pprod = pprod'.aux.sig_trd_spp_aux
and sig_trd_spp_pprod = pprod'.aux.sig_trd_spp
and spair_sigs_spp_pprod = pprod'.aux.spair_sigs_spp
and is_pred_syz_pprod = pprod'.aux.is_pred_syz
and is_rewritable_spp_pprod = pprod'.aux.is_rewritable_spp
and sig_crit_spp_pprod = pprod'.aux.sig_crit_spp
and spair_spp_pprod = pprod'.aux.spair_spp
and spp_of_pair_pprod = pprod'.aux.spp_of_pair
and pair_ord_spp_pprod = pprod'.aux.pair_ord_spp
and sig_of_pair_spp_pprod = pprod'.aux.sig_of_pair_spp
and new_spairs_spp_pprod = pprod'.aux.new_spairs_spp
and is_regular_spair_spp_pprod = pprod'.aux.is_regular_spair_spp
and add_spairs_spp_pprod = pprod'.aux.add_spairs_spp
and is_pot_ord_pprod = pprod'.is_pot_ord
and new_syz_sigs_spp_pprod = pprod'.aux.new_syz_sigs_spp
and rb_spp_body_pprod = pprod'.aux.rb_spp_body
and rb_spp_aux_pprod = pprod'.aux.rb_spp_aux
and gb_sig_z_pprod' = pprod'.aux.gb_sig_z
and gb_sig_pprod' = pprod'.aux.gb_sig
and rw_rat_strict_pprod = pprod'.aux.rw_rat_strict
and rw_add_strict_pprod = pprod'.aux.rw_add_strict
subgoal by (rule qpm_nat_inf_term.intro, fact gd_nat_term_id)
subgoal by (fact pprod_pp_of_term)
subgoal by (fact pprod_component_of_term)
subgoal by (simp only: splus_pprod_def)
subgoal by (simp only: adds_term_pprod_def)
subgoal by (simp only: monom_mult_punit_def)
subgoal by (simp only: lt_punit_def)
subgoal by (simp only: lc_punit_def)
subgoal by (simp only: tail_punit_def)
done
subsubsection \<open>More Lemmas and Definitions\<close>
lemma compute_adds_term_pprod [code]:
"adds_term_pprod u v = (snd u = snd v \<and> adds_pp_add_linorder (fst u) (fst v))"
by (simp add: adds_term_pprod_def pprod.adds_term_def adds_pp_add_linorder_def)
lemma compute_splus_pprod [code]: "splus_pprod t (s, i) = (t + s, i)"
by (simp add: splus_pprod_def pprod.splus_def)
lemma compute_sig_trd_spp_body_pprod [code]:
"sig_trd_spp_body_pprod to bs v (p, r) =
(case find_sig_reducer_pprod to bs v (lt_punit (proj_ord to) p) 0 of
None \<Rightarrow> (tail_punit (proj_ord to) p, plus_monomial_less r (lc_punit (proj_ord to) p) (lt_punit (proj_ord to) p))
| Some i \<Rightarrow> let b = snd (bs ! i) in
(tail_punit (proj_ord to) p - monom_mult_punit (lc_punit (proj_ord to) p / lc_punit (proj_ord to) b)
(lt_punit (proj_ord to) p - lt_punit (proj_ord to) b) (tail_punit (proj_ord to) b), r))"
by (simp add: plus_monomial_less_def split: option.split)
lemma compute_sig_trd_spp_pprod [code]:
"sig_trd_spp_pprod to bs (v, p) \<equiv> (v, sig_trd_spp_aux_pprod to bs v (p, change_ord (proj_ord to) 0))"
by (simp add: change_ord_def)
lemmas [code] = conversep_iff
lemma compute_is_pot_ord [code]:
"is_pot_ord_pprod (LEX::(('a::nat, 'b::nat) pp \<times> nat) nat_term_order) = False"
(is "is_pot_ord_pprod ?lex = _")
"is_pot_ord_pprod (DRLEX::(('a::nat, 'b::nat) pp \<times> nat) nat_term_order) = False"
(is "is_pot_ord_pprod ?drlex = _")
"is_pot_ord_pprod (DEG (to::(('a::nat, 'b::nat) pp \<times> nat) nat_term_order)) = False"
"is_pot_ord_pprod (POT (to::(('a::nat, 'b::nat) pp \<times> nat) nat_term_order)) = True"
proof -
have eq1: "snd ((Term_Order.of_exps a b i)::('a, 'b) pp \<times> nat) = i" for a b and i::nat
proof -
have "snd ((Term_Order.of_exps a b i)::('a, 'b) pp \<times> nat) =
snd (rep_nat_term ((Term_Order.of_exps a b i)::('a, 'b) pp \<times> nat))"
by (simp add: rep_nat_term_prod_def rep_nat_nat_def)
also have "... = i"
proof (rule snd_of_exps)
show "snd (rep_nat_term (undefined, i)) = i" by (simp add: rep_nat_term_prod_def rep_nat_nat_def)
qed
finally show ?thesis .
qed
let ?u = "(Term_Order.of_exps 1 0 0)::('a, 'b) pp \<times> nat"
let ?v = "(Term_Order.of_exps 0 0 1)::('a, 'b) pp \<times> nat"
have "\<not> is_pot_ord_pprod ?lex"
proof
assume "is_pot_ord_pprod ?lex"
moreover have "le_of_nat_term_order ?lex ?v ?u"
by (simp add: le_of_nat_term_order nat_term_compare_LEX lex_comp lex_comp_aux_def
comp_of_ord_def lex_pp_of_exps eq_of_exps)
ultimately have "snd ?v \<le> snd ?u" by (rule pprod'.is_pot_ordD2)
thus False by (simp add: eq1)
qed
thus "is_pot_ord_pprod ?lex = False" by simp
have "\<not> is_pot_ord_pprod ?drlex"
proof
assume "is_pot_ord_pprod ?drlex"
moreover have "le_of_nat_term_order ?drlex ?v ?u"
by (simp add: le_of_nat_term_order nat_term_compare_DRLEX deg_comp comparator_of_def)
ultimately have "snd ?v \<le> snd ?u" by (rule pprod'.is_pot_ordD2)
thus False by (simp add: eq1)
qed
thus "is_pot_ord_pprod ?drlex = False" by simp
have "\<not> is_pot_ord_pprod (DEG to)"
proof
assume "is_pot_ord_pprod (DEG to)"
moreover have "le_of_nat_term_order (DEG to) ?v ?u"
by (simp add: le_of_nat_term_order nat_term_compare_DEG deg_comp comparator_of_def)
ultimately have "snd ?v \<le> snd ?u" by (rule pprod'.is_pot_ordD2)
thus False by (simp add: eq1)
qed
thus "is_pot_ord_pprod (DEG to) = False" by simp
have "is_pot_ord_pprod (POT to)"
by (rule pprod'.is_pot_ordI, simp add: lt_of_nat_term_order nat_term_compare_POT pot_comp rep_nat_term_prod_def,
simp add: comparator_of_def)
thus "is_pot_ord_pprod (POT to) = True" by simp
qed
corollary is_pot_ord_POT: "is_pot_ord_pprod (POT to)"
by (simp only: compute_is_pot_ord)
definition "gb_sig_z_pprod to rword_strict fs \<equiv>
(let res = gb_sig_z_pprod' to (rword_strict to) (map (change_ord (proj_ord to)) fs) in
(length (fst res), snd res))"
definition "gb_sig_pprod to rword_strict fs \<equiv> gb_sig_pprod' to (rword_strict to) (map (change_ord (proj_ord to)) fs)"
lemma snd_gb_sig_z_pprod'_eq_gb_sig_z_pprod:
"snd (gb_sig_z_pprod' to (rword_strict to) fs) = snd (gb_sig_z_pprod to rword_strict fs)"
by (simp add: gb_sig_z_pprod_def change_ord_def Let_def)
lemma gb_sig_pprod'_eq_gb_sig_pprod:
"gb_sig_pprod' to (rword_strict to) fs = gb_sig_pprod to rword_strict fs"
by (simp add: gb_sig_pprod_def change_ord_def)
thm pprod'.aux.gb_sig_isGB[OF pprod'.aux.rw_rat_strict_is_strict_rewrite_ord, simplified gb_sig_pprod'_eq_gb_sig_pprod]
thm pprod'.aux.gb_sig_no_zero_red[OF pprod'.aux.rw_rat_strict_is_strict_rewrite_ord is_pot_ord_POT, simplified snd_gb_sig_z_pprod'_eq_gb_sig_z_pprod]
subsection \<open>Computations\<close>
experiment begin interpretation trivariate\<^sub>0_rat .
lemma
"gb_sig_pprod DRLEX rw_rat_strict_pprod [X\<^sup>2 * Z ^ 3 + 3 * X\<^sup>2 * Y, X * Y * Z + 2 * Y\<^sup>2] =
[C\<^sub>0 (3 / 4) * X ^ 3 * Y\<^sup>2 - 2 * Y ^ 4, - 4 * Y ^ 3 * Z - 3 * X\<^sup>2 * Y\<^sup>2, X * Y * Z + 2 * Y\<^sup>2, X\<^sup>2 * Z ^ 3 + 3 * X\<^sup>2 * Y]"
by eval
end
text \<open>Recall that the first return value of @{const gb_sig_z_pprod} is the size of the computed
Gr\"obner basis, and the second return value is the total number of useless zero-reductions:\<close>
lemma
"gb_sig_z_pprod (POT DRLEX) rw_rat_strict_pprod ((cyclic DRLEX 6)::(_ \<Rightarrow>\<^sub>0 rat) list) = (155, 8)"
by eval
lemma
"gb_sig_z_pprod (POT DRLEX) rw_rat_strict_pprod ((katsura DRLEX 5)::(_ \<Rightarrow>\<^sub>0 rat) list) = (29, 0)"
by eval
lemma
"gb_sig_z_pprod (POT DRLEX) rw_rat_strict_pprod ((eco DRLEX 8)::(_ \<Rightarrow>\<^sub>0 rat) list) = (76, 0)"
by eval
lemma
"gb_sig_z_pprod (POT DRLEX) rw_rat_strict_pprod ((noon DRLEX 5)::(_ \<Rightarrow>\<^sub>0 rat) list) = (83, 0)"
by eval
end (* theory *)
diff --git a/thys/Signature_Groebner/Signature_Groebner.thy b/thys/Signature_Groebner/Signature_Groebner.thy
--- a/thys/Signature_Groebner/Signature_Groebner.thy
+++ b/thys/Signature_Groebner/Signature_Groebner.thy
@@ -1,9304 +1,9305 @@
(* Author: Alexander Maletzky *)
section \<open>Signature-Based Algorithms for Computing Gr\"obner Bases\<close>
theory Signature_Groebner
imports More_MPoly Groebner_Bases.Syzygy Polynomials.Quasi_PM_Power_Products
begin
text \<open>First, we develop the whole theory for elements of the module $K[X]^r$, i.\,e. objects of type
@{typ "'t \<Rightarrow>\<^sub>0 'b"}. Later, we transfer all algorithms defined on such objects to algorithms
efficiently operating on sig-poly-pairs, i.\,e. objects of type @{typ "'t \<times> ('a \<Rightarrow>\<^sub>0 'b)"}.\<close>
subsection \<open>More Preliminaries\<close>
lemma (in gd_term) lt_spoly_less_lcs:
assumes "p \<noteq> 0" and "q \<noteq> 0" and "spoly p q \<noteq> 0"
shows "lt (spoly p q) \<prec>\<^sub>t term_of_pair (lcs (lp p) (lp q), component_of_term (lt p))"
proof -
let ?l = "lcs (lp p) (lp q)"
let ?p = "monom_mult (1 / lc p) (?l - lp p) p"
let ?q = "monom_mult (1 / lc q) (?l - lp q) q"
from assms(3) have eq1: "component_of_term (lt p) = component_of_term (lt q)"
and eq2: "spoly p q = ?p - ?q"
by (simp_all add: spoly_def Let_def lc_def split: if_split_asm)
from \<open>p \<noteq> 0\<close> have "lc p \<noteq> 0" by (rule lc_not_0)
with assms(1) have "lt ?p = (?l - lp p) \<oplus> lt p" and "lc ?p = 1" by (simp_all add: lt_monom_mult)
from this(1) have lt_p: "lt ?p = term_of_pair (?l, component_of_term (lt p))"
by (simp add: splus_def adds_minus adds_lcs)
from \<open>q \<noteq> 0\<close> have "lc q \<noteq> 0" by (rule lc_not_0)
with assms(2) have "lt ?q = (?l - lp q) \<oplus> lt q" and "lc ?q = 1" by (simp_all add: lt_monom_mult)
from this(1) have lt_q: "lt ?q = term_of_pair (?l, component_of_term (lt p))"
by (simp add: eq1 splus_def adds_minus adds_lcs_2)
from assms(3) have "?p - ?q \<noteq> 0" by (simp add: eq2)
moreover have "lt ?q = lt ?p" by (simp only: lt_p lt_q)
moreover have "lc ?q = lc ?p" by (simp only: \<open>lc ?p = 1\<close> \<open>lc ?q = 1\<close>)
ultimately have "lt (?p - ?q) \<prec>\<^sub>t lt ?p" by (rule lt_minus_lessI)
thus ?thesis by (simp only: eq2 lt_p)
qed
subsection \<open>Module Polynomials\<close>
locale qpm_inf_term =
gd_term pair_of_term term_of_pair ord ord_strict ord_term ord_term_strict
for pair_of_term::"'t \<Rightarrow> ('a::quasi_pm_powerprod \<times> nat)"
and term_of_pair::"('a \<times> nat) \<Rightarrow> 't"
and ord::"'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<preceq>" 50)
and ord_strict (infixl "\<prec>" 50)
and ord_term::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<preceq>\<^sub>t" 50)
and ord_term_strict::"'t \<Rightarrow> 't \<Rightarrow> bool" (infixl "\<prec>\<^sub>t" 50)
begin
lemma in_idealE_rep_dgrad_p_set:
assumes "hom_grading d" and "B \<subseteq> punit.dgrad_p_set d m" and "p \<in> punit.dgrad_p_set d m" and "p \<in> ideal B"
obtains r where "keys r \<subseteq> B" and "Poly_Mapping.range r \<subseteq> punit.dgrad_p_set d m" and "p = ideal.rep r"
proof -
from assms obtain A q where "finite A" and "A \<subseteq> B" and 0: "\<And>b. q b \<in> punit.dgrad_p_set d m"
and p: "p = (\<Sum>a\<in>A. q a * a)" by (rule punit.in_pmdlE_dgrad_p_set[simplified], blast)
define r where "r = Abs_poly_mapping (\<lambda>k. q k when k \<in> A)"
have 1: "lookup r = (\<lambda>k. q k when k \<in> A)" unfolding r_def
by (rule Abs_poly_mapping_inverse, simp add: \<open>finite A\<close>)
have 2: "keys r \<subseteq> A" by (auto simp: in_keys_iff 1)
show ?thesis
proof
show "Poly_Mapping.range r \<subseteq> punit.dgrad_p_set d m"
proof
fix f
assume "f \<in> Poly_Mapping.range r"
then obtain b where "b \<in> keys r" and f: "f = lookup r b" by (rule poly_mapping_rangeE)
from this(1) 2 have "b \<in> A" ..
hence "f = q b" by (simp add: f 1)
show "f \<in> punit.dgrad_p_set d m" unfolding \<open>f = q b\<close> by (rule 0)
qed
next
have "p = (\<Sum>a\<in>A. lookup r a * a)" unfolding p by (rule sum.cong, simp_all add: 1)
also from \<open>finite A\<close> 2 have "... = (\<Sum>a\<in>keys r. lookup r a * a)"
proof (rule sum.mono_neutral_right)
show "\<forall>a\<in>A - keys r. lookup r a * a = 0"
by (simp add: in_keys_iff)
qed
finally show "p = ideal.rep r" by (simp only: ideal.rep_def)
next
from 2 \<open>A \<subseteq> B\<close> show "keys r \<subseteq> B" by (rule subset_trans)
qed
qed
context fixes fs :: "('a \<Rightarrow>\<^sub>0 'b::field) list"
begin
definition sig_inv_set' :: "nat \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set"
where "sig_inv_set' j = {r. keys (vectorize_poly r) \<subseteq> {0..<j}}"
abbreviation "sig_inv_set \<equiv> sig_inv_set' (length fs)"
definition rep_list :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b)"
where "rep_list r = ideal.rep (pm_of_idx_pm fs (vectorize_poly r))"
lemma sig_inv_setI: "keys (vectorize_poly r) \<subseteq> {0..<j} \<Longrightarrow> r \<in> sig_inv_set' j"
by (simp add: sig_inv_set'_def)
lemma sig_inv_setD: "r \<in> sig_inv_set' j \<Longrightarrow> keys (vectorize_poly r) \<subseteq> {0..<j}"
by (simp add: sig_inv_set'_def)
lemma sig_inv_setI':
assumes "\<And>v. v \<in> keys r \<Longrightarrow> component_of_term v < j"
shows "r \<in> sig_inv_set' j"
proof (rule sig_inv_setI, rule)
fix k
assume "k \<in> keys (vectorize_poly r)"
then obtain v where "v \<in> keys r" and k: "k = component_of_term v" unfolding keys_vectorize_poly ..
from this(1) have "k < j" unfolding k by (rule assms)
thus "k \<in> {0..<j}" by simp
qed
lemma sig_inv_setD':
assumes "r \<in> sig_inv_set' j" and "v \<in> keys r"
shows "component_of_term v < j"
proof -
from assms(2) have "component_of_term v \<in> component_of_term ` keys r" by (rule imageI)
also have "... = keys (vectorize_poly r)" by (simp only: keys_vectorize_poly)
also from assms(1) have "... \<subseteq> {0..<j}" by (rule sig_inv_setD)
finally show ?thesis by simp
qed
corollary sig_inv_setD_lt:
assumes "r \<in> sig_inv_set' j" and "r \<noteq> 0"
shows "component_of_term (lt r) < j"
by (rule sig_inv_setD', fact, rule lt_in_keys, fact)
lemma sig_inv_set_mono:
assumes "i \<le> j"
shows "sig_inv_set' i \<subseteq> sig_inv_set' j"
proof
fix r
assume "r \<in> sig_inv_set' i"
hence "keys (vectorize_poly r) \<subseteq> {0..<i}" by (rule sig_inv_setD)
also from assms have "... \<subseteq> {0..<j}" by fastforce
finally show "r \<in> sig_inv_set' j" by (rule sig_inv_setI)
qed
lemma sig_inv_set_zero: "0 \<in> sig_inv_set' j"
by (rule sig_inv_setI', simp)
lemma sig_inv_set_closed_uminus: "r \<in> sig_inv_set' j \<Longrightarrow> - r \<in> sig_inv_set' j"
by (auto dest!: sig_inv_setD' intro!: sig_inv_setI' simp: keys_uminus)
lemma sig_inv_set_closed_plus:
assumes "r \<in> sig_inv_set' j" and "s \<in> sig_inv_set' j"
shows "r + s \<in> sig_inv_set' j"
proof (rule sig_inv_setI')
fix v
assume "v \<in> keys (r + s)"
hence "v \<in> keys r \<union> keys s" using Poly_Mapping.keys_add ..
thus "component_of_term v < j"
proof
assume "v \<in> keys r"
with assms(1) show ?thesis by (rule sig_inv_setD')
next
assume "v \<in> keys s"
with assms(2) show ?thesis by (rule sig_inv_setD')
qed
qed
lemma sig_inv_set_closed_minus:
assumes "r \<in> sig_inv_set' j" and "s \<in> sig_inv_set' j"
shows "r - s \<in> sig_inv_set' j"
proof (rule sig_inv_setI')
fix v
assume "v \<in> keys (r - s)"
hence "v \<in> keys r \<union> keys s" using keys_minus ..
thus "component_of_term v < j"
proof
assume "v \<in> keys r"
with assms(1) show ?thesis by (rule sig_inv_setD')
next
assume "v \<in> keys s"
with assms(2) show ?thesis by (rule sig_inv_setD')
qed
qed
lemma sig_inv_set_closed_monom_mult:
assumes "r \<in> sig_inv_set' j"
shows "monom_mult c t r \<in> sig_inv_set' j"
proof (rule sig_inv_setI')
fix v
assume "v \<in> keys (monom_mult c t r)"
hence "v \<in> (\<oplus>) t ` keys r" using keys_monom_mult_subset ..
then obtain u where "u \<in> keys r" and v: "v = t \<oplus> u" ..
from assms this(1) have "component_of_term u < j" by (rule sig_inv_setD')
thus "component_of_term v < j" by (simp add: v term_simps)
qed
lemma sig_inv_set_closed_mult_scalar:
assumes "r \<in> sig_inv_set' j"
shows "p \<odot> r \<in> sig_inv_set' j"
proof (rule sig_inv_setI')
fix v
assume "v \<in> keys (p \<odot> r)"
then obtain t u where "u \<in> keys r" and v: "v = t \<oplus> u" by (rule in_keys_mult_scalarE)
from assms this(1) have "component_of_term u < j" by (rule sig_inv_setD')
thus "component_of_term v < j" by (simp add: v term_simps)
qed
lemma rep_list_zero: "rep_list 0 = 0"
by (simp add: rep_list_def vectorize_zero)
lemma rep_list_uminus: "rep_list (- r) = - rep_list r"
by (simp add: rep_list_def vectorize_uminus pm_of_idx_pm_uminus)
lemma rep_list_plus: "rep_list (r + s) = rep_list r + rep_list s"
by (simp add: rep_list_def vectorize_plus pm_of_idx_pm_plus ideal.rep_plus)
lemma rep_list_minus: "rep_list (r - s) = rep_list r - rep_list s"
by (simp add: rep_list_def vectorize_minus pm_of_idx_pm_minus ideal.rep_minus)
lemma vectorize_mult_scalar:
"vectorize_poly (p \<odot> q) = MPoly_Type_Class.punit.monom_mult p 0 (vectorize_poly q)"
by (rule poly_mapping_eqI, simp add: lookup_vectorize_poly MPoly_Type_Class.punit.lookup_monom_mult_zero proj_mult_scalar)
lemma rep_list_mult_scalar: "rep_list (c \<odot> r) = c * rep_list r"
by (simp add: rep_list_def vectorize_mult_scalar pm_of_idx_pm_monom_mult punit.rep_mult_scalar[simplified])
lemma rep_list_monom_mult: "rep_list (monom_mult c t r) = punit.monom_mult c t (rep_list r)"
unfolding mult_scalar_monomial[symmetric] times_monomial_left[symmetric] by (rule rep_list_mult_scalar)
lemma rep_list_monomial:
assumes "distinct fs"
shows "rep_list (monomial c u) =
(punit.monom_mult c (pp_of_term u) (fs ! (component_of_term u))
when component_of_term u < length fs)"
by (simp add: rep_list_def vectorize_monomial pm_of_idx_pm_monomial[OF assms] when_def times_monomial_left)
lemma rep_list_in_ideal_sig_inv_set:
assumes "r \<in> sig_inv_set' j"
shows "rep_list r \<in> ideal (set (take j fs))"
proof -
let ?fs = "take j fs"
from assms have "keys (vectorize_poly r) \<subseteq> {0..<j}" by (rule sig_inv_setD)
hence eq: "pm_of_idx_pm fs (vectorize_poly r) = pm_of_idx_pm ?fs (vectorize_poly r)"
by (simp only: pm_of_idx_pm_take)
have "rep_list r \<in> ideal (keys (pm_of_idx_pm fs (vectorize_poly r)))"
unfolding rep_list_def by (rule ideal.rep_in_span)
also have "... = ideal (keys (pm_of_idx_pm ?fs (vectorize_poly r)))" by (simp only: eq)
also from keys_pm_of_idx_pm_subset have "... \<subseteq> ideal (set ?fs)" by (rule ideal.span_mono)
finally show ?thesis .
qed
corollary rep_list_subset_ideal_sig_inv_set:
"B \<subseteq> sig_inv_set' j \<Longrightarrow> rep_list ` B \<subseteq> ideal (set (take j fs))"
by (auto dest: rep_list_in_ideal_sig_inv_set)
lemma rep_list_in_ideal: "rep_list r \<in> ideal (set fs)"
proof -
have "rep_list r \<in> ideal (keys (pm_of_idx_pm fs (vectorize_poly r)))"
unfolding rep_list_def by (rule ideal.rep_in_span)
also from keys_pm_of_idx_pm_subset have "... \<subseteq> ideal (set fs)" by (rule ideal.span_mono)
finally show ?thesis .
qed
corollary rep_list_subset_ideal: "rep_list ` B \<subseteq> ideal (set fs)"
by (auto intro: rep_list_in_ideal)
lemma in_idealE_rep_list:
assumes "p \<in> ideal (set fs)"
obtains r where "p = rep_list r" and "r \<in> sig_inv_set"
proof -
from assms obtain r0 where r0: "keys r0 \<subseteq> set fs" and p: "p = ideal.rep r0"
by (rule ideal.spanE_rep)
show ?thesis
proof
show "p = rep_list (atomize_poly (idx_pm_of_pm fs r0))"
by (simp add: rep_list_def vectorize_atomize_poly pm_of_idx_pm_of_pm[OF r0] p)
next
show "atomize_poly (idx_pm_of_pm fs r0) \<in> sig_inv_set"
by (rule sig_inv_setI, simp add: vectorize_atomize_poly keys_idx_pm_of_pm_subset)
qed
qed
lemma keys_rep_list_subset:
assumes "t \<in> keys (rep_list r)"
obtains v s where "v \<in> keys r" and "s \<in> Keys (set fs)" and "t = pp_of_term v + s"
proof -
from assms obtain v0 s where v0: "v0 \<in> Keys (Poly_Mapping.range (pm_of_idx_pm fs (vectorize_poly r)))"
and s: "s \<in> Keys (keys (pm_of_idx_pm fs (vectorize_poly r)))" and t: "t = v0 + s"
unfolding rep_list_def by (rule punit.keys_rep_subset[simplified])
note s
also from keys_pm_of_idx_pm_subset have "Keys (keys (pm_of_idx_pm fs (vectorize_poly r))) \<subseteq> Keys (set fs)"
by (rule Keys_mono)
finally have "s \<in> Keys (set fs)" .
note v0
also from range_pm_of_idx_pm_subset'
have "Keys (Poly_Mapping.range (pm_of_idx_pm fs (vectorize_poly r))) \<subseteq> Keys (Poly_Mapping.range (vectorize_poly r))"
by (rule Keys_mono)
also have "... = pp_of_term ` keys r" by (fact Keys_range_vectorize_poly)
finally obtain v where "v \<in> keys r" and "v0 = pp_of_term v" ..
from this(2) have "t = pp_of_term v + s" by (simp only: t)
with \<open>v \<in> keys r\<close> \<open>s \<in> Keys (set fs)\<close> show ?thesis ..
qed
lemma dgrad_p_set_le_rep_list:
assumes "dickson_grading d" and "dgrad_set_le d (pp_of_term ` keys r) (Keys (set fs))"
shows "punit.dgrad_p_set_le d {rep_list r} (set fs)"
proof (simp add: punit.dgrad_p_set_le_def Keys_insert, rule dgrad_set_leI)
fix t
assume "t \<in> keys (rep_list r)"
then obtain v s1 where "v \<in> keys r" and "s1 \<in> Keys (set fs)" and t: "t = pp_of_term v + s1"
by (rule keys_rep_list_subset)
from this(1) have "pp_of_term v \<in> pp_of_term ` keys r" by fastforce
with assms(2) obtain s2 where "s2 \<in> Keys (set fs)" and "d (pp_of_term v) \<le> d s2"
by (rule dgrad_set_leE)
from assms(1) have "d t = ord_class.max (d (pp_of_term v)) (d s1)" unfolding t
by (rule dickson_gradingD1)
hence "d t = d (pp_of_term v) \<or> d t = d s1" by (simp add: ord_class.max_def)
thus "\<exists>s\<in>Keys (set fs). d t \<le> d s"
proof
assume "d t = d (pp_of_term v)"
with \<open>d (pp_of_term v) \<le> d s2\<close> have "d t \<le> d s2" by simp
with \<open>s2 \<in> Keys (set fs)\<close> show ?thesis ..
next
assume "d t = d s1"
hence "d t \<le> d s1" by simp
with \<open>s1 \<in> Keys (set fs)\<close> show ?thesis ..
qed
qed
corollary dgrad_p_set_le_rep_list_image:
assumes "dickson_grading d" and "dgrad_set_le d (pp_of_term ` Keys F) (Keys (set fs))"
shows "punit.dgrad_p_set_le d (rep_list ` F) (set fs)"
proof (rule punit.dgrad_p_set_leI, elim imageE, simp)
fix f
assume "f \<in> F"
have "pp_of_term ` keys f \<subseteq> pp_of_term ` Keys F" by (rule image_mono, rule keys_subset_Keys, fact)
hence "dgrad_set_le d (pp_of_term ` keys f) (pp_of_term ` Keys F)" by (rule dgrad_set_le_subset)
hence "dgrad_set_le d (pp_of_term ` keys f) (Keys (set fs))" using assms(2) by (rule dgrad_set_le_trans)
with assms(1) show "punit.dgrad_p_set_le d {rep_list f} (set fs)" by (rule dgrad_p_set_le_rep_list)
qed
term Max
definition dgrad_max :: "('a \<Rightarrow> nat) \<Rightarrow> nat"
where "dgrad_max d = (Max (d ` (insert 0 (Keys (set fs)))))"
abbreviation "dgrad_max_set d \<equiv> dgrad_p_set d (dgrad_max d)"
abbreviation "punit_dgrad_max_set d \<equiv> punit.dgrad_p_set d (dgrad_max d)"
lemma dgrad_max_0: "d 0 \<le> dgrad_max d"
proof -
from finite_Keys have "finite (d ` insert 0 (Keys (set fs)))" by auto
moreover have "d 0 \<in> d ` insert 0 (Keys (set fs))" by blast
ultimately show ?thesis unfolding dgrad_max_def by (rule Max_ge)
qed
lemma dgrad_max_1: "set fs \<subseteq> punit_dgrad_max_set d"
proof (cases "Keys (set fs) = {}")
case True
show ?thesis
proof (rule, rule punit.dgrad_p_setI[simplified])
fix f v
assume "f \<in> set fs" and "v \<in> keys f"
with True show "d v \<le> dgrad_max d" by (auto simp: Keys_def)
qed
next
case False
show ?thesis
proof (rule subset_trans)
from finite_set show "set fs \<subseteq> punit.dgrad_p_set d (Max (d ` (Keys (set fs))))"
by (rule punit.dgrad_p_set_exhaust_expl[simplified])
next
from finite_set have "finite (Keys (set fs))" by (rule finite_Keys)
hence "finite (d ` Keys (set fs))" by (rule finite_imageI)
moreover from False have 2: "d ` Keys (set fs) \<noteq> {}" by simp
ultimately have "dgrad_max d = ord_class.max (d 0) (Max (d ` Keys (set fs)))"
by (simp add: dgrad_max_def)
hence "Max (d ` (Keys (set fs))) \<le> dgrad_max d" by simp
thus "punit.dgrad_p_set d (Max (d ` (Keys (set fs)))) \<subseteq> punit_dgrad_max_set d"
by (rule punit.dgrad_p_set_subset)
qed
qed
lemma dgrad_max_2:
assumes "dickson_grading d" and "r \<in> dgrad_max_set d"
shows "rep_list r \<in> punit_dgrad_max_set d"
proof (rule punit.dgrad_p_setI[simplified])
fix t
assume "t \<in> keys (rep_list r)"
then obtain v s where "v \<in> keys r" and "s \<in> Keys (set fs)" and t: "t = pp_of_term v + s"
by (rule keys_rep_list_subset)
from assms(2) \<open>v \<in> keys r\<close> have "d (pp_of_term v) \<le> dgrad_max d" by (rule dgrad_p_setD)
moreover have "d s \<le> dgrad_max d" by (simp add: \<open>s \<in> Keys (set fs)\<close> dgrad_max_def finite_Keys)
ultimately show "d t \<le> dgrad_max d" by (simp add: t dickson_gradingD1[OF assms(1)])
qed
corollary dgrad_max_3:
assumes "dickson_grading d" and "F \<subseteq> dgrad_max_set d"
shows "rep_list ` F \<subseteq> punit_dgrad_max_set d"
proof (rule, elim imageE, simp)
fix f
assume "f \<in> F"
hence "f \<in> dgrad_p_set d (dgrad_max d)" using assms(2) ..
with assms(1) show "rep_list f \<in> punit.dgrad_p_set d (dgrad_max d)" by (rule dgrad_max_2)
qed
lemma punit_dgrad_max_set_subset_dgrad_p_set:
assumes "dickson_grading d" and "set fs \<subseteq> punit.dgrad_p_set d m" and "\<not> set fs \<subseteq> {0}"
shows "punit_dgrad_max_set d \<subseteq> punit.dgrad_p_set d m"
proof (rule punit.dgrad_p_set_subset)
show "dgrad_max d \<le> m" unfolding dgrad_max_def
proof (rule Max.boundedI)
show "finite (d ` insert 0 (Keys (set fs)))" by (simp add: finite_Keys)
next
show "d ` insert 0 (Keys (set fs)) \<noteq> {}" by simp
next
fix a
assume "a \<in> d ` insert 0 (Keys (set fs))"
then obtain t where "t \<in> insert 0 (Keys (set fs))" and "a = d t" ..
from this(1) show "a \<le> m" unfolding \<open>a = d t\<close>
proof
assume "t = 0"
from assms(3) obtain f where "f \<in> set fs" and "f \<noteq> 0" by auto
from this(1) assms(2) have "f \<in> punit.dgrad_p_set d m" ..
from \<open>f \<noteq> 0\<close> have "keys f \<noteq> {}" by simp
then obtain s where "s \<in> keys f" by blast
have "d s = d (t + s)" by (simp add: \<open>t = 0\<close>)
also from assms(1) have "... = ord_class.max (d t) (d s)" by (rule dickson_gradingD1)
finally have "d t \<le> d s" by (simp add: max_def)
also from \<open>f \<in> punit.dgrad_p_set d m\<close> \<open>s \<in> keys f\<close> have "... \<le> m"
by (rule punit.dgrad_p_setD[simplified])
finally show "d t \<le> m" .
next
assume "t \<in> Keys (set fs)"
then obtain f where "f \<in> set fs" and "t \<in> keys f" by (rule in_KeysE)
from this(1) assms(2) have "f \<in> punit.dgrad_p_set d m" ..
thus "d t \<le> m" using \<open>t \<in> keys f\<close> by (rule punit.dgrad_p_setD[simplified])
qed
qed
qed
definition dgrad_sig_set' :: "nat \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set"
where "dgrad_sig_set' j d = dgrad_max_set d \<inter> sig_inv_set' j"
abbreviation "dgrad_sig_set \<equiv> dgrad_sig_set' (length fs)"
lemma dgrad_sig_set_set_mono: "i \<le> j \<Longrightarrow> dgrad_sig_set' i d \<subseteq> dgrad_sig_set' j d"
by (auto simp: dgrad_sig_set'_def dest: sig_inv_set_mono)
lemma dgrad_sig_set_closed_uminus: "r \<in> dgrad_sig_set' j d \<Longrightarrow> - r \<in> dgrad_sig_set' j d"
unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_uminus sig_inv_set_closed_uminus)
lemma dgrad_sig_set_closed_plus:
"r \<in> dgrad_sig_set' j d \<Longrightarrow> s \<in> dgrad_sig_set' j d \<Longrightarrow> r + s \<in> dgrad_sig_set' j d"
unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_plus sig_inv_set_closed_plus)
lemma dgrad_sig_set_closed_minus:
"r \<in> dgrad_sig_set' j d \<Longrightarrow> s \<in> dgrad_sig_set' j d \<Longrightarrow> r - s \<in> dgrad_sig_set' j d"
unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_minus sig_inv_set_closed_minus)
lemma dgrad_sig_set_closed_monom_mult:
assumes "dickson_grading d" and "d t \<le> dgrad_max d"
shows "p \<in> dgrad_sig_set' j d \<Longrightarrow> monom_mult c t p \<in> dgrad_sig_set' j d"
unfolding dgrad_sig_set'_def by (auto intro: assms dgrad_p_set_closed_monom_mult sig_inv_set_closed_monom_mult)
lemma dgrad_sig_set_closed_monom_mult_zero:
"p \<in> dgrad_sig_set' j d \<Longrightarrow> monom_mult c 0 p \<in> dgrad_sig_set' j d"
unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_monom_mult_zero sig_inv_set_closed_monom_mult)
lemma dgrad_sig_set_closed_mult_scalar:
"dickson_grading d \<Longrightarrow> p \<in> punit_dgrad_max_set d \<Longrightarrow> r \<in> dgrad_sig_set' j d \<Longrightarrow> p \<odot> r \<in> dgrad_sig_set' j d"
unfolding dgrad_sig_set'_def by (auto intro: dgrad_p_set_closed_mult_scalar sig_inv_set_closed_mult_scalar)
lemma dgrad_sig_set_closed_monomial:
assumes "d (pp_of_term u) \<le> dgrad_max d" and "component_of_term u < j"
shows "monomial c u \<in> dgrad_sig_set' j d"
proof (simp add: dgrad_sig_set'_def, rule)
show "monomial c u \<in> dgrad_max_set d"
proof (rule dgrad_p_setI)
fix v
assume "v \<in> keys (monomial c u)"
also have "... \<subseteq> {u}" by simp
finally show "d (pp_of_term v) \<le> dgrad_max d" using assms(1) by simp
qed
next
show "monomial c u \<in> sig_inv_set' j"
proof (rule sig_inv_setI')
fix v
assume "v \<in> keys (monomial c u)"
also have "... \<subseteq> {u}" by simp
finally show "component_of_term v < j" using assms(2) by simp
qed
qed
lemma rep_list_in_ideal_dgrad_sig_set:
"r \<in> dgrad_sig_set' j d \<Longrightarrow> rep_list r \<in> ideal (set (take j fs))"
by (auto simp: dgrad_sig_set'_def dest: rep_list_in_ideal_sig_inv_set)
lemma in_idealE_rep_list_dgrad_sig_set_take:
assumes "hom_grading d" and "p \<in> punit_dgrad_max_set d" and "p \<in> ideal (set (take j fs))"
obtains r where "r \<in> dgrad_sig_set d" and "r \<in> dgrad_sig_set' j d" and "p = rep_list r"
proof -
let ?fs = "take j fs"
from set_take_subset dgrad_max_1 have "set ?fs \<subseteq> punit_dgrad_max_set d"
by (rule subset_trans)
with assms(1) obtain r0 where r0: "keys r0 \<subseteq> set ?fs"
and 1: "Poly_Mapping.range r0 \<subseteq> punit_dgrad_max_set d" and p: "p = ideal.rep r0"
using assms(2, 3) by (rule in_idealE_rep_dgrad_p_set)
define q where "q = idx_pm_of_pm ?fs r0"
have "keys q \<subseteq> {0..<length ?fs}" unfolding q_def by (rule keys_idx_pm_of_pm_subset)
also have "... \<subseteq> {0..<j}" by fastforce
finally have keys_q: "keys q \<subseteq> {0..<j}" .
have *: "atomize_poly q \<in> dgrad_max_set d"
proof
fix v
assume "v \<in> keys (atomize_poly q)"
then obtain i where i: "i \<in> keys q"
and v_in: "v \<in> (\<lambda>t. term_of_pair (t, i)) ` keys (lookup q i)"
unfolding keys_atomize_poly ..
from i keys_idx_pm_of_pm_subset[of ?fs r0] have "i < length ?fs" by (auto simp: q_def)
from v_in obtain t where "t \<in> keys (lookup q i)" and v: "v = term_of_pair (t, i)" ..
from this(1) \<open>i < length ?fs\<close> have t: "t \<in> keys (lookup r0 (?fs ! i))"
by (simp add: lookup_idx_pm_of_pm q_def)
hence "lookup r0 (?fs ! i) \<noteq> 0" by fastforce
hence "lookup r0 (?fs ! i) \<in> Poly_Mapping.range r0" by (simp add: in_keys_iff)
hence "lookup r0 (?fs ! i) \<in> punit_dgrad_max_set d" using 1 ..
hence "d t \<le> dgrad_max d" using t by (rule punit.dgrad_p_setD[simplified])
thus "d (pp_of_term v) \<le> dgrad_max d" by (simp add: v pp_of_term_of_pair)
qed
show ?thesis
proof
have "atomize_poly q \<in> sig_inv_set' j"
by (rule sig_inv_setI, simp add: vectorize_atomize_poly keys_q)
with * show "atomize_poly q \<in> dgrad_sig_set' j d" unfolding dgrad_sig_set'_def ..
next
from \<open>keys q \<subseteq> {0..<length ?fs}\<close> have keys_q': "keys q \<subseteq> {0..<length fs}" by auto
have "atomize_poly q \<in> sig_inv_set"
by (rule sig_inv_setI, simp add: vectorize_atomize_poly keys_q')
with * show "atomize_poly q \<in> dgrad_sig_set d" unfolding dgrad_sig_set'_def ..
next
from keys_q have "pm_of_idx_pm fs q = pm_of_idx_pm ?fs q" by (simp only: pm_of_idx_pm_take)
thus "p = rep_list (atomize_poly q)"
by (simp add: rep_list_def vectorize_atomize_poly pm_of_idx_pm_of_pm[OF r0] p q_def)
qed
qed
corollary in_idealE_rep_list_dgrad_sig_set:
assumes "hom_grading d" and "p \<in> punit_dgrad_max_set d" and "p \<in> ideal (set fs)"
obtains r where "r \<in> dgrad_sig_set d" and "p = rep_list r"
proof -
from assms(3) have "p \<in> ideal (set (take (length fs) fs))" by simp
with assms(1, 2) obtain r where "r \<in> dgrad_sig_set d" and "p = rep_list r"
by (rule in_idealE_rep_list_dgrad_sig_set_take)
thus ?thesis ..
qed
lemma dgrad_sig_setD_lp:
assumes "p \<in> dgrad_sig_set' j d"
shows "d (lp p) \<le> dgrad_max d"
proof (cases "p = 0")
case True
show ?thesis by (simp add: True min_term_def pp_of_term_of_pair dgrad_max_0)
next
case False
from assms have "p \<in> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
thus ?thesis using False by (rule dgrad_p_setD_lp)
qed
lemma dgrad_sig_setD_lt:
assumes "p \<in> dgrad_sig_set' j d" and "p \<noteq> 0"
shows "component_of_term (lt p) < j"
proof -
from assms have "p \<in> sig_inv_set' j" by (simp add: dgrad_sig_set'_def)
thus ?thesis using assms(2) by (rule sig_inv_setD_lt)
qed
lemma dgrad_sig_setD_rep_list_lt:
assumes "dickson_grading d" and "p \<in> dgrad_sig_set' j d"
shows "d (punit.lt (rep_list p)) \<le> dgrad_max d"
proof (cases "rep_list p = 0")
case True
show ?thesis by (simp add: True dgrad_max_0)
next
case False
from assms(2) have "p \<in> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
with assms(1) have "rep_list p \<in> punit_dgrad_max_set d" by (rule dgrad_max_2)
thus ?thesis using False by (rule punit.dgrad_p_setD_lp[simplified])
qed
definition spp_of :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))"
where "spp_of r = (lt r, rep_list r)"
text \<open>``spp'' stands for ``sig-poly-pair''.\<close>
lemma fst_spp_of: "fst (spp_of r) = lt r"
by (simp add: spp_of_def)
lemma snd_spp_of: "snd (spp_of r) = rep_list r"
by (simp add: spp_of_def)
subsubsection \<open>Signature Reduction\<close>
lemma term_is_le_rel_canc_left:
assumes "ord_term_lin.is_le_rel rel"
shows "rel (t \<oplus> u) (t \<oplus> v) \<longleftrightarrow> rel u v"
using assms
by (rule ord_term_lin.is_le_relE,
auto simp: splus_left_canc dest: ord_term_canc ord_term_strict_canc splus_mono splus_mono_strict)
lemma term_is_le_rel_minus:
assumes "ord_term_lin.is_le_rel rel" and "s adds t"
shows "rel ((t - s) \<oplus> u) v \<longleftrightarrow> rel (t \<oplus> u) (s \<oplus> v)"
proof -
from assms(2) have eq: "s + (t - s) = t" unfolding add.commute[of s] by (rule adds_minus)
from assms(1) have "rel ((t - s) \<oplus> u) v = rel (s \<oplus> ((t - s) \<oplus> u)) (s \<oplus> v)"
by (simp only: term_is_le_rel_canc_left)
also have "... = rel (t \<oplus> u) (s \<oplus> v)" by (simp only: splus_assoc[symmetric] eq)
finally show ?thesis .
qed
lemma term_is_le_rel_minus_minus:
assumes "ord_term_lin.is_le_rel rel" and "a adds t" and "b adds t"
shows "rel ((t - a) \<oplus> u) ((t - b) \<oplus> v) \<longleftrightarrow> rel (b \<oplus> u) (a \<oplus> v)"
proof -
from assms(2) have eq1: "a + (t - a) = t" unfolding add.commute[of a] by (rule adds_minus)
from assms(3) have eq2: "b + (t - b) = t" unfolding add.commute[of b] by (rule adds_minus)
from assms(1) have "rel ((t - a) \<oplus> u) ((t - b) \<oplus> v) = rel ((a + b) \<oplus> ((t - a) \<oplus> u)) ((a + b) \<oplus> ((t - b) \<oplus> v))"
by (simp only: term_is_le_rel_canc_left)
also have "... = rel ((t + b) \<oplus> u) ((t + a) \<oplus> v)" unfolding splus_assoc[symmetric]
by (metis (no_types, lifting) add.assoc add.commute eq1 eq2)
also from assms(1) have "... = rel (b \<oplus> u) (a \<oplus> v)" by (simp only: splus_assoc term_is_le_rel_canc_left)
finally show ?thesis .
qed
lemma pp_is_le_rel_canc_right:
assumes "ordered_powerprod_lin.is_le_rel rel"
shows "rel (s + u) (t + u) \<longleftrightarrow> rel s t"
using assms
by (rule ordered_powerprod_lin.is_le_relE, auto dest: ord_canc ord_strict_canc plus_monotone plus_monotone_strict)
lemma pp_is_le_rel_canc_left: "ordered_powerprod_lin.is_le_rel rel \<Longrightarrow> rel (t + u) (t + v) \<longleftrightarrow> rel u v"
by (simp add: add.commute[of t] pp_is_le_rel_canc_right)
definition sig_red_single :: "('t \<Rightarrow> 't \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 'a \<Rightarrow> bool"
where "sig_red_single sing_reg top_tail p q f t \<longleftrightarrow>
(rep_list f \<noteq> 0 \<and> lookup (rep_list p) (t + punit.lt (rep_list f)) \<noteq> 0 \<and>
q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f \<and>
ord_term_lin.is_le_rel sing_reg \<and> ordered_powerprod_lin.is_le_rel top_tail \<and>
sing_reg (t \<oplus> lt f) (lt p) \<and> top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p)))"
text \<open>The first two parameters of @{const sig_red_single}, \<open>sing_reg\<close> and \<open>top_tail\<close>, specify whether
the reduction is a singular/regular/arbitrary top/tail/arbitrary signature-reduction.
\<^item> If \<open>sing_reg\<close> is @{const HOL.eq}, the reduction is singular.
\<^item> If \<open>sing_reg\<close> is @{term "(\<prec>\<^sub>t)"}, the reduction is regular.
\<^item> If \<open>sing_reg\<close> is @{term "(\<preceq>\<^sub>t)"}, the reduction is an arbitrary signature-reduction.
\<^item> If \<open>top_tail\<close> is @{const HOL.eq}, it is a top reduction.
\<^item> If \<open>top_tail\<close> is @{term "(\<prec>)"}, it is a tail reduction.
\<^item> If \<open>top_tail\<close> is @{term "(\<preceq>)"}, the reduction is an arbitrary signature-reduction.\<close>
definition sig_red :: "('t \<Rightarrow> 't \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "sig_red sing_reg top_tail F p q \<longleftrightarrow> (\<exists>f\<in>F. \<exists>t. sig_red_single sing_reg top_tail p q f t)"
definition is_sig_red :: "('t \<Rightarrow> 't \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "is_sig_red sing_reg top_tail F p \<longleftrightarrow> (\<exists>q. sig_red sing_reg top_tail F p q)"
lemma sig_red_singleI:
assumes "rep_list f \<noteq> 0" and "t + punit.lt (rep_list f) \<in> keys (rep_list p)"
and "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
and "ord_term_lin.is_le_rel sing_reg" and "ordered_powerprod_lin.is_le_rel top_tail"
and "sing_reg (t \<oplus> lt f) (lt p)"
and "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
shows "sig_red_single sing_reg top_tail p q f t"
unfolding sig_red_single_def using assms by blast
lemma sig_red_singleD1:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "rep_list f \<noteq> 0"
using assms unfolding sig_red_single_def by blast
lemma sig_red_singleD2:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "t + punit.lt (rep_list f) \<in> keys (rep_list p)"
using assms unfolding sig_red_single_def by (simp add: in_keys_iff)
lemma sig_red_singleD3:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
using assms unfolding sig_red_single_def by blast
lemma sig_red_singleD4:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "ord_term_lin.is_le_rel sing_reg"
using assms unfolding sig_red_single_def by blast
lemma sig_red_singleD5:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "ordered_powerprod_lin.is_le_rel top_tail"
using assms unfolding sig_red_single_def by blast
lemma sig_red_singleD6:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "sing_reg (t \<oplus> lt f) (lt p)"
using assms unfolding sig_red_single_def by blast
lemma sig_red_singleD7:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
using assms unfolding sig_red_single_def by blast
lemma sig_red_singleD8:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "t \<oplus> lt f \<preceq>\<^sub>t lt p"
proof -
from assms have "ord_term_lin.is_le_rel sing_reg" and "sing_reg (t \<oplus> lt f) (lt p)"
by (rule sig_red_singleD4, rule sig_red_singleD6)
thus ?thesis by (rule ord_term_lin.is_le_rel_le)
qed
lemma sig_red_singleD9:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "t + punit.lt (rep_list f) \<preceq> punit.lt (rep_list p)"
proof -
from assms have "ordered_powerprod_lin.is_le_rel top_tail"
and "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
by (rule sig_red_singleD5, rule sig_red_singleD7)
thus ?thesis by (rule ordered_powerprod_lin.is_le_rel_le)
qed
lemmas sig_red_singleD = sig_red_singleD1 sig_red_singleD2 sig_red_singleD3 sig_red_singleD4
sig_red_singleD5 sig_red_singleD6 sig_red_singleD7 sig_red_singleD8 sig_red_singleD9
lemma sig_red_single_red_single:
"sig_red_single sing_reg top_tail p q f t \<Longrightarrow> punit.red_single (rep_list p) (rep_list q) (rep_list f) t"
by (simp add: sig_red_single_def punit.red_single_def rep_list_minus rep_list_monom_mult)
lemma sig_red_single_regular_lt:
assumes "sig_red_single (\<prec>\<^sub>t) top_tail p q f t"
shows "lt q = lt p"
proof -
let ?f = "monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
from assms have lt: "t \<oplus> lt f \<prec>\<^sub>t lt p" and q: "q = p - ?f"
by (rule sig_red_singleD6, rule sig_red_singleD3)
from lt_monom_mult_le lt have "lt ?f \<prec>\<^sub>t lt p" by (rule ord_term_lin.order.strict_trans1)
thus ?thesis unfolding q by (rule lt_minus_eqI_2)
qed
lemma sig_red_single_regular_lc:
assumes "sig_red_single (\<prec>\<^sub>t) top_tail p q f t"
shows "lc q = lc p"
proof -
from assms have "lt q = lt p" by (rule sig_red_single_regular_lt)
from assms have lt: "t \<oplus> lt f \<prec>\<^sub>t lt p"
and q: "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
(is "_ = _ - ?f") by (rule sig_red_singleD6, rule sig_red_singleD3)
from lt_monom_mult_le lt have "lt ?f \<prec>\<^sub>t lt p" by (rule ord_term_lin.order.strict_trans1)
hence "lookup ?f (lt p) = 0" using lt_max ord_term_lin.leD by blast
thus ?thesis unfolding lc_def \<open>lt q = lt p\<close> by (simp add: q lookup_minus)
qed
lemma sig_red_single_lt:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "lt q \<preceq>\<^sub>t lt p"
proof -
from assms have lt: "t \<oplus> lt f \<preceq>\<^sub>t lt p"
and "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
by (rule sig_red_singleD8, rule sig_red_singleD3)
from this(2) have q: "q = p + monom_mult (- (lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
(is "_ = _ + ?f") by (simp add: monom_mult_uminus_left)
from lt_monom_mult_le lt have 1: "lt ?f \<preceq>\<^sub>t lt p" by (rule ord_term_lin.order.trans)
have "lt q \<preceq>\<^sub>t ord_term_lin.max (lt p) (lt ?f)" unfolding q by (fact lt_plus_le_max)
also from 1 have "ord_term_lin.max (lt p) (lt ?f) = lt p" by (rule ord_term_lin.max.absorb1)
finally show ?thesis .
qed
lemma sig_red_single_lt_rep_list:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "punit.lt (rep_list q) \<preceq> punit.lt (rep_list p)"
proof -
from assms have "punit.red_single (rep_list p) (rep_list q) (rep_list f) t"
by (rule sig_red_single_red_single)
hence "punit.ord_strict_p (rep_list q) (rep_list p)" by (rule punit.red_single_ord)
hence "punit.ord_p (rep_list q) (rep_list p)" by simp
thus ?thesis by (rule punit.ord_p_lt)
qed
lemma sig_red_single_tail_lt_in_keys_rep_list:
assumes "sig_red_single sing_reg (\<prec>) p q f t"
shows "punit.lt (rep_list p) \<in> keys (rep_list q)"
proof -
from assms have "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
by (rule sig_red_singleD3)
hence q: "q = p + monom_mult (- (lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
by (simp add: monom_mult_uminus_left)
show ?thesis unfolding q rep_list_plus rep_list_monom_mult
proof (rule in_keys_plusI1)
from assms have "t + punit.lt (rep_list f) \<in> keys (rep_list p)" by (rule sig_red_singleD2)
hence "rep_list p \<noteq> 0" by auto
thus "punit.lt (rep_list p) \<in> keys (rep_list p)" by (rule punit.lt_in_keys)
next
show "punit.lt (rep_list p) \<notin>
keys (punit.monom_mult (- lookup (rep_list p) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t (rep_list f))"
(is "_ \<notin> keys ?f")
proof
assume "punit.lt (rep_list p) \<in> keys ?f"
hence "punit.lt (rep_list p) \<preceq> punit.lt ?f" by (rule punit.lt_max_keys)
also have "... \<preceq> t + punit.lt (rep_list f)" by (fact punit.lt_monom_mult_le[simplified])
also from assms have "... \<prec> punit.lt (rep_list p)" by (rule sig_red_singleD7)
finally show False by simp
qed
qed
qed
corollary sig_red_single_tail_lt_rep_list:
assumes "sig_red_single sing_reg (\<prec>) p q f t"
shows "punit.lt (rep_list q) = punit.lt (rep_list p)"
-proof (rule ordered_powerprod_lin.antisym)
+proof (rule ordered_powerprod_lin.order_antisym)
from assms show "punit.lt (rep_list q) \<preceq> punit.lt (rep_list p)" by (rule sig_red_single_lt_rep_list)
next
from assms have "punit.lt (rep_list p) \<in> keys (rep_list q)" by (rule sig_red_single_tail_lt_in_keys_rep_list)
thus "punit.lt (rep_list p) \<preceq> punit.lt (rep_list q)" by (rule punit.lt_max_keys)
qed
lemma sig_red_single_tail_lc_rep_list:
assumes "sig_red_single sing_reg (\<prec>) p q f t"
shows "punit.lc (rep_list q) = punit.lc (rep_list p)"
proof -
from assms have *: "punit.lt (rep_list q) = punit.lt (rep_list p)"
by (rule sig_red_single_tail_lt_rep_list)
from assms have lt: "t + punit.lt (rep_list f) \<prec> punit.lt (rep_list p)"
and q: "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
(is "_ = _ - ?f") by (rule sig_red_singleD7, rule sig_red_singleD3)
from punit.lt_monom_mult_le[simplified] lt have "punit.lt (rep_list ?f) \<prec> punit.lt (rep_list p)"
unfolding rep_list_monom_mult by (rule ordered_powerprod_lin.order.strict_trans1)
hence "lookup (rep_list ?f) (punit.lt (rep_list p)) = 0"
using punit.lt_max ordered_powerprod_lin.leD by blast
thus ?thesis unfolding punit.lc_def * by (simp add: q lookup_minus rep_list_minus punit.lc_def)
qed
lemma sig_red_single_top_lt_rep_list:
assumes "sig_red_single sing_reg (=) p q f t" and "rep_list q \<noteq> 0"
shows "punit.lt (rep_list q) \<prec> punit.lt (rep_list p)"
proof -
from assms(1) have "rep_list f \<noteq> 0" and in_keys: "t + punit.lt (rep_list f) \<in> keys (rep_list p)"
and lt: "t + punit.lt (rep_list f) = punit.lt (rep_list p)"
and "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
by (rule sig_red_singleD)+
from this(4) have q: "q = p + monom_mult (- (lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
(is "_ = _ + monom_mult ?c _ _") by (simp add: monom_mult_uminus_left)
from \<open>rep_list f \<noteq> 0\<close> have "punit.lc (rep_list f) \<noteq> 0" by (rule punit.lc_not_0)
from assms(2) have *: "rep_list p + punit.monom_mult ?c t (rep_list f) \<noteq> 0"
by (simp add: q rep_list_plus rep_list_monom_mult)
from in_keys have "lookup (rep_list p) (t + punit.lt (rep_list f)) \<noteq> 0"
by (simp add: in_keys_iff)
moreover from \<open>rep_list f \<noteq> 0\<close> have "punit.lc (rep_list f) \<noteq> 0" by (rule punit.lc_not_0)
ultimately have "?c \<noteq> 0" by simp
hence "punit.lt (punit.monom_mult ?c t (rep_list f)) = t + punit.lt (rep_list f)"
using \<open>rep_list f \<noteq> 0\<close> by (rule lp_monom_mult)
hence "punit.lt (punit.monom_mult ?c t (rep_list f)) = punit.lt (rep_list p)" by (simp only: lt)
moreover have "punit.lc (punit.monom_mult ?c t (rep_list f)) = - punit.lc (rep_list p)"
by (simp add: lt punit.lc_def[symmetric] \<open>punit.lc (rep_list f) \<noteq> 0\<close>)
ultimately show ?thesis unfolding rep_list_plus rep_list_monom_mult q by (rule punit.lt_plus_lessI[OF *])
qed
lemma sig_red_single_monom_mult:
assumes "sig_red_single sing_reg top_tail p q f t" and "c \<noteq> 0"
shows "sig_red_single sing_reg top_tail (monom_mult c s p) (monom_mult c s q) f (s + t)"
proof -
from assms(1) have a: "ord_term_lin.is_le_rel sing_reg" and b: "ordered_powerprod_lin.is_le_rel top_tail"
by (rule sig_red_singleD4, rule sig_red_singleD5)
have eq1: "(s + t) \<oplus> lt f = s \<oplus> (t \<oplus> lt f)" by (simp only: splus_assoc)
from assms(1) have 1: "t + punit.lt (rep_list f) \<in> keys (rep_list p)" by (rule sig_red_singleD2)
hence "rep_list p \<noteq> 0" by auto
hence "p \<noteq> 0" by (auto simp: rep_list_zero)
with assms(2) have eq2: "lt (monom_mult c s p) = s \<oplus> lt p" by (rule lt_monom_mult)
show ?thesis
proof (rule sig_red_singleI)
from assms(1) show "rep_list f \<noteq> 0" by (rule sig_red_singleD1)
next
show "s + t + punit.lt (rep_list f) \<in> keys (rep_list (monom_mult c s p))"
by (auto simp: rep_list_monom_mult punit.keys_monom_mult[OF assms(2)] ac_simps intro: 1)
next
from assms(1) have q: "q = p - monom_mult ((lookup (rep_list p) (t + punit.lt (rep_list f))) / punit.lc (rep_list f)) t f"
by (rule sig_red_singleD3)
show "monom_mult c s q =
monom_mult c s p -
monom_mult (lookup (rep_list (monom_mult c s p)) (s + t + punit.lt (rep_list f)) / punit.lc (rep_list f)) (s + t) f"
by (simp add: q monom_mult_dist_right_minus ac_simps rep_list_monom_mult
punit.lookup_monom_mult_plus[simplified] monom_mult_assoc)
next
from assms(1) have "sing_reg (t \<oplus> lt f) (lt p)" by (rule sig_red_singleD6)
thus "sing_reg ((s + t) \<oplus> lt f) (lt (monom_mult c s p))"
by (simp only: eq1 eq2 term_is_le_rel_canc_left[OF a])
next
from assms(1) have "top_tail (t + punit.lt (rep_list f)) (punit.lt (rep_list p))"
by (rule sig_red_singleD7)
thus "top_tail (s + t + punit.lt (rep_list f)) (punit.lt (rep_list (monom_mult c s p)))"
by (simp add: rep_list_monom_mult punit.lt_monom_mult[OF assms(2) \<open>rep_list p \<noteq> 0\<close>] add.assoc pp_is_le_rel_canc_left[OF b])
qed (fact a, fact b)
qed
lemma sig_red_single_sing_reg_cases:
"sig_red_single (\<preceq>\<^sub>t) top_tail p q f t = (sig_red_single (=) top_tail p q f t \<or> sig_red_single (\<prec>\<^sub>t) top_tail p q f t)"
by (auto simp: sig_red_single_def)
corollary sig_red_single_sing_regI:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "sig_red_single (\<preceq>\<^sub>t) top_tail p q f t"
proof -
from assms have "ord_term_lin.is_le_rel sing_reg" by (rule sig_red_singleD)
with assms show ?thesis unfolding ord_term_lin.is_le_rel_def
by (auto simp: sig_red_single_sing_reg_cases)
qed
lemma sig_red_single_top_tail_cases:
"sig_red_single sing_reg (\<preceq>) p q f t = (sig_red_single sing_reg (=) p q f t \<or> sig_red_single sing_reg (\<prec>) p q f t)"
by (auto simp: sig_red_single_def)
corollary sig_red_single_top_tailI:
assumes "sig_red_single sing_reg top_tail p q f t"
shows "sig_red_single sing_reg (\<preceq>) p q f t"
proof -
from assms have "ordered_powerprod_lin.is_le_rel top_tail" by (rule sig_red_singleD)
with assms show ?thesis unfolding ordered_powerprod_lin.is_le_rel_def
by (auto simp: sig_red_single_top_tail_cases)
qed
lemma dgrad_max_set_closed_sig_red_single:
assumes "dickson_grading d" and "p \<in> dgrad_max_set d" and "f \<in> dgrad_max_set d"
and "sig_red_single sing_red top_tail p q f t"
shows "q \<in> dgrad_max_set d"
proof -
let ?f = "monom_mult (lookup (rep_list p) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t f"
from assms(4) have t: "t + punit.lt (rep_list f) \<in> keys (rep_list p)" and q: "q = p - ?f"
by (rule sig_red_singleD2, rule sig_red_singleD3)
from assms(1, 2) have "rep_list p \<in> punit_dgrad_max_set d" by (rule dgrad_max_2)
show ?thesis unfolding q using assms(2)
proof (rule dgrad_p_set_closed_minus)
from assms(1) _ assms(3) show "?f \<in> dgrad_max_set d"
proof (rule dgrad_p_set_closed_monom_mult)
from assms(1) have "d t \<le> d (t + punit.lt (rep_list f))" by (simp add: dickson_gradingD1)
also from \<open>rep_list p \<in> punit_dgrad_max_set d\<close> t have "... \<le> dgrad_max d"
by (rule punit.dgrad_p_setD[simplified])
finally show "d t \<le> dgrad_max d" .
qed
qed
qed
lemma sig_inv_set_closed_sig_red_single:
assumes "p \<in> sig_inv_set" and "f \<in> sig_inv_set" and "sig_red_single sing_red top_tail p q f t"
shows "q \<in> sig_inv_set"
proof -
let ?f = "monom_mult (lookup (rep_list p) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t f"
from assms(3) have t: "t + punit.lt (rep_list f) \<in> keys (rep_list p)" and q: "q = p - ?f"
by (rule sig_red_singleD2, rule sig_red_singleD3)
show ?thesis unfolding q using assms(1)
proof (rule sig_inv_set_closed_minus)
from assms(2) show "?f \<in> sig_inv_set" by (rule sig_inv_set_closed_monom_mult)
qed
qed
corollary dgrad_sig_set_closed_sig_red_single:
assumes "dickson_grading d" and "p \<in> dgrad_sig_set d" and "f \<in> dgrad_sig_set d"
and "sig_red_single sing_red top_tail p q f t"
shows "q \<in> dgrad_sig_set d"
using assms unfolding dgrad_sig_set'_def
by (auto intro: dgrad_max_set_closed_sig_red_single sig_inv_set_closed_sig_red_single)
lemma sig_red_regular_lt: "sig_red (\<prec>\<^sub>t) top_tail F p q \<Longrightarrow> lt q = lt p"
by (auto simp: sig_red_def intro: sig_red_single_regular_lt)
lemma sig_red_regular_lc: "sig_red (\<prec>\<^sub>t) top_tail F p q \<Longrightarrow> lc q = lc p"
by (auto simp: sig_red_def intro: sig_red_single_regular_lc)
lemma sig_red_lt: "sig_red sing_reg top_tail F p q \<Longrightarrow> lt q \<preceq>\<^sub>t lt p"
by (auto simp: sig_red_def intro: sig_red_single_lt)
lemma sig_red_tail_lt_rep_list: "sig_red sing_reg (\<prec>) F p q \<Longrightarrow> punit.lt (rep_list q) = punit.lt (rep_list p)"
by (auto simp: sig_red_def intro: sig_red_single_tail_lt_rep_list)
lemma sig_red_tail_lc_rep_list: "sig_red sing_reg (\<prec>) F p q \<Longrightarrow> punit.lc (rep_list q) = punit.lc (rep_list p)"
by (auto simp: sig_red_def intro: sig_red_single_tail_lc_rep_list)
lemma sig_red_top_lt_rep_list:
"sig_red sing_reg (=) F p q \<Longrightarrow> rep_list q \<noteq> 0 \<Longrightarrow> punit.lt (rep_list q) \<prec> punit.lt (rep_list p)"
by (auto simp: sig_red_def intro: sig_red_single_top_lt_rep_list)
lemma sig_red_lt_rep_list: "sig_red sing_reg top_tail F p q \<Longrightarrow> punit.lt (rep_list q) \<preceq> punit.lt (rep_list p)"
by (auto simp: sig_red_def intro: sig_red_single_lt_rep_list)
lemma sig_red_red: "sig_red sing_reg top_tail F p q \<Longrightarrow> punit.red (rep_list ` F) (rep_list p) (rep_list q)"
by (auto simp: sig_red_def punit.red_def dest: sig_red_single_red_single)
lemma sig_red_monom_mult:
"sig_red sing_reg top_tail F p q \<Longrightarrow> c \<noteq> 0 \<Longrightarrow> sig_red sing_reg top_tail F (monom_mult c s p) (monom_mult c s q)"
by (auto simp: sig_red_def punit.red_def dest: sig_red_single_monom_mult)
lemma sig_red_sing_reg_cases:
"sig_red (\<preceq>\<^sub>t) top_tail F p q = (sig_red (=) top_tail F p q \<or> sig_red (\<prec>\<^sub>t) top_tail F p q)"
by (auto simp: sig_red_def sig_red_single_sing_reg_cases)
corollary sig_red_sing_regI: "sig_red sing_reg top_tail F p q \<Longrightarrow> sig_red (\<preceq>\<^sub>t) top_tail F p q"
by (auto simp: sig_red_def intro: sig_red_single_sing_regI)
lemma sig_red_top_tail_cases:
"sig_red sing_reg (\<preceq>) F p q = (sig_red sing_reg (=) F p q \<or> sig_red sing_reg (\<prec>) F p q)"
by (auto simp: sig_red_def sig_red_single_top_tail_cases)
corollary sig_red_top_tailI: "sig_red sing_reg top_tail F p q \<Longrightarrow> sig_red sing_reg (\<preceq>) F p q"
by (auto simp: sig_red_def intro: sig_red_single_top_tailI)
lemma sig_red_wf_dgrad_max_set:
assumes "dickson_grading d" and "F \<subseteq> dgrad_max_set d"
shows "wfP (sig_red sing_reg top_tail F)\<inverse>\<inverse>"
proof -
from assms have "rep_list ` F \<subseteq> punit_dgrad_max_set d" by (rule dgrad_max_3)
with assms(1) have "wfP (punit.red (rep_list ` F))\<inverse>\<inverse>" by (rule punit.red_wf_dgrad_p_set)
hence *: "\<nexists>f. \<forall>i. (punit.red (rep_list ` F))\<inverse>\<inverse> (f (Suc i)) (f i)"
by (simp add: wf_iff_no_infinite_down_chain[to_pred])
show ?thesis unfolding wf_iff_no_infinite_down_chain[to_pred]
proof (rule, elim exE)
fix seq
assume "\<forall>i. (sig_red sing_reg top_tail F)\<inverse>\<inverse> (seq (Suc i)) (seq i)"
hence "sig_red sing_reg top_tail F (seq i) (seq (Suc i))" for i by simp
hence "punit.red (rep_list ` F) ((rep_list \<circ> seq) i) ((rep_list \<circ> seq) (Suc i))" for i
by (auto intro: sig_red_red)
hence "\<forall>i. (punit.red (rep_list ` F))\<inverse>\<inverse> ((rep_list \<circ> seq) (Suc i)) ((rep_list \<circ> seq) i)" by simp
hence "\<exists>f. \<forall>i. (punit.red (rep_list ` F))\<inverse>\<inverse> (f (Suc i)) (f i)" by blast
with * show False ..
qed
qed
lemma dgrad_sig_set_closed_sig_red:
assumes "dickson_grading d" and "F \<subseteq> dgrad_sig_set d" and "p \<in> dgrad_sig_set d"
and "sig_red sing_red top_tail F p q"
shows "q \<in> dgrad_sig_set d"
using assms by (auto simp: sig_red_def intro: dgrad_sig_set_closed_sig_red_single)
lemma sig_red_mono: "sig_red sing_reg top_tail F p q \<Longrightarrow> F \<subseteq> F' \<Longrightarrow> sig_red sing_reg top_tail F' p q"
by (auto simp: sig_red_def)
lemma sig_red_Un:
"sig_red sing_reg top_tail (A \<union> B) p q \<longleftrightarrow> (sig_red sing_reg top_tail A p q \<or> sig_red sing_reg top_tail B p q)"
by (auto simp: sig_red_def)
lemma sig_red_subset:
assumes "sig_red sing_reg top_tail F p q" and "sing_reg = (\<preceq>\<^sub>t) \<or> sing_reg = (\<prec>\<^sub>t)"
shows "sig_red sing_reg top_tail {f\<in>F. sing_reg (lt f) (lt p)} p q"
proof -
from assms(1) obtain f t where "f \<in> F" and *: "sig_red_single sing_reg top_tail p q f t"
unfolding sig_red_def by blast
have "lt f = 0 \<oplus> lt f" by (simp only: term_simps)
also from zero_min have "... \<preceq>\<^sub>t t \<oplus> lt f" by (rule splus_mono_left)
finally have 1: "lt f \<preceq>\<^sub>t t \<oplus> lt f" .
from * have 2: "sing_reg (t \<oplus> lt f) (lt p)" by (rule sig_red_singleD6)
from assms(2) have "sing_reg (lt f) (lt p)"
proof
assume "sing_reg = (\<preceq>\<^sub>t)"
with 1 2 show ?thesis by simp
next
assume "sing_reg = (\<prec>\<^sub>t)"
with 1 2 show ?thesis by simp
qed
with \<open>f \<in> F\<close> have "f \<in> {f\<in>F. sing_reg (lt f) (lt p)}" by simp
thus ?thesis using * unfolding sig_red_def by blast
qed
lemma sig_red_regular_rtrancl_lt:
assumes "(sig_red (\<prec>\<^sub>t) top_tail F)\<^sup>*\<^sup>* p q"
shows "lt q = lt p"
using assms by (induct, auto dest: sig_red_regular_lt)
lemma sig_red_regular_rtrancl_lc:
assumes "(sig_red (\<prec>\<^sub>t) top_tail F)\<^sup>*\<^sup>* p q"
shows "lc q = lc p"
using assms by (induct, auto dest: sig_red_regular_lc)
lemma sig_red_rtrancl_lt:
assumes "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q"
shows "lt q \<preceq>\<^sub>t lt p"
using assms by (induct, auto dest: sig_red_lt)
lemma sig_red_tail_rtrancl_lt_rep_list:
assumes "(sig_red sing_reg (\<prec>) F)\<^sup>*\<^sup>* p q"
shows "punit.lt (rep_list q) = punit.lt (rep_list p)"
using assms by (induct, auto dest: sig_red_tail_lt_rep_list)
lemma sig_red_tail_rtrancl_lc_rep_list:
assumes "(sig_red sing_reg (\<prec>) F)\<^sup>*\<^sup>* p q"
shows "punit.lc (rep_list q) = punit.lc (rep_list p)"
using assms by (induct, auto dest: sig_red_tail_lc_rep_list)
lemma sig_red_rtrancl_lt_rep_list:
assumes "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q"
shows "punit.lt (rep_list q) \<preceq> punit.lt (rep_list p)"
using assms by (induct, auto dest: sig_red_lt_rep_list)
lemma sig_red_red_rtrancl:
assumes "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q"
shows "(punit.red (rep_list ` F))\<^sup>*\<^sup>* (rep_list p) (rep_list q)"
using assms by (induct, auto dest: sig_red_red)
lemma sig_red_rtrancl_monom_mult:
assumes "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q"
shows "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* (monom_mult c s p) (monom_mult c s q)"
proof (cases "c = 0")
case True
thus ?thesis by simp
next
case False
from assms(1) show ?thesis
proof induct
case base
show ?case ..
next
case (step y z)
from step(2) False have "sig_red sing_reg top_tail F (monom_mult c s y) (monom_mult c s z)"
by (rule sig_red_monom_mult)
with step(3) show ?case ..
qed
qed
lemma sig_red_rtrancl_sing_regI: "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q \<Longrightarrow> (sig_red (\<preceq>\<^sub>t) top_tail F)\<^sup>*\<^sup>* p q"
by (induct rule: rtranclp_induct, auto dest: sig_red_sing_regI)
lemma sig_red_rtrancl_top_tailI: "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q \<Longrightarrow> (sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* p q"
by (induct rule: rtranclp_induct, auto dest: sig_red_top_tailI)
lemma dgrad_sig_set_closed_sig_red_rtrancl:
assumes "dickson_grading d" and "F \<subseteq> dgrad_sig_set d" and "p \<in> dgrad_sig_set d"
and "(sig_red sing_red top_tail F)\<^sup>*\<^sup>* p q"
shows "q \<in> dgrad_sig_set d"
using assms(4, 1, 2, 3) by (induct, auto intro: dgrad_sig_set_closed_sig_red)
lemma sig_red_rtrancl_mono:
assumes "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q" and "F \<subseteq> F'"
shows "(sig_red sing_reg top_tail F')\<^sup>*\<^sup>* p q"
using assms(1) by (induct rule: rtranclp_induct, auto dest: sig_red_mono[OF _ assms(2)])
lemma sig_red_rtrancl_subset:
assumes "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q" and "sing_reg = (\<preceq>\<^sub>t) \<or> sing_reg = (\<prec>\<^sub>t)"
shows "(sig_red sing_reg top_tail {f\<in>F. sing_reg (lt f) (lt p)})\<^sup>*\<^sup>* p q"
using assms(1)
proof (induct rule: rtranclp_induct)
case base
show ?case by (fact rtranclp.rtrancl_refl)
next
case (step y z)
from step(2) assms(2) have "sig_red sing_reg top_tail {f \<in> F. sing_reg (lt f) (lt y)} y z"
by (rule sig_red_subset)
moreover have "{f \<in> F. sing_reg (lt f) (lt y)} \<subseteq> {f \<in> F. sing_reg (lt f) (lt p)}"
proof
fix f
assume "f \<in> {f \<in> F. sing_reg (lt f) (lt y)}"
hence "f \<in> F" and 1: "sing_reg (lt f) (lt y)" by simp_all
from step(1) have 2: "lt y \<preceq>\<^sub>t lt p" by (rule sig_red_rtrancl_lt)
from assms(2) have "sing_reg (lt f) (lt p)"
proof
assume "sing_reg = (\<preceq>\<^sub>t)"
with 1 2 show ?thesis by simp
next
assume "sing_reg = (\<prec>\<^sub>t)"
with 1 2 show ?thesis by simp
qed
with \<open>f \<in> F\<close> show "f \<in> {f \<in> F. sing_reg (lt f) (lt p)}" by simp
qed
ultimately have "sig_red sing_reg top_tail {f \<in> F. sing_reg (lt f) (lt p)} y z"
by (rule sig_red_mono)
with step(3) show ?case ..
qed
lemma is_sig_red_is_red: "is_sig_red sing_reg top_tail F p \<Longrightarrow> punit.is_red (rep_list ` F) (rep_list p)"
by (auto simp: is_sig_red_def punit.is_red_alt dest: sig_red_red)
lemma is_sig_red_monom_mult:
assumes "is_sig_red sing_reg top_tail F p" and "c \<noteq> 0"
shows "is_sig_red sing_reg top_tail F (monom_mult c s p)"
proof -
from assms(1) obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
hence "sig_red sing_reg top_tail F (monom_mult c s p) (monom_mult c s q)"
using assms(2) by (rule sig_red_monom_mult)
thus ?thesis unfolding is_sig_red_def ..
qed
lemma is_sig_red_sing_reg_cases:
"is_sig_red (\<preceq>\<^sub>t) top_tail F p = (is_sig_red (=) top_tail F p \<or> is_sig_red (\<prec>\<^sub>t) top_tail F p)"
by (auto simp: is_sig_red_def sig_red_sing_reg_cases)
corollary is_sig_red_sing_regI: "is_sig_red sing_reg top_tail F p \<Longrightarrow> is_sig_red (\<preceq>\<^sub>t) top_tail F p"
by (auto simp: is_sig_red_def intro: sig_red_sing_regI)
lemma is_sig_red_top_tail_cases:
"is_sig_red sing_reg (\<preceq>) F p = (is_sig_red sing_reg (=) F p \<or> is_sig_red sing_reg (\<prec>) F p)"
by (auto simp: is_sig_red_def sig_red_top_tail_cases)
corollary is_sig_red_top_tailI: "is_sig_red sing_reg top_tail F p \<Longrightarrow> is_sig_red sing_reg (\<preceq>) F p"
by (auto simp: is_sig_red_def intro: sig_red_top_tailI)
lemma is_sig_red_singletonI:
assumes "is_sig_red sing_reg top_tail F r"
obtains f where "f \<in> F" and "is_sig_red sing_reg top_tail {f} r"
proof -
from assms obtain r' where "sig_red sing_reg top_tail F r r'" unfolding is_sig_red_def ..
then obtain f t where "f \<in> F" and t: "sig_red_single sing_reg top_tail r r' f t"
by (auto simp: sig_red_def)
have "is_sig_red sing_reg top_tail {f} r" unfolding is_sig_red_def sig_red_def
proof (intro exI bexI)
show "f \<in> {f}" by simp
qed fact
with \<open>f \<in> F\<close> show ?thesis ..
qed
lemma is_sig_red_singletonD:
assumes "is_sig_red sing_reg top_tail {f} r" and "f \<in> F"
shows "is_sig_red sing_reg top_tail F r"
proof -
from assms(1) obtain r' where "sig_red sing_reg top_tail {f} r r'" unfolding is_sig_red_def ..
then obtain t where "sig_red_single sing_reg top_tail r r' f t" by (auto simp: sig_red_def)
show ?thesis unfolding is_sig_red_def sig_red_def by (intro exI bexI, fact+)
qed
lemma is_sig_redD1:
assumes "is_sig_red sing_reg top_tail F p"
shows "ord_term_lin.is_le_rel sing_reg"
proof -
from assms obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
then obtain f s where "f \<in> F" and "sig_red_single sing_reg top_tail p q f s" unfolding sig_red_def by blast
from this(2) show ?thesis by (rule sig_red_singleD)
qed
lemma is_sig_redD2:
assumes "is_sig_red sing_reg top_tail F p"
shows "ordered_powerprod_lin.is_le_rel top_tail"
proof -
from assms obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
then obtain f s where "f \<in> F" and "sig_red_single sing_reg top_tail p q f s" unfolding sig_red_def by blast
from this(2) show ?thesis by (rule sig_red_singleD)
qed
lemma is_sig_red_addsI:
assumes "f \<in> F" and "t \<in> keys (rep_list p)" and "rep_list f \<noteq> 0" and "punit.lt (rep_list f) adds t"
and "ord_term_lin.is_le_rel sing_reg" and "ordered_powerprod_lin.is_le_rel top_tail"
and "sing_reg (t \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)" and "top_tail t (punit.lt (rep_list p))"
shows "is_sig_red sing_reg top_tail F p"
unfolding is_sig_red_def
proof
let ?q = "p - monom_mult ((lookup (rep_list p) t) / punit.lc (rep_list f)) (t - punit.lt (rep_list f)) f"
show "sig_red sing_reg top_tail F p ?q" unfolding sig_red_def
proof (intro bexI exI)
from assms(4) have eq: "(t - punit.lt (rep_list f)) + punit.lt (rep_list f) = t" by (rule adds_minus)
from assms(4, 5, 7) have "sing_reg ((t - punit.lt (rep_list f)) \<oplus> lt f) (lt p)"
by (simp only: term_is_le_rel_minus)
thus "sig_red_single sing_reg top_tail p ?q f (t - punit.lt (rep_list f))"
by (simp add: assms eq sig_red_singleI)
qed fact
qed
lemma is_sig_red_addsE:
assumes "is_sig_red sing_reg top_tail F p"
obtains f t where "f \<in> F" and "t \<in> keys (rep_list p)" and "rep_list f \<noteq> 0"
and "punit.lt (rep_list f) adds t"
and "sing_reg (t \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
and "top_tail t (punit.lt (rep_list p))"
proof -
from assms have *: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
from assms obtain q where "sig_red sing_reg top_tail F p q" unfolding is_sig_red_def ..
then obtain f s where "f \<in> F" and "sig_red_single sing_reg top_tail p q f s" unfolding sig_red_def by blast
from this(2) have 1: "rep_list f \<noteq> 0" and 2: "s + punit.lt (rep_list f) \<in> keys (rep_list p)"
and 3: "sing_reg (s \<oplus> lt f) (lt p)" and 4: "top_tail (s + punit.lt (rep_list f)) (punit.lt (rep_list p))"
by (rule sig_red_singleD)+
note \<open>f \<in> F\<close> 2 1
moreover have "punit.lt (rep_list f) adds s + punit.lt (rep_list f)" by simp
moreover from 3 have "sing_reg ((s + punit.lt (rep_list f)) \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
by (simp add: add.commute[of s] splus_assoc term_is_le_rel_canc_left[OF *])
moreover from 4 have "top_tail (s + punit.lt (rep_list f)) (punit.lt (rep_list p))" by simp
ultimately show ?thesis ..
qed
lemma is_sig_red_top_addsI:
assumes "f \<in> F" and "rep_list f \<noteq> 0" and "rep_list p \<noteq> 0"
and "punit.lt (rep_list f) adds punit.lt (rep_list p)" and "ord_term_lin.is_le_rel sing_reg"
and "sing_reg (punit.lt (rep_list p) \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
shows "is_sig_red sing_reg (=) F p"
proof -
note assms(1)
moreover from assms(3) have "punit.lt (rep_list p) \<in> keys (rep_list p)" by (rule punit.lt_in_keys)
moreover note assms(2, 4, 5) ordered_powerprod_lin.is_le_relI(1) assms(6) refl
ultimately show ?thesis by (rule is_sig_red_addsI)
qed
lemma is_sig_red_top_addsE:
assumes "is_sig_red sing_reg (=) F p"
obtains f where "f \<in> F" and "rep_list f \<noteq> 0" and "rep_list p \<noteq> 0"
and "punit.lt (rep_list f) adds punit.lt (rep_list p)"
and "sing_reg (punit.lt (rep_list p) \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
proof -
from assms obtain f t where 1: "f \<in> F" and 2: "t \<in> keys (rep_list p)" and 3: "rep_list f \<noteq> 0"
and 4: "punit.lt (rep_list f) adds t"
and 5: "sing_reg (t \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
and t: "t = punit.lt (rep_list p)" by (rule is_sig_red_addsE)
note 1 3
moreover from 2 have "rep_list p \<noteq> 0" by auto
moreover from 4 have "punit.lt (rep_list f) adds punit.lt (rep_list p)" by (simp only: t)
moreover from 5 have "sing_reg (punit.lt (rep_list p) \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
by (simp only: t)
ultimately show ?thesis ..
qed
lemma is_sig_red_top_plusE:
assumes "is_sig_red sing_reg (=) F p" and "is_sig_red sing_reg (=) F q"
and "lt p \<preceq>\<^sub>t lt (p + q)" and "lt q \<preceq>\<^sub>t lt (p + q)" and "sing_reg = (\<preceq>\<^sub>t) \<or> sing_reg = (\<prec>\<^sub>t)"
assumes 1: "is_sig_red sing_reg (=) F (p + q) \<Longrightarrow> thesis"
assumes 2: "punit.lt (rep_list p) = punit.lt (rep_list q) \<Longrightarrow> punit.lc (rep_list p) + punit.lc (rep_list q) = 0 \<Longrightarrow> thesis"
shows thesis
proof -
from assms(1) obtain f1 where "f1 \<in> F" and "rep_list f1 \<noteq> 0" and "rep_list p \<noteq> 0"
and a: "punit.lt (rep_list f1) adds punit.lt (rep_list p)"
and b: "sing_reg (punit.lt (rep_list p) \<oplus> lt f1) (punit.lt (rep_list f1) \<oplus> lt p)"
by (rule is_sig_red_top_addsE)
from assms(2) obtain f2 where "f2 \<in> F" and "rep_list f2 \<noteq> 0" and "rep_list q \<noteq> 0"
and c: "punit.lt (rep_list f2) adds punit.lt (rep_list q)"
and d: "sing_reg (punit.lt (rep_list q) \<oplus> lt f2) (punit.lt (rep_list f2) \<oplus> lt q)"
by (rule is_sig_red_top_addsE)
show ?thesis
proof (cases "punit.lt (rep_list p) = punit.lt (rep_list q) \<and> punit.lc (rep_list p) + punit.lc (rep_list q) = 0")
case True
hence "punit.lt (rep_list p) = punit.lt (rep_list q)" and "punit.lc (rep_list p) + punit.lc (rep_list q) = 0"
by simp_all
thus ?thesis by (rule 2)
next
case False
hence disj: "punit.lt (rep_list p) \<noteq> punit.lt (rep_list q) \<or> punit.lc (rep_list p) + punit.lc (rep_list q) \<noteq> 0"
by simp
from assms(5) have "ord_term_lin.is_le_rel sing_reg" by (simp add: ord_term_lin.is_le_rel_def)
have "rep_list (p + q) \<noteq> 0" unfolding rep_list_plus
proof
assume eq: "rep_list p + rep_list q = 0"
have eq2: "punit.lt (rep_list p) = punit.lt (rep_list q)"
proof (rule ordered_powerprod_lin.linorder_cases)
assume *: "punit.lt (rep_list p) \<prec> punit.lt (rep_list q)"
hence "punit.lt (rep_list p + rep_list q) = punit.lt (rep_list q)" by (rule punit.lt_plus_eqI)
with * zero_min[of "punit.lt (rep_list p)"] show ?thesis by (simp add: eq)
next
assume *: "punit.lt (rep_list q) \<prec> punit.lt (rep_list p)"
hence "punit.lt (rep_list p + rep_list q) = punit.lt (rep_list p)" by (rule punit.lt_plus_eqI_2)
with * zero_min[of "punit.lt (rep_list q)"] show ?thesis by (simp add: eq)
qed
with disj have "punit.lc (rep_list p) + punit.lc (rep_list q) \<noteq> 0" by simp
thus False by (simp add: punit.lc_def eq2 lookup_add[symmetric] eq)
qed
have "punit.lt (rep_list (p + q)) = ordered_powerprod_lin.max (punit.lt (rep_list p)) (punit.lt (rep_list q))"
unfolding rep_list_plus
proof (rule punit.lt_plus_eq_maxI)
assume "punit.lt (rep_list p) = punit.lt (rep_list q)"
with disj show "punit.lc (rep_list p) + punit.lc (rep_list q) \<noteq> 0" by simp
qed
hence "punit.lt (rep_list (p + q)) = punit.lt (rep_list p) \<or> punit.lt (rep_list (p + q)) = punit.lt (rep_list q)"
by (simp add: ordered_powerprod_lin.max_def)
thus ?thesis
proof
assume eq: "punit.lt (rep_list (p + q)) = punit.lt (rep_list p)"
show ?thesis
proof (rule 1, rule is_sig_red_top_addsI)
from a show "punit.lt (rep_list f1) adds punit.lt (rep_list (p + q))" by (simp only: eq)
next
from b have "sing_reg (punit.lt (rep_list (p + q)) \<oplus> lt f1) (punit.lt (rep_list f1) \<oplus> lt p)"
by (simp only: eq)
moreover from assms(3) have "... \<preceq>\<^sub>t punit.lt (rep_list f1) \<oplus> lt (p + q)" by (rule splus_mono)
ultimately show "sing_reg (punit.lt (rep_list (p + q)) \<oplus> lt f1) (punit.lt (rep_list f1) \<oplus> lt (p + q))"
using assms(5) by auto
qed fact+
next
assume eq: "punit.lt (rep_list (p + q)) = punit.lt (rep_list q)"
show ?thesis
proof (rule 1, rule is_sig_red_top_addsI)
from c show "punit.lt (rep_list f2) adds punit.lt (rep_list (p + q))" by (simp only: eq)
next
from d have "sing_reg (punit.lt (rep_list (p + q)) \<oplus> lt f2) (punit.lt (rep_list f2) \<oplus> lt q)"
by (simp only: eq)
moreover from assms(4) have "... \<preceq>\<^sub>t punit.lt (rep_list f2) \<oplus> lt (p + q)" by (rule splus_mono)
ultimately show "sing_reg (punit.lt (rep_list (p + q)) \<oplus> lt f2) (punit.lt (rep_list f2) \<oplus> lt (p + q))"
using assms(5) by auto
qed fact+
qed
qed
qed
lemma is_sig_red_singleton_monom_multD:
assumes "is_sig_red sing_reg top_tail {monom_mult c t f} p"
shows "is_sig_red sing_reg top_tail {f} p"
proof -
let ?f = "monom_mult c t f"
from assms obtain s where "s \<in> keys (rep_list p)" and 2: "rep_list ?f \<noteq> 0"
and 3: "punit.lt (rep_list ?f) adds s"
and 4: "sing_reg (s \<oplus> lt ?f) (punit.lt (rep_list ?f) \<oplus> lt p)"
and "top_tail s (punit.lt (rep_list p))"
by (auto elim: is_sig_red_addsE)
from 2 have "c \<noteq> 0" and "rep_list f \<noteq> 0"
by (simp_all add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
hence "f \<noteq> 0" by (auto simp: rep_list_zero)
with \<open>c \<noteq> 0\<close> have eq1: "lt ?f = t \<oplus> lt f" by (simp add: lt_monom_mult)
from \<open>c \<noteq> 0\<close> \<open>rep_list f \<noteq> 0\<close> have eq2: "punit.lt (rep_list ?f) = t + punit.lt (rep_list f)"
by (simp add: rep_list_monom_mult punit.lt_monom_mult)
from assms have *: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
show ?thesis
proof (rule is_sig_red_addsI)
show "f \<in> {f}" by simp
next
have "punit.lt (rep_list f) adds t + punit.lt (rep_list f)" by (rule adds_triv_right)
also from 3 have "... adds s" by (simp only: eq2)
finally show "punit.lt (rep_list f) adds s" .
next
from 4 have "sing_reg (t \<oplus> (s \<oplus> lt f)) (t \<oplus> (punit.lt (rep_list f) \<oplus> lt p))"
by (simp add: eq1 eq2 splus_assoc splus_left_commute)
with * show "sing_reg (s \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
by (simp add: term_is_le_rel_canc_left)
next
from assms show "ordered_powerprod_lin.is_le_rel top_tail" by (rule is_sig_redD2)
qed fact+
qed
lemma is_sig_red_top_singleton_monom_multI:
assumes "is_sig_red sing_reg (=) {f} p" and "c \<noteq> 0"
and "t adds punit.lt (rep_list p) - punit.lt (rep_list f)"
shows "is_sig_red sing_reg (=) {monom_mult c t f} p"
proof -
let ?f = "monom_mult c t f"
from assms have 2: "rep_list f \<noteq> 0" and "rep_list p \<noteq> 0"
and 3: "punit.lt (rep_list f) adds punit.lt (rep_list p)"
and 4: "sing_reg (punit.lt (rep_list p) \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
by (auto elim: is_sig_red_top_addsE)
hence "f \<noteq> 0" by (auto simp: rep_list_zero)
with \<open>c \<noteq> 0\<close> have eq1: "lt ?f = t \<oplus> lt f" by (simp add: lt_monom_mult)
from \<open>c \<noteq> 0\<close> \<open>rep_list f \<noteq> 0\<close> have eq2: "punit.lt (rep_list ?f) = t + punit.lt (rep_list f)"
by (simp add: rep_list_monom_mult punit.lt_monom_mult)
from assms(1) have *: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
show ?thesis
proof (rule is_sig_red_top_addsI)
show "?f \<in> {?f}" by simp
next
from \<open>c \<noteq> 0\<close> \<open>rep_list f \<noteq> 0\<close> show "rep_list ?f \<noteq> 0"
by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
next
from assms(3) have "t + punit.lt (rep_list f) adds
(punit.lt (rep_list p) - punit.lt (rep_list f)) + punit.lt (rep_list f)"
by (simp only: adds_canc)
also from 3 have "... = punit.lt (rep_list p)" by (rule adds_minus)
finally show "punit.lt (rep_list ?f) adds punit.lt (rep_list p)" by (simp only: eq2)
next
from 4 * show "sing_reg (punit.lt (rep_list p) \<oplus> lt ?f) (punit.lt (rep_list ?f) \<oplus> lt p)"
by (simp add: eq1 eq2 term_is_le_rel_canc_left splus_assoc splus_left_commute)
qed fact+
qed
lemma is_sig_red_cong':
assumes "is_sig_red sing_reg top_tail F p" and "lt p = lt q" and "rep_list p = rep_list q"
shows "is_sig_red sing_reg top_tail F q"
proof -
from assms(1) have 1: "ord_term_lin.is_le_rel sing_reg" and 2: "ordered_powerprod_lin.is_le_rel top_tail"
by (rule is_sig_redD1, rule is_sig_redD2)
from assms(1) obtain f t where "f \<in> F" and "t \<in> keys (rep_list p)" and "rep_list f \<noteq> 0"
and "punit.lt (rep_list f) adds t"
and "sing_reg (t \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
and "top_tail t (punit.lt (rep_list p))" by (rule is_sig_red_addsE)
from this(1-4) 1 2 this(5, 6) show ?thesis unfolding assms(2, 3) by (rule is_sig_red_addsI)
qed
lemma is_sig_red_cong:
"lt p = lt q \<Longrightarrow> rep_list p = rep_list q \<Longrightarrow>
is_sig_red sing_reg top_tail F p \<longleftrightarrow> is_sig_red sing_reg top_tail F q"
by (auto intro: is_sig_red_cong')
lemma is_sig_red_top_cong:
assumes "is_sig_red sing_reg (=) F p" and "rep_list q \<noteq> 0" and "lt p = lt q"
and "punit.lt (rep_list p) = punit.lt (rep_list q)"
shows "is_sig_red sing_reg (=) F q"
proof -
from assms(1) have 1: "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
from assms(1) obtain f where "f \<in> F" and "rep_list f \<noteq> 0" and "rep_list p \<noteq> 0"
and "punit.lt (rep_list f) adds punit.lt (rep_list p)"
and "sing_reg (punit.lt (rep_list p) \<oplus> lt f) (punit.lt (rep_list f) \<oplus> lt p)"
by (rule is_sig_red_top_addsE)
from this(1, 2) assms(2) this(4) 1 this(5) show ?thesis
unfolding assms(3, 4) by (rule is_sig_red_top_addsI)
qed
lemma sig_irredE_dgrad_max_set:
assumes "dickson_grading d" and "F \<subseteq> dgrad_max_set d"
obtains q where "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q" and "\<not> is_sig_red sing_reg top_tail F q"
proof -
let ?Q = "{q. (sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q}"
from assms have "wfP (sig_red sing_reg top_tail F)\<inverse>\<inverse>" by (rule sig_red_wf_dgrad_max_set)
moreover have "p \<in> ?Q" by simp
ultimately obtain q where "q \<in> ?Q" and "\<And>x. (sig_red sing_reg top_tail F)\<inverse>\<inverse> x q \<Longrightarrow> x \<notin> ?Q"
by (rule wfE_min[to_pred], blast)
hence 1: "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p q"
and 2: "\<And>x. sig_red sing_reg top_tail F q x \<Longrightarrow> \<not> (sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p x"
by simp_all
show ?thesis
proof
show "\<not> is_sig_red sing_reg top_tail F q"
proof
assume "is_sig_red sing_reg top_tail F q"
then obtain x where 3: "sig_red sing_reg top_tail F q x" unfolding is_sig_red_def ..
hence "\<not> (sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p x" by (rule 2)
moreover from 1 3 have "(sig_red sing_reg top_tail F)\<^sup>*\<^sup>* p x" ..
ultimately show False ..
qed
qed fact
qed
lemma is_sig_red_mono:
"is_sig_red sing_reg top_tail F p \<Longrightarrow> F \<subseteq> F' \<Longrightarrow> is_sig_red sing_reg top_tail F' p"
by (auto simp: is_sig_red_def dest: sig_red_mono)
lemma is_sig_red_Un:
"is_sig_red sing_reg top_tail (A \<union> B) p \<longleftrightarrow> (is_sig_red sing_reg top_tail A p \<or> is_sig_red sing_reg top_tail B p)"
by (auto simp: is_sig_red_def sig_red_Un)
lemma is_sig_redD_lt:
assumes "is_sig_red (\<preceq>\<^sub>t) top_tail {f} p"
shows "lt f \<preceq>\<^sub>t lt p"
proof -
from assms obtain s where "rep_list f \<noteq> 0" and "s \<in> keys (rep_list p)"
and 1: "punit.lt (rep_list f) adds s" and 2: "s \<oplus> lt f \<preceq>\<^sub>t punit.lt (rep_list f) \<oplus> lt p"
by (auto elim!: is_sig_red_addsE)
from 1 obtain t where eq: "s = punit.lt (rep_list f) + t" by (rule addsE)
hence "punit.lt (rep_list f) \<oplus> (t \<oplus> lt f) = s \<oplus> lt f" by (simp add: splus_assoc)
also note 2
finally have "t \<oplus> lt f \<preceq>\<^sub>t lt p" by (rule ord_term_canc)
have "0 \<preceq> t" by (fact zero_min)
hence "0 \<oplus> lt f \<preceq>\<^sub>t t \<oplus> lt f" by (rule splus_mono_left)
hence "lt f \<preceq>\<^sub>t t \<oplus> lt f" by (simp add: term_simps)
thus ?thesis using \<open>t \<oplus> lt f \<preceq>\<^sub>t lt p\<close> by simp
qed
lemma is_sig_red_regularD_lt:
assumes "is_sig_red (\<prec>\<^sub>t) top_tail {f} p"
shows "lt f \<prec>\<^sub>t lt p"
proof -
from assms obtain s where "rep_list f \<noteq> 0" and "s \<in> keys (rep_list p)"
and 1: "punit.lt (rep_list f) adds s" and 2: "s \<oplus> lt f \<prec>\<^sub>t punit.lt (rep_list f) \<oplus> lt p"
by (auto elim!: is_sig_red_addsE)
from 1 obtain t where eq: "s = punit.lt (rep_list f) + t" by (rule addsE)
hence "punit.lt (rep_list f) \<oplus> (t \<oplus> lt f) = s \<oplus> lt f" by (simp add: splus_assoc)
also note 2
finally have "t \<oplus> lt f \<prec>\<^sub>t lt p" by (rule ord_term_strict_canc)
have "0 \<preceq> t" by (fact zero_min)
hence "0 \<oplus> lt f \<preceq>\<^sub>t t \<oplus> lt f" by (rule splus_mono_left)
hence "lt f \<preceq>\<^sub>t t \<oplus> lt f" by (simp add: term_simps)
thus ?thesis using \<open>t \<oplus> lt f \<prec>\<^sub>t lt p\<close> by (rule ord_term_lin.le_less_trans)
qed
lemma sig_irred_regular_self: "\<not> is_sig_red (\<prec>\<^sub>t) top_tail {p} p"
by (auto dest: is_sig_red_regularD_lt)
subsubsection \<open>Signature Gr\"obner Bases\<close>
definition sig_red_zero :: "('t \<Rightarrow>'t \<Rightarrow> bool) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "sig_red_zero sing_reg F r \<longleftrightarrow> (\<exists>s. (sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* r s \<and> rep_list s = 0)"
definition is_sig_GB_in :: "('a \<Rightarrow> nat) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> 't \<Rightarrow> bool"
where "is_sig_GB_in d G u \<longleftrightarrow> (\<forall>r. lt r = u \<longrightarrow> r \<in> dgrad_sig_set d \<longrightarrow> sig_red_zero (\<preceq>\<^sub>t) G r)"
definition is_sig_GB_upt :: "('a \<Rightarrow> nat) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> 't \<Rightarrow> bool"
where "is_sig_GB_upt d G u \<longleftrightarrow>
(G \<subseteq> dgrad_sig_set d \<and> (\<forall>v. v \<prec>\<^sub>t u \<longrightarrow> d (pp_of_term v) \<le> dgrad_max d \<longrightarrow>
component_of_term v < length fs \<longrightarrow> is_sig_GB_in d G v))"
definition is_min_sig_GB :: "('a \<Rightarrow> nat) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> bool"
where "is_min_sig_GB d G \<longleftrightarrow> G \<subseteq> dgrad_sig_set d \<and>
(\<forall>u. d (pp_of_term u) \<le> dgrad_max d \<longrightarrow> component_of_term u < length fs \<longrightarrow>
is_sig_GB_in d G u) \<and>
(\<forall>g\<in>G. \<not> is_sig_red (\<preceq>\<^sub>t) (=) (G - {g}) g)"
definition is_syz_sig :: "('a \<Rightarrow> nat) \<Rightarrow> 't \<Rightarrow> bool"
where "is_syz_sig d u \<longleftrightarrow> (\<exists>s\<in>dgrad_sig_set d. s \<noteq> 0 \<and> lt s = u \<and> rep_list s = 0)"
lemma sig_red_zeroI:
assumes "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* r s" and "rep_list s = 0"
shows "sig_red_zero sing_reg F r"
unfolding sig_red_zero_def using assms by blast
lemma sig_red_zeroE:
assumes "sig_red_zero sing_reg F r"
obtains s where "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* r s" and "rep_list s = 0"
using assms unfolding sig_red_zero_def by blast
lemma sig_red_zero_monom_mult:
assumes "sig_red_zero sing_reg F r"
shows "sig_red_zero sing_reg F (monom_mult c t r)"
proof -
from assms obtain s where "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* r s" and "rep_list s = 0"
by (rule sig_red_zeroE)
from this(1) have "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* (monom_mult c t r) (monom_mult c t s)"
by (rule sig_red_rtrancl_monom_mult)
moreover have "rep_list (monom_mult c t s) = 0" by (simp add: rep_list_monom_mult \<open>rep_list s = 0\<close>)
ultimately show ?thesis by (rule sig_red_zeroI)
qed
lemma sig_red_zero_sing_regI:
assumes "sig_red_zero sing_reg G p"
shows "sig_red_zero (\<preceq>\<^sub>t) G p"
proof -
from assms obtain s where "(sig_red sing_reg (\<preceq>) G)\<^sup>*\<^sup>* p s" and "rep_list s = 0"
by (rule sig_red_zeroE)
from this(1) have "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* p s" by (rule sig_red_rtrancl_sing_regI)
thus ?thesis using \<open>rep_list s = 0\<close> by (rule sig_red_zeroI)
qed
lemma sig_red_zero_nonzero:
assumes "sig_red_zero sing_reg F r" and "rep_list r \<noteq> 0" and "sing_reg = (\<preceq>\<^sub>t) \<or> sing_reg = (\<prec>\<^sub>t)"
shows "is_sig_red sing_reg (=) F r"
proof -
from assms(1) obtain s where "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* r s" and "rep_list s = 0"
by (rule sig_red_zeroE)
from this(1) assms(2) show ?thesis
proof (induct rule: converse_rtranclp_induct)
case base
thus ?case using \<open>rep_list s = 0\<close> ..
next
case (step y z)
from step(1) obtain f t where "f \<in> F" and *: "sig_red_single sing_reg (\<preceq>) y z f t"
unfolding sig_red_def by blast
from this(2) have 1: "rep_list f \<noteq> 0" and 2: "t + punit.lt (rep_list f) \<in> keys (rep_list y)"
and 3: "z = y - monom_mult (lookup (rep_list y) (t + punit.lt (rep_list f)) / punit.lc (rep_list f)) t f"
and 4: "ord_term_lin.is_le_rel sing_reg" and 5: "sing_reg (t \<oplus> lt f) (lt y)"
by (rule sig_red_singleD)+
show ?case
proof (cases "t + punit.lt (rep_list f) = punit.lt (rep_list y)")
case True
show ?thesis unfolding is_sig_red_def
proof
show "sig_red sing_reg (=) F y z" unfolding sig_red_def
proof (intro bexI exI)
from 1 2 3 4 ordered_powerprod_lin.is_le_relI(1) 5 True
show "sig_red_single sing_reg (=) y z f t" by (rule sig_red_singleI)
qed fact
qed
next
case False
from 2 have "t + punit.lt (rep_list f) \<preceq> punit.lt (rep_list y)" by (rule punit.lt_max_keys)
with False have "t + punit.lt (rep_list f) \<prec> punit.lt (rep_list y)" by simp
with 1 2 3 4 ordered_powerprod_lin.is_le_relI(3) 5 have "sig_red_single sing_reg (\<prec>) y z f t"
by (rule sig_red_singleI)
hence "punit.lt (rep_list y) \<in> keys (rep_list z)"
and lt_z: "punit.lt (rep_list z) = punit.lt (rep_list y)"
by (rule sig_red_single_tail_lt_in_keys_rep_list, rule sig_red_single_tail_lt_rep_list)
from this(1) have "rep_list z \<noteq> 0" by auto
hence "is_sig_red sing_reg (=) F z" by (rule step(3))
then obtain g where "g \<in> F" and "rep_list g \<noteq> 0"
and "punit.lt (rep_list g) adds punit.lt (rep_list z)"
and a: "sing_reg (punit.lt (rep_list z) \<oplus> lt g) (punit.lt (rep_list g) \<oplus> lt z)"
by (rule is_sig_red_top_addsE)
from this(3) have "punit.lt (rep_list g) adds punit.lt (rep_list y)" by (simp only: lt_z)
with \<open>g \<in> F\<close> \<open>rep_list g \<noteq> 0\<close> step(4) show ?thesis
proof (rule is_sig_red_top_addsI)
from \<open>is_sig_red sing_reg (=) F z\<close> show "ord_term_lin.is_le_rel sing_reg" by (rule is_sig_redD1)
next
from \<open>sig_red_single sing_reg (\<prec>) y z f t\<close> have "lt z \<preceq>\<^sub>t lt y" by (rule sig_red_single_lt)
from assms(3) show "sing_reg (punit.lt (rep_list y) \<oplus> lt g) (punit.lt (rep_list g) \<oplus> lt y)"
proof
assume "sing_reg = (\<preceq>\<^sub>t)"
from a have "punit.lt (rep_list y) \<oplus> lt g \<preceq>\<^sub>t punit.lt (rep_list g) \<oplus> lt z"
by (simp only: lt_z \<open>sing_reg = (\<preceq>\<^sub>t)\<close>)
also from \<open>lt z \<preceq>\<^sub>t lt y\<close> have "... \<preceq>\<^sub>t punit.lt (rep_list g) \<oplus> lt y" by (rule splus_mono)
finally show ?thesis by (simp only: \<open>sing_reg = (\<preceq>\<^sub>t)\<close>)
next
assume "sing_reg = (\<prec>\<^sub>t)"
from a have "punit.lt (rep_list y) \<oplus> lt g \<prec>\<^sub>t punit.lt (rep_list g) \<oplus> lt z"
by (simp only: lt_z \<open>sing_reg = (\<prec>\<^sub>t)\<close>)
also from \<open>lt z \<preceq>\<^sub>t lt y\<close> have "... \<preceq>\<^sub>t punit.lt (rep_list g) \<oplus> lt y" by (rule splus_mono)
finally show ?thesis by (simp only: \<open>sing_reg = (\<prec>\<^sub>t)\<close>)
qed
qed
qed
qed
qed
lemma sig_red_zero_mono: "sig_red_zero sing_reg F p \<Longrightarrow> F \<subseteq> F' \<Longrightarrow> sig_red_zero sing_reg F' p"
by (auto simp: sig_red_zero_def dest: sig_red_rtrancl_mono)
lemma sig_red_zero_subset:
assumes "sig_red_zero sing_reg F p" and "sing_reg = (\<preceq>\<^sub>t) \<or> sing_reg = (\<prec>\<^sub>t)"
shows "sig_red_zero sing_reg {f\<in>F. sing_reg (lt f) (lt p)} p"
proof -
from assms(1) obtain s where "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* p s" and "rep_list s = 0"
by (rule sig_red_zeroE)
from this(1) assms(2) have "(sig_red sing_reg (\<preceq>) {f\<in>F. sing_reg (lt f) (lt p)})\<^sup>*\<^sup>* p s"
by (rule sig_red_rtrancl_subset)
thus ?thesis using \<open>rep_list s = 0\<close> by (rule sig_red_zeroI)
qed
lemma sig_red_zero_idealI:
assumes "sig_red_zero sing_reg F p"
shows "rep_list p \<in> ideal (rep_list ` F)"
proof -
from assms obtain s where "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* p s" and "rep_list s = 0" by (rule sig_red_zeroE)
from this(1) have "(punit.red (rep_list ` F))\<^sup>*\<^sup>* (rep_list p) (rep_list s)" by (rule sig_red_red_rtrancl)
hence "(punit.red (rep_list ` F))\<^sup>*\<^sup>* (rep_list p) 0" by (simp only: \<open>rep_list s = 0\<close>)
thus ?thesis by (rule punit.red_rtranclp_0_in_pmdl[simplified])
qed
lemma is_sig_GB_inI:
assumes "\<And>r. lt r = u \<Longrightarrow> r \<in> dgrad_sig_set d \<Longrightarrow> sig_red_zero (\<preceq>\<^sub>t) G r"
shows "is_sig_GB_in d G u"
unfolding is_sig_GB_in_def using assms by blast
lemma is_sig_GB_inD:
assumes "is_sig_GB_in d G u" and "r \<in> dgrad_sig_set d" and "lt r = u"
shows "sig_red_zero (\<preceq>\<^sub>t) G r"
using assms unfolding is_sig_GB_in_def by blast
lemma is_sig_GB_inI_triv:
assumes "\<not> d (pp_of_term u) \<le> dgrad_max d \<or> \<not> component_of_term u < length fs"
shows "is_sig_GB_in d G u"
proof (rule is_sig_GB_inI)
fix r::"'t \<Rightarrow>\<^sub>0 'b"
assume "lt r = u" and "r \<in> dgrad_sig_set d"
show "sig_red_zero (\<preceq>\<^sub>t) G r"
proof (cases "r = 0")
case True
hence "rep_list r = 0" by (simp only: rep_list_zero)
with rtrancl_refl[to_pred] show ?thesis by (rule sig_red_zeroI)
next
case False
from \<open>r \<in> dgrad_sig_set d\<close> have "d (lp r) \<le> dgrad_max d" by (rule dgrad_sig_setD_lp)
moreover from \<open>r \<in> dgrad_sig_set d\<close> False have "component_of_term (lt r) < length fs"
by (rule dgrad_sig_setD_lt)
ultimately show ?thesis using assms by (simp add: \<open>lt r = u\<close>)
qed
qed
lemma is_sig_GB_in_mono: "is_sig_GB_in d G u \<Longrightarrow> G \<subseteq> G' \<Longrightarrow> is_sig_GB_in d G' u"
by (auto simp: is_sig_GB_in_def dest: sig_red_zero_mono)
lemma is_sig_GB_uptI:
assumes "G \<subseteq> dgrad_sig_set d"
and "\<And>v. v \<prec>\<^sub>t u \<Longrightarrow> d (pp_of_term v) \<le> dgrad_max d \<Longrightarrow> component_of_term v < length fs \<Longrightarrow>
is_sig_GB_in d G v"
shows "is_sig_GB_upt d G u"
unfolding is_sig_GB_upt_def using assms by blast
lemma is_sig_GB_uptD1:
assumes "is_sig_GB_upt d G u"
shows "G \<subseteq> dgrad_sig_set d"
using assms unfolding is_sig_GB_upt_def by blast
lemma is_sig_GB_uptD2:
assumes "is_sig_GB_upt d G u" and "v \<prec>\<^sub>t u"
shows "is_sig_GB_in d G v"
using assms is_sig_GB_inI_triv unfolding is_sig_GB_upt_def by blast
lemma is_sig_GB_uptD3:
assumes "is_sig_GB_upt d G u" and "r \<in> dgrad_sig_set d" and "lt r \<prec>\<^sub>t u"
shows "sig_red_zero (\<preceq>\<^sub>t) G r"
by (rule is_sig_GB_inD, rule is_sig_GB_uptD2, fact+, fact refl)
lemma is_sig_GB_upt_le:
assumes "is_sig_GB_upt d G u" and "v \<preceq>\<^sub>t u"
shows "is_sig_GB_upt d G v"
proof (rule is_sig_GB_uptI)
from assms(1) show "G \<subseteq> dgrad_sig_set d" by (rule is_sig_GB_uptD1)
next
fix w
assume "w \<prec>\<^sub>t v"
hence "w \<prec>\<^sub>t u" using assms(2) by (rule ord_term_lin.less_le_trans)
with assms(1) show "is_sig_GB_in d G w" by (rule is_sig_GB_uptD2)
qed
lemma is_sig_GB_upt_mono:
"is_sig_GB_upt d G u \<Longrightarrow> G \<subseteq> G' \<Longrightarrow> G' \<subseteq> dgrad_sig_set d \<Longrightarrow> is_sig_GB_upt d G' u"
by (auto simp: is_sig_GB_upt_def dest!: is_sig_GB_in_mono)
lemma is_sig_GB_upt_is_Groebner_basis:
assumes "dickson_grading d" and "hom_grading d" and "G \<subseteq> dgrad_sig_set' j d"
and "\<And>u. component_of_term u < j \<Longrightarrow> is_sig_GB_in d G u"
shows "punit.is_Groebner_basis (rep_list ` G)"
using assms(1)
proof (rule punit.weak_GB_is_strong_GB_dgrad_p_set[simplified])
from assms(3) have "G \<subseteq> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
with assms(1) show "rep_list ` G \<subseteq> punit_dgrad_max_set d" by (rule dgrad_max_3)
next
fix f::"'a \<Rightarrow>\<^sub>0 'b"
assume "f \<in> punit_dgrad_max_set d"
from assms(3) have G_sub: "G \<subseteq> sig_inv_set' j" by (simp add: dgrad_sig_set'_def)
assume "f \<in> ideal (rep_list ` G)"
also from rep_list_subset_ideal_sig_inv_set[OF G_sub] have "... \<subseteq> ideal (set (take j fs))"
by (rule ideal.span_subset_spanI)
finally have "f \<in> ideal (set (take j fs))" .
with assms(2) \<open>f \<in> punit_dgrad_max_set d\<close> obtain r where "r \<in> dgrad_sig_set d"
and "r \<in> dgrad_sig_set' j d" and f: "f = rep_list r"
by (rule in_idealE_rep_list_dgrad_sig_set_take)
from this(2) have "r \<in> sig_inv_set' j" by (simp add: dgrad_sig_set'_def)
show "(punit.red (rep_list ` G))\<^sup>*\<^sup>* f 0"
proof (cases "r = 0")
case True
thus ?thesis by (simp add: f rep_list_zero)
next
case False
hence "lt r \<in> keys r" by (rule lt_in_keys)
with \<open>r \<in> sig_inv_set' j\<close> have "component_of_term (lt r) < j" by (rule sig_inv_setD')
hence "is_sig_GB_in d G (lt r)" by (rule assms(4))
hence "sig_red_zero (\<preceq>\<^sub>t) G r" using \<open>r \<in> dgrad_sig_set d\<close> refl by (rule is_sig_GB_inD)
then obtain s where "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* r s" and s: "rep_list s = 0" by (rule sig_red_zeroE)
from this(1) have "(punit.red (rep_list ` G))\<^sup>*\<^sup>* (rep_list r) (rep_list s)"
by (rule sig_red_red_rtrancl)
thus ?thesis by (simp only: f s)
qed
qed
lemma is_sig_GB_is_Groebner_basis:
assumes "dickson_grading d" and "hom_grading d" and "G \<subseteq> dgrad_max_set d" and "\<And>u. is_sig_GB_in d G u"
shows "punit.is_Groebner_basis (rep_list ` G)"
using assms(1)
proof (rule punit.weak_GB_is_strong_GB_dgrad_p_set[simplified])
from assms(1, 3) show "rep_list ` G \<subseteq> punit_dgrad_max_set d" by (rule dgrad_max_3)
next
fix f::"'a \<Rightarrow>\<^sub>0 'b"
assume "f \<in> punit_dgrad_max_set d"
assume "f \<in> ideal (rep_list ` G)"
also from rep_list_subset_ideal have "... \<subseteq> ideal (set fs)" by (rule ideal.span_subset_spanI)
finally have "f \<in> ideal (set fs)" .
with assms(2) \<open>f \<in> punit_dgrad_max_set d\<close> obtain r where "r \<in> dgrad_sig_set d" and f: "f = rep_list r"
by (rule in_idealE_rep_list_dgrad_sig_set)
from assms(4) this(1) refl have "sig_red_zero (\<preceq>\<^sub>t) G r" by (rule is_sig_GB_inD)
then obtain s where "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* r s" and s: "rep_list s = 0" by (rule sig_red_zeroE)
from this(1) have "(punit.red (rep_list ` G))\<^sup>*\<^sup>* (rep_list r) (rep_list s)"
by (rule sig_red_red_rtrancl)
thus "(punit.red (rep_list ` G))\<^sup>*\<^sup>* f 0" by (simp only: f s)
qed
lemma sig_red_zero_is_red:
assumes "sig_red_zero sing_reg F r" and "rep_list r \<noteq> 0"
shows "is_sig_red sing_reg (\<preceq>) F r"
proof -
from assms(1) obtain s where *: "(sig_red sing_reg (\<preceq>) F)\<^sup>*\<^sup>* r s" and "rep_list s = 0"
by (rule sig_red_zeroE)
from this(2) assms(2) have "r \<noteq> s" by auto
with * show ?thesis by (induct rule: converse_rtranclp_induct, auto simp: is_sig_red_def)
qed
lemma is_sig_red_sing_top_is_red_zero:
assumes "dickson_grading d" and "is_sig_GB_upt d G u" and "a \<in> dgrad_sig_set d" and "lt a = u"
and "is_sig_red (=) (=) G a" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G a"
shows "sig_red_zero (\<preceq>\<^sub>t) G a"
proof -
from assms(5) obtain g where "g \<in> G" and "rep_list g \<noteq> 0" and "rep_list a \<noteq> 0"
and 1: "punit.lt (rep_list g) adds punit.lt (rep_list a)"
and 2: "punit.lt (rep_list a) \<oplus> lt g = punit.lt (rep_list g) \<oplus> lt a"
by (rule is_sig_red_top_addsE)
from this(2, 3) have "g \<noteq> 0" and "a \<noteq> 0" by (auto simp: rep_list_zero)
hence "lc g \<noteq> 0" and "lc a \<noteq> 0" using lc_not_0 by blast+
from 1 have 3: "(punit.lt (rep_list a) - punit.lt (rep_list g)) \<oplus> lt g = lt a"
by (simp add: term_is_le_rel_minus 2)
define g' where "g' = monom_mult (lc a / lc g) (punit.lt (rep_list a) - punit.lt (rep_list g)) g"
from \<open>g \<noteq> 0\<close> \<open>lc a \<noteq> 0\<close> \<open>lc g \<noteq> 0\<close> have lt_g': "lt g' = lt a" by (simp add: g'_def lt_monom_mult 3)
from \<open>lc g \<noteq> 0\<close> have lc_g': "lc g' = lc a" by (simp add: g'_def)
from assms(1) have "g' \<in> dgrad_sig_set d" unfolding g'_def
proof (rule dgrad_sig_set_closed_monom_mult)
from assms(1) 1 have "d (punit.lt (rep_list a) - punit.lt (rep_list g)) \<le> d (punit.lt (rep_list a))"
by (rule dickson_grading_minus)
also from assms(1, 3) have "... \<le> dgrad_max d" by (rule dgrad_sig_setD_rep_list_lt)
finally show "d (punit.lt (rep_list a) - punit.lt (rep_list g)) \<le> dgrad_max d" .
next
from assms(2) have "G \<subseteq> dgrad_sig_set d" by (rule is_sig_GB_uptD1)
with \<open>g \<in> G\<close> show "g \<in> dgrad_sig_set d" ..
qed
with assms(3) have b_in: "a - g' \<in> dgrad_sig_set d" (is "?b \<in> _")
by (rule dgrad_sig_set_closed_minus)
from 1 have 4: "punit.lt (rep_list a) - punit.lt (rep_list g) + punit.lt (rep_list g) =
punit.lt (rep_list a)"
by (rule adds_minus)
show ?thesis
proof (cases "lc a / lc g = punit.lc (rep_list a) / punit.lc (rep_list g)")
case True
have "sig_red_single (=) (=) a ?b g (punit.lt (rep_list a) - punit.lt (rep_list g))"
proof (rule sig_red_singleI)
show "punit.lt (rep_list a) - punit.lt (rep_list g) + punit.lt (rep_list g) \<in> keys (rep_list a)"
unfolding 4 using \<open>rep_list a \<noteq> 0\<close> by (rule punit.lt_in_keys)
next
show "?b =
a - monom_mult
(lookup (rep_list a) (punit.lt (rep_list a) - punit.lt (rep_list g) + punit.lt (rep_list g)) /
punit.lc (rep_list g))
(punit.lt (rep_list a) - punit.lt (rep_list g)) g"
by (simp add: g'_def 4 punit.lc_def True)
qed (simp_all add: 3 4 \<open>rep_list g \<noteq> 0\<close>)
hence "sig_red (=) (=) G a ?b" unfolding sig_red_def using \<open>g \<in> G\<close> by blast
hence "sig_red (\<preceq>\<^sub>t) (\<preceq>) G a ?b" by (auto dest: sig_red_sing_regI sig_red_top_tailI)
hence 5: "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* a ?b" ..
show ?thesis
proof (cases "?b = 0")
case True
hence "rep_list ?b = 0" by (simp only: rep_list_zero)
with 5 show ?thesis by (rule sig_red_zeroI)
next
case False
hence "lt ?b \<prec>\<^sub>t lt a" using lt_g' lc_g' by (rule lt_minus_lessI)
hence "lt ?b \<prec>\<^sub>t u" by (simp only: assms(4))
with assms(2) b_in have "sig_red_zero (\<preceq>\<^sub>t) G ?b" by (rule is_sig_GB_uptD3)
then obtain s where "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* ?b s" and "rep_list s = 0" by (rule sig_red_zeroE)
from 5 this(1) have "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* a s" by (rule rtranclp_trans)
thus ?thesis using \<open>rep_list s = 0\<close> by (rule sig_red_zeroI)
qed
next
case False
from \<open>rep_list g \<noteq> 0\<close> \<open>lc g \<noteq> 0\<close> \<open>lc a \<noteq> 0\<close> have 5: "punit.lt (rep_list g') = punit.lt (rep_list a)"
by (simp add: g'_def rep_list_monom_mult punit.lt_monom_mult 4)
have 6: "punit.lc (rep_list g') = (lc a / lc g) * punit.lc (rep_list g)"
by (simp add: g'_def rep_list_monom_mult)
also have 7: "... \<noteq> punit.lc (rep_list a)"
proof
assume "lc a / lc g * punit.lc (rep_list g) = punit.lc (rep_list a)"
moreover from \<open>rep_list g \<noteq> 0\<close> have "punit.lc (rep_list g) \<noteq> 0" by (rule punit.lc_not_0)
ultimately have "lc a / lc g = punit.lc (rep_list a) / punit.lc (rep_list g)"
by (simp add: field_simps)
with False show False ..
qed
finally have "punit.lc (rep_list g') \<noteq> punit.lc (rep_list a)" .
with 5 have 8: "punit.lt (rep_list ?b) = punit.lt (rep_list a)" unfolding rep_list_minus
by (rule punit.lt_minus_eqI_3)
hence "punit.lc (rep_list ?b) = punit.lc (rep_list a) - (lc a / lc g) * punit.lc (rep_list g)"
unfolding 6[symmetric] by (simp only: punit.lc_def lookup_minus rep_list_minus 5)
also have "... \<noteq> 0"
proof
assume "punit.lc (rep_list a) - lc a / lc g * punit.lc (rep_list g) = 0"
hence "lc a / lc g * punit.lc (rep_list g) = punit.lc (rep_list a)" by simp
with 7 show False ..
qed
finally have "rep_list ?b \<noteq> 0" by (simp add: punit.lc_eq_zero_iff)
hence "?b \<noteq> 0" by (auto simp: rep_list_zero)
hence "lt ?b \<prec>\<^sub>t lt a" using lt_g' lc_g' by (rule lt_minus_lessI)
hence "lt ?b \<prec>\<^sub>t u" by (simp only: assms(4))
with assms(2) b_in have "sig_red_zero (\<preceq>\<^sub>t) G ?b" by (rule is_sig_GB_uptD3)
moreover note \<open>rep_list ?b \<noteq> 0\<close>
moreover have "(\<preceq>\<^sub>t) = (\<preceq>\<^sub>t) \<or> (\<preceq>\<^sub>t) = (\<prec>\<^sub>t)" by simp
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) G ?b" by (rule sig_red_zero_nonzero)
then obtain g0 where "g0 \<in> G" and "rep_list g0 \<noteq> 0"
and 9: "punit.lt (rep_list g0) adds punit.lt (rep_list ?b)"
and 10: "punit.lt (rep_list ?b) \<oplus> lt g0 \<preceq>\<^sub>t punit.lt (rep_list g0) \<oplus> lt ?b"
by (rule is_sig_red_top_addsE)
from 9 have "punit.lt (rep_list g0) adds punit.lt (rep_list a)" by (simp only: 8)
from 10 have "punit.lt (rep_list a) \<oplus> lt g0 \<preceq>\<^sub>t punit.lt (rep_list g0) \<oplus> lt ?b" by (simp only: 8)
also from \<open>lt ?b \<prec>\<^sub>t lt a\<close> have "... \<prec>\<^sub>t punit.lt (rep_list g0) \<oplus> lt a" by (rule splus_mono_strict)
finally have "punit.lt (rep_list a) \<oplus> lt g0 \<prec>\<^sub>t punit.lt (rep_list g0) \<oplus> lt a" .
have "is_sig_red (\<prec>\<^sub>t) (=) G a"
proof (rule is_sig_red_top_addsI)
show "ord_term_lin.is_le_rel (\<prec>\<^sub>t)" by simp
qed fact+
with assms(6) show ?thesis ..
qed
qed
lemma sig_regular_reduced_unique:
assumes "is_sig_GB_upt d G (lt q)" and "p \<in> dgrad_sig_set d" and "q \<in> dgrad_sig_set d"
and "lt p = lt q" and "lc p = lc q" and "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G p" and "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G q"
shows "rep_list p = rep_list q"
proof (rule ccontr)
assume "rep_list p \<noteq> rep_list q"
hence "rep_list (p - q) \<noteq> 0" by (auto simp: rep_list_minus)
hence "p - q \<noteq> 0" by (auto simp: rep_list_zero)
hence "p + (- q) \<noteq> 0" by simp
moreover from assms(4) have "lt (- q) = lt p" by simp
moreover from assms(5) have "lc (- q) = - lc p" by simp
ultimately have "lt (p + (- q)) \<prec>\<^sub>t lt p" by (rule lt_plus_lessI)
hence "lt (p - q) \<prec>\<^sub>t lt q" using assms(4) by simp
with assms(1) have "is_sig_GB_in d G (lt (p - q))" by (rule is_sig_GB_uptD2)
moreover from assms(2, 3) have "p - q \<in> dgrad_sig_set d" by (rule dgrad_sig_set_closed_minus)
ultimately have "sig_red_zero (\<preceq>\<^sub>t) G (p - q)" using refl by (rule is_sig_GB_inD)
hence "is_sig_red (\<preceq>\<^sub>t) (\<preceq>) G (p - q)" using \<open>rep_list (p - q) \<noteq> 0\<close> by (rule sig_red_zero_is_red)
then obtain g t where "g \<in> G" and t: "t \<in> keys (rep_list (p - q))" and "rep_list g \<noteq> 0"
and adds: "punit.lt (rep_list g) adds t" and "t \<oplus> lt g \<preceq>\<^sub>t punit.lt (rep_list g) \<oplus> lt (p - q)"
by (rule is_sig_red_addsE)
note this(5)
also from \<open>lt (p - q) \<prec>\<^sub>t lt q\<close> have "punit.lt (rep_list g) \<oplus> lt (p - q) \<prec>\<^sub>t punit.lt (rep_list g) \<oplus> lt q"
by (rule splus_mono_strict)
finally have 1: "t \<oplus> lt g \<prec>\<^sub>t punit.lt (rep_list g) \<oplus> lt q" .
hence 2: "t \<oplus> lt g \<prec>\<^sub>t punit.lt (rep_list g) \<oplus> lt p" by (simp only: assms(4))
from t keys_minus have "t \<in> keys (rep_list p) \<union> keys (rep_list q)" unfolding rep_list_minus ..
thus False
proof
assume t_in: "t \<in> keys (rep_list p)"
hence "t \<preceq> punit.lt (rep_list p)" by (rule punit.lt_max_keys)
with \<open>g \<in> G\<close> t_in \<open>rep_list g \<noteq> 0\<close> adds ord_term_lin.is_le_relI(3) ordered_powerprod_lin.is_le_relI(2) 2
have "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G p" by (rule is_sig_red_addsI)
with assms(6) show False ..
next
assume t_in: "t \<in> keys (rep_list q)"
hence "t \<preceq> punit.lt (rep_list q)" by (rule punit.lt_max_keys)
with \<open>g \<in> G\<close> t_in \<open>rep_list g \<noteq> 0\<close> adds ord_term_lin.is_le_relI(3) ordered_powerprod_lin.is_le_relI(2) 1
have "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G q" by (rule is_sig_red_addsI)
with assms(7) show False ..
qed
qed
corollary sig_regular_reduced_unique':
assumes "is_sig_GB_upt d G (lt q)" and "p \<in> dgrad_sig_set d" and "q \<in> dgrad_sig_set d"
and "lt p = lt q" and "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G p" and "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G q"
shows "punit.monom_mult (lc q) 0 (rep_list p) = punit.monom_mult (lc p) 0 (rep_list q)"
proof (cases "p = 0 \<or> q = 0")
case True
thus ?thesis by (auto simp: rep_list_zero)
next
case False
hence "p \<noteq> 0" and "q \<noteq> 0" by simp_all
hence "lc p \<noteq> 0" and "lc q \<noteq> 0" by (simp_all add: lc_not_0)
let ?p = "monom_mult (lc q) 0 p"
let ?q = "monom_mult (lc p) 0 q"
have "lt ?q = lt q" by (simp add: lt_monom_mult[OF \<open>lc p \<noteq> 0\<close> \<open>q \<noteq> 0\<close>] splus_zero)
with assms(1) have "is_sig_GB_upt d G (lt ?q)" by simp
moreover from assms(2) have "?p \<in> dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult_zero)
moreover from assms(3) have "?q \<in> dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult_zero)
moreover from \<open>lt ?q = lt q\<close> have "lt ?p = lt ?q"
by (simp add: lt_monom_mult[OF \<open>lc q \<noteq> 0\<close> \<open>p \<noteq> 0\<close>] splus_zero assms(4))
moreover have "lc ?p = lc ?q" by simp
moreover have "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G ?p"
proof
assume "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G ?p"
moreover from \<open>lc q \<noteq> 0\<close> have "1 / (lc q) \<noteq> 0" by simp
ultimately have "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G (monom_mult (1 / lc q) 0 ?p)" by (rule is_sig_red_monom_mult)
hence "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G p" by (simp add: monom_mult_assoc \<open>lc q \<noteq> 0\<close>)
with assms(5) show False ..
qed
moreover have "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G ?q"
proof
assume "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G ?q"
moreover from \<open>lc p \<noteq> 0\<close> have "1 / (lc p) \<noteq> 0" by simp
ultimately have "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G (monom_mult (1 / lc p) 0 ?q)" by (rule is_sig_red_monom_mult)
hence "is_sig_red (\<prec>\<^sub>t) (\<preceq>) G q" by (simp add: monom_mult_assoc \<open>lc p \<noteq> 0\<close>)
with assms(6) show False ..
qed
ultimately have "rep_list ?p = rep_list ?q" by (rule sig_regular_reduced_unique)
thus ?thesis by (simp only: rep_list_monom_mult)
qed
lemma sig_regular_top_reduced_lt_lc_unique:
assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)" and "p \<in> dgrad_sig_set d" and "q \<in> dgrad_sig_set d"
and "lt p = lt q" and "(p = 0) \<longleftrightarrow> (q = 0)" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G p" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G q"
shows "punit.lt (rep_list p) = punit.lt (rep_list q) \<and> lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)"
proof (cases "p = 0")
case True
with assms(6) have "q = 0" by simp
thus ?thesis by (simp add: True)
next
case False
with assms(6) have "q \<noteq> 0" by simp
from False have "lc p \<noteq> 0" by (rule lc_not_0)
from \<open>q \<noteq> 0\<close> have "lc q \<noteq> 0" by (rule lc_not_0)
from assms(2) have G_sub: "G \<subseteq> dgrad_sig_set d" by (rule is_sig_GB_uptD1)
hence "G \<subseteq> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
with assms(1) obtain p' where p'_red: "(sig_red (\<prec>\<^sub>t) (\<prec>) G)\<^sup>*\<^sup>* p p'" and "\<not> is_sig_red (\<prec>\<^sub>t) (\<prec>) G p'"
by (rule sig_irredE_dgrad_max_set)
from this(1) have lt_p': "lt p' = lt p" and lt_p'': "punit.lt (rep_list p') = punit.lt (rep_list p)"
and lc_p': "lc p' = lc p" and lc_p'': "punit.lc (rep_list p') = punit.lc (rep_list p)"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_tail_rtrancl_lt_rep_list,
rule sig_red_regular_rtrancl_lc, rule sig_red_tail_rtrancl_lc_rep_list)
have "\<not> is_sig_red (\<prec>\<^sub>t) (=) G p'"
proof
assume a: "is_sig_red (\<prec>\<^sub>t) (=) G p'"
hence "rep_list p' \<noteq> 0" using is_sig_red_top_addsE by blast
hence "rep_list p \<noteq> 0" using \<open>(sig_red (\<prec>\<^sub>t) (\<prec>) G)\<^sup>*\<^sup>* p p'\<close>
by (auto simp: punit.rtrancl_0 dest!: sig_red_red_rtrancl)
with a have "is_sig_red (\<prec>\<^sub>t) (=) G p" using lt_p' lt_p'' by (rule is_sig_red_top_cong)
with assms(7) show False ..
qed
with \<open>\<not> is_sig_red (\<prec>\<^sub>t) (\<prec>) G p'\<close> have 1: "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G p'" by (simp add: is_sig_red_top_tail_cases)
from assms(1) \<open>G \<subseteq> dgrad_max_set d\<close> obtain q' where q'_red: "(sig_red (\<prec>\<^sub>t) (\<prec>) G)\<^sup>*\<^sup>* q q'"
and "\<not> is_sig_red (\<prec>\<^sub>t) (\<prec>) G q'" by (rule sig_irredE_dgrad_max_set)
from this(1) have lt_q': "lt q' = lt q" and lt_q'': "punit.lt (rep_list q') = punit.lt (rep_list q)"
and lc_q': "lc q' = lc q" and lc_q'': "punit.lc (rep_list q') = punit.lc (rep_list q)"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_tail_rtrancl_lt_rep_list,
rule sig_red_regular_rtrancl_lc, rule sig_red_tail_rtrancl_lc_rep_list)
have "\<not> is_sig_red (\<prec>\<^sub>t) (=) G q'"
proof
assume a: "is_sig_red (\<prec>\<^sub>t) (=) G q'"
hence "rep_list q' \<noteq> 0" using is_sig_red_top_addsE by blast
hence "rep_list q \<noteq> 0" using \<open>(sig_red (\<prec>\<^sub>t) (\<prec>) G)\<^sup>*\<^sup>* q q'\<close>
by (auto simp: punit.rtrancl_0 dest!: sig_red_red_rtrancl)
with a have "is_sig_red (\<prec>\<^sub>t) (=) G q" using lt_q' lt_q'' by (rule is_sig_red_top_cong)
with assms(8) show False ..
qed
with \<open>\<not> is_sig_red (\<prec>\<^sub>t) (\<prec>) G q'\<close> have 2: "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G q'" by (simp add: is_sig_red_top_tail_cases)
from assms(2) have "is_sig_GB_upt d G (lt q')" by (simp only: lt_q')
moreover from assms(1) G_sub assms(3) p'_red have "p' \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_sig_red_rtrancl)
moreover from assms(1) G_sub assms(4) q'_red have "q' \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_sig_red_rtrancl)
moreover have "lt p' = lt q'" by (simp only: lt_p' lt_q' assms(5))
ultimately have eq: "punit.monom_mult (lc q') 0 (rep_list p') = punit.monom_mult (lc p') 0 (rep_list q')"
using 1 2 by (rule sig_regular_reduced_unique')
have "lc q * punit.lc (rep_list p) = lc q * punit.lc (rep_list p')" by (simp only: lc_p'')
also from \<open>lc q \<noteq> 0\<close> have "... = punit.lc (punit.monom_mult (lc q') 0 (rep_list p'))"
by (simp add: lc_q')
also have "... = punit.lc (punit.monom_mult (lc p') 0 (rep_list q'))" by (simp only: eq)
also from \<open>lc p \<noteq> 0\<close> have "... = lc p * punit.lc (rep_list q')" by (simp add: lc_p')
also have "... = lc p * punit.lc (rep_list q)" by (simp only: lc_q'')
finally have *: "lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)" .
have "punit.lt (rep_list p) = punit.lt (rep_list p')" by (simp only: lt_p'')
also from \<open>lc q \<noteq> 0\<close> have "... = punit.lt (punit.monom_mult (lc q') 0 (rep_list p'))"
by (simp add: lc_q' punit.lt_monom_mult_zero)
also have "... = punit.lt (punit.monom_mult (lc p') 0 (rep_list q'))" by (simp only: eq)
also from \<open>lc p \<noteq> 0\<close> have "... = punit.lt (rep_list q')" by (simp add: lc_p' punit.lt_monom_mult_zero)
also have "... = punit.lt (rep_list q)" by (fact lt_q'')
finally show ?thesis using * ..
qed
corollary sig_regular_top_reduced_lt_unique:
assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)" and "p \<in> dgrad_sig_set d"
and "q \<in> dgrad_sig_set d" and "lt p = lt q" and "p \<noteq> 0" and "q \<noteq> 0"
and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G p" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G q"
shows "punit.lt (rep_list p) = punit.lt (rep_list q)"
proof -
from assms(6, 7) have "(p = 0) \<longleftrightarrow> (q = 0)" by simp
with assms(1, 2, 3, 4, 5)
have "punit.lt (rep_list p) = punit.lt (rep_list q) \<and> lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)"
using assms(8, 9) by (rule sig_regular_top_reduced_lt_lc_unique)
thus ?thesis ..
qed
corollary sig_regular_top_reduced_lc_unique:
assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)" and "p \<in> dgrad_sig_set d" and "q \<in> dgrad_sig_set d"
and "lt p = lt q" and "lc p = lc q" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G p" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G q"
shows "punit.lc (rep_list p) = punit.lc (rep_list q)"
proof (cases "p = 0")
case True
with assms(6) have "q = 0" by (simp add: lc_eq_zero_iff)
with True show ?thesis by simp
next
case False
hence "lc p \<noteq> 0" by (rule lc_not_0)
hence "lc q \<noteq> 0" by (simp add: assms(6))
hence "q \<noteq> 0" by (simp add: lc_eq_zero_iff)
with False have "(p = 0) \<longleftrightarrow> (q = 0)" by simp
with assms(1, 2, 3, 4, 5)
have "punit.lt (rep_list p) = punit.lt (rep_list q) \<and> lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)"
using assms(7, 8) by (rule sig_regular_top_reduced_lt_lc_unique)
hence "lc q * punit.lc (rep_list p) = lc p * punit.lc (rep_list q)" ..
also have "... = lc q * punit.lc (rep_list q)" by (simp only: assms(6))
finally show ?thesis using \<open>lc q \<noteq> 0\<close> by simp
qed
text \<open>Minimal signature Gr\"obner bases are indeed minimal, at least up to sig-lead-pairs:\<close>
lemma is_min_sig_GB_minimal:
assumes "is_min_sig_GB d G" and "G' \<subseteq> dgrad_sig_set d"
and "\<And>u. d (pp_of_term u) \<le> dgrad_max d \<Longrightarrow> component_of_term u < length fs \<Longrightarrow> is_sig_GB_in d G' u"
and "g \<in> G" and "rep_list g \<noteq> 0"
obtains g' where "g' \<in> G'" and "rep_list g' \<noteq> 0" and "lt g' = lt g"
and "punit.lt (rep_list g') = punit.lt (rep_list g)"
proof -
from assms(1) have "G \<subseteq> dgrad_sig_set d"
and 1: "\<And>u. d (pp_of_term u) \<le> dgrad_max d \<Longrightarrow> component_of_term u < length fs \<Longrightarrow> is_sig_GB_in d G u"
and 2: "\<And>g0. g0 \<in> G \<Longrightarrow> \<not> is_sig_red (\<preceq>\<^sub>t) (=) (G - {g0}) g0"
by (simp_all add: is_min_sig_GB_def)
from assms(4) have 3: "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (G - {g}) g" by (rule 2)
from assms(5) have "g \<noteq> 0" by (auto simp: rep_list_zero)
from assms(4) \<open>G \<subseteq> dgrad_sig_set d\<close> have "g \<in> dgrad_sig_set d" ..
hence "d (lp g) \<le> dgrad_max d" and "component_of_term (lt g) < length fs"
by (rule dgrad_sig_setD_lp, rule dgrad_sig_setD_lt[OF _ \<open>g \<noteq> 0\<close>])
hence "is_sig_GB_in d G' (lt g)" by (rule assms(3))
hence "sig_red_zero (\<preceq>\<^sub>t) G' g" using \<open>g \<in> dgrad_sig_set d\<close> refl by (rule is_sig_GB_inD)
moreover note assms(5)
moreover have "(\<preceq>\<^sub>t) = (\<preceq>\<^sub>t) \<or> (\<preceq>\<^sub>t) = (\<prec>\<^sub>t)" by simp
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) G' g" by (rule sig_red_zero_nonzero)
then obtain g' where "g' \<in> G'" and "rep_list g' \<noteq> 0"
and adds1: "punit.lt (rep_list g') adds punit.lt (rep_list g)"
and le1: "punit.lt (rep_list g) \<oplus> lt g' \<preceq>\<^sub>t punit.lt (rep_list g') \<oplus> lt g"
by (rule is_sig_red_top_addsE)
from \<open>rep_list g' \<noteq> 0\<close> have "g' \<noteq> 0" by (auto simp: rep_list_zero)
from \<open>g' \<in> G'\<close> assms(2) have "g' \<in> dgrad_sig_set d" ..
hence "d (lp g') \<le> dgrad_max d" and "component_of_term (lt g') < length fs"
by (rule dgrad_sig_setD_lp, rule dgrad_sig_setD_lt[OF _ \<open>g' \<noteq> 0\<close>])
hence "is_sig_GB_in d G (lt g')" by (rule 1)
hence "sig_red_zero (\<preceq>\<^sub>t) G g'" using \<open>g' \<in> dgrad_sig_set d\<close> refl by (rule is_sig_GB_inD)
moreover note \<open>rep_list g' \<noteq> 0\<close>
moreover have "(\<preceq>\<^sub>t) = (\<preceq>\<^sub>t) \<or> (\<preceq>\<^sub>t) = (\<prec>\<^sub>t)" by simp
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) G g'" by (rule sig_red_zero_nonzero)
then obtain g0 where "g0 \<in> G" and "rep_list g0 \<noteq> 0"
and adds2: "punit.lt (rep_list g0) adds punit.lt (rep_list g')"
and le2: "punit.lt (rep_list g') \<oplus> lt g0 \<preceq>\<^sub>t punit.lt (rep_list g0) \<oplus> lt g'"
by (rule is_sig_red_top_addsE)
have eq1: "g0 = g"
proof (rule ccontr)
assume "g0 \<noteq> g"
with \<open>g0 \<in> G\<close> have "g0 \<in> G - {g}" by simp
moreover note \<open>rep_list g0 \<noteq> 0\<close> assms(5)
moreover from adds2 adds1 have "punit.lt (rep_list g0) adds punit.lt (rep_list g)"
by (rule adds_trans)
moreover have "ord_term_lin.is_le_rel (\<preceq>\<^sub>t)" by simp
moreover have "punit.lt (rep_list g) \<oplus> lt g0 \<preceq>\<^sub>t punit.lt (rep_list g0) \<oplus> lt g"
proof (rule ord_term_canc)
have "punit.lt (rep_list g') \<oplus> (punit.lt (rep_list g) \<oplus> lt g0) =
punit.lt (rep_list g) \<oplus> (punit.lt (rep_list g') \<oplus> lt g0)" by (fact splus_left_commute)
also from le2 have "... \<preceq>\<^sub>t punit.lt (rep_list g) \<oplus> (punit.lt (rep_list g0) \<oplus> lt g')"
by (rule splus_mono)
also have "... = punit.lt (rep_list g0) \<oplus> (punit.lt (rep_list g) \<oplus> lt g')"
by (fact splus_left_commute)
also from le1 have "... \<preceq>\<^sub>t punit.lt (rep_list g0) \<oplus> (punit.lt (rep_list g') \<oplus> lt g)"
by (rule splus_mono)
also have "... = punit.lt (rep_list g') \<oplus> (punit.lt (rep_list g0) \<oplus> lt g)"
by (fact splus_left_commute)
finally show "punit.lt (rep_list g') \<oplus> (punit.lt (rep_list g) \<oplus> lt g0) \<preceq>\<^sub>t
punit.lt (rep_list g') \<oplus> (punit.lt (rep_list g0) \<oplus> lt g)" .
qed
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) (G - {g}) g" by (rule is_sig_red_top_addsI)
with 3 show False ..
qed
from adds2 adds1 have eq2: "punit.lt (rep_list g') = punit.lt (rep_list g)" by (simp add: eq1 adds_antisym)
with le1 le2 have "punit.lt (rep_list g) \<oplus> lt g' = punit.lt (rep_list g) \<oplus> lt g" by (simp add: eq1)
hence "lt g' = lt g" by (simp only: splus_left_canc)
with \<open>g' \<in> G'\<close> \<open>rep_list g' \<noteq> 0\<close> show ?thesis using eq2 ..
qed
lemma sig_red_zero_regularI_adds:
assumes "dickson_grading d" and "is_sig_GB_upt d G (lt q)"
and "p \<in> dgrad_sig_set d" and "q \<in> dgrad_sig_set d" and "p \<noteq> 0" and "sig_red_zero (\<prec>\<^sub>t) G p"
and "lt p adds\<^sub>t lt q"
shows "sig_red_zero (\<prec>\<^sub>t) G q"
proof (cases "q = 0")
case True
hence "rep_list q = 0" by (simp only: rep_list_zero)
with rtrancl_refl[to_pred] show ?thesis by (rule sig_red_zeroI)
next
case False
hence "lc q \<noteq> 0" by (rule lc_not_0)
moreover from assms(5) have "lc p \<noteq> 0" by (rule lc_not_0)
ultimately have "lc q / lc p \<noteq> 0" by simp
from assms(7) have eq1: "(lp q - lp p) \<oplus> lt p = lt q"
by (metis add_diff_cancel_right' adds_termE pp_of_term_splus)
from assms(7) have "lp p adds lp q" by (simp add: adds_term_def)
with assms(1) have "d (lp q - lp p) \<le> d (lp q)" by (rule dickson_grading_minus)
also from assms(4) have "... \<le> dgrad_max d" by (rule dgrad_sig_setD_lp)
finally have "d (lp q - lp p) \<le> dgrad_max d" .
from assms(2) have G_sub: "G \<subseteq> dgrad_sig_set d" by (rule is_sig_GB_uptD1)
hence "G \<subseteq> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
let ?mult = "\<lambda>r. monom_mult (lc q / lc p) (lp q - lp p) r"
from assms(6) obtain p' where p_red: "(sig_red (\<prec>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* p p'" and "rep_list p' = 0"
by (rule sig_red_zeroE)
from p_red have "lt p' = lt p" and "lc p' = lc p"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
hence "p' \<noteq> 0" using \<open>lc p \<noteq> 0\<close> by auto
with \<open>lc q / lc p \<noteq> 0\<close> have "?mult p' \<noteq> 0" by (simp add: monom_mult_eq_zero_iff)
from \<open>lc q / lc p \<noteq> 0\<close> \<open>p' \<noteq> 0\<close> have "lt (?mult p') = lt q"
by (simp add: lt_monom_mult \<open>lt p' = lt p\<close> eq1)
from \<open>lc p \<noteq> 0\<close> have "lc (?mult p') = lc q" by (simp add: \<open>lc p' = lc p\<close>)
from p_red have mult_p_red: "(sig_red (\<prec>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* (?mult p) (?mult p')"
by (rule sig_red_rtrancl_monom_mult)
have "rep_list (?mult p') = 0" by (simp add: rep_list_monom_mult \<open>rep_list p' = 0\<close>)
hence mult_p'_irred: "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G (?mult p')"
using is_sig_red_addsE by fastforce
from assms(1) G_sub assms(3) p_red have "p' \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_sig_red_rtrancl)
with assms(1) \<open>d (lp q - lp p) \<le> dgrad_max d\<close> have "?mult p' \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_monom_mult)
from assms(1) \<open>G \<subseteq> dgrad_max_set d\<close> obtain q' where q_red: "(sig_red (\<prec>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* q q'"
and q'_irred: "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G q'" by (rule sig_irredE_dgrad_max_set)
from q_red have "lt q' = lt q" and "lc q' = lc q"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
hence "q' \<noteq> 0" using \<open>lc q \<noteq> 0\<close> by auto
from assms(2) have "is_sig_GB_upt d G (lt (?mult p'))" by (simp only: \<open>lt (?mult p') = lt q\<close>)
moreover from assms(1) G_sub assms(4) q_red have "q' \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_sig_red_rtrancl)
moreover note \<open>?mult p' \<in> dgrad_sig_set d\<close>
moreover have "lt q' = lt (?mult p')" by (simp only: \<open>lt (?mult p') = lt q\<close> \<open>lt q' = lt q\<close>)
moreover have "lc q' = lc (?mult p')" by (simp only: \<open>lc (?mult p') = lc q\<close> \<open>lc q' = lc q\<close>)
ultimately have "rep_list q' = rep_list (?mult p')" using q'_irred mult_p'_irred
by (rule sig_regular_reduced_unique)
with \<open>rep_list (?mult p') = 0\<close> have "rep_list q' = 0" by simp
with q_red show ?thesis by (rule sig_red_zeroI)
qed
lemma is_syz_sigI:
assumes "s \<noteq> 0" and "lt s = u" and "s \<in> dgrad_sig_set d" and "rep_list s = 0"
shows "is_syz_sig d u"
unfolding is_syz_sig_def using assms by blast
lemma is_syz_sigE:
assumes "is_syz_sig d u"
obtains r where "r \<noteq> 0" and "lt r = u" and "r \<in> dgrad_sig_set d" and "rep_list r = 0"
using assms unfolding is_syz_sig_def by blast
lemma is_syz_sig_adds:
assumes "dickson_grading d" and "is_syz_sig d u" and "u adds\<^sub>t v"
and "d (pp_of_term v) \<le> dgrad_max d"
shows "is_syz_sig d v"
proof -
from assms(2) obtain s where "s \<noteq> 0" and "lt s = u" and "s \<in> dgrad_sig_set d"
and "rep_list s = 0" by (rule is_syz_sigE)
from assms(3) obtain t where v: "v = t \<oplus> u" by (rule adds_termE)
show ?thesis
proof (rule is_syz_sigI)
from \<open>s \<noteq> 0\<close> show "monom_mult 1 t s \<noteq> 0" by (simp add: monom_mult_eq_zero_iff)
next
from \<open>s \<noteq> 0\<close> show "lt (monom_mult 1 t s) = v" by (simp add: lt_monom_mult v \<open>lt s = u\<close>)
next
from assms(4) have "d (t + pp_of_term u) \<le> dgrad_max d" by (simp add: v term_simps)
with assms(1) have "d t \<le> dgrad_max d" by (simp add: dickson_gradingD1)
with assms(1) show "monom_mult 1 t s \<in> dgrad_sig_set d" using \<open>s \<in> dgrad_sig_set d\<close>
by (rule dgrad_sig_set_closed_monom_mult)
next
show "rep_list (monom_mult 1 t s) = 0" by (simp add: \<open>rep_list s = 0\<close> rep_list_monom_mult)
qed
qed
lemma syzygy_crit:
assumes "dickson_grading d" and "is_sig_GB_upt d G u" and "is_syz_sig d u"
and "p \<in> dgrad_sig_set d" and "lt p = u"
shows "sig_red_zero (\<prec>\<^sub>t) G p"
proof -
from assms(3) obtain s where "s \<noteq> 0" and "lt s = u" and "s \<in> dgrad_sig_set d"
and "rep_list s = 0" by (rule is_syz_sigE)
note assms(1)
moreover from assms(2) have "is_sig_GB_upt d G (lt p)" by (simp only: assms(5))
moreover note \<open>s \<in> dgrad_sig_set d\<close> assms(4) \<open>s \<noteq> 0\<close>
moreover from rtranclp.rtrancl_refl \<open>rep_list s = 0\<close> have "sig_red_zero (\<prec>\<^sub>t) G s"
by (rule sig_red_zeroI)
moreover have "lt s adds\<^sub>t lt p" by (simp only: assms(5) \<open>lt s = u\<close> adds_term_refl)
ultimately show ?thesis by (rule sig_red_zero_regularI_adds)
qed
lemma lemma_21:
assumes "dickson_grading d" and "is_sig_GB_upt d G (lt p)" and "p \<in> dgrad_sig_set d" and "g \<in> G"
and "rep_list p \<noteq> 0" and "rep_list g \<noteq> 0" and "lt g adds\<^sub>t lt p"
and "punit.lt (rep_list g) adds punit.lt (rep_list p)"
shows "is_sig_red (\<preceq>\<^sub>t) (=) G p"
proof -
let ?lp = "punit.lt (rep_list p)"
define s where "s = ?lp - punit.lt (rep_list g)"
from assms(8) have s: "?lp = s + punit.lt (rep_list g)" by (simp add: s_def minus_plus)
from assms(7) obtain t where lt_p: "lt p = t \<oplus> lt g" by (rule adds_termE)
show ?thesis
proof (cases "s \<oplus> lt g \<preceq>\<^sub>t lt p")
case True
hence "?lp \<oplus> lt g \<preceq>\<^sub>t punit.lt (rep_list g) \<oplus> lt p"
by (simp add: s splus_assoc splus_left_commute[of s] splus_mono)
with assms(4, 6, 5, 8) ord_term_lin.is_le_relI(2) show ?thesis
by (rule is_sig_red_top_addsI)
next
case False
hence "lt p \<prec>\<^sub>t s \<oplus> lt g" by simp
hence "t \<prec> s" by (simp add: lt_p ord_term_strict_canc_left)
hence "t + punit.lt (rep_list g) \<prec> s + punit.lt (rep_list g)" by (rule plus_monotone_strict)
hence "t + punit.lt (rep_list g) \<prec> ?lp" by (simp only: s)
from assms(5) have "p \<noteq> 0" by (auto simp: rep_list_zero)
hence "lc p \<noteq> 0" by (rule lc_not_0)
from assms(6) have "g \<noteq> 0" by (auto simp: rep_list_zero)
hence "lc g \<noteq> 0" by (rule lc_not_0)
with \<open>lc p \<noteq> 0\<close> have 1: "lc p / lc g \<noteq> 0" by simp
let ?g = "monom_mult (lc p / lc g) t g"
from 1 \<open>g \<noteq> 0\<close> have "lt ?g = lt p" unfolding lt_p by (rule lt_monom_mult)
from \<open>lc g \<noteq> 0\<close> have "lc ?g = lc p" by simp
have "punit.lt (rep_list ?g) = t + punit.lt (rep_list g)"
unfolding rep_list_monom_mult using 1 assms(6) by (rule punit.lt_monom_mult[simplified])
also have "... \<prec> ?lp" by fact
finally have "punit.lt (rep_list ?g) \<prec> ?lp" .
hence lt_pg: "punit.lt (rep_list (p - ?g)) = ?lp" and "rep_list p \<noteq> rep_list ?g"
by (auto simp: rep_list_minus punit.lt_minus_eqI_2)
from this(2) have "rep_list (p - ?g) \<noteq> 0" and "p - ?g \<noteq> 0"
by (auto simp: rep_list_minus rep_list_zero)
from assms(2) have "G \<subseteq> dgrad_sig_set d" by (rule is_sig_GB_uptD1)
note assms(1)
moreover have "d t \<le> dgrad_max d"
proof (rule le_trans)
have "lp p = t + lp g" by (simp add: lt_p term_simps)
with assms(1) show "d t \<le> d (lp p)" by (simp add: dickson_grading_adds_imp_le)
next
from assms(3) show "d (lp p) \<le> dgrad_max d" by (rule dgrad_sig_setD_lp)
qed
moreover from assms(4) \<open>G \<subseteq> dgrad_sig_set d\<close> have "g \<in> dgrad_sig_set d" ..
ultimately have "?g \<in> dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult)
note assms(2)
moreover from assms(3) \<open>?g \<in> dgrad_sig_set d\<close> have "p - ?g \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_minus)
moreover from \<open>p - ?g \<noteq> 0\<close> \<open>lt ?g = lt p\<close> \<open>lc ?g = lc p\<close> have "lt (p - ?g) \<prec>\<^sub>t lt p"
by (rule lt_minus_lessI)
ultimately have "sig_red_zero (\<preceq>\<^sub>t) G (p - ?g)"
by (rule is_sig_GB_uptD3)
moreover note \<open>rep_list (p - ?g) \<noteq> 0\<close>
moreover have "(\<preceq>\<^sub>t) = (\<preceq>\<^sub>t) \<or> (\<preceq>\<^sub>t) = (\<prec>\<^sub>t)" by simp
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) G (p - ?g)" by (rule sig_red_zero_nonzero)
then obtain g1 where "g1 \<in> G" and "rep_list g1 \<noteq> 0"
and 2: "punit.lt (rep_list g1) adds punit.lt (rep_list (p - ?g))"
and 3: "punit.lt (rep_list (p - ?g)) \<oplus> lt g1 \<preceq>\<^sub>t punit.lt (rep_list g1) \<oplus> lt (p - ?g)"
by (rule is_sig_red_top_addsE)
from \<open>g1 \<in> G\<close> \<open>rep_list g1 \<noteq> 0\<close> assms(5) show ?thesis
proof (rule is_sig_red_top_addsI)
from 2 show "punit.lt (rep_list g1) adds punit.lt (rep_list p)" by (simp only: lt_pg)
next
have "?lp \<oplus> lt g1 = punit.lt (rep_list (p - ?g)) \<oplus> lt g1" by (simp only: lt_pg)
also have "... \<preceq>\<^sub>t punit.lt (rep_list g1) \<oplus> lt (p - ?g)" by (fact 3)
also from \<open>lt (p - ?g) \<prec>\<^sub>t lt p\<close> have "... \<prec>\<^sub>t punit.lt (rep_list g1) \<oplus> lt p"
by (rule splus_mono_strict)
finally show "?lp \<oplus> lt g1 \<preceq>\<^sub>t punit.lt (rep_list g1) \<oplus> lt p" by (rule ord_term_lin.less_imp_le)
qed simp
qed
qed
subsubsection \<open>Rewrite Bases\<close>
definition is_rewrite_ord :: "(('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool) \<Rightarrow> bool"
where "is_rewrite_ord rword \<longleftrightarrow> (reflp rword \<and> transp rword \<and> (\<forall>a b. rword a b \<or> rword b a) \<and>
(\<forall>a b. rword a b \<longrightarrow> rword b a \<longrightarrow> fst a = fst b) \<and>
(\<forall>d G a b. dickson_grading d \<longrightarrow> is_sig_GB_upt d G (lt b) \<longrightarrow>
a \<in> G \<longrightarrow> b \<in> G \<longrightarrow> a \<noteq> 0 \<longrightarrow> b \<noteq> 0 \<longrightarrow> lt a adds\<^sub>t lt b \<longrightarrow>
\<not> is_sig_red (\<prec>\<^sub>t) (=) G b \<longrightarrow> rword (spp_of a) (spp_of b)))"
definition is_canon_rewriter :: "(('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> 't \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "is_canon_rewriter rword A u p \<longleftrightarrow>
(p \<in> A \<and> p \<noteq> 0 \<and> lt p adds\<^sub>t u \<and> (\<forall>a\<in>A. a \<noteq> 0 \<longrightarrow> lt a adds\<^sub>t u \<longrightarrow> rword (spp_of a) (spp_of p)))"
definition is_RB_in :: "('a \<Rightarrow> nat) \<Rightarrow> (('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> 't \<Rightarrow> bool"
where "is_RB_in d rword G u \<longleftrightarrow>
((\<exists>g. is_canon_rewriter rword G u g \<and> \<not> is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term u - lp g) g)) \<or>
is_syz_sig d u)"
definition is_RB_upt :: "('a \<Rightarrow> nat) \<Rightarrow> (('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) set \<Rightarrow> 't \<Rightarrow> bool"
where "is_RB_upt d rword G u \<longleftrightarrow>
(G \<subseteq> dgrad_sig_set d \<and> (\<forall>v. v \<prec>\<^sub>t u \<longrightarrow> d (pp_of_term v) \<le> dgrad_max d \<longrightarrow>
component_of_term v < length fs \<longrightarrow> is_RB_in d rword G v))"
lemma is_rewrite_ordI:
assumes "reflp rword" and "transp rword" and "\<And>a b. rword a b \<or> rword b a"
and "\<And>a b. rword a b \<Longrightarrow> rword b a \<Longrightarrow> fst a = fst b"
and "\<And>d G a b. dickson_grading d \<Longrightarrow> is_sig_GB_upt d G (lt b) \<Longrightarrow> a \<in> G \<Longrightarrow> b \<in> G \<Longrightarrow>
a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> lt a adds\<^sub>t lt b \<Longrightarrow> \<not> is_sig_red (\<prec>\<^sub>t) (=) G b \<Longrightarrow> rword (spp_of a) (spp_of b)"
shows "is_rewrite_ord rword"
unfolding is_rewrite_ord_def using assms by blast
lemma is_rewrite_ordD1: "is_rewrite_ord rword \<Longrightarrow> rword a a"
by (simp add: is_rewrite_ord_def reflpD)
lemma is_rewrite_ordD2: "is_rewrite_ord rword \<Longrightarrow> rword a b \<Longrightarrow> rword b c \<Longrightarrow> rword a c"
by (auto simp: is_rewrite_ord_def dest: transpD)
lemma is_rewrite_ordD3:
assumes "is_rewrite_ord rword"
and "rword a b \<Longrightarrow> thesis"
and "\<not> rword a b \<Longrightarrow> rword b a \<Longrightarrow> thesis"
shows thesis
proof -
from assms(1) have disj: "rword a b \<or> rword b a"
by (simp add: is_rewrite_ord_def del: split_paired_All)
show ?thesis
proof (cases "rword a b")
case True
thus ?thesis by (rule assms(2))
next
case False
moreover from this disj have "rword b a" by simp
ultimately show ?thesis by (rule assms(3))
qed
qed
lemma is_rewrite_ordD4:
assumes "is_rewrite_ord rword" and "rword a b" and "rword b a"
shows "fst a = fst b"
using assms unfolding is_rewrite_ord_def by blast
lemma is_rewrite_ordD4':
assumes "is_rewrite_ord rword" and "rword (spp_of a) (spp_of b)" and "rword (spp_of b) (spp_of a)"
shows "lt a = lt b"
proof -
from assms have "fst (spp_of a) = fst (spp_of b)" by (rule is_rewrite_ordD4)
thus ?thesis by (simp add: spp_of_def)
qed
lemma is_rewrite_ordD5:
assumes "is_rewrite_ord rword" and "dickson_grading d" and "is_sig_GB_upt d G (lt b)"
and "a \<in> G" and "b \<in> G" and "a \<noteq> 0" and "b \<noteq> 0" and "lt a adds\<^sub>t lt b"
and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G b"
shows "rword (spp_of a) (spp_of b)"
using assms unfolding is_rewrite_ord_def by blast
lemma is_canon_rewriterI:
assumes "p \<in> A" and "p \<noteq> 0" and "lt p adds\<^sub>t u"
and "\<And>a. a \<in> A \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> lt a adds\<^sub>t u \<Longrightarrow> rword (spp_of a) (spp_of p)"
shows "is_canon_rewriter rword A u p"
unfolding is_canon_rewriter_def using assms by blast
lemma is_canon_rewriterD1: "is_canon_rewriter rword A u p \<Longrightarrow> p \<in> A"
by (simp add: is_canon_rewriter_def)
lemma is_canon_rewriterD2: "is_canon_rewriter rword A u p \<Longrightarrow> p \<noteq> 0"
by (simp add: is_canon_rewriter_def)
lemma is_canon_rewriterD3: "is_canon_rewriter rword A u p \<Longrightarrow> lt p adds\<^sub>t u"
by (simp add: is_canon_rewriter_def)
lemma is_canon_rewriterD4:
"is_canon_rewriter rword A u p \<Longrightarrow> a \<in> A \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> lt a adds\<^sub>t u \<Longrightarrow> rword (spp_of a) (spp_of p)"
by (simp add: is_canon_rewriter_def)
lemmas is_canon_rewriterD = is_canon_rewriterD1 is_canon_rewriterD2 is_canon_rewriterD3 is_canon_rewriterD4
lemma is_rewrite_ord_finite_canon_rewriterE:
assumes "is_rewrite_ord rword" and "finite A" and "a \<in> A" and "a \<noteq> 0" and "lt a adds\<^sub>t u"
obtains p where "is_canon_rewriter rword A u p"
proof -
let ?A = "{x. x \<in> A \<and> x \<noteq> 0 \<and> lt x adds\<^sub>t u}"
let ?rel = "\<lambda>x y. strict rword (spp_of y) (spp_of x)"
have "finite ?A"
proof (rule finite_subset)
show "?A \<subseteq> A" by blast
qed fact
moreover have "?A \<noteq> {}"
proof
from assms(3, 4, 5) have "a \<in> ?A" by simp
also assume "?A = {}"
finally show False by simp
qed
moreover have "irreflp ?rel"
proof -
from assms(1) have "reflp rword" by (simp add: is_rewrite_ord_def)
thus ?thesis by (simp add: reflp_def irreflp_def)
qed
moreover have "transp ?rel"
proof -
from assms(1) have "transp rword" by (simp add: is_rewrite_ord_def)
thus ?thesis by (auto simp: transp_def simp del: split_paired_All)
qed
ultimately obtain p where "p \<in> ?A" and *: "\<And>b. ?rel b p \<Longrightarrow> b \<notin> ?A" by (rule finite_minimalE, blast)
from this(1) have "p \<in> A" and "p \<noteq> 0" and "lt p adds\<^sub>t u" by simp_all
show ?thesis
proof (rule, rule is_canon_rewriterI)
fix q
assume "q \<in> A" and "q \<noteq> 0" and "lt q adds\<^sub>t u"
hence "q \<in> ?A" by simp
with * have "\<not> ?rel q p" by blast
hence disj: "\<not> rword (spp_of p) (spp_of q) \<or> rword (spp_of q) (spp_of p)" by simp
from assms(1) show "rword (spp_of q) (spp_of p)"
proof (rule is_rewrite_ordD3)
assume "\<not> rword (spp_of q) (spp_of p)" and "rword (spp_of p) (spp_of q)"
with disj show ?thesis by simp
qed
qed fact+
qed
lemma is_rewrite_ord_canon_rewriterD1:
assumes "is_rewrite_ord rword" and "is_canon_rewriter rword A u p" and "is_canon_rewriter rword A v q"
and "lt p adds\<^sub>t v" and "lt q adds\<^sub>t u"
shows "lt p = lt q"
proof -
from assms(2) have "p \<in> A" and "p \<noteq> 0"
and 1: "\<And>a. a \<in> A \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> lt a adds\<^sub>t u \<Longrightarrow> rword (spp_of a) (spp_of p)"
by (rule is_canon_rewriterD)+
from assms(3) have "q \<in> A" and "q \<noteq> 0"
and 2: "\<And>a. a \<in> A \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> lt a adds\<^sub>t v \<Longrightarrow> rword (spp_of a) (spp_of q)"
by (rule is_canon_rewriterD)+
note assms(1)
moreover from \<open>p \<in> A\<close> \<open>p \<noteq> 0\<close> assms(4) have "rword (spp_of p) (spp_of q)" by (rule 2)
moreover from \<open>q \<in> A\<close> \<open>q \<noteq> 0\<close> assms(5) have "rword (spp_of q) (spp_of p)" by (rule 1)
ultimately show ?thesis by (rule is_rewrite_ordD4')
qed
corollary is_rewrite_ord_canon_rewriterD2:
assumes "is_rewrite_ord rword" and "is_canon_rewriter rword A u p" and "is_canon_rewriter rword A u q"
shows "lt p = lt q"
using assms
proof (rule is_rewrite_ord_canon_rewriterD1)
from assms(2) show "lt p adds\<^sub>t u" by (rule is_canon_rewriterD)
next
from assms(3) show "lt q adds\<^sub>t u" by (rule is_canon_rewriterD)
qed
lemma is_rewrite_ord_canon_rewriterD3:
assumes "is_rewrite_ord rword" and "dickson_grading d" and "is_canon_rewriter rword A u p"
and "a \<in> A" and "a \<noteq> 0" and "lt a adds\<^sub>t u" and "is_sig_GB_upt d A (lt a)"
and "lt p adds\<^sub>t lt a" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) A a"
shows "lt p = lt a"
proof -
note assms(1)
moreover from assms(1, 2, 7) _ assms(4) _ assms(5, 8, 9) have "rword (spp_of p) (spp_of a)"
proof (rule is_rewrite_ordD5)
from assms(3) show "p \<in> A" and "p \<noteq> 0" by (rule is_canon_rewriterD)+
qed
moreover from assms(3, 4, 5, 6) have "rword (spp_of a) (spp_of p)" by (rule is_canon_rewriterD4)
ultimately show ?thesis by (rule is_rewrite_ordD4')
qed
lemma is_RB_inI1:
assumes "is_canon_rewriter rword G u g" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term u - lp g) g)"
shows "is_RB_in d rword G u"
unfolding is_RB_in_def using assms is_canon_rewriterD1 by blast
lemma is_RB_inI2:
assumes "is_syz_sig d u"
shows "is_RB_in d rword G u"
unfolding is_RB_in_def Let_def using assms by blast
lemma is_RB_inE:
assumes "is_RB_in d rword G u"
and "is_syz_sig d u \<Longrightarrow> thesis"
and "\<And>g. \<not> is_syz_sig d u \<Longrightarrow> is_canon_rewriter rword G u g \<Longrightarrow>
\<not> is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term u - lp g) g) \<Longrightarrow> thesis"
shows thesis
using assms unfolding is_RB_in_def by blast
lemma is_RB_inD:
assumes "dickson_grading d" and "G \<subseteq> dgrad_sig_set d" and "is_RB_in d rword G u"
and "\<not> is_syz_sig d u" and "d (pp_of_term u) \<le> dgrad_max d"
and "is_canon_rewriter rword G u g"
shows "rep_list g \<noteq> 0"
proof
assume a: "rep_list g = 0"
from assms(1) have "is_syz_sig d u"
proof (rule is_syz_sig_adds)
show "is_syz_sig d (lt g)"
proof (rule is_syz_sigI)
from assms(6) show "g \<noteq> 0" by (rule is_canon_rewriterD2)
next
from assms(6) have "g \<in> G" by (rule is_canon_rewriterD1)
thus "g \<in> dgrad_sig_set d" using assms(2) ..
qed (fact refl, fact a)
next
from assms(6) show "lt g adds\<^sub>t u" by (rule is_canon_rewriterD3)
qed fact
with assms(4) show False ..
qed
lemma is_RB_uptI:
assumes "G \<subseteq> dgrad_sig_set d"
and "\<And>v. v \<prec>\<^sub>t u \<Longrightarrow> d (pp_of_term v) \<le> dgrad_max d \<Longrightarrow> component_of_term v < length fs \<Longrightarrow>
is_RB_in d canon G v"
shows "is_RB_upt d canon G u"
unfolding is_RB_upt_def using assms by blast
lemma is_RB_uptD1:
assumes "is_RB_upt d canon G u"
shows "G \<subseteq> dgrad_sig_set d"
using assms unfolding is_RB_upt_def by blast
lemma is_RB_uptD2:
assumes "is_RB_upt d canon G u" and "v \<prec>\<^sub>t u" and "d (pp_of_term v) \<le> dgrad_max d"
and "component_of_term v < length fs"
shows "is_RB_in d canon G v"
using assms unfolding is_RB_upt_def by blast
lemma is_RB_in_UnI:
assumes "is_RB_in d rword G u" and "\<And>h. h \<in> H \<Longrightarrow> u \<prec>\<^sub>t lt h"
shows "is_RB_in d rword (H \<union> G) u"
using assms(1)
proof (rule is_RB_inE)
assume "is_syz_sig d u"
thus "is_RB_in d rword (H \<union> G) u" by (rule is_RB_inI2)
next
fix g'
assume crw: "is_canon_rewriter rword G u g'"
and irred: "\<not> is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term u - lp g') g')"
from crw have "g' \<in> G" and "g' \<noteq> 0" and "lt g' adds\<^sub>t u"
and max: "\<And>a. a \<in> G \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> lt a adds\<^sub>t u \<Longrightarrow> rword (spp_of a) (spp_of g')"
by (rule is_canon_rewriterD)+
show "is_RB_in d rword (H \<union> G) u"
proof (rule is_RB_inI1)
show "is_canon_rewriter rword (H \<union> G) u g'"
proof (rule is_canon_rewriterI)
from \<open>g' \<in> G\<close> show "g' \<in> H \<union> G" by simp
next
fix a
assume "a \<in> H \<union> G" and "a \<noteq> 0" and "lt a adds\<^sub>t u"
from this(1) show "rword (spp_of a) (spp_of g')"
proof
assume "a \<in> H"
with \<open>lt a adds\<^sub>t u\<close> have "lt a adds\<^sub>t u" by simp
hence "lt a \<preceq>\<^sub>t u" by (rule ord_adds_term)
moreover from \<open>a \<in> H\<close> have "u \<prec>\<^sub>t lt a" by (rule assms(2))
ultimately show ?thesis by simp
next
assume "a \<in> G"
thus ?thesis using \<open>a \<noteq> 0\<close> \<open>lt a adds\<^sub>t u\<close> by (rule max)
qed
qed fact+
next
show "\<not> is_sig_red (\<prec>\<^sub>t) (=) (H \<union> G) (monom_mult 1 (pp_of_term u - lp g') g')"
(is "\<not> is_sig_red _ _ _ ?g")
proof
assume "is_sig_red (\<prec>\<^sub>t) (=) (H \<union> G) ?g"
with irred have "is_sig_red (\<prec>\<^sub>t) (=) H ?g" by (simp add: is_sig_red_Un del: Un_insert_left)
then obtain h where "h \<in> H" and "is_sig_red (\<prec>\<^sub>t) (=) {h} ?g" by (rule is_sig_red_singletonI)
from this(2) have "lt h \<prec>\<^sub>t lt ?g" by (rule is_sig_red_regularD_lt)
also from \<open>g' \<noteq> 0\<close> \<open>lt g' adds\<^sub>t u\<close> have "... = u"
by (auto simp: lt_monom_mult adds_term_alt pp_of_term_splus)
finally have "lt h \<prec>\<^sub>t u" .
moreover from \<open>h \<in> H\<close> have "u \<prec>\<^sub>t lt h" by (rule assms(2))
ultimately show False by simp
qed
qed
qed
corollary is_RB_in_insertI:
assumes "is_RB_in d rword G u" and "u \<prec>\<^sub>t lt g"
shows "is_RB_in d rword (insert g G) u"
proof -
from assms(1) have "is_RB_in d rword ({g} \<union> G) u"
proof (rule is_RB_in_UnI)
fix h
assume "h \<in> {g}"
with assms(2) show "u \<prec>\<^sub>t lt h" by simp
qed
thus ?thesis by simp
qed
corollary is_RB_upt_UnI:
assumes "is_RB_upt d rword G u" and "H \<subseteq> dgrad_sig_set d" and "\<And>h. h \<in> H \<Longrightarrow> u \<preceq>\<^sub>t lt h"
shows "is_RB_upt d rword (H \<union> G) u"
proof (rule is_RB_uptI)
from assms(1) have "G \<subseteq> dgrad_sig_set d" by (rule is_RB_uptD1)
with assms(2) show "H \<union> G \<subseteq> dgrad_sig_set d" by (rule Un_least)
next
fix v
assume "v \<prec>\<^sub>t u" and "d (pp_of_term v) \<le> dgrad_max d" and "component_of_term v < length fs"
with assms(1) have "is_RB_in d rword G v" by (rule is_RB_uptD2)
moreover from \<open>v \<prec>\<^sub>t u\<close> assms(3) have "\<And>h. h \<in> H \<Longrightarrow> v \<prec>\<^sub>t lt h" by (rule ord_term_lin.less_le_trans)
ultimately show "is_RB_in d rword (H \<union> G) v" by (rule is_RB_in_UnI)
qed
corollary is_RB_upt_insertI:
assumes "is_RB_upt d rword G u" and "g \<in> dgrad_sig_set d" and "u \<preceq>\<^sub>t lt g"
shows "is_RB_upt d rword (insert g G) u"
proof -
from assms(1) have "is_RB_upt d rword ({g} \<union> G) u"
proof (rule is_RB_upt_UnI)
from assms(2) show "{g} \<subseteq> dgrad_sig_set d" by simp
next
fix h
assume "h \<in> {g}"
with assms(3) show "u \<preceq>\<^sub>t lt h" by simp
qed
thus ?thesis by simp
qed
lemma is_RB_upt_is_sig_GB_upt:
assumes "dickson_grading d" and "is_RB_upt d rword G u"
shows "is_sig_GB_upt d G u"
proof (rule ccontr)
let ?Q = "{v. v \<prec>\<^sub>t u \<and> d (pp_of_term v) \<le> dgrad_max d \<and> component_of_term v < length fs \<and> \<not> is_sig_GB_in d G v}"
have Q_sub: "pp_of_term ` ?Q \<subseteq> dgrad_set d (dgrad_max d)" by blast
from assms(2) have G_sub: "G \<subseteq> dgrad_sig_set d" by (rule is_RB_uptD1)
hence "G \<subseteq> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
assume "\<not> is_sig_GB_upt d G u"
with G_sub obtain v' where "v' \<in> ?Q" unfolding is_sig_GB_upt_def by blast
with assms(1) obtain v where "v \<in> ?Q" and min: "\<And>y. y \<prec>\<^sub>t v \<Longrightarrow> y \<notin> ?Q" using Q_sub
by (rule ord_term_minimum_dgrad_set, blast)
from \<open>v \<in> ?Q\<close> have "v \<prec>\<^sub>t u" and "d (pp_of_term v) \<le> dgrad_max d" and "component_of_term v < length fs"
and "\<not> is_sig_GB_in d G v" by simp_all
from assms(2) this(1, 2, 3) have "is_RB_in d rword G v" by (rule is_RB_uptD2)
from \<open>\<not> is_sig_GB_in d G v\<close> obtain r where "lt r = v" and "r \<in> dgrad_sig_set d" and "\<not> sig_red_zero (\<preceq>\<^sub>t) G r"
unfolding is_sig_GB_in_def by blast
from this(3) have "rep_list r \<noteq> 0" by (auto simp: sig_red_zero_def)
hence "r \<noteq> 0" by (auto simp: rep_list_zero)
hence "lc r \<noteq> 0" by (rule lc_not_0)
from G_sub have "is_sig_GB_upt d G v"
proof (rule is_sig_GB_uptI)
fix w
assume dw: "d (pp_of_term w) \<le> dgrad_max d" and cp: "component_of_term w < length fs"
assume "w \<prec>\<^sub>t v"
hence "w \<notin> ?Q" by (rule min)
hence "\<not> w \<prec>\<^sub>t u \<or> is_sig_GB_in d G w" by (simp add: dw cp)
thus "is_sig_GB_in d G w"
proof
assume "\<not> w \<prec>\<^sub>t u"
moreover from \<open>w \<prec>\<^sub>t v\<close> \<open>v \<prec>\<^sub>t u\<close> have "w \<prec>\<^sub>t u" by (rule ord_term_lin.less_trans)
ultimately show ?thesis ..
qed
qed
from \<open>is_RB_in d rword G v\<close> have "sig_red_zero (\<preceq>\<^sub>t) G r"
proof (rule is_RB_inE)
assume "is_syz_sig d v"
have "sig_red_zero (\<prec>\<^sub>t) G r" by (rule syzygy_crit, fact+)
thus ?thesis by (rule sig_red_zero_sing_regI)
next
fix g
assume a: "\<not> is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term v - lp g) g)"
assume "is_canon_rewriter rword G v g"
hence "g \<in> G" and "g \<noteq> 0" and "lt g adds\<^sub>t v" by (rule is_canon_rewriterD)+
assume "\<not> is_syz_sig d v"
from \<open>g \<in> G\<close> G_sub have "g \<in> dgrad_sig_set d" ..
from \<open>g \<noteq> 0\<close> have "lc g \<noteq> 0" by (rule lc_not_0)
with \<open>lc r \<noteq> 0\<close> have "lc r / lc g \<noteq> 0" by simp
from \<open>lt g adds\<^sub>t v\<close> have "lt g adds\<^sub>t lt r" by (simp only: \<open>lt r = v\<close>)
hence eq1: "(lp r - lp g) \<oplus> lt g = lt r" by (metis add_implies_diff adds_termE pp_of_term_splus)
let ?h = "monom_mult (lc r / lc g) (lp r - lp g) g"
from \<open>lc g \<noteq> 0\<close> \<open>lc r \<noteq> 0\<close> \<open>g \<noteq> 0\<close> have "?h \<noteq> 0" by (simp add: monom_mult_eq_zero_iff)
have h_irred: "\<not> is_sig_red (\<prec>\<^sub>t) (=) G ?h"
proof
assume "is_sig_red (\<prec>\<^sub>t) (=) G ?h"
moreover from \<open>lc g \<noteq> 0\<close> \<open>lc r \<noteq> 0\<close> have "lc g / lc r \<noteq> 0" by simp
ultimately have "is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult (lc g / lc r) 0 ?h)" by (rule is_sig_red_monom_mult)
with \<open>lc g \<noteq> 0\<close> \<open>lc r \<noteq> 0\<close> have "is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term v - lp g) g)"
by (simp add: monom_mult_assoc \<open>lt r = v\<close>)
with a show False ..
qed
from \<open>lc r / lc g \<noteq> 0\<close> \<open>g \<noteq> 0\<close> have "lt ?h = lt r" by (simp add: lt_monom_mult eq1)
hence "lt ?h = v" by (simp only: \<open>lt r = v\<close>)
from \<open>lc g \<noteq> 0\<close> have "lc ?h = lc r" by simp
from assms(1) _ \<open>g \<in> dgrad_sig_set d\<close> have "?h \<in> dgrad_sig_set d"
proof (rule dgrad_sig_set_closed_monom_mult)
from \<open>lt g adds\<^sub>t lt r\<close> have "lp g adds lp r" by (simp add: adds_term_def)
with assms(1) have "d (lp r - lp g) \<le> d (lp r)" by (rule dickson_grading_minus)
also from \<open>r \<in> dgrad_sig_set d\<close> have "... \<le> dgrad_max d" by (rule dgrad_sig_setD_lp)
finally show "d (lp r - lp g) \<le> dgrad_max d" .
qed
have "rep_list ?h \<noteq> 0"
proof
assume "rep_list ?h = 0"
with \<open>?h \<noteq> 0\<close> \<open>lt ?h = v\<close> \<open>?h \<in> dgrad_sig_set d\<close> have "is_syz_sig d v" by (rule is_syz_sigI)
with \<open>\<not> is_syz_sig d v\<close> show False ..
qed
hence "rep_list g \<noteq> 0" by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
hence "punit.lc (rep_list g) \<noteq> 0" by (rule punit.lc_not_0)
from assms(1) \<open>G \<subseteq> dgrad_max_set d\<close> obtain s where s_red: "(sig_red (\<prec>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* r s"
and s_irred: "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) G s" by (rule sig_irredE_dgrad_max_set)
from s_red have s_red': "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* r s" by (rule sig_red_rtrancl_sing_regI)
have "rep_list s \<noteq> 0"
proof
assume "rep_list s = 0"
with s_red' have "sig_red_zero (\<preceq>\<^sub>t) G r" by (rule sig_red_zeroI)
with \<open>\<not> sig_red_zero (\<preceq>\<^sub>t) G r\<close> show False ..
qed
from assms(1) G_sub \<open>r \<in> dgrad_sig_set d\<close> s_red have "s \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_sig_red_rtrancl)
from s_red have "lt s = lt r" and "lc s = lc r"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
hence "lt ?h = lt s" and "lc ?h = lc s" and "s \<noteq> 0"
using \<open>lc r \<noteq> 0\<close> by (auto simp: \<open>lt ?h = lt r\<close> \<open>lc ?h = lc r\<close> simp del: lc_monom_mult)
from s_irred have "\<not> is_sig_red (\<prec>\<^sub>t) (=) G s" by (simp add: is_sig_red_top_tail_cases)
from \<open>is_sig_GB_upt d G v\<close> have "is_sig_GB_upt d G (lt s)" by (simp only: \<open>lt s = lt r\<close> \<open>lt r = v\<close>)
have "punit.lt (rep_list ?h) = punit.lt (rep_list s)"
by (rule sig_regular_top_reduced_lt_unique, fact+)
hence eq2: "lp r - lp g + punit.lt (rep_list g) = punit.lt (rep_list s)"
using \<open>lc r / lc g \<noteq> 0\<close> \<open>rep_list g \<noteq> 0\<close> by (simp add: rep_list_monom_mult punit.lt_monom_mult)
have "punit.lc (rep_list ?h) = punit.lc (rep_list s)"
by (rule sig_regular_top_reduced_lc_unique, fact+)
hence eq3: "lc r / lc g = punit.lc (rep_list s) / punit.lc (rep_list g)"
using \<open>punit.lc (rep_list g) \<noteq> 0\<close> by (simp add: rep_list_monom_mult field_simps)
have "sig_red_single (=) (=) s (s - ?h) g (lp r - lp g)"
by (rule sig_red_singleI, auto simp: eq1 eq2 eq3 punit.lc_def[symmetric] \<open>lt s = lt r\<close>
\<open>rep_list g \<noteq> 0\<close> \<open>rep_list s \<noteq> 0\<close> intro!: punit.lt_in_keys)
with \<open>g \<in> G\<close> have "sig_red (=) (=) G s (s - ?h)" unfolding sig_red_def by blast
hence "sig_red (\<preceq>\<^sub>t) (\<preceq>) G s (s - ?h)" by (auto dest: sig_red_sing_regI sig_red_top_tailI)
with s_red' have r_red: "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* r (s - ?h)" ..
show ?thesis
proof (cases "s - ?h = 0")
case True
hence "rep_list (s - ?h) = 0" by (simp only: rep_list_zero)
with r_red show ?thesis by (rule sig_red_zeroI)
next
case False
note \<open>is_sig_GB_upt d G (lt s)\<close>
moreover from \<open>s \<in> dgrad_sig_set d\<close> \<open>?h \<in> dgrad_sig_set d\<close> have "s - ?h \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_minus)
moreover from False \<open>lt ?h = lt s\<close> \<open>lc ?h = lc s\<close> have "lt (s - ?h) \<prec>\<^sub>t lt s" by (rule lt_minus_lessI)
ultimately have "sig_red_zero (\<preceq>\<^sub>t) G (s - ?h)" by (rule is_sig_GB_uptD3)
then obtain s' where "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* (s - ?h) s'" and "rep_list s' = 0"
by (rule sig_red_zeroE)
from r_red this(1) have "(sig_red (\<preceq>\<^sub>t) (\<preceq>) G)\<^sup>*\<^sup>* r s'" by simp
thus ?thesis using \<open>rep_list s' = 0\<close> by (rule sig_red_zeroI)
qed
qed
with \<open>\<not> sig_red_zero (\<preceq>\<^sub>t) G r\<close> show False ..
qed
corollary is_RB_upt_is_syz_sigD:
assumes "dickson_grading d" and "is_RB_upt d rword G u"
and "is_syz_sig d u" and "p \<in> dgrad_sig_set d" and "lt p = u"
shows "sig_red_zero (\<prec>\<^sub>t) G p"
proof -
note assms(1)
moreover from assms(1, 2) have "is_sig_GB_upt d G u" by (rule is_RB_upt_is_sig_GB_upt)
ultimately show ?thesis using assms(3, 4, 5) by (rule syzygy_crit)
qed
subsubsection \<open>S-Pairs\<close>
definition spair :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b)"
where "spair p q = (let t1 = punit.lt (rep_list p); t2 = punit.lt (rep_list q); l = lcs t1 t2 in
(monom_mult (1 / punit.lc (rep_list p)) (l - t1) p) -
(monom_mult (1 / punit.lc (rep_list q)) (l - t2) q))"
definition is_regular_spair :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "is_regular_spair p q \<longleftrightarrow>
(rep_list p \<noteq> 0 \<and> rep_list q \<noteq> 0 \<and>
(let t1 = punit.lt (rep_list p); t2 = punit.lt (rep_list q); l = lcs t1 t2 in
(l - t1) \<oplus> lt p \<noteq> (l - t2) \<oplus> lt q))"
lemma rep_list_spair: "rep_list (spair p q) = punit.spoly (rep_list p) (rep_list q)"
by (simp add: spair_def punit.spoly_def Let_def rep_list_minus rep_list_monom_mult punit.lc_def)
lemma spair_comm: "spair p q = - spair q p"
by (simp add: spair_def Let_def lcs_comm)
lemma dgrad_sig_set_closed_spair:
assumes "dickson_grading d" and "p \<in> dgrad_sig_set d" and "q \<in> dgrad_sig_set d"
shows "spair p q \<in> dgrad_sig_set d"
proof -
define t1 where "t1 = punit.lt (rep_list p)"
define t2 where "t2 = punit.lt (rep_list q)"
let ?l = "lcs t1 t2"
have "d t1 \<le> dgrad_max d"
proof (cases "rep_list p = 0")
case True
show ?thesis by (simp add: t1_def True dgrad_max_0)
next
case False
from assms(2) have "p \<in> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
with assms(1) have "rep_list p \<in> punit_dgrad_max_set d" by (rule dgrad_max_2)
thus ?thesis unfolding t1_def using False by (rule punit.dgrad_p_setD_lp[simplified])
qed
moreover have "d t2 \<le> dgrad_max d"
proof (cases "rep_list q = 0")
case True
show ?thesis by (simp add: t2_def True dgrad_max_0)
next
case False
from assms(3) have "q \<in> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
with assms(1) have "rep_list q \<in> punit_dgrad_max_set d" by (rule dgrad_max_2)
thus ?thesis unfolding t2_def using False by (rule punit.dgrad_p_setD_lp[simplified])
qed
ultimately have "ord_class.max (d t1) (d t2) \<le> dgrad_max d" by simp
moreover from assms(1) have "d ?l \<le> ord_class.max (d t1) (d t2)" by (rule dickson_grading_lcs)
ultimately have *: "d ?l \<le> dgrad_max d" by auto
thm dickson_grading_minus
show ?thesis
proof (simp add: spair_def Let_def t1_def[symmetric] t2_def[symmetric],
intro dgrad_sig_set_closed_minus dgrad_sig_set_closed_monom_mult[OF assms(1)])
from assms(1) adds_lcs have "d (?l - t1) \<le> d ?l" by (rule dickson_grading_minus)
thus "d (?l - t1) \<le> dgrad_max d" using * by (rule le_trans)
next
from assms(1) adds_lcs_2 have "d (?l - t2) \<le> d ?l" by (rule dickson_grading_minus)
thus "d (?l - t2) \<le> dgrad_max d" using * by (rule le_trans)
qed fact+
qed
lemma lt_spair:
assumes "rep_list p \<noteq> 0" and "punit.lt (rep_list p) \<oplus> lt q \<prec>\<^sub>t punit.lt (rep_list q) \<oplus> lt p"
shows "lt (spair p q) = (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p)) \<oplus> lt p"
proof -
define l where "l = lcs (punit.lt (rep_list p)) (punit.lt (rep_list q))"
have 1: "punit.lt (rep_list p) adds l" and 2: "punit.lt (rep_list q) adds l"
unfolding l_def by (rule adds_lcs, rule adds_lcs_2)
have eq1: "spair p q = (monom_mult (1 / punit.lc (rep_list p)) (l - punit.lt (rep_list p)) p) +
(monom_mult (- 1 / punit.lc (rep_list q)) (l - punit.lt (rep_list q)) q)"
by (simp add: spair_def Let_def l_def monom_mult_uminus_left)
from assms(1) have "punit.lc (rep_list p) \<noteq> 0" by (rule punit.lc_not_0)
hence "1 / punit.lc (rep_list p) \<noteq> 0" by simp
moreover from assms(1) have "p \<noteq> 0" by (auto simp: rep_list_zero)
ultimately have eq2: "lt (monom_mult (1 / punit.lc (rep_list p)) (l - punit.lt (rep_list p)) p) =
(l - punit.lt (rep_list p)) \<oplus> lt p"
by (rule lt_monom_mult)
have "lt (monom_mult (- 1 / punit.lc (rep_list q)) (l - punit.lt (rep_list q)) q) \<preceq>\<^sub>t
(l - punit.lt (rep_list q)) \<oplus> lt q"
by (fact lt_monom_mult_le)
also from assms(2) have "... \<prec>\<^sub>t (l - punit.lt (rep_list p)) \<oplus> lt p"
by (simp add: term_is_le_rel_minus_minus[OF _ 2 1])
finally show ?thesis unfolding eq2[symmetric] eq1 l_def[symmetric] by (rule lt_plus_eqI_2)
qed
lemma lt_spair':
assumes "rep_list p \<noteq> 0" and "a + punit.lt (rep_list p) = b + punit.lt (rep_list q)" and "b \<oplus> lt q \<prec>\<^sub>t a \<oplus> lt p"
shows "lt (spair p q) = (a - gcs a b) \<oplus> lt p"
proof -
from assms(3) have "punit.lt (rep_list p) \<oplus> (b \<oplus> lt q) \<prec>\<^sub>t punit.lt (rep_list p) \<oplus> (a \<oplus> lt p)"
by (fact splus_mono_strict)
hence "(b + punit.lt (rep_list p)) \<oplus> lt q \<prec>\<^sub>t (b + punit.lt (rep_list q)) \<oplus> lt p"
by (simp only: splus_assoc[symmetric] add.commute assms(2))
hence "punit.lt (rep_list p) \<oplus> lt q \<prec>\<^sub>t punit.lt (rep_list q) \<oplus> lt p"
by (simp only: splus_assoc ord_term_strict_canc)
with assms(1)
have "lt (spair p q) = (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p)) \<oplus> lt p"
by (rule lt_spair)
with assms(2) show ?thesis by (simp add: lcs_minus_1)
qed
lemma lt_rep_list_spair:
assumes "rep_list p \<noteq> 0" and "rep_list q \<noteq> 0" and "rep_list (spair p q) \<noteq> 0"
and "a + punit.lt (rep_list p) = b + punit.lt (rep_list q)"
shows "punit.lt (rep_list (spair p q)) \<prec> (a - gcs a b) + punit.lt (rep_list p)"
proof -
from assms(1) have 1: "punit.lc (rep_list p) \<noteq> 0" by (rule punit.lc_not_0)
from assms(2) have 2: "punit.lc (rep_list q) \<noteq> 0" by (rule punit.lc_not_0)
define l where "l = lcs (punit.lt (rep_list p)) (punit.lt (rep_list q))"
have eq: "rep_list (spair p q) = (punit.monom_mult (1 / punit.lc (rep_list p)) (l - punit.lt (rep_list p)) (rep_list p)) +
(punit.monom_mult (- 1 / punit.lc (rep_list q)) (l - punit.lt (rep_list q)) (rep_list q))"
(is "_ = ?a + ?b")
by (simp add: spair_def Let_def rep_list_minus rep_list_monom_mult punit.monom_mult_uminus_left l_def)
have "?a + ?b \<noteq> 0" unfolding eq[symmetric] by (fact assms(3))
moreover from 1 2 assms(1, 2) have "punit.lt ?b = punit.lt ?a"
by (simp add: lp_monom_mult l_def minus_plus adds_lcs adds_lcs_2)
moreover have "punit.lc ?b = - punit.lc ?a" by (simp add: 1 2)
ultimately have "punit.lt (rep_list (spair p q)) \<prec> punit.lt ?a" unfolding eq by (rule punit.lt_plus_lessI)
also from 1 assms(1) have "... = (l - punit.lt (rep_list p)) + punit.lt (rep_list p)" by (simp add: lp_monom_mult)
also have "... = l" by (simp add: l_def minus_plus adds_lcs)
also have "... = (a + punit.lt (rep_list p)) - gcs a b" unfolding l_def using assms(4) by (rule lcs_alt_1)
also have "... = (a - gcs a b) + punit.lt (rep_list p)" by (simp add: minus_plus gcs_adds)
finally show ?thesis .
qed
lemma is_regular_spair_sym: "is_regular_spair p q \<Longrightarrow> is_regular_spair q p"
by (auto simp: is_regular_spair_def Let_def lcs_comm)
lemma is_regular_spairI:
assumes "rep_list p \<noteq> 0" and "rep_list q \<noteq> 0"
and "punit.lt (rep_list q) \<oplus> lt p \<noteq> punit.lt (rep_list p) \<oplus> lt q"
shows "is_regular_spair p q"
proof -
have *: "(lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p)) \<oplus> lt p \<noteq>
(lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list q)) \<oplus> lt q"
(is "?l \<noteq> ?r")
proof
assume "?l = ?r"
hence "punit.lt (rep_list q) \<oplus> lt p = punit.lt (rep_list p) \<oplus> lt q"
by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
with assms(3) show False ..
qed
with assms(1, 2) show ?thesis by (simp add: is_regular_spair_def)
qed
lemma is_regular_spairI':
assumes "rep_list p \<noteq> 0" and "rep_list q \<noteq> 0"
and "a + punit.lt (rep_list p) = b + punit.lt (rep_list q)" and "a \<oplus> lt p \<noteq> b \<oplus> lt q"
shows "is_regular_spair p q"
proof -
have "punit.lt (rep_list q) \<oplus> lt p \<noteq> punit.lt (rep_list p) \<oplus> lt q"
proof
assume "punit.lt (rep_list q) \<oplus> lt p = punit.lt (rep_list p) \<oplus> lt q"
hence "a \<oplus> (punit.lt (rep_list q) \<oplus> lt p) = a \<oplus> (punit.lt (rep_list p) \<oplus> lt q)" by (simp only:)
hence "(a + punit.lt (rep_list q)) \<oplus> lt p = (b + punit.lt (rep_list q)) \<oplus> lt q"
by (simp add: splus_assoc[symmetric] assms(3))
hence "punit.lt (rep_list q) \<oplus> (a \<oplus> lt p) = punit.lt (rep_list q) \<oplus> (b \<oplus> lt q)"
by (simp only: add.commute[of _ "punit.lt (rep_list q)"] splus_assoc)
hence "a \<oplus> lt p = b \<oplus> lt q" by (simp only: splus_left_canc)
with assms(4) show False ..
qed
with assms(1, 2) show ?thesis by (rule is_regular_spairI)
qed
lemma is_regular_spairD1: "is_regular_spair p q \<Longrightarrow> rep_list p \<noteq> 0"
by (simp add: is_regular_spair_def)
lemma is_regular_spairD2: "is_regular_spair p q \<Longrightarrow> rep_list q \<noteq> 0"
by (simp add: is_regular_spair_def)
lemma is_regular_spairD3:
fixes p q
defines "t1 \<equiv> punit.lt (rep_list p)"
defines "t2 \<equiv> punit.lt (rep_list q)"
assumes "is_regular_spair p q"
shows "t2 \<oplus> lt p \<noteq> t1 \<oplus> lt q" (is ?thesis1)
and "lt (monom_mult (1 / punit.lc (rep_list p)) (lcs t1 t2 - t1) p) \<noteq>
lt (monom_mult (1 / punit.lc (rep_list q)) (lcs t1 t2 - t2) q)" (is "?l \<noteq> ?r")
proof -
from assms(3) have "rep_list p \<noteq> 0" by (rule is_regular_spairD1)
hence "punit.lc (rep_list p) \<noteq> 0" and "p \<noteq> 0" by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
from assms(3) have "rep_list q \<noteq> 0" by (rule is_regular_spairD2)
hence "punit.lc (rep_list q) \<noteq> 0" and "q \<noteq> 0" by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
have "?l = (lcs t1 t2 - t1) \<oplus> lt p"
using \<open>punit.lc (rep_list p) \<noteq> 0\<close> \<open>p \<noteq> 0\<close> by (simp add: lt_monom_mult)
also from assms(3) have *: "... \<noteq> (lcs t1 t2 - t2) \<oplus> lt q"
by (simp add: is_regular_spair_def t1_def t2_def Let_def)
also have "(lcs t1 t2 - t2) \<oplus> lt q = ?r"
using \<open>punit.lc (rep_list q) \<noteq> 0\<close> \<open>q \<noteq> 0\<close> by (simp add: lt_monom_mult)
finally show "?l \<noteq> ?r" .
show ?thesis1
proof
assume "t2 \<oplus> lt p = t1 \<oplus> lt q"
hence "(lcs t1 t2 - t1) \<oplus> lt p = (lcs t1 t2 - t2) \<oplus> lt q"
by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
with * show False ..
qed
qed
lemma is_regular_spair_nonzero: "is_regular_spair p q \<Longrightarrow> spair p q \<noteq> 0"
by (auto simp: spair_def Let_def dest: is_regular_spairD3)
lemma is_regular_spair_lt:
assumes "is_regular_spair p q"
shows "lt (spair p q) = ord_term_lin.max
((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p)) \<oplus> lt p)
((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list q)) \<oplus> lt q)"
proof -
let ?t1 = "punit.lt (rep_list p)"
let ?t2 = "punit.lt (rep_list q)"
let ?l = "lcs ?t1 ?t2"
show ?thesis
proof (rule ord_term_lin.linorder_cases)
assume a: "?t2 \<oplus> lt p \<prec>\<^sub>t ?t1 \<oplus> lt q"
hence "(?l - ?t1) \<oplus> lt p \<prec>\<^sub>t (?l - ?t2) \<oplus> lt q"
by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
hence le: "(?l - ?t1) \<oplus> lt p \<preceq>\<^sub>t (?l - ?t2) \<oplus> lt q" by (rule ord_term_lin.less_imp_le)
from assms have "rep_list q \<noteq> 0" by (rule is_regular_spairD2)
have "lt (spair p q) = lt (spair q p)" by (simp add: spair_comm[of p])
also from \<open>rep_list q \<noteq> 0\<close> a have "... = (lcs ?t2 ?t1 - ?t2) \<oplus> lt q" by (rule lt_spair)
also have "... = (?l - ?t2) \<oplus> lt q" by (simp only: lcs_comm)
finally show ?thesis using le by (simp add: ord_term_lin.max_def)
next
assume a: "?t1 \<oplus> lt q \<prec>\<^sub>t ?t2 \<oplus> lt p"
hence "(?l - ?t2) \<oplus> lt q \<prec>\<^sub>t (?l - ?t1) \<oplus> lt p"
by (simp add: term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
hence le: "\<not> ((?l - ?t1) \<oplus> lt p \<preceq>\<^sub>t (?l - ?t2) \<oplus> lt q)" by simp
from assms have "rep_list p \<noteq> 0" by (rule is_regular_spairD1)
hence "lt (spair p q) = (lcs ?t1 ?t2 - ?t1) \<oplus> lt p" using a by (rule lt_spair)
thus ?thesis using le by (simp add: ord_term_lin.max_def)
next
from assms have "?t2 \<oplus> lt p \<noteq> ?t1 \<oplus> lt q" by (rule is_regular_spairD3)
moreover assume "?t2 \<oplus> lt p = ?t1 \<oplus> lt q"
ultimately show ?thesis ..
qed
qed
lemma is_regular_spair_lt_ge_1:
assumes "is_regular_spair p q"
shows "lt p \<preceq>\<^sub>t lt (spair p q)"
proof -
have "lt p = 0 \<oplus> lt p" by (simp only: term_simps)
also from zero_min have "... \<preceq>\<^sub>t (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p)) \<oplus> lt p"
by (rule splus_mono_left)
also have "... \<preceq>\<^sub>t ord_term_lin.max
((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p)) \<oplus> lt p)
((lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list q)) \<oplus> lt q)"
by (rule ord_term_lin.max.cobounded1)
also from assms have "... = lt (spair p q)" by (simp only: is_regular_spair_lt)
finally show ?thesis .
qed
corollary is_regular_spair_lt_ge_2:
assumes "is_regular_spair p q"
shows "lt q \<preceq>\<^sub>t lt (spair p q)"
proof -
from assms have "is_regular_spair q p" by (rule is_regular_spair_sym)
hence "lt q \<preceq>\<^sub>t lt (spair q p)" by (rule is_regular_spair_lt_ge_1)
also have "... = lt (spair p q)" by (simp add: spair_comm[of q])
finally show ?thesis .
qed
lemma is_regular_spair_component_lt_cases:
assumes "is_regular_spair p q"
shows "component_of_term (lt (spair p q)) = component_of_term (lt p) \<or>
component_of_term (lt (spair p q)) = component_of_term (lt q)"
proof (rule ord_term_lin.linorder_cases)
from assms have "rep_list q \<noteq> 0" by (rule is_regular_spairD2)
moreover assume "punit.lt (rep_list q) \<oplus> lt p \<prec>\<^sub>t punit.lt (rep_list p) \<oplus> lt q"
ultimately have "lt (spair q p) = (lcs (punit.lt (rep_list q)) (punit.lt (rep_list p)) - punit.lt (rep_list q)) \<oplus> lt q"
by (rule lt_spair)
thus ?thesis by (simp add: spair_comm[of p] term_simps)
next
from assms have "rep_list p \<noteq> 0" by (rule is_regular_spairD1)
moreover assume "punit.lt (rep_list p) \<oplus> lt q \<prec>\<^sub>t punit.lt (rep_list q) \<oplus> lt p"
ultimately have "lt (spair p q) = (lcs (punit.lt (rep_list p)) (punit.lt (rep_list q)) - punit.lt (rep_list p)) \<oplus> lt p"
by (rule lt_spair)
thus ?thesis by (simp add: term_simps)
next
from assms have "punit.lt (rep_list q) \<oplus> lt p \<noteq> punit.lt (rep_list p) \<oplus> lt q"
by (rule is_regular_spairD3)
moreover assume "punit.lt (rep_list q) \<oplus> lt p = punit.lt (rep_list p) \<oplus> lt q"
ultimately show ?thesis ..
qed
lemma lemma_9:
assumes "dickson_grading d" and "is_rewrite_ord rword" and "is_RB_upt d rword G u"
and "inj_on lt G" and "\<not> is_syz_sig d u" and "is_canon_rewriter rword G u g1" and "h \<in> G"
and "is_sig_red (\<prec>\<^sub>t) (=) {h} (monom_mult 1 (pp_of_term u - lp g1) g1)"
and "d (pp_of_term u) \<le> dgrad_max d"
shows "lcs (punit.lt (rep_list g1)) (punit.lt (rep_list h)) - punit.lt (rep_list g1) =
pp_of_term u - lp g1" (is ?thesis1)
and "lcs (punit.lt (rep_list g1)) (punit.lt (rep_list h)) - punit.lt (rep_list h) =
((pp_of_term u - lp g1) + punit.lt (rep_list g1)) - punit.lt (rep_list h)" (is ?thesis2)
and "is_regular_spair g1 h" (is ?thesis3)
and "lt (spair g1 h) = u" (is ?thesis4)
proof -
from assms(8) have "rep_list (monom_mult 1 (pp_of_term u - lp g1) g1) \<noteq> 0"
using is_sig_red_top_addsE by fastforce
hence "rep_list g1 \<noteq> 0" by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
hence "g1 \<noteq> 0" by (auto simp: rep_list_zero)
from assms(6) have "g1 \<in> G" and "lt g1 adds\<^sub>t u" by (rule is_canon_rewriterD)+
from assms(3) have "G \<subseteq> dgrad_sig_set d" by (rule is_RB_uptD1)
with \<open>g1 \<in> G\<close> have "g1 \<in> dgrad_sig_set d" ..
hence "component_of_term (lt g1) < length fs" using \<open>g1 \<noteq> 0\<close> by (rule dgrad_sig_setD_lt)
with \<open>lt g1 adds\<^sub>t u\<close> have "component_of_term u < length fs" by (simp add: adds_term_def)
from \<open>lt g1 adds\<^sub>t u\<close> obtain a where u: "u = a \<oplus> lt g1" by (rule adds_termE)
hence a: "a = pp_of_term u - lp g1" by (simp add: term_simps)
from assms(8) have "is_sig_red (\<prec>\<^sub>t) (=) {h} (monom_mult 1 a g1)" by (simp only: a)
hence "rep_list h \<noteq> 0" and "rep_list (monom_mult 1 a g1) \<noteq> 0" and
2: "punit.lt (rep_list h) adds punit.lt (rep_list (monom_mult 1 a g1))" and
3: "punit.lt (rep_list (monom_mult 1 a g1)) \<oplus> lt h \<prec>\<^sub>t punit.lt (rep_list h) \<oplus> lt (monom_mult 1 a g1)"
by (auto elim: is_sig_red_top_addsE)
from this(2) have "rep_list g1 \<noteq> 0" by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
hence "g1 \<noteq> 0" by (auto simp: rep_list_zero)
from \<open>rep_list h \<noteq> 0\<close> have "h \<noteq> 0" by (auto simp: rep_list_zero)
from 2 \<open>rep_list g1 \<noteq> 0\<close> have "punit.lt (rep_list h) adds a + punit.lt (rep_list g1)"
by (simp add: rep_list_monom_mult punit.lt_monom_mult)
then obtain b where eq1: "a + punit.lt (rep_list g1) = b + punit.lt (rep_list h)"
by (auto elim: addsE simp: add.commute)
hence b: "b = ((pp_of_term u - lp g1) + punit.lt (rep_list g1)) - punit.lt (rep_list h)"
by (simp add: a)
define g where "g = gcs a b"
have "g = 0"
proof (rule ccontr)
assume "g \<noteq> 0"
have "g adds a" unfolding g_def by (fact gcs_adds)
also have "... adds\<^sub>p u" unfolding u by (fact adds_pp_triv)
finally obtain v where u2: "u = g \<oplus> v" by (rule adds_ppE)
hence v: "v = u \<ominus> g" by (simp add: term_simps)
from u2 have "v adds\<^sub>t u" by (rule adds_termI)
hence "v \<preceq>\<^sub>t u" by (rule ord_adds_term)
moreover have "v \<noteq> u"
proof
assume "v = u"
hence "g \<oplus> v = 0 \<oplus> v" by (simp add: u2 term_simps)
hence "g = 0" by (simp only: splus_right_canc)
with \<open>g \<noteq> 0\<close> show False ..
qed
ultimately have "v \<prec>\<^sub>t u" by simp
note assms(3) \<open>v \<prec>\<^sub>t u\<close>
moreover have "d (pp_of_term v) \<le> dgrad_max d"
proof (rule le_trans)
from assms(1) show "d (pp_of_term v) \<le> d (pp_of_term u)"
by (simp add: u2 term_simps dickson_gradingD1)
qed fact
moreover from \<open>component_of_term u < length fs\<close> have "component_of_term v < length fs"
by (simp only: v term_simps)
ultimately have "is_RB_in d rword G v" by (rule is_RB_uptD2)
thus False
proof (rule is_RB_inE)
assume "is_syz_sig d v"
with assms(1) have "is_syz_sig d u" using \<open>v adds\<^sub>t u\<close> assms(9) by (rule is_syz_sig_adds)
with assms(5) show False ..
next
fix g2
assume *: "\<not> is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term v - lp g2) g2)"
assume "is_canon_rewriter rword G v g2"
hence "g2 \<in> G" and "g2 \<noteq> 0" and "lt g2 adds\<^sub>t v" by (rule is_canon_rewriterD)+
assume "\<not> is_syz_sig d v"
note assms(2) \<open>is_canon_rewriter rword G v g2\<close> assms(6)
moreover from \<open>lt g2 adds\<^sub>t v\<close> \<open>v adds\<^sub>t u\<close> have "lt g2 adds\<^sub>t u" by (rule adds_term_trans)
moreover from \<open>g adds a\<close> have "lt g1 adds\<^sub>t v" by (simp add: v u minus_splus[symmetric] adds_termI)
ultimately have "lt g2 = lt g1" by (rule is_rewrite_ord_canon_rewriterD1)
with assms(4) have "g2 = g1" using \<open>g2 \<in> G\<close> \<open>g1 \<in> G\<close> by (rule inj_onD)
have "pp_of_term v - lp g1 = a - g" by (simp add: u v term_simps diff_diff_add)
have "is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term v - lp g2) g2)"
unfolding \<open>g2 = g1\<close> \<open>pp_of_term v - lp g1 = a - g\<close> using assms(7) \<open>rep_list h \<noteq> 0\<close>
proof (rule is_sig_red_top_addsI)
from \<open>rep_list g1 \<noteq> 0\<close> show "rep_list (monom_mult 1 (a - g) g1) \<noteq> 0"
by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
next
have eq3: "(a - g) + punit.lt (rep_list g1) = lcs (punit.lt (rep_list g1)) (punit.lt (rep_list h))"
by (simp add: g_def lcs_minus_1[OF eq1, symmetric] adds_minus adds_lcs)
from \<open>rep_list g1 \<noteq> 0\<close>
show "punit.lt (rep_list h) adds punit.lt (rep_list (monom_mult 1 (a - g) g1))"
by (simp add: rep_list_monom_mult punit.lt_monom_mult eq3 adds_lcs_2)
next
from 3 \<open>rep_list g1 \<noteq> 0\<close> \<open>g1 \<noteq> 0\<close>
show "punit.lt (rep_list (monom_mult 1 (a - g) g1)) \<oplus> lt h \<prec>\<^sub>t
punit.lt (rep_list h) \<oplus> lt (monom_mult 1 (a - g) g1)"
by (auto simp: rep_list_monom_mult punit.lt_monom_mult lt_monom_mult splus_assoc splus_left_commute
dest!: ord_term_strict_canc intro: splus_mono_strict)
next
show "ord_term_lin.is_le_rel (\<prec>\<^sub>t)" by (fact ord_term_lin.is_le_relI)
qed
with * show False ..
qed
qed
thus ?thesis1 and ?thesis2 by (simp_all add: a b lcs_minus_1[OF eq1] lcs_minus_2[OF eq1] g_def)
hence eq3: "spair g1 h = monom_mult (1 / punit.lc (rep_list g1)) a g1 -
monom_mult (1 / punit.lc (rep_list h)) b h"
by (simp add: spair_def Let_def a b)
from 3 \<open>rep_list g1 \<noteq> 0\<close> \<open>g1 \<noteq> 0\<close> have "b \<oplus> lt h \<prec>\<^sub>t a \<oplus> lt g1"
by (auto simp: rep_list_monom_mult punit.lt_monom_mult lt_monom_mult eq1 splus_assoc
splus_left_commute[of b] dest!: ord_term_strict_canc)
hence "a \<oplus> lt g1 \<noteq> b \<oplus> lt h" by simp
with \<open>rep_list g1 \<noteq> 0\<close> \<open>rep_list h \<noteq> 0\<close> eq1 show ?thesis3
by (rule is_regular_spairI')
have "lt (monom_mult (1 / punit.lc (rep_list h)) b h) = b \<oplus> lt h"
proof (rule lt_monom_mult)
from \<open>rep_list h \<noteq> 0\<close> show "1 / punit.lc (rep_list h) \<noteq> 0" by (simp add: punit.lc_eq_zero_iff)
qed fact
also have "... \<prec>\<^sub>t a \<oplus> lt g1" by fact
also have "... = lt (monom_mult (1 / punit.lc (rep_list g1)) a g1)"
proof (rule HOL.sym, rule lt_monom_mult)
from \<open>rep_list g1 \<noteq> 0\<close> show "1 / punit.lc (rep_list g1) \<noteq> 0" by (simp add: punit.lc_eq_zero_iff)
qed fact
finally have "lt (spair g1 h) = lt (monom_mult (1 / punit.lc (rep_list g1)) a g1)"
unfolding eq3 by (rule lt_minus_eqI_2)
also have "... = a \<oplus> lt g1" by (rule HOL.sym, fact)
finally show ?thesis4 by (simp only: u)
qed
lemma is_RB_upt_finite:
assumes "dickson_grading d" and "is_rewrite_ord rword" and "G \<subseteq> dgrad_sig_set d" and "inj_on lt G"
and "finite G"
and "\<And>g1 g2. g1 \<in> G \<Longrightarrow> g2 \<in> G \<Longrightarrow> is_regular_spair g1 g2 \<Longrightarrow> lt (spair g1 g2) \<prec>\<^sub>t u \<Longrightarrow>
is_RB_in d rword G (lt (spair g1 g2))"
and "\<And>i. i < length fs \<Longrightarrow> term_of_pair (0, i) \<prec>\<^sub>t u \<Longrightarrow> is_RB_in d rword G (term_of_pair (0, i))"
shows "is_RB_upt d rword G u"
proof (rule ccontr)
let ?Q = "{v. v \<prec>\<^sub>t u \<and> d (pp_of_term v) \<le> dgrad_max d \<and> component_of_term v < length fs \<and> \<not> is_RB_in d rword G v}"
have Q_sub: "pp_of_term ` ?Q \<subseteq> dgrad_set d (dgrad_max d)" by blast
from assms(3) have "G \<subseteq> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
assume "\<not> is_RB_upt d rword G u"
with assms(3) obtain v' where "v' \<in> ?Q" unfolding is_RB_upt_def by blast
with assms(1) obtain v where "v \<in> ?Q" and min: "\<And>y. y \<prec>\<^sub>t v \<Longrightarrow> y \<notin> ?Q" using Q_sub
by (rule ord_term_minimum_dgrad_set, blast)
from \<open>v \<in> ?Q\<close> have "v \<prec>\<^sub>t u" and "d (pp_of_term v) \<le> dgrad_max d" and "component_of_term v < length fs"
and "\<not> is_RB_in d rword G v" by simp_all
from this(4)
have impl: "\<And>g. g \<in> G \<Longrightarrow> is_canon_rewriter rword G v g \<Longrightarrow>
is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term v - lp g) g)"
and "\<not> is_syz_sig d v" by (simp_all add: is_RB_in_def Let_def)
from assms(3) have "is_RB_upt d rword G v"
proof (rule is_RB_uptI)
fix w
assume dw: "d (pp_of_term w) \<le> dgrad_max d" and cp: "component_of_term w < length fs"
assume "w \<prec>\<^sub>t v"
hence "w \<notin> ?Q" by (rule min)
hence "\<not> w \<prec>\<^sub>t u \<or> is_RB_in d rword G w" by (simp add: dw cp)
thus "is_RB_in d rword G w"
proof
assume "\<not> w \<prec>\<^sub>t u"
moreover from \<open>w \<prec>\<^sub>t v\<close> \<open>v \<prec>\<^sub>t u\<close> have "w \<prec>\<^sub>t u" by (rule ord_term_lin.less_trans)
ultimately show ?thesis ..
qed
qed
show False
proof (cases "\<exists>g\<in>G. g \<noteq> 0 \<and> lt g adds\<^sub>t v")
case False
hence x: "\<And>g. g \<in> G \<Longrightarrow> lt g adds\<^sub>t v \<Longrightarrow> g = 0" by blast
let ?w = "term_of_pair (0, component_of_term v)"
have "?w adds\<^sub>t v" by (simp add: adds_term_def term_simps)
hence "?w \<preceq>\<^sub>t v" by (rule ord_adds_term)
also have "... \<prec>\<^sub>t u" by fact
finally have "?w \<prec>\<^sub>t u" .
with \<open>component_of_term v < length fs\<close> have "is_RB_in d rword G ?w" by (rule assms(7))
thus ?thesis
proof (rule is_RB_inE)
assume "is_syz_sig d ?w"
with assms(1) have "is_syz_sig d v" using \<open>?w adds\<^sub>t v\<close> \<open>d (pp_of_term v) \<le> dgrad_max d\<close>
by (rule is_syz_sig_adds)
with \<open>\<not> is_syz_sig d v\<close> show ?thesis ..
next
fix g1
assume "is_canon_rewriter rword G ?w g1"
hence "g1 \<noteq> 0" and "g1 \<in> G" and "lt g1 adds\<^sub>t ?w" by (rule is_canon_rewriterD)+
from this(3) have "lt g1 adds\<^sub>t v" using \<open>?w adds\<^sub>t v\<close> by (rule adds_term_trans)
with \<open>g1 \<in> G\<close> have "g1 = 0" by (rule x)
with \<open>g1 \<noteq> 0\<close> show ?thesis ..
qed
next
case True
then obtain g' where "g' \<in> G" and "g' \<noteq> 0" and "lt g' adds\<^sub>t v" by blast
with assms(2, 5) obtain g1 where crw: "is_canon_rewriter rword G v g1"
by (rule is_rewrite_ord_finite_canon_rewriterE)
hence "g1 \<in> G" by (rule is_canon_rewriterD1)
hence "is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term v - lp g1) g1)" using crw by (rule impl)
then obtain h where "h \<in> G" and "is_sig_red (\<prec>\<^sub>t) (=) {h} (monom_mult 1 (pp_of_term v - lp g1) g1)"
by (rule is_sig_red_singletonI)
with assms(1, 2) \<open>is_RB_upt d rword G v\<close> assms(4) \<open>\<not> is_syz_sig d v\<close> crw
have "is_regular_spair g1 h" and eq: "lt (spair g1 h) = v"
using \<open>d (pp_of_term v) \<le> dgrad_max d\<close> by (rule lemma_9)+
from \<open>v \<prec>\<^sub>t u\<close> have "lt (spair g1 h) \<prec>\<^sub>t u" by (simp only: eq)
with \<open>g1 \<in> G\<close> \<open>h \<in> G\<close> \<open>is_regular_spair g1 h\<close> have "is_RB_in d rword G (lt (spair g1 h))"
by (rule assms(6))
hence "is_RB_in d rword G v" by (simp only: eq)
with \<open>\<not> is_RB_in d rword G v\<close> show ?thesis ..
qed
qed
text \<open>Note that the following lemma actually holds for @{emph \<open>all\<close>} regularly reducible power-products
in @{term "rep_list p"}, not just for the leading power-product.\<close>
lemma lemma_11:
assumes "dickson_grading d" and "is_rewrite_ord rword" and "is_RB_upt d rword G (lt p)"
and "p \<in> dgrad_sig_set d" and "is_sig_red (\<prec>\<^sub>t) (=) G p"
obtains u g where "u \<prec>\<^sub>t lt p" and "d (pp_of_term u) \<le> dgrad_max d" and "component_of_term u < length fs"
and "\<not> is_syz_sig d u" and "is_canon_rewriter rword G u g"
and "u = (punit.lt (rep_list p) - punit.lt (rep_list g)) \<oplus> lt g" and "is_sig_red (\<prec>\<^sub>t) (=) {g} p"
proof -
from assms(3) have G_sub: "G \<subseteq> dgrad_sig_set d" by (rule is_RB_uptD1)
from assms(5) have "rep_list p \<noteq> 0" using is_sig_red_addsE by fastforce
hence "p \<noteq> 0" by (auto simp: rep_list_zero)
let ?lc = "punit.lc (rep_list p)"
let ?lp = "punit.lt (rep_list p)"
from \<open>rep_list p \<noteq> 0\<close> have "?lc \<noteq> 0" by (rule punit.lc_not_0)
from assms(4) have "p \<in> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
from assms(4) have "d (lp p) \<le> dgrad_max d" by (rule dgrad_sig_setD_lp)
from assms(4) \<open>p \<noteq> 0\<close> have "component_of_term (lt p) < length fs" by (rule dgrad_sig_setD_lt)
from assms(1) \<open>p \<in> dgrad_max_set d\<close> have "rep_list p \<in> punit_dgrad_max_set d" by (rule dgrad_max_2)
hence "d ?lp \<le> dgrad_max d" using \<open>rep_list p \<noteq> 0\<close> by (rule punit.dgrad_p_setD_lp[simplified])
from assms(5) obtain g0 where "g0 \<in> G" and "is_sig_red (\<prec>\<^sub>t) (=) {g0} p"
by (rule is_sig_red_singletonI)
from \<open>g0 \<in> G\<close> G_sub have "g0 \<in> dgrad_sig_set d" ..
let ?g0 = "monom_mult (?lc / punit.lc (rep_list g0)) (?lp - punit.lt (rep_list g0)) g0"
define M where "M = {monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g |
g. g \<in> dgrad_sig_set d \<and> is_sig_red (\<prec>\<^sub>t) (=) {g} p}"
from \<open>g0 \<in> dgrad_sig_set d\<close> \<open>is_sig_red (\<prec>\<^sub>t) (=) {g0} p\<close> have "?g0 \<in> M" by (auto simp: M_def)
have "0 \<notin> rep_list ` M"
proof
assume "0 \<in> rep_list ` M"
then obtain g where 1: "is_sig_red (\<prec>\<^sub>t) (=) {g} p"
and 2: "rep_list (monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g) = 0"
unfolding M_def by fastforce
from 1 have "rep_list g \<noteq> 0" using is_sig_red_addsE by fastforce
moreover from this have "punit.lc (rep_list g) \<noteq> 0" by (rule punit.lc_not_0)
ultimately have "rep_list (monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g) \<noteq> 0"
using \<open>?lc \<noteq> 0\<close> by (simp add: rep_list_monom_mult punit.monom_mult_eq_zero_iff)
thus False using 2 ..
qed
with rep_list_zero have "0 \<notin> M" by auto
have "M \<subseteq> dgrad_sig_set d"
proof
fix m
assume "m \<in> M"
then obtain g where "g \<in> dgrad_sig_set d" and 1: "is_sig_red (\<prec>\<^sub>t) (=) {g} p"
and m: "m = monom_mult (?lc / punit.lc (rep_list g)) (?lp - punit.lt (rep_list g)) g"
unfolding M_def by fastforce
from 1 have "punit.lt (rep_list g) adds ?lp" using is_sig_red_top_addsE by fastforce
note assms(1)
thm dickson_grading_minus
moreover have "d (?lp - punit.lt (rep_list g)) \<le> dgrad_max d"
by (rule le_trans, rule dickson_grading_minus, fact+)
ultimately show "m \<in> dgrad_sig_set d" unfolding m using \<open>g \<in> dgrad_sig_set d\<close>
by (rule dgrad_sig_set_closed_monom_mult)
qed
hence "M \<subseteq> sig_inv_set" by (simp add: dgrad_sig_set'_def)
let ?M = "lt ` M"
note assms(1)
moreover from \<open>?g0 \<in> M\<close> have "lt ?g0 \<in> ?M" by (rule imageI)
moreover from \<open>M \<subseteq> dgrad_sig_set d\<close> have "pp_of_term ` ?M \<subseteq> dgrad_set d (dgrad_max d)"
by (auto intro!: dgrad_sig_setD_lp)
ultimately obtain u where "u \<in> ?M" and min: "\<And>v. v \<prec>\<^sub>t u \<Longrightarrow> v \<notin> ?M"
by (rule ord_term_minimum_dgrad_set, blast)
from \<open>u \<in> ?M\<close> obtain m where "m \<in> M" and u': "u = lt m" ..
from this(1) obtain g1 where "g1 \<in> dgrad_sig_set d" and 1: "is_sig_red (\<prec>\<^sub>t) (=) {g1} p"
and m: "m = monom_mult (?lc / punit.lc (rep_list g1)) (?lp - punit.lt (rep_list g1)) g1"
unfolding M_def by fastforce
from 1 have adds: "punit.lt (rep_list g1) adds ?lp" and "?lp \<oplus> lt g1 \<prec>\<^sub>t punit.lt (rep_list g1) \<oplus> lt p"
and "rep_list g1 \<noteq> 0" using is_sig_red_top_addsE by fastforce+
from this(3) have lc_g1: "punit.lc (rep_list g1) \<noteq> 0" by (rule punit.lc_not_0)
from \<open>m \<in> M\<close> \<open>0 \<notin> rep_list ` M\<close> have "rep_list m \<noteq> 0" by fastforce
from \<open>m \<in> M\<close> \<open>0 \<notin> M\<close> have "m \<noteq> 0" by blast
hence "lc m \<noteq> 0" by (rule lc_not_0)
from lc_g1 have eq0: "punit.lc (rep_list m) = ?lc" by (simp add: m rep_list_monom_mult)
from \<open>?lc \<noteq> 0\<close> \<open>rep_list g1 \<noteq> 0\<close> adds have eq1: "punit.lt (rep_list m) = ?lp"
by (simp add: m rep_list_monom_mult punit.lt_monom_mult punit.lc_eq_zero_iff adds_minus)
from \<open>m \<in> M\<close> \<open>M \<subseteq> dgrad_sig_set d\<close> have "m \<in> dgrad_sig_set d" ..
from \<open>rep_list g1 \<noteq> 0\<close> have "punit.lc (rep_list g1) \<noteq> 0" and "g1 \<noteq> 0"
by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
with \<open>?lc \<noteq> 0\<close> have u: "u = (?lp - punit.lt (rep_list g1)) \<oplus> lt g1"
by (simp add: u' m lt_monom_mult lc_eq_zero_iff)
hence "punit.lt (rep_list g1) \<oplus> u = punit.lt (rep_list g1) \<oplus> ((?lp - punit.lt (rep_list g1)) \<oplus> lt g1)"
by simp
also from adds have "... = ?lp \<oplus> lt g1"
by (simp only: splus_assoc[symmetric], metis add.commute adds_minus)
also have "... \<prec>\<^sub>t punit.lt (rep_list g1) \<oplus> lt p" by fact
finally have "u \<prec>\<^sub>t lt p" by (rule ord_term_strict_canc)
from \<open>u \<in> ?M\<close> have "pp_of_term u \<in> pp_of_term ` ?M" by (rule imageI)
also have "... \<subseteq> dgrad_set d (dgrad_max d)" by fact
finally have "d (pp_of_term u) \<le> dgrad_max d" by (rule dgrad_setD)
from \<open>u \<in> ?M\<close> have "component_of_term u \<in> component_of_term ` ?M" by (rule imageI)
also from \<open>M \<subseteq> sig_inv_set\<close> \<open>0 \<notin> M\<close> sig_inv_setD_lt have "... \<subseteq> {0..<length fs}" by fastforce
finally have "component_of_term u < length fs" by simp
have "\<not> is_syz_sig d u"
proof
assume "is_syz_sig d u"
then obtain s where "s \<noteq> 0" and "lt s = u" and "s \<in> dgrad_sig_set d" and "rep_list s = 0"
by (rule is_syz_sigE)
let ?s = "monom_mult (lc m / lc s) 0 s"
have "rep_list ?s = 0" by (simp add: rep_list_monom_mult \<open>rep_list s = 0\<close>)
from \<open>s \<noteq> 0\<close> have "lc s \<noteq> 0" by (rule lc_not_0)
hence "lc m / lc s \<noteq> 0" using \<open>lc m \<noteq> 0\<close> by simp
have "m - ?s \<noteq> 0"
proof
assume "m - ?s = 0"
hence "m = ?s" by simp
with \<open>rep_list ?s = 0\<close> have "rep_list m = 0" by simp
with \<open>rep_list m \<noteq> 0\<close> show False ..
qed
moreover from \<open>lc m / lc s \<noteq> 0\<close> have "lt ?s = lt m" by (simp add: lt_monom_mult_zero \<open>lt s = u\<close> u')
moreover from \<open>lc s \<noteq> 0\<close> have "lc ?s = lc m" by simp
ultimately have "lt (m - ?s) \<prec>\<^sub>t u" unfolding u' by (rule lt_minus_lessI)
hence "lt (m - ?s) \<notin> ?M" by (rule min)
hence "m - ?s \<notin> M" by blast
moreover have "m - ?s \<in> M"
proof -
have "?s = monom_mult (?lc / lc s) 0 (monom_mult (lc g1 / punit.lc (rep_list g1)) 0 s)"
by (simp add: m monom_mult_assoc mult.commute)
define m' where "m' = m - ?s"
have eq: "rep_list m' = rep_list m" by (simp add: m'_def rep_list_minus \<open>rep_list ?s = 0\<close>)
from \<open>?lc \<noteq> 0\<close> have "m' = monom_mult (?lc / punit.lc (rep_list m')) (?lp - punit.lt (rep_list m')) m'"
by (simp add: eq eq0 eq1)
also have "... \<in> M" unfolding M_def
proof (rule, intro exI conjI)
from \<open>s \<in> dgrad_sig_set d\<close> have "?s \<in> dgrad_sig_set d"
by (rule dgrad_sig_set_closed_monom_mult_zero)
with \<open>m \<in> dgrad_sig_set d\<close> show "m' \<in> dgrad_sig_set d" unfolding m'_def
by (rule dgrad_sig_set_closed_minus)
next
show "is_sig_red (\<prec>\<^sub>t) (=) {m'} p"
proof (rule is_sig_red_top_addsI)
show "m' \<in> {m'}" by simp
next
from \<open>rep_list m \<noteq> 0\<close> show "rep_list m' \<noteq> 0" by (simp add: eq)
next
show "punit.lt (rep_list m') adds punit.lt (rep_list p)" by (simp add: eq eq1)
next
have "punit.lt (rep_list p) \<oplus> lt m' \<prec>\<^sub>t punit.lt (rep_list p) \<oplus> u"
by (rule splus_mono_strict, simp only: m'_def \<open>lt (m - ?s) \<prec>\<^sub>t u\<close>)
also have "... \<prec>\<^sub>t punit.lt (rep_list m') \<oplus> lt p"
unfolding eq eq1 using \<open>u \<prec>\<^sub>t lt p\<close> by (rule splus_mono_strict)
finally show "punit.lt (rep_list p) \<oplus> lt m' \<prec>\<^sub>t punit.lt (rep_list m') \<oplus> lt p" .
next
show "ord_term_lin.is_le_rel (\<prec>\<^sub>t)" by simp
qed fact
qed (fact refl)
finally show ?thesis by (simp only: m'_def)
qed
ultimately show False ..
qed
have "is_RB_in d rword G u" by (rule is_RB_uptD2, fact+)
thus ?thesis
proof (rule is_RB_inE)
assume "is_syz_sig d u"
with \<open>\<not> is_syz_sig d u\<close> show ?thesis ..
next
fix g
assume "is_canon_rewriter rword G u g"
hence "g \<in> G" and "g \<noteq> 0" and adds': "lt g adds\<^sub>t u" by (rule is_canon_rewriterD)+
assume irred: "\<not> is_sig_red (\<prec>\<^sub>t) (=) G (monom_mult 1 (pp_of_term u - lp g) g)"
define b where "b = monom_mult 1 (pp_of_term u - lp g) g"
note assms(1)
moreover have "is_sig_GB_upt d G (lt m)" unfolding u'[symmetric]
by (rule is_sig_GB_upt_le, rule is_RB_upt_is_sig_GB_upt, fact+, rule ord_term_lin.less_imp_le, fact)
moreover from assms(1) have "b \<in> dgrad_sig_set d" unfolding b_def
proof (rule dgrad_sig_set_closed_monom_mult)
from adds' have "lp g adds pp_of_term u" by (simp add: adds_term_def)
with assms(1) have "d (pp_of_term u - lp g) \<le> d (pp_of_term u)" by (rule dickson_grading_minus)
thus "d (pp_of_term u - lp g) \<le> dgrad_max d" using \<open>d (pp_of_term u) \<le> dgrad_max d\<close>
by (rule le_trans)
next
from \<open>g \<in> G\<close> G_sub show "g \<in> dgrad_sig_set d" ..
qed
moreover note \<open>m \<in> dgrad_sig_set d\<close>
moreover from \<open>g \<noteq> 0\<close> have "lt b = lt m"
by (simp add: b_def u'[symmetric] lt_monom_mult,
metis adds' add_diff_cancel_right' adds_termE pp_of_term_splus)
moreover from \<open>g \<noteq> 0\<close> have "b \<noteq> 0" by (simp add: b_def monom_mult_eq_zero_iff)
moreover note \<open>m \<noteq> 0\<close>
moreover from irred have "\<not> is_sig_red (\<prec>\<^sub>t) (=) G b" by (simp add: b_def)
moreover have "\<not> is_sig_red (\<prec>\<^sub>t) (=) G m"
proof
assume "is_sig_red (\<prec>\<^sub>t) (=) G m"
then obtain g2 where 1: "g2 \<in> G" and 2: "rep_list g2 \<noteq> 0"
and 3: "punit.lt (rep_list g2) adds punit.lt (rep_list m)"
and 4: "punit.lt (rep_list m) \<oplus> lt g2 \<prec>\<^sub>t punit.lt (rep_list g2) \<oplus> lt m"
by (rule is_sig_red_top_addsE)
from 2 have "g2 \<noteq> 0" and "punit.lc (rep_list g2) \<noteq> 0" by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
with 3 4 have "lt (monom_mult (?lc / punit.lc (rep_list g2)) (?lp - punit.lt (rep_list g2)) g2) \<prec>\<^sub>t u"
(is "lt ?g2 \<prec>\<^sub>t u")
using \<open>?lc \<noteq> 0\<close> by (simp add: term_is_le_rel_minus u' eq1 lt_monom_mult)
hence "lt ?g2 \<notin> ?M" by (rule min)
hence "?g2 \<notin> M" by blast
hence "g2 \<notin> dgrad_sig_set d \<or> \<not> is_sig_red (\<prec>\<^sub>t) (=) {g2} p" by (simp add: M_def)
thus False
proof
assume "g2 \<notin> dgrad_sig_set d"
moreover from \<open>g2 \<in> G\<close> G_sub have "g2 \<in> dgrad_sig_set d" ..
ultimately show ?thesis ..
next
assume "\<not> is_sig_red (\<prec>\<^sub>t) (=) {g2} p"
moreover have "is_sig_red (\<prec>\<^sub>t) (=) {g2} p"
proof (rule is_sig_red_top_addsI)
show "g2 \<in> {g2}" by simp
next
from 3 show "punit.lt (rep_list g2) adds punit.lt (rep_list p)" by (simp only: eq1)
next
from 4 have "?lp \<oplus> lt g2 \<prec>\<^sub>t punit.lt (rep_list g2) \<oplus> u" by (simp only: eq1 u')
also from \<open>u \<prec>\<^sub>t lt p\<close> have "... \<prec>\<^sub>t punit.lt (rep_list g2) \<oplus> lt p" by (rule splus_mono_strict)
finally show "?lp \<oplus> lt g2 \<prec>\<^sub>t punit.lt (rep_list g2) \<oplus> lt p" .
next
show "ord_term_lin.is_le_rel (\<prec>\<^sub>t)" by simp
qed fact+
ultimately show ?thesis ..
qed
qed
ultimately have eq2: "punit.lt (rep_list b) = punit.lt (rep_list m)"
by (rule sig_regular_top_reduced_lt_unique)
have "rep_list g \<noteq> 0" by (rule is_RB_inD, fact+)
moreover from adds' have "lp g adds pp_of_term u" and "component_of_term (lt g) = component_of_term u"
by (simp_all add: adds_term_def)
ultimately have "u = (?lp - punit.lt (rep_list g)) \<oplus> lt g"
by (simp add: eq1[symmetric] eq2[symmetric] b_def rep_list_monom_mult punit.lt_monom_mult
splus_def adds_minus term_simps)
have "is_sig_red (\<prec>\<^sub>t) (=) {b} p"
proof (rule is_sig_red_top_addsI)
show "b \<in> {b}" by simp
next
from \<open>rep_list g \<noteq> 0\<close> show "rep_list b \<noteq> 0"
by (simp add: b_def rep_list_monom_mult punit.monom_mult_eq_zero_iff)
next
show "punit.lt (rep_list b) adds punit.lt (rep_list p)" by (simp add: eq1 eq2)
next
show "punit.lt (rep_list p) \<oplus> lt b \<prec>\<^sub>t punit.lt (rep_list b) \<oplus> lt p"
by (simp add: eq1 eq2 \<open>lt b = lt m\<close> u'[symmetric] \<open>u \<prec>\<^sub>t lt p\<close> splus_mono_strict)
next
show "ord_term_lin.is_le_rel (\<prec>\<^sub>t)" by simp
qed fact
hence "is_sig_red (\<prec>\<^sub>t) (=) {g} p" unfolding b_def by (rule is_sig_red_singleton_monom_multD)
show ?thesis by (rule, fact+)
qed
qed
subsubsection \<open>Termination\<close>
definition term_pp_rel :: "('t \<Rightarrow> 't \<Rightarrow> bool) \<Rightarrow> ('t \<times> 'a) \<Rightarrow> ('t \<times> 'a) \<Rightarrow> bool"
where "term_pp_rel r a b \<longleftrightarrow> r (snd b \<oplus> fst a) (snd a \<oplus> fst b)"
definition canon_term_pp_pair :: "('t \<times> 'a) \<Rightarrow> bool"
where "canon_term_pp_pair a \<longleftrightarrow> (gcs (pp_of_term (fst a)) (snd a) = 0)"
definition cancel_term_pp_pair :: "('t \<times> 'a) \<Rightarrow> ('t \<times> 'a)"
where "cancel_term_pp_pair a = (fst a \<ominus> (gcs (pp_of_term (fst a)) (snd a)), snd a - (gcs (pp_of_term (fst a)) (snd a)))"
lemma term_pp_rel_refl: "reflp r \<Longrightarrow> term_pp_rel r a a"
by (simp add: term_pp_rel_def reflp_def)
lemma term_pp_rel_irrefl: "irreflp r \<Longrightarrow> \<not> term_pp_rel r a a"
by (simp add: term_pp_rel_def irreflp_def)
lemma term_pp_rel_sym: "symp r \<Longrightarrow> term_pp_rel r a b \<Longrightarrow> term_pp_rel r b a"
by (auto simp: term_pp_rel_def symp_def)
lemma term_pp_rel_trans:
assumes "ord_term_lin.is_le_rel r" and "term_pp_rel r a b" and "term_pp_rel r b c"
shows "term_pp_rel r a c"
proof -
from assms(1) have "transp r" by (rule ord_term_lin.is_le_relE, auto)
from assms(2) have 1: "r (snd b \<oplus> fst a) (snd a \<oplus> fst b)" by (simp only: term_pp_rel_def)
from assms(3) have 2: "r (snd c \<oplus> fst b) (snd b \<oplus> fst c)" by (simp only: term_pp_rel_def)
have "snd b \<oplus> (snd c \<oplus> fst a) = snd c \<oplus> (snd b \<oplus> fst a)" by (rule splus_left_commute)
also from assms(1) 1 have "r ... (snd a \<oplus> (snd c \<oplus> fst b))"
by (simp add: splus_left_commute[of "snd a"] term_is_le_rel_canc_left)
also from assms(1) 2 have "r ... (snd b \<oplus> (snd a \<oplus> fst c))"
by (simp add: splus_left_commute[of "snd b"] term_is_le_rel_canc_left)
finally(transpD[OF \<open>transp r\<close>]) show ?thesis using assms(1)
by (simp only: term_pp_rel_def term_is_le_rel_canc_left)
qed
lemma term_pp_rel_trans_eq_left:
assumes "ord_term_lin.is_le_rel r" and "term_pp_rel (=) a b" and "term_pp_rel r b c"
shows "term_pp_rel r a c"
proof -
from assms(1) have "transp r" by (rule ord_term_lin.is_le_relE, auto)
from assms(2) have 1: "snd b \<oplus> fst a = snd a \<oplus> fst b" by (simp only: term_pp_rel_def)
from assms(3) have 2: "r (snd c \<oplus> fst b) (snd b \<oplus> fst c)" by (simp only: term_pp_rel_def)
have "snd b \<oplus> (snd c \<oplus> fst a) = snd c \<oplus> (snd b \<oplus> fst a)" by (rule splus_left_commute)
also from assms(1) 1 have "... = (snd a \<oplus> (snd c \<oplus> fst b))"
by (simp add: splus_left_commute[of "snd a"])
finally have eq: "snd b \<oplus> (snd c \<oplus> fst a) = snd a \<oplus> (snd c \<oplus> fst b)" .
from assms(1) 2 have "r (snd b \<oplus> (snd c \<oplus> fst a)) (snd b \<oplus> (snd a \<oplus> fst c))"
unfolding eq by (simp add: splus_left_commute[of "snd b"] term_is_le_rel_canc_left)
thus ?thesis using assms(1) by (simp only: term_pp_rel_def term_is_le_rel_canc_left)
qed
lemma term_pp_rel_trans_eq_right:
assumes "ord_term_lin.is_le_rel r" and "term_pp_rel r a b" and "term_pp_rel (=) b c"
shows "term_pp_rel r a c"
proof -
from assms(1) have "transp r" by (rule ord_term_lin.is_le_relE, auto)
from assms(2) have 1: "r (snd b \<oplus> fst a) (snd a \<oplus> fst b)" by (simp only: term_pp_rel_def)
from assms(3) have 2: "snd c \<oplus> fst b = snd b \<oplus> fst c" by (simp only: term_pp_rel_def)
have "snd b \<oplus> (snd a \<oplus> fst c) = snd a \<oplus> (snd b \<oplus> fst c)" by (rule splus_left_commute)
also from assms(1) 2 have "... = (snd a \<oplus> (snd c \<oplus> fst b))"
by (simp add: splus_left_commute[of "snd a"])
finally have eq: "snd b \<oplus> (snd a \<oplus> fst c) = snd a \<oplus> (snd c \<oplus> fst b)" .
from assms(1) 1 have "r (snd b \<oplus> (snd c \<oplus> fst a)) (snd b \<oplus> (snd a \<oplus> fst c))"
unfolding eq by (simp add: splus_left_commute[of _ "snd c"] term_is_le_rel_canc_left)
thus ?thesis using assms(1) by (simp only: term_pp_rel_def term_is_le_rel_canc_left)
qed
lemma canon_term_pp_cancel: "canon_term_pp_pair (cancel_term_pp_pair a)"
by (simp add: cancel_term_pp_pair_def canon_term_pp_pair_def gcs_minus_gcs term_simps)
lemma term_pp_rel_cancel:
assumes "reflp r"
shows "term_pp_rel r a (cancel_term_pp_pair a)"
proof -
obtain u s where a: "a = (u, s)" by (rule prod.exhaust)
show ?thesis
proof (simp add: a cancel_term_pp_pair_def)
let ?g = "gcs (pp_of_term u) s"
have "?g adds s" by (fact gcs_adds_2)
hence "(s - ?g) \<oplus> (u \<ominus> 0) = s \<oplus> u \<ominus> (?g + 0)" using zero_adds_pp
by (rule minus_splus_sminus)
also have "... = s \<oplus> (u \<ominus> ?g)"
by (metis add.left_neutral add.right_neutral adds_pp_def diff_zero gcs_adds_2 gcs_comm
minus_splus_sminus zero_adds)
finally have "r ((s - ?g) \<oplus> u) (s \<oplus> (u \<ominus> ?g))" using assms by (simp add: term_simps reflp_def)
thus "term_pp_rel r (u, s) (u \<ominus> ?g, s - ?g)" by (simp add: a term_pp_rel_def)
qed
qed
lemma canon_term_pp_rel_id:
assumes "term_pp_rel (=) a b" and "canon_term_pp_pair a" and "canon_term_pp_pair b"
shows "a = b"
proof -
obtain u s where a: "a = (u, s)" by (rule prod.exhaust)
obtain v t where b: "b = (v, t)" by (rule prod.exhaust)
from assms(1) have "t \<oplus> u = s \<oplus> v" by (simp add: term_pp_rel_def a b)
hence 1: "t + pp_of_term u = s + pp_of_term v" by (metis pp_of_term_splus)
from assms(2) have 2: "gcs (pp_of_term u) s = 0" by (simp add: canon_term_pp_pair_def a)
from assms(3) have 3: "gcs (pp_of_term v) t = 0" by (simp add: canon_term_pp_pair_def b)
have "t = t + gcs (pp_of_term u) s" by (simp add: 2)
also have "... = gcs (t + pp_of_term u) (t + s)" by (simp only: gcs_plus_left)
also have "... = gcs (s + pp_of_term v) (s + t)" by (simp only: 1 add.commute)
also have "... = s + gcs (pp_of_term v) t" by (simp only: gcs_plus_left)
also have "... = s" by (simp add: 3)
finally have "t = s" .
moreover from \<open>t \<oplus> u = s \<oplus> v\<close> have "u = v" by (simp only: \<open>t = s\<close> splus_left_canc)
ultimately show ?thesis by (simp add: a b)
qed
lemma min_set_finite:
fixes seq :: "nat \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b::field)"
assumes "dickson_grading d" and "range seq \<subseteq> dgrad_sig_set d" and "0 \<notin> rep_list ` range seq"
and "\<And>i j. i < j \<Longrightarrow> lt (seq i) \<prec>\<^sub>t lt (seq j)"
shows "finite {i. \<not> (\<exists>j<i. lt (seq j) adds\<^sub>t lt (seq i) \<and>
punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq i)))}"
proof -
have "inj (\<lambda>i. lt (seq i))"
proof
fix i j
assume eq: "lt (seq i) = lt (seq j)"
show "i = j"
proof (rule linorder_cases)
assume "i < j"
hence "lt (seq i) \<prec>\<^sub>t lt (seq j)" by (rule assms(4))
thus ?thesis by (simp add: eq)
next
assume "j < i"
hence "lt (seq j) \<prec>\<^sub>t lt (seq i)" by (rule assms(4))
thus ?thesis by (simp add: eq)
qed
qed
hence "inj seq" unfolding comp_def[symmetric] by (rule inj_on_imageI2)
let ?P1 = "\<lambda>p q. lt p adds\<^sub>t lt q"
let ?P2 = "\<lambda>p q. punit.lt (rep_list p) adds punit.lt (rep_list q)"
let ?P = "\<lambda>p q. ?P1 p q \<and> ?P2 p q"
have "reflp ?P" by (simp add: reflp_def adds_term_refl)
have "almost_full_on ?P1 (range seq)"
proof (rule almost_full_on_map)
let ?B = "{t. pp_of_term t \<in> dgrad_set d (dgrad_max d) \<and> component_of_term t \<in> {0..<length fs}}"
from assms(1) finite_atLeastLessThan show "almost_full_on (adds\<^sub>t) ?B" by (rule Dickson_term)
show "lt ` range seq \<subseteq> ?B"
proof
fix v
assume "v \<in> lt ` range seq"
then obtain p where "p \<in> range seq" and v: "v = lt p" ..
from this(1) assms(3) have "rep_list p \<noteq> 0" by auto
hence "p \<noteq> 0" by (auto simp: rep_list_zero)
from \<open>p \<in> range seq\<close> assms(2) have "p \<in> dgrad_sig_set d" ..
hence "d (lp p) \<le> dgrad_max d" by (rule dgrad_sig_setD_lp)
hence "lp p \<in> dgrad_set d (dgrad_max d)" by (simp add: dgrad_set_def)
moreover from \<open>p \<in> dgrad_sig_set d\<close> \<open>p \<noteq> 0\<close> have "component_of_term (lt p) < length fs"
by (rule dgrad_sig_setD_lt)
ultimately show "v \<in> ?B" by (simp add: v)
qed
qed
moreover have "almost_full_on ?P2 (range seq)"
proof (rule almost_full_on_map)
let ?B = "dgrad_set d (dgrad_max d)"
from assms(1) show "almost_full_on (adds) ?B" by (rule dickson_gradingD_dgrad_set)
show "(\<lambda>p. punit.lt (rep_list p)) ` range seq \<subseteq> ?B"
proof
fix t
assume "t \<in> (\<lambda>p. punit.lt (rep_list p)) ` range seq"
then obtain p where "p \<in> range seq" and t: "t = punit.lt (rep_list p)" ..
from this(1) assms(3) have "rep_list p \<noteq> 0" by auto
from \<open>p \<in> range seq\<close> assms(2) have "p \<in> dgrad_sig_set d" ..
hence "p \<in> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
with assms(1) have "rep_list p \<in> punit_dgrad_max_set d" by (rule dgrad_max_2)
from this \<open>rep_list p \<noteq> 0\<close> have "d (punit.lt (rep_list p)) \<le> dgrad_max d"
by (rule punit.dgrad_p_setD_lp[simplified])
thus "t \<in> ?B" by (simp add: t dgrad_set_def)
qed
qed
ultimately have "almost_full_on ?P (range seq)" by (rule almost_full_on_same)
with \<open>reflp ?P\<close> obtain T where "finite T" and "T \<subseteq> range seq" and *: "\<And>p. p \<in> range seq \<Longrightarrow> (\<exists>q\<in>T. ?P q p)"
by (rule almost_full_on_finite_subsetE, blast)
from \<open>T \<subseteq> range seq\<close> obtain I where T: "T = seq ` I" by (meson subset_image_iff)
have "{i. \<not> (\<exists>j<i. ?P (seq j) (seq i))} \<subseteq> I"
proof
fix i
assume "i \<in> {i. \<not> (\<exists>j<i. ?P (seq j) (seq i))}"
hence x: "\<not> (\<exists>j<i. ?P (seq j) (seq i))" by simp
obtain j where "j \<in> I" and "?P (seq j) (seq i)"
proof -
have "seq i \<in> range seq" by simp
hence "\<exists>q\<in>T. ?P q (seq i)" by (rule *)
then obtain q where "q \<in> T" and "?P q (seq i)" ..
from this(1) obtain j where "j \<in> I" and "q = seq j" unfolding T ..
from this(1) \<open>?P q (seq i)\<close> show ?thesis unfolding \<open>q = seq j\<close> ..
qed
from this(2) x have "i \<le> j" by auto
moreover have "\<not> i < j"
proof
assume "i < j"
hence "lt (seq i) \<prec>\<^sub>t lt (seq j)" by (rule assms(4))
hence "\<not> ?P1 (seq j) (seq i)" using ord_adds_term ord_term_lin.leD by blast
with \<open>?P (seq j) (seq i)\<close> show False by simp
qed
ultimately show "i \<in> I" using \<open>j \<in> I\<close> by simp
qed
moreover from \<open>inj seq\<close> \<open>finite T\<close> have "finite I" by (simp add: finite_image_iff inj_on_subset T)
ultimately show ?thesis by (rule finite_subset)
qed
lemma rb_termination:
fixes seq :: "nat \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b::field)"
assumes "dickson_grading d" and "range seq \<subseteq> dgrad_sig_set d" and "0 \<notin> rep_list ` range seq"
and "\<And>i j. i < j \<Longrightarrow> lt (seq i) \<prec>\<^sub>t lt (seq j)"
and "\<And>i. \<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (seq ` {0..<i}) (seq i)"
and "\<And>i. (\<exists>j<length fs. lt (seq i) = lt (monomial (1::'b) (term_of_pair (0, j))) \<and>
punit.lt (rep_list (seq i)) \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) \<or>
(\<exists>j k. is_regular_spair (seq j) (seq k) \<and> rep_list (spair (seq j) (seq k)) \<noteq> 0 \<and>
lt (seq i) = lt (spair (seq j) (seq k)) \<and>
punit.lt (rep_list (seq i)) \<preceq> punit.lt (rep_list (spair (seq j) (seq k))))"
and "\<And>i. is_sig_GB_upt d (seq ` {0..<i}) (lt (seq i))"
shows thesis
proof -
from assms(3) have "0 \<notin> range seq" using rep_list_zero by auto
have "ord_term_lin.is_le_rel (=)" and "ord_term_lin.is_le_rel (\<prec>\<^sub>t)" by (rule ord_term_lin.is_le_relI)+
have "reflp (=)" and "symp (=)" by (simp_all add: symp_def)
have "irreflp (\<prec>\<^sub>t)" by (simp add: irreflp_def)
have "inj (\<lambda>i. lt (seq i))"
proof
fix i j
assume eq: "lt (seq i) = lt (seq j)"
show "i = j"
proof (rule linorder_cases)
assume "i < j"
hence "lt (seq i) \<prec>\<^sub>t lt (seq j)" by (rule assms(4))
thus ?thesis by (simp add: eq)
next
assume "j < i"
hence "lt (seq j) \<prec>\<^sub>t lt (seq i)" by (rule assms(4))
thus ?thesis by (simp add: eq)
qed
qed
hence "inj seq" unfolding comp_def[symmetric] by (rule inj_on_imageI2)
define R where "R = (\<lambda>x. {i. term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) x})"
let ?A = "{x. canon_term_pp_pair x \<and> R x \<noteq> {}}"
have "finite ?A"
proof -
define min_set where "min_set = {i. \<not> (\<exists>j<i. lt (seq j) adds\<^sub>t lt (seq i) \<and>
punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq i)))}"
have "?A \<subseteq> (\<lambda>i. cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i)))) ` min_set"
proof
fix u t
assume "(u, t) \<in> ?A"
hence "canon_term_pp_pair (u, t)" and "R (u, t) \<noteq> {}" by simp_all
from this(2) obtain i where x: "term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) (u, t)"
by (auto simp: R_def)
let ?equiv = "(\<lambda>i j. term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) (lt (seq j), punit.lt (rep_list (seq j))))"
obtain j where "j \<in> min_set" and "?equiv j i"
proof (cases "i \<in> min_set")
case True
moreover have "?equiv i i" by (simp add: term_pp_rel_refl)
ultimately show ?thesis ..
next
case False
let ?Q = "{seq j | j. j < i \<and> is_sig_red (=) (=) {seq j} (seq i)}"
have "?Q \<subseteq> range seq" by blast
also have "... \<subseteq> dgrad_sig_set d" by (fact assms(2))
finally have "?Q \<subseteq> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
moreover from \<open>?Q \<subseteq> range seq\<close> \<open>0 \<notin> range seq\<close> have "0 \<notin> ?Q" by blast
ultimately have Q_sub: "pp_of_term ` lt ` ?Q \<subseteq> dgrad_set d (dgrad_max d)"
unfolding image_image by (smt CollectI dgrad_p_setD_lp dgrad_set_def image_subset_iff subsetCE)
have *: "\<exists>g\<in>seq ` {0..<k}. is_sig_red (=) (=) {g} (seq k)" if "k \<notin> min_set" for k
proof -
from that obtain j where "j < k" and a: "lt (seq j) adds\<^sub>t lt (seq k)"
and b: "punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq k))" by (auto simp: min_set_def)
note assms(1, 7)
moreover from assms(2) have "seq k \<in> dgrad_sig_set d" by fastforce
moreover from \<open>j < k\<close> have "seq j \<in> seq ` {0..<k}" by simp
moreover from assms(3) have "rep_list (seq k) \<noteq> 0" and "rep_list (seq j) \<noteq> 0" by fastforce+
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) (seq ` {0..<k}) (seq k)" using a b by (rule lemma_21)
moreover from assms(5)[of k] have "\<not> is_sig_red (\<prec>\<^sub>t) (=) (seq ` {0..<k}) (seq k)"
by (simp add: is_sig_red_top_tail_cases)
ultimately have "is_sig_red (=) (=) (seq ` {0..<k}) (seq k)"
by (simp add: is_sig_red_sing_reg_cases)
then obtain g0 where "g0 \<in> seq ` {0..<k}" and "is_sig_red (=) (=) {g0} (seq k)"
by (rule is_sig_red_singletonI)
thus ?thesis ..
qed
from this[OF False] obtain g0 where "g0 \<in> seq ` {0..<i}" and "is_sig_red (=) (=) {g0} (seq i)" ..
hence "g0 \<in> ?Q" by fastforce
hence "lt g0 \<in> lt ` ?Q" by (rule imageI)
with assms(1) obtain v where "v \<in> lt ` ?Q" and min: "\<And>v'. v' \<prec>\<^sub>t v \<Longrightarrow> v' \<notin> lt ` ?Q"
using Q_sub by (rule ord_term_minimum_dgrad_set, blast)
from this(1) obtain j where "j < i" and "is_sig_red (=) (=) {seq j} (seq i)"
and v: "v = lt (seq j)" by fastforce
hence 1: "punit.lt (rep_list (seq j)) adds punit.lt (rep_list (seq i))"
and 2: "punit.lt (rep_list (seq i)) \<oplus> lt (seq j) = punit.lt (rep_list (seq j)) \<oplus> lt (seq i)"
by (auto elim: is_sig_red_top_addsE)
show ?thesis
proof
show "?equiv j i" by (simp add: term_pp_rel_def 2)
next
show "j \<in> min_set"
proof (rule ccontr)
assume "j \<notin> min_set"
from *[OF this] obtain g1 where "g1 \<in> seq ` {0..<j}" and red: "is_sig_red (=) (=) {g1} (seq j)" ..
from this(1) obtain j0 where "j0 < j" and "g1 = seq j0" by fastforce+
from red have 3: "punit.lt (rep_list (seq j0)) adds punit.lt (rep_list (seq j))"
and 4: "punit.lt (rep_list (seq j)) \<oplus> lt (seq j0) = punit.lt (rep_list (seq j0)) \<oplus> lt (seq j)"
by (auto simp: \<open>g1 = seq j0\<close> elim: is_sig_red_top_addsE)
from \<open>j0 < j\<close> \<open>j < i\<close> have "j0 < i" by simp
from \<open>j0 < j\<close> have "lt (seq j0) \<prec>\<^sub>t v" unfolding v by (rule assms(4))
hence "lt (seq j0) \<notin> lt `?Q" by (rule min)
with \<open>j0 < i\<close> have "\<not> is_sig_red (=) (=) {seq j0} (seq i)" by blast
moreover have "is_sig_red (=) (=) {seq j0} (seq i)"
proof (rule is_sig_red_top_addsI)
from assms(3) show "rep_list (seq j0) \<noteq> 0" by fastforce
next
from assms(3) show "rep_list (seq i) \<noteq> 0" by fastforce
next
from 3 1 show "punit.lt (rep_list (seq j0)) adds punit.lt (rep_list (seq i))"
by (rule adds_trans)
next
from 4 have "?equiv j0 j" by (simp add: term_pp_rel_def)
also from 2 have "?equiv j i" by (simp add: term_pp_rel_def)
finally(term_pp_rel_trans[OF \<open>ord_term_lin.is_le_rel (=)\<close>])
show "punit.lt (rep_list (seq i)) \<oplus> lt (seq j0) = punit.lt (rep_list (seq j0)) \<oplus> lt (seq i)"
by (simp add: term_pp_rel_def)
next
show "ord_term_lin.is_le_rel (=)" by simp
qed simp_all
ultimately show False ..
qed
qed
qed
have "term_pp_rel (=) (cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j)))) (lt (seq j), punit.lt (rep_list (seq j)))"
by (rule term_pp_rel_sym, fact \<open>symp (=)\<close>, rule term_pp_rel_cancel, fact \<open>reflp (=)\<close>)
also note \<open>?equiv j i\<close>
also(term_pp_rel_trans[OF \<open>ord_term_lin.is_le_rel (=)\<close>]) note x
finally(term_pp_rel_trans[OF \<open>ord_term_lin.is_le_rel (=)\<close>])
have "term_pp_rel (=) (cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j)))) (u, t)" .
with \<open>symp (=)\<close> have "term_pp_rel (=) (u, t) (cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j))))"
by (rule term_pp_rel_sym)
hence "(u, t) = cancel_term_pp_pair (lt (seq j), punit.lt (rep_list (seq j)))"
using \<open>canon_term_pp_pair (u, t)\<close> canon_term_pp_cancel by (rule canon_term_pp_rel_id)
with \<open>j \<in> min_set\<close> show "(u, t) \<in> (\<lambda>i. cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i)))) ` min_set"
by fastforce
qed
moreover have "finite ((\<lambda>i. cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i)))) ` min_set)"
proof (rule finite_imageI)
show "finite min_set" unfolding min_set_def using assms(1-4) by (rule min_set_finite)
qed
ultimately show ?thesis by (rule finite_subset)
qed
have "range seq \<subseteq> seq ` (\<Union> (R ` ?A))"
proof (rule image_mono, rule)
fix i
show "i \<in> (\<Union> (R ` ?A))"
proof
show "i \<in> R (cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i))))"
by (simp add: R_def term_pp_rel_cancel)
thus "cancel_term_pp_pair (lt (seq i), punit.lt (rep_list (seq i))) \<in> ?A"
using canon_term_pp_cancel by blast
qed
qed
moreover from \<open>inj seq\<close> have "infinite (range seq)" by (rule range_inj_infinite)
ultimately have "infinite (seq ` (\<Union> (R ` ?A)))" by (rule infinite_super)
moreover have "finite (seq ` (\<Union> (R ` ?A)))"
proof (rule finite_imageI, rule finite_UN_I)
fix x
assume "x \<in> ?A"
let ?rel = "term_pp_rel (\<prec>\<^sub>t)"
have "irreflp ?rel" by (rule irreflpI, rule term_pp_rel_irrefl, fact)
moreover have "transp ?rel" by (rule transpI, drule term_pp_rel_trans[OF \<open>ord_term_lin.is_le_rel (\<prec>\<^sub>t)\<close>])
ultimately have "wfp_on ?rel ?A" using \<open>finite ?A\<close> by (rule wfp_on_finite)
thus "finite (R x)" using \<open>x \<in> ?A\<close>
proof (induct rule: wfp_on_induct)
case (less x)
from less(1) have "canon_term_pp_pair x" by simp
define R' where " R' = \<Union> (R ` ({x. canon_term_pp_pair x \<and> R x \<noteq> {}} \<inter> {z. term_pp_rel (\<prec>\<^sub>t) z x}))"
define red_set where "red_set = (\<lambda>p::'t \<Rightarrow>\<^sub>0 'b. {k. lt (seq k) = lt p \<and>
punit.lt (rep_list (seq k)) \<preceq> punit.lt (rep_list p)})"
have finite_red_set: "finite (red_set p)" for p
proof (cases "red_set p = {}")
case True
thus ?thesis by simp
next
case False
then obtain k where lt_k: "lt (seq k) = lt p" by (auto simp: red_set_def)
have "red_set p \<subseteq> {k}"
proof
fix k'
assume "k' \<in> red_set p"
hence "lt (seq k') = lt p" by (simp add: red_set_def)
hence "lt (seq k') = lt (seq k)" by (simp only: lt_k)
with \<open>inj (\<lambda>i. lt (seq i))\<close> have "k' = k" by (rule injD)
thus "k' \<in> {k}" by simp
qed
thus ?thesis using infinite_super by auto
qed
have "R x \<subseteq> (\<Union>i\<in>R'. \<Union>j\<in>R'. red_set (spair (seq i) (seq j))) \<union>
(\<Union>j\<in>{0..<length fs}. red_set (monomial 1 (term_of_pair (0, j))))"
(is "_ \<subseteq> ?B \<union> ?C")
proof
fix i
assume "i \<in> R x"
hence i_x: "term_pp_rel (=) (lt (seq i), punit.lt (rep_list (seq i))) x"
by (simp add: R_def term_pp_rel_def)
from assms(6)[of i] show "i \<in> ?B \<union> ?C"
proof (elim disjE exE conjE)
fix j
assume "j < length fs"
hence "j \<in> {0..<length fs}" by simp
assume "lt (seq i) = lt (monomial (1::'b) (term_of_pair (0, j)))"
and "punit.lt (rep_list (seq i)) \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))"
hence "i \<in> red_set (monomial 1 (term_of_pair (0, j)))" by (simp add: red_set_def)
with \<open>j \<in> {0..<length fs}\<close> have "i \<in> ?C" ..
thus ?thesis ..
next
fix j k
let ?li = "punit.lt (rep_list (seq i))"
let ?lj = "punit.lt (rep_list (seq j))"
let ?lk = "punit.lt (rep_list (seq k))"
assume lt_i: "lt (seq i) = lt (spair (seq j) (seq k))"
and lt_i': "?li \<preceq> punit.lt (rep_list (spair (seq j) (seq k)))"
and spair_0: "rep_list (spair (seq j) (seq k)) \<noteq> 0"
hence "i \<in> red_set (spair (seq j) (seq k))" by (simp add: red_set_def)
from assms(3) have i_0: "rep_list (seq i) \<noteq> 0" and j_0: "rep_list (seq j) \<noteq> 0"
and k_0: "rep_list (seq k) \<noteq> 0" by fastforce+
have R'I: "a \<in> R'" if "term_pp_rel (\<prec>\<^sub>t) (lt (seq a), punit.lt (rep_list (seq a))) x" for a
proof -
let ?x = "cancel_term_pp_pair (lt (seq a), punit.lt (rep_list (seq a)))"
show ?thesis unfolding R'_def
proof (rule UN_I, simp, intro conjI)
show "a \<in> R ?x" by (simp add: R_def term_pp_rel_cancel)
thus "R ?x \<noteq> {}" by blast
next
note \<open>ord_term_lin.is_le_rel (\<prec>\<^sub>t)\<close>
moreover have "term_pp_rel (=) ?x (lt (seq a), punit.lt (rep_list (seq a)))"
by (rule term_pp_rel_sym, fact, rule term_pp_rel_cancel, fact)
ultimately show "term_pp_rel (\<prec>\<^sub>t) ?x x" using that by (rule term_pp_rel_trans_eq_left)
qed (fact canon_term_pp_cancel)
qed
assume "is_regular_spair (seq j) (seq k)"
hence "?lk \<oplus> lt (seq j) \<noteq> ?lj \<oplus> lt (seq k)" by (rule is_regular_spairD3)
hence "term_pp_rel (\<prec>\<^sub>t) (lt (seq j), ?lj) x \<and> term_pp_rel (\<prec>\<^sub>t) (lt (seq k), ?lk) x"
proof (rule ord_term_lin.neqE)
assume c: "?lk \<oplus> lt (seq j) \<prec>\<^sub>t ?lj \<oplus> lt (seq k)"
hence j_k: "term_pp_rel (\<prec>\<^sub>t) (lt (seq j), ?lj) (lt (seq k), ?lk)"
by (simp add: term_pp_rel_def)
note \<open>ord_term_lin.is_le_rel (\<prec>\<^sub>t)\<close>
moreover have "term_pp_rel (\<prec>\<^sub>t) (lt (seq k), ?lk) (lt (seq i), ?li)"
proof (simp add: term_pp_rel_def)
from lt_i' have "?li \<oplus> lt (seq k) \<preceq>\<^sub>t
punit.lt (rep_list (spair (seq j) (seq k))) \<oplus> lt (seq k)"
by (rule splus_mono_left)
also have "... \<prec>\<^sub>t (?lk - gcs ?lk ?lj + ?lj) \<oplus> lt (seq k)"
by (rule splus_mono_strict_left, rule lt_rep_list_spair, fact+, simp only: add.commute)
also have "... = ((?lk + ?lj) - gcs ?lj ?lk) \<oplus> lt (seq k)"
by (simp add: minus_plus gcs_adds_2 gcs_comm)
also have "... = ?lk \<oplus> ((?lj - gcs ?lj ?lk) \<oplus> lt (seq k))"
by (simp add: minus_plus' gcs_adds splus_assoc[symmetric])
also have "... = ?lk \<oplus> lt (seq i)"
by (simp add: lt_spair'[OF k_0 _ c] add.commute spair_comm[of "seq j"] lt_i)
finally show "?li \<oplus> lt (seq k) \<prec>\<^sub>t ?lk \<oplus> lt (seq i)" .
qed
ultimately have "term_pp_rel (\<prec>\<^sub>t) (lt (seq k), ?lk) x" using i_x
by (rule term_pp_rel_trans_eq_right)
moreover from \<open>ord_term_lin.is_le_rel (\<prec>\<^sub>t)\<close> j_k this
have "term_pp_rel (\<prec>\<^sub>t) (lt (seq j), ?lj) x" by (rule term_pp_rel_trans)
ultimately show ?thesis by simp
next
assume c: "?lj \<oplus> lt (seq k) \<prec>\<^sub>t ?lk \<oplus> lt (seq j)"
hence j_k: "term_pp_rel (\<prec>\<^sub>t) (lt (seq k), ?lk) (lt (seq j), ?lj)"
by (simp add: term_pp_rel_def)
note \<open>ord_term_lin.is_le_rel (\<prec>\<^sub>t)\<close>
moreover have "term_pp_rel (\<prec>\<^sub>t) (lt (seq j), ?lj) (lt (seq i), ?li)"
proof (simp add: term_pp_rel_def)
from lt_i' have "?li \<oplus> lt (seq j) \<preceq>\<^sub>t
punit.lt (rep_list (spair (seq j) (seq k))) \<oplus> lt (seq j)"
by (rule splus_mono_left)
thm lt_rep_list_spair
also have "... \<prec>\<^sub>t (?lk - gcs ?lk ?lj + ?lj) \<oplus> lt (seq j)"
by (rule splus_mono_strict_left, rule lt_rep_list_spair, fact+, simp only: add.commute)
also have "... = ((?lk + ?lj) - gcs ?lk ?lj) \<oplus> lt (seq j)"
by (simp add: minus_plus gcs_adds_2 gcs_comm)
also have "... = ?lj \<oplus> ((?lk - gcs ?lk ?lj) \<oplus> lt (seq j))"
by (simp add: minus_plus' gcs_adds splus_assoc[symmetric] add.commute)
also have "... = ?lj \<oplus> lt (seq i)" by (simp add: lt_spair'[OF j_0 _ c] lt_i add.commute)
finally show "?li \<oplus> lt (seq j) \<prec>\<^sub>t ?lj \<oplus> lt (seq i)" .
qed
ultimately have "term_pp_rel (\<prec>\<^sub>t) (lt (seq j), ?lj) x" using i_x
by (rule term_pp_rel_trans_eq_right)
moreover from \<open>ord_term_lin.is_le_rel (\<prec>\<^sub>t)\<close> j_k this
have "term_pp_rel (\<prec>\<^sub>t) (lt (seq k), ?lk) x" by (rule term_pp_rel_trans)
ultimately show ?thesis by simp
qed
with \<open>i \<in> red_set (spair (seq j) (seq k))\<close> have "i \<in> ?B" using R'I by blast
thus ?thesis ..
qed
qed
moreover have "finite (?B \<union> ?C)"
proof (rule finite_UnI)
have "finite R'" unfolding R'_def
proof (rule finite_UN_I)
from \<open>finite ?A\<close> show "finite (?A \<inter> {z. term_pp_rel (\<prec>\<^sub>t) z x})" by simp
next
fix y
assume "y \<in> ?A \<inter> {z. term_pp_rel (\<prec>\<^sub>t) z x}"
hence "y \<in> ?A" and "term_pp_rel (\<prec>\<^sub>t) y x" by simp_all
thus "finite (R y)" by (rule less(2))
qed
show "finite ?B" by (intro finite_UN_I \<open>finite R'\<close> finite_red_set)
next
show "finite ?C" by (intro finite_UN_I finite_atLeastLessThan finite_red_set)
qed
ultimately show ?case by (rule finite_subset)
qed
qed fact
ultimately show ?thesis ..
qed
subsubsection \<open>Concrete Rewrite Orders\<close>
definition is_strict_rewrite_ord :: "(('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool) \<Rightarrow> bool"
where "is_strict_rewrite_ord rel \<longleftrightarrow> is_rewrite_ord (\<lambda>x y. \<not> rel y x)"
lemma is_strict_rewrite_ordI: "is_rewrite_ord (\<lambda>x y. \<not> rel y x) \<Longrightarrow> is_strict_rewrite_ord rel"
unfolding is_strict_rewrite_ord_def by blast
lemma is_strict_rewrite_ordD: "is_strict_rewrite_ord rel \<Longrightarrow> is_rewrite_ord (\<lambda>x y. \<not> rel y x)"
unfolding is_strict_rewrite_ord_def by blast
lemma is_strict_rewrite_ord_antisym:
assumes "is_strict_rewrite_ord rel" and "\<not> rel x y" and "\<not> rel y x"
shows "fst x = fst y"
by (rule is_rewrite_ordD4, rule is_strict_rewrite_ordD, fact+)
lemma is_strict_rewrite_ord_asym:
assumes "is_strict_rewrite_ord rel" and "rel x y"
shows "\<not> rel y x"
proof -
from assms(1) have "is_rewrite_ord (\<lambda>x y. \<not> rel y x)" by (rule is_strict_rewrite_ordD)
thus ?thesis
proof (rule is_rewrite_ordD3)
assume "\<not> \<not> rel y x"
assume "\<not> rel x y"
thus ?thesis using \<open>rel x y\<close> ..
qed
qed
lemma is_strict_rewrite_ord_irrefl: "is_strict_rewrite_ord rel \<Longrightarrow> \<not> rel x x"
using is_strict_rewrite_ord_asym by blast
definition rw_rat :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool"
where "rw_rat p q \<longleftrightarrow> (let u = punit.lt (snd q) \<oplus> fst p; v = punit.lt (snd p) \<oplus> fst q in
u \<prec>\<^sub>t v \<or> (u = v \<and> fst p \<preceq>\<^sub>t fst q))"
definition rw_rat_strict :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool"
where "rw_rat_strict p q \<longleftrightarrow> (let u = punit.lt (snd q) \<oplus> fst p; v = punit.lt (snd p) \<oplus> fst q in
u \<prec>\<^sub>t v \<or> (u = v \<and> fst p \<prec>\<^sub>t fst q))"
definition rw_add :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool"
where "rw_add p q \<longleftrightarrow> (fst p \<preceq>\<^sub>t fst q)"
definition rw_add_strict :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool"
where "rw_add_strict p q \<longleftrightarrow> (fst p \<prec>\<^sub>t fst q)"
lemma rw_rat_alt: "rw_rat = (\<lambda>p q. \<not> rw_rat_strict q p)"
by (intro ext, auto simp: rw_rat_def rw_rat_strict_def Let_def)
lemma rw_rat_is_rewrite_ord: "is_rewrite_ord rw_rat"
proof (rule is_rewrite_ordI)
show "reflp rw_rat" by (simp add: reflp_def rw_rat_def)
next
have 1: "ord_term_lin.is_le_rel (\<prec>\<^sub>t)" and 2: "ord_term_lin.is_le_rel (=)"
by (rule ord_term_lin.is_le_relI)+
have "rw_rat p q \<longleftrightarrow> (term_pp_rel (\<prec>\<^sub>t) (fst p, punit.lt (snd p)) (fst q, punit.lt (snd q)) \<or>
(term_pp_rel (=) (fst p, punit.lt (snd p)) (fst q, punit.lt (snd q)) \<and>
fst p \<preceq>\<^sub>t fst q))"
for p q
by (simp add: rw_rat_def term_pp_rel_def Let_def)
thus "transp rw_rat"
by (auto simp: transp_def dest: term_pp_rel_trans[OF 1] term_pp_rel_trans_eq_left[OF 1]
term_pp_rel_trans_eq_right[OF 1] term_pp_rel_trans[OF 2])
next
fix p q
show "rw_rat p q \<or> rw_rat q p" by (auto simp: rw_rat_def Let_def)
next
fix p q
assume "rw_rat p q" and "rw_rat q p"
thus "fst p = fst q" by (auto simp: rw_rat_def Let_def)
next
fix d G p q
assume d: "dickson_grading d" and gb: "is_sig_GB_upt d G (lt q)" and "p \<in> G" and "q \<in> G"
and "p \<noteq> 0" and "q \<noteq> 0" and "lt p adds\<^sub>t lt q" and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G q"
let ?u = "punit.lt (rep_list q) \<oplus> lt p"
let ?v = "punit.lt (rep_list p) \<oplus> lt q"
from \<open>lt p adds\<^sub>t lt q\<close> obtain t where lt_q: "lt q = t \<oplus> lt p" by (rule adds_termE)
from gb have "G \<subseteq> dgrad_sig_set d" by (rule is_sig_GB_uptD1)
hence "G \<subseteq> dgrad_max_set d" by (simp add: dgrad_sig_set'_def)
with d obtain p' where red: "(sig_red (\<prec>\<^sub>t) (=) G)\<^sup>*\<^sup>* (monom_mult 1 t p) p'"
and "\<not> is_sig_red (\<prec>\<^sub>t) (=) G p'" by (rule sig_irredE_dgrad_max_set)
from red have "lt p' = lt (monom_mult 1 t p)" and "lc p' = lc (monom_mult 1 t p)"
and 2: "punit.lt (rep_list p') \<preceq> punit.lt (rep_list (monom_mult 1 t p))"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc, rule sig_red_rtrancl_lt_rep_list)
with \<open>p \<noteq> 0\<close> have "lt p' = lt q" and "lc p' = lc p" by (simp_all add: lt_q lt_monom_mult)
from 2 punit.lt_monom_mult_le[simplified] have 3: "punit.lt (rep_list p') \<preceq> t + punit.lt (rep_list p)"
unfolding rep_list_monom_mult by (rule ordered_powerprod_lin.order_trans)
have "punit.lt (rep_list p') = punit.lt (rep_list q)"
proof (rule sig_regular_top_reduced_lt_unique)
show "p' \<in> dgrad_sig_set d"
proof (rule dgrad_sig_set_closed_sig_red_rtrancl)
note d
moreover have "d t \<le> dgrad_max d"
proof (rule le_trans)
have "t adds lp q" by (simp add: lt_q term_simps)
with d show "d t \<le> d (lp q)" by (rule dickson_grading_adds_imp_le)
next
from \<open>q \<in> G\<close> \<open>G \<subseteq> dgrad_max_set d\<close> have "q \<in> dgrad_max_set d" ..
thus "d (lp q) \<le> dgrad_max d" using \<open>q \<noteq> 0\<close> by (rule dgrad_p_setD_lp)
qed
moreover from \<open>p \<in> G\<close> \<open>G \<subseteq> dgrad_sig_set d\<close> have "p \<in> dgrad_sig_set d" ..
ultimately show "monom_mult 1 t p \<in> dgrad_sig_set d" by (rule dgrad_sig_set_closed_monom_mult)
qed fact+
next
from \<open>q \<in> G\<close> \<open>G \<subseteq> dgrad_sig_set d\<close> show "q \<in> dgrad_sig_set d" ..
next
from \<open>p \<noteq> 0\<close> \<open>lc p' = lc p\<close> show "p' \<noteq> 0" by (auto simp: lc_eq_zero_iff)
qed fact+
with 3 have "punit.lt (rep_list q) \<preceq> t + punit.lt (rep_list p)" by simp
hence "?u \<preceq>\<^sub>t (t + punit.lt (rep_list p)) \<oplus> lt p" by (rule splus_mono_left)
also have "... = ?v" by (simp add: lt_q splus_assoc splus_left_commute)
finally have "?u \<preceq>\<^sub>t ?v" by (simp only: rel_def)
moreover from \<open>lt p adds\<^sub>t lt q\<close> have "lt p \<preceq>\<^sub>t lt q" by (rule ord_adds_term)
ultimately show "rw_rat (spp_of p) (spp_of q)" by (auto simp: rw_rat_def Let_def spp_of_def)
qed
lemma rw_rat_strict_is_strict_rewrite_ord: "is_strict_rewrite_ord rw_rat_strict"
proof (rule is_strict_rewrite_ordI)
show "is_rewrite_ord (\<lambda>x y. \<not> rw_rat_strict y x)"
unfolding rw_rat_alt[symmetric] by (fact rw_rat_is_rewrite_ord)
qed
lemma rw_add_alt: "rw_add = (\<lambda>p q. \<not> rw_add_strict q p)"
by (intro ext, auto simp: rw_add_def rw_add_strict_def)
lemma rw_add_is_rewrite_ord: "is_rewrite_ord rw_add"
proof (rule is_rewrite_ordI)
show "reflp rw_add" by (simp add: reflp_def rw_add_def)
next
show "transp rw_add" by (auto simp: transp_def rw_add_def)
next
fix p q
show "rw_add p q \<or> rw_add q p" by (simp only: rw_add_def ord_term_lin.linear)
next
fix p q
assume "rw_add p q" and "rw_add q p"
- thus "fst p = fst q" unfolding rw_add_def by (rule ord_term_lin.antisym)
+ thus "fst p = fst q" unfolding rw_add_def
+ by simp
next
fix p q :: "'t \<Rightarrow>\<^sub>0 'b"
assume "lt p adds\<^sub>t lt q"
thus "rw_add (spp_of p) (spp_of q)" unfolding rw_add_def spp_of_def fst_conv by (rule ord_adds_term)
qed
lemma rw_add_strict_is_strict_rewrite_ord: "is_strict_rewrite_ord rw_add_strict"
proof (rule is_strict_rewrite_ordI)
show "is_rewrite_ord (\<lambda>x y. \<not> rw_add_strict y x)"
unfolding rw_add_alt[symmetric] by (fact rw_add_is_rewrite_ord)
qed
subsubsection \<open>Preparations for Sig-Poly-Pairs\<close>
context
fixes dgrad :: "'a \<Rightarrow> nat"
begin
definition spp_rel :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
where "spp_rel sp r \<longleftrightarrow> (r \<noteq> 0 \<and> r \<in> dgrad_sig_set dgrad \<and> lt r = fst sp \<and> rep_list r = snd sp)"
definition spp_inv :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool"
where "spp_inv sp \<longleftrightarrow> Ex (spp_rel sp)"
definition vec_of :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b)"
where "vec_of sp = (if spp_inv sp then Eps (spp_rel sp) else 0)"
lemma spp_inv_spp_of:
assumes "r \<noteq> 0" and "r \<in> dgrad_sig_set dgrad"
shows "spp_inv (spp_of r)"
unfolding spp_inv_def spp_rel_def
proof (intro exI conjI)
show "lt r = fst (spp_of r)" by (simp add: spp_of_def)
next
show "rep_list r = snd (spp_of r)" by (simp add: spp_of_def)
qed fact+
context
fixes sp :: "'t \<times> ('a \<Rightarrow>\<^sub>0 'b)"
assumes spi: "spp_inv sp"
begin
lemma sig_poly_rel_vec_of: "spp_rel sp (vec_of sp)"
proof -
from spi have eq: "vec_of sp = Eps (spp_rel sp)" by (simp add: vec_of_def)
from spi show ?thesis unfolding eq spp_inv_def by (rule someI_ex)
qed
lemma vec_of_nonzero: "vec_of sp \<noteq> 0"
using sig_poly_rel_vec_of by (simp add: spp_rel_def)
lemma lt_vec_of: "lt (vec_of sp) = fst sp"
using sig_poly_rel_vec_of by (simp add: spp_rel_def)
lemma rep_list_vec_of: "rep_list (vec_of sp) = snd sp"
using sig_poly_rel_vec_of by (simp add: spp_rel_def)
lemma spp_of_vec_of: "spp_of (vec_of sp) = sp"
by (simp add: spp_of_def lt_vec_of rep_list_vec_of)
end
lemma map_spp_of_vec_of:
assumes "list_all spp_inv sps"
shows "map (spp_of \<circ> vec_of) sps = sps"
proof (rule map_idI)
fix sp
assume "sp \<in> set sps"
with assms have "spp_inv sp" by (simp add: list_all_def)
hence "spp_of (vec_of sp) = sp" by (rule spp_of_vec_of)
thus "(spp_of \<circ> vec_of) sp = sp" by simp
qed
lemma vec_of_dgrad_sig_set: "vec_of sp \<in> dgrad_sig_set dgrad"
proof (cases "spp_inv sp")
case True
hence "spp_rel sp (vec_of sp)" by (rule sig_poly_rel_vec_of)
thus ?thesis by (simp add: spp_rel_def)
next
case False
moreover have "0 \<in> dgrad_sig_set dgrad" unfolding dgrad_sig_set'_def
proof
show "0 \<in> dgrad_max_set dgrad" by (rule dgrad_p_setI) simp
next
show "0 \<in> sig_inv_set" by (rule sig_inv_setI) (simp add: term_simps)
qed
ultimately show ?thesis by (simp add: vec_of_def)
qed
lemma spp_invD_fst:
assumes "spp_inv sp"
shows "dgrad (pp_of_term (fst sp)) \<le> dgrad_max dgrad" and "component_of_term (fst sp) < length fs"
proof -
from vec_of_dgrad_sig_set have "dgrad (lp (vec_of sp)) \<le> dgrad_max dgrad" by (rule dgrad_sig_setD_lp)
with assms show "dgrad (pp_of_term (fst sp)) \<le> dgrad_max dgrad" by (simp add: lt_vec_of)
from vec_of_dgrad_sig_set vec_of_nonzero[OF assms] have "component_of_term (lt (vec_of sp)) < length fs"
by (rule dgrad_sig_setD_lt)
with assms show "component_of_term (fst sp) < length fs" by (simp add: lt_vec_of)
qed
lemma spp_invD_snd:
assumes "dickson_grading dgrad" and "spp_inv sp"
shows "snd sp \<in> punit_dgrad_max_set dgrad"
proof -
from vec_of_dgrad_sig_set[of sp] have "vec_of sp \<in> dgrad_max_set dgrad" by (simp add: dgrad_sig_set'_def)
with assms(1) have "rep_list (vec_of sp) \<in> punit_dgrad_max_set dgrad" by (rule dgrad_max_2)
with assms(2) show ?thesis by (simp add: rep_list_vec_of)
qed
lemma vec_of_inj:
assumes "spp_inv sp" and "vec_of sp = vec_of sp'"
shows "sp = sp'"
proof -
from assms(1) have "vec_of sp \<noteq> 0" by (rule vec_of_nonzero)
hence "vec_of sp' \<noteq> 0" by (simp add: assms(2))
hence "spp_inv sp'" by (simp add: vec_of_def split: if_split_asm)
from assms(1) have "sp = spp_of (vec_of sp)" by (simp only: spp_of_vec_of)
also have "... = spp_of (vec_of sp')" by (simp only: assms(2))
also from \<open>spp_inv sp'\<close> have "... = sp'" by (rule spp_of_vec_of)
finally show ?thesis .
qed
lemma spp_inv_alt: "spp_inv sp \<longleftrightarrow> (vec_of sp \<noteq> 0)"
proof -
have "spp_inv sp" if "vec_of sp \<noteq> 0"
proof (rule ccontr)
assume "\<not> spp_inv sp"
hence "vec_of sp = 0" by (simp add: vec_of_def)
with that show False ..
qed
thus ?thesis by (auto dest: vec_of_nonzero)
qed
lemma spp_of_vec_of_spp_of:
assumes "p \<in> dgrad_sig_set dgrad"
shows "spp_of (vec_of (spp_of p)) = spp_of p"
proof (cases "p = 0")
case True
show ?thesis
proof (cases "spp_inv (spp_of p)")
case True
thus ?thesis by (rule spp_of_vec_of)
next
case False
hence "vec_of (spp_of p) = 0" by (simp add: spp_inv_alt)
thus ?thesis by (simp only: True)
qed
next
case False
have "spp_inv (spp_of p)" unfolding spp_inv_def
proof
from False assms show "spp_rel (spp_of p) p" by (simp add: spp_rel_def spp_of_def)
qed
thus ?thesis by (rule spp_of_vec_of)
qed
subsubsection \<open>Total Reduction\<close>
primrec find_sig_reducer :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<Rightarrow> 't \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> nat option" where
"find_sig_reducer [] _ _ _ = None"|
"find_sig_reducer (b # bs) u t i =
(if snd b \<noteq> 0 \<and> punit.lt (snd b) adds t \<and> (t - punit.lt (snd b)) \<oplus> fst b \<prec>\<^sub>t u then Some i
else find_sig_reducer bs u t (Suc i))"
lemma find_sig_reducer_SomeD_aux:
assumes "find_sig_reducer bs u t i = Some j"
shows "i \<le> j" and "j - i < length bs"
proof -
from assms have "i \<le> j \<and> j - i < length bs"
proof (induct bs arbitrary: i)
case Nil
thus ?case by simp
next
case (Cons b bs)
from Cons(2) show ?case
proof (simp split: if_split_asm)
assume "find_sig_reducer bs u t (Suc i) = Some j"
hence "Suc i \<le> j \<and> j - Suc i < length bs" by (rule Cons(1))
thus "i \<le> j \<and> j - i < Suc (length bs)" by auto
qed
qed
thus "i \<le> j" and "j - i < length bs" by simp_all
qed
lemma find_sig_reducer_SomeD':
assumes "find_sig_reducer bs u t i = Some j" and "b = bs ! (j - i)"
shows "b \<in> set bs" and "snd b \<noteq> 0" and "punit.lt (snd b) adds t" and "(t - punit.lt (snd b)) \<oplus> fst b \<prec>\<^sub>t u"
proof -
from assms(1) have "j - i < length bs" by (rule find_sig_reducer_SomeD_aux)
thus "b \<in> set bs" unfolding assms(2) by (rule nth_mem)
next
from assms have "snd b \<noteq> 0 \<and> punit.lt (snd b) adds t \<and> (t - punit.lt (snd b)) \<oplus> fst b \<prec>\<^sub>t u"
proof (induct bs arbitrary: i)
case Nil
from Nil(1) show ?case by simp
next
case (Cons a bs)
from Cons(2) show ?case
proof (simp split: if_split_asm)
assume "i = j"
with Cons(3) have "b = a" by simp
moreover assume "snd a \<noteq> 0" and "punit.lt (snd a) adds t" and "(t - punit.lt (snd a)) \<oplus> fst a \<prec>\<^sub>t u"
ultimately show ?case by simp
next
assume *: "find_sig_reducer bs u t (Suc i) = Some j"
hence "Suc i \<le> j" by (rule find_sig_reducer_SomeD_aux)
note Cons(3)
also from \<open>Suc i \<le> j\<close> have "(a # bs) ! (j - i) = bs ! (j - Suc i)" by simp
finally have "b = bs ! (j - Suc i)" .
with * show ?case by (rule Cons(1))
qed
qed
thus "snd b \<noteq> 0" and "punit.lt (snd b) adds t" and "(t - punit.lt (snd b)) \<oplus> fst b \<prec>\<^sub>t u" by simp_all
qed
corollary find_sig_reducer_SomeD:
assumes "find_sig_reducer (map spp_of bs) u t 0 = Some i"
shows "i < length bs" and "rep_list (bs ! i) \<noteq> 0" and "punit.lt (rep_list (bs ! i)) adds t"
and "(t - punit.lt (rep_list (bs ! i))) \<oplus> lt (bs ! i) \<prec>\<^sub>t u"
proof -
from assms have "i - 0 < length (map spp_of bs)" by (rule find_sig_reducer_SomeD_aux)
thus "i < length bs" by simp
hence "spp_of (bs ! i) = (map spp_of bs) ! (i - 0)" by simp
with assms have "snd (spp_of (bs ! i)) \<noteq> 0" and "punit.lt (snd (spp_of (bs ! i))) adds t"
and "(t - punit.lt (snd (spp_of (bs ! i)))) \<oplus> fst (spp_of (bs ! i)) \<prec>\<^sub>t u"
by (rule find_sig_reducer_SomeD')+
thus "rep_list (bs ! i) \<noteq> 0" and "punit.lt (rep_list (bs ! i)) adds t"
and "(t - punit.lt (rep_list (bs ! i))) \<oplus> lt (bs ! i) \<prec>\<^sub>t u" by (simp_all add: fst_spp_of snd_spp_of)
qed
lemma find_sig_reducer_NoneE:
assumes "find_sig_reducer bs u t i = None" and "b \<in> set bs"
assumes "snd b = 0 \<Longrightarrow> thesis" and "snd b \<noteq> 0 \<Longrightarrow> \<not> punit.lt (snd b) adds t \<Longrightarrow> thesis"
and "snd b \<noteq> 0 \<Longrightarrow> punit.lt (snd b) adds t \<Longrightarrow> \<not> (t - punit.lt (snd b)) \<oplus> fst b \<prec>\<^sub>t u \<Longrightarrow> thesis"
shows thesis
using assms
proof (induct bs arbitrary: thesis i)
case Nil
from Nil(2) show ?case by simp
next
case (Cons a bs)
from Cons(2) have 1: "snd a = 0 \<or> \<not> punit.lt (snd a) adds t \<or> \<not> (t - punit.lt (snd a)) \<oplus> fst a \<prec>\<^sub>t u"
and eq: "find_sig_reducer bs u t (Suc i) = None" by (simp_all split: if_splits)
from Cons(3) have "b = a \<or> b \<in> set bs" by simp
thus ?case
proof
assume "b = a"
show ?thesis
proof (cases "snd a = 0")
case True
show ?thesis by (rule Cons(4), simp add: \<open>b = a\<close> True)
next
case False
with 1 have 2: "\<not> punit.lt (snd a) adds t \<or> \<not> (t - punit.lt (snd a)) \<oplus> fst a \<prec>\<^sub>t u" by simp
show ?thesis
proof (cases "punit.lt (snd a) adds t")
case True
with 2 have 3: "\<not> (t - punit.lt (snd a)) \<oplus> fst a \<prec>\<^sub>t u" by simp
show ?thesis by (rule Cons(6), simp_all add: \<open>b = a\<close> \<open>snd a \<noteq> 0\<close> True 3)
next
case False
show ?thesis by (rule Cons(5), simp_all add: \<open>b = a\<close> \<open>snd a \<noteq> 0\<close> False)
qed
qed
next
assume "b \<in> set bs"
with eq show ?thesis
proof (rule Cons(1))
assume "snd b = 0"
thus ?thesis by (rule Cons(4))
next
assume "snd b \<noteq> 0" and "\<not> punit.lt (snd b) adds t"
thus ?thesis by (rule Cons(5))
next
assume "snd b \<noteq> 0" and "punit.lt (snd b) adds t" and "\<not> (t - punit.lt (snd b)) \<oplus> fst b \<prec>\<^sub>t u"
thus ?thesis by (rule Cons(6))
qed
qed
qed
lemma find_sig_reducer_SomeD_red_single:
assumes "t \<in> keys (rep_list p)" and "find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i"
shows "sig_red_single (\<prec>\<^sub>t) (\<preceq>) p (p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i)) (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
proof -
from assms(2) have "punit.lt (rep_list (bs ! i)) adds t" and 1: "rep_list (bs ! i) \<noteq> 0"
and 2: "(t - punit.lt (rep_list (bs ! i))) \<oplus> lt (bs ! i) \<prec>\<^sub>t lt p"
by (rule find_sig_reducer_SomeD)+
from this(1) have eq: "t - punit.lt (rep_list (bs ! i)) + punit.lt (rep_list (bs ! i)) = t"
by (rule adds_minus)
from assms(1) have 3: "t \<preceq> punit.lt (rep_list p)" by (rule punit.lt_max_keys)
show ?thesis by (rule sig_red_singleI, simp_all add: eq 1 2 3 assms(1))
qed
corollary find_sig_reducer_SomeD_red:
assumes "t \<in> keys (rep_list p)" and "find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i"
shows "sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p (p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i))"
unfolding sig_red_def
proof (intro bexI exI, rule find_sig_reducer_SomeD_red_single)
from assms(2) have "i - 0 < length (map spp_of bs)" by (rule find_sig_reducer_SomeD_aux)
hence "i < length bs" by simp
thus "bs ! i \<in> set bs" by (rule nth_mem)
qed fact+
context
fixes bs :: "('t \<Rightarrow>\<^sub>0 'b) list"
begin
definition sig_trd_term :: "('a \<Rightarrow> nat) \<Rightarrow> (('a \<times> ('t \<Rightarrow>\<^sub>0 'b)) \<times> ('a \<times> ('t \<Rightarrow>\<^sub>0 'b))) set"
where "sig_trd_term d = {(x, y). punit.dgrad_p_set_le d {rep_list (snd x)}
(insert (rep_list (snd y)) (rep_list ` set bs)) \<and>
fst x \<in> keys (rep_list (snd x)) \<and> fst y \<in> keys (rep_list (snd y)) \<and>
fst x \<prec> fst y}"
lemma sig_trd_term_wf:
assumes "dickson_grading d"
shows "wf (sig_trd_term d)"
proof (rule wfI_min)
fix x :: "'a \<times> ('t \<Rightarrow>\<^sub>0 'b)" and Q
assume "x \<in> Q"
show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> sig_trd_term d \<longrightarrow> y \<notin> Q"
proof (cases "fst x \<in> keys (rep_list (snd x))")
case True
define X where "X = rep_list ` set bs"
let ?A = "insert (rep_list (snd x)) X"
have "finite X" unfolding X_def by simp
hence "finite ?A" by (simp only: finite_insert)
then obtain m where A: "?A \<subseteq> punit.dgrad_p_set d m" by (rule punit.dgrad_p_set_exhaust)
hence x: "rep_list (snd x) \<in> punit.dgrad_p_set d m" and X: "X \<subseteq> punit.dgrad_p_set d m"
by simp_all
let ?Q = "{q \<in> Q. rep_list (snd q) \<in> punit.dgrad_p_set d m \<and> fst q \<in> keys (rep_list (snd q))}"
from \<open>x \<in> Q\<close> x True have "x \<in> ?Q" by simp
have "\<forall>Q x. x \<in> Q \<and> Q \<subseteq> {q. d q \<le> m} \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. y \<prec> z \<longrightarrow> y \<notin> Q)"
by (rule wfp_on_imp_minimal, rule wfp_on_ord_strict, fact assms)
hence 1: "fst x \<in> fst ` ?Q \<Longrightarrow> fst ` ?Q \<subseteq> {q. d q \<le> m} \<Longrightarrow> (\<exists>z\<in>fst ` ?Q. \<forall>y. y \<prec> z \<longrightarrow> y \<notin> fst ` ?Q)"
by meson
have "fst x \<in> fst ` ?Q" by (rule, fact refl, fact)
moreover have "fst ` ?Q \<subseteq> {q. d q \<le> m}"
proof -
{
fix q
assume a: "rep_list (snd q) \<in> punit.dgrad_p_set d m" and b: "fst q \<in> keys (rep_list (snd q))"
from a have "keys (rep_list (snd q)) \<subseteq> dgrad_set d m" by (simp add: punit.dgrad_p_set_def)
with b have "fst q \<in> dgrad_set d m" ..
hence "d (fst q) \<le> m" by (simp add: dgrad_set_def)
}
thus ?thesis by auto
qed
ultimately have "\<exists>z\<in>fst ` ?Q. \<forall>y. y \<prec> z \<longrightarrow> y \<notin> fst ` ?Q" by (rule 1)
then obtain z0 where "z0 \<in> fst ` ?Q" and 2: "\<And>y. y \<prec> z0 \<Longrightarrow> y \<notin> fst ` ?Q" by blast
from this(1) obtain z where "z \<in> ?Q" and z0: "z0 = fst z" ..
hence "z \<in> Q" and z: "rep_list (snd z) \<in> punit.dgrad_p_set d m" by simp_all
from this(1) show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> sig_trd_term d \<longrightarrow> y \<notin> Q"
proof
show "\<forall>y. (y, z) \<in> sig_trd_term d \<longrightarrow> y \<notin> Q"
proof (intro allI impI)
fix y
assume "(y, z) \<in> sig_trd_term d"
hence 3: "punit.dgrad_p_set_le d {rep_list (snd y)} (insert (rep_list (snd z)) X)"
and 4: "fst y \<in> keys (rep_list (snd y))" and "fst y \<prec> z0"
by (simp_all add: sig_trd_term_def X_def z0)
from this(3) have "fst y \<notin> fst ` ?Q" by (rule 2)
hence "y \<notin> Q \<or> rep_list (snd y) \<notin> punit.dgrad_p_set d m \<or> fst y \<notin> keys (rep_list (snd y))"
by auto
thus "y \<notin> Q"
proof (elim disjE)
assume 5: "rep_list (snd y) \<notin> punit.dgrad_p_set d m"
from z X have "insert (rep_list (snd z)) X \<subseteq> punit.dgrad_p_set d m" by simp
with 3 have "{rep_list (snd y)} \<subseteq> punit.dgrad_p_set d m" by (rule punit.dgrad_p_set_le_dgrad_p_set)
hence "rep_list (snd y) \<in> punit.dgrad_p_set d m" by simp
with 5 show ?thesis ..
next
assume "fst y \<notin> keys (rep_list (snd y))"
thus ?thesis using 4 ..
qed
qed
qed
next
case False
from \<open>x \<in> Q\<close> show ?thesis
proof
show "\<forall>y. (y, x) \<in> sig_trd_term d \<longrightarrow> y \<notin> Q"
proof (intro allI impI)
fix y
assume "(y, x) \<in> sig_trd_term d"
hence "fst x \<in> keys (rep_list (snd x))" by (simp add: sig_trd_term_def)
with False show "y \<notin> Q" ..
qed
qed
qed
qed
function (domintros) sig_trd_aux :: "('a \<times> ('t \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b)" where
"sig_trd_aux (t, p) =
(let p' =
(case find_sig_reducer (map spp_of bs) (lt p) t 0 of
None \<Rightarrow> p
| Some i \<Rightarrow> p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i));
p'' = punit.lower (rep_list p') t in
if p'' = 0 then p' else sig_trd_aux (punit.lt p'', p'))"
by auto
lemma sig_trd_aux_domI:
assumes "fst args0 \<in> keys (rep_list (snd args0))"
shows "sig_trd_aux_dom args0"
proof -
from ex_hgrad obtain d::"'a \<Rightarrow> nat" where "dickson_grading d \<and> hom_grading d" ..
hence dg: "dickson_grading d" ..
hence "wf (sig_trd_term d)" by (rule sig_trd_term_wf)
thus ?thesis using assms
proof (induct args0)
case (less args)
obtain t p where args: "args = (t, p)" using prod.exhaust by blast
with less(1) have 1: "\<And>s q. ((s, q), (t, p)) \<in> sig_trd_term d \<Longrightarrow> s \<in> keys (rep_list q) \<Longrightarrow> sig_trd_aux_dom (s, q)"
using prod.exhaust by auto
from less(2) have "t \<in> keys (rep_list p)" by (simp add: args)
show ?case unfolding args
proof (rule sig_trd_aux.domintros)
define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
None \<Rightarrow> p
| Some i \<Rightarrow> p -
monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i))"
define p'' where "p'' = punit.lower (rep_list p') t"
assume "p'' \<noteq> 0"
from \<open>p'' \<noteq> 0\<close> have "punit.lt p'' \<in> keys p''" by (rule punit.lt_in_keys)
also have "... \<subseteq> keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
finally have "punit.lt p'' \<in> keys (rep_list p')" .
with _ show "sig_trd_aux_dom (punit.lt p'', p')"
proof (rule 1)
have "punit.dgrad_p_set_le d {rep_list p'} (insert (rep_list p) (rep_list ` set bs))"
proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
case None
hence "p' = p" by (simp add: p'_def)
hence "{rep_list p'} \<subseteq> insert (rep_list p) (rep_list ` set bs)" by simp
thus ?thesis by (rule punit.dgrad_p_set_le_subset)
next
case (Some i)
hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
have "sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p p'" unfolding p' using \<open>t \<in> keys (rep_list p)\<close> Some
by (rule find_sig_reducer_SomeD_red)
hence "punit.red (rep_list ` set bs) (rep_list p) (rep_list p')" by (rule sig_red_red)
with dg show ?thesis by (rule punit.dgrad_p_set_le_red)
qed
moreover note \<open>punit.lt p'' \<in> keys (rep_list p')\<close> \<open>t \<in> keys (rep_list p)\<close>
moreover from \<open>p'' \<noteq> 0\<close> have "punit.lt p'' \<prec> t" unfolding p''_def by (rule punit.lt_lower_less)
ultimately show "((punit.lt p'', p'), t, p) \<in> sig_trd_term d" by (simp add: sig_trd_term_def)
qed
qed
qed
qed
definition sig_trd :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b)"
where "sig_trd p = (if rep_list p = 0 then p else sig_trd_aux (punit.lt (rep_list p), p))"
lemma sig_trd_aux_red_rtrancl:
assumes "fst args0 \<in> keys (rep_list (snd args0))"
shows "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (snd args0) (sig_trd_aux args0)"
proof -
from assms have "sig_trd_aux_dom args0" by (rule sig_trd_aux_domI)
thus ?thesis using assms
proof (induct args0 rule: sig_trd_aux.pinduct)
case (1 t p)
define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
None \<Rightarrow> p
| Some i \<Rightarrow> p -
monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i))"
define p'' where "p'' = punit.lower (rep_list p') t"
from 1(3) have "t \<in> keys (rep_list p)" by simp
have *: "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* p p'"
proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
case None
hence "p' = p" by (simp add: p'_def)
thus ?thesis by simp
next
case (Some i)
hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
have "sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p p'" unfolding p' using \<open>t \<in> keys (rep_list p)\<close> Some
by (rule find_sig_reducer_SomeD_red)
thus ?thesis ..
qed
show ?case
proof (simp add: sig_trd_aux.psimps[OF 1(1)] Let_def p'_def[symmetric] p''_def[symmetric] *, intro impI)
assume "p'' \<noteq> 0"
from * have "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* p (snd (punit.lt p'', p'))" by (simp only: snd_conv)
moreover have "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (snd (punit.lt p'', p')) (sig_trd_aux (punit.lt p'', p'))"
using p'_def p''_def \<open>p'' \<noteq> 0\<close>
proof (rule 1(2))
from \<open>p'' \<noteq> 0\<close> have "punit.lt p'' \<in> keys p''" by (rule punit.lt_in_keys)
also have "... \<subseteq> keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
finally show "fst (punit.lt p'', p') \<in> keys (rep_list (snd (punit.lt p'', p')))" by simp
qed
ultimately show "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* p (sig_trd_aux (punit.lt p'', p'))"
by (rule rtranclp_trans)
qed
qed
qed
corollary sig_trd_red_rtrancl: "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* p (sig_trd p)"
unfolding sig_trd_def
proof (split if_split, intro conjI impI rtranclp.rtrancl_refl)
let ?args = "(punit.lt (rep_list p), p)"
assume "rep_list p \<noteq> 0"
hence "punit.lt (rep_list p) \<in> keys (rep_list p)" by (rule punit.lt_in_keys)
hence "fst (punit.lt (rep_list p), p) \<in> keys (rep_list (snd (punit.lt (rep_list p), p)))"
by (simp only: fst_conv snd_conv)
hence "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (snd ?args) (sig_trd_aux ?args)" by (rule sig_trd_aux_red_rtrancl)
thus "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* p (sig_trd_aux (punit.lt (rep_list p), p))" by (simp only: snd_conv)
qed
lemma sig_trd_aux_irred:
assumes "fst args0 \<in> keys (rep_list (snd args0))"
and "\<And>b s. b \<in> set bs \<Longrightarrow> rep_list b \<noteq> 0 \<Longrightarrow> fst args0 \<prec> s + punit.lt (rep_list b) \<Longrightarrow>
s \<oplus> lt b \<prec>\<^sub>t lt (snd (args0)) \<Longrightarrow> lookup (rep_list (snd args0)) (s + punit.lt (rep_list b)) = 0"
shows "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) (sig_trd_aux args0)"
proof -
from assms(1) have "sig_trd_aux_dom args0" by (rule sig_trd_aux_domI)
thus ?thesis using assms
proof (induct args0 rule: sig_trd_aux.pinduct)
case (1 t p)
define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
None \<Rightarrow> p
| Some i \<Rightarrow> p -
monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i))"
define p'' where "p'' = punit.lower (rep_list p') t"
from 1(3) have "t \<in> keys (rep_list p)" by simp
from 1(4) have a: "b \<in> set bs \<Longrightarrow> rep_list b \<noteq> 0 \<Longrightarrow> t \<prec> s + punit.lt (rep_list b) \<Longrightarrow>
s \<oplus> lt b \<prec>\<^sub>t lt p \<Longrightarrow> lookup (rep_list p) (s + punit.lt (rep_list b)) = 0"
for b s by (simp only: fst_conv snd_conv)
have "lt p' = lt p \<and> (\<forall>s. t \<prec> s \<longrightarrow> lookup (rep_list p') s = lookup (rep_list p) s)"
proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
case None
thus ?thesis by (simp add: p'_def)
next
case (Some i)
hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
have "sig_red_single (\<prec>\<^sub>t) (\<preceq>) p p' (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
unfolding p' using \<open>t \<in> keys (rep_list p)\<close> Some by (rule find_sig_reducer_SomeD_red_single)
hence r: "punit.red_single (rep_list p) (rep_list p') (rep_list (bs ! i)) (t - punit.lt (rep_list (bs ! i)))"
and "lt p' = lt p" by (rule sig_red_single_red_single, rule sig_red_single_regular_lt)
have "\<forall>s. t \<prec> s \<longrightarrow> lookup (rep_list p') s = lookup (rep_list p) s"
proof (intro allI impI)
fix s
assume "t \<prec> s"
from Some have "punit.lt (rep_list (bs ! i)) adds t" by (rule find_sig_reducer_SomeD)
hence eq0: "(t - punit.lt (rep_list (bs ! i))) + punit.lt (rep_list (bs ! i)) = t" (is "?t = t")
by (rule adds_minus)
from \<open>t \<prec> s\<close> have "lookup (rep_list p') s = lookup (punit.higher (rep_list p') ?t) s"
by (simp add: eq0 punit.lookup_higher_when)
also from r have "... = lookup (punit.higher (rep_list p) ?t) s"
by (simp add: punit.red_single_higher[simplified])
also from \<open>t \<prec> s\<close> have "... = lookup (rep_list p) s" by (simp add: eq0 punit.lookup_higher_when)
finally show "lookup (rep_list p') s = lookup (rep_list p) s" .
qed
with \<open>lt p' = lt p\<close> show ?thesis ..
qed
hence lt_p': "lt p' = lt p" and b: "\<And>s. t \<prec> s \<Longrightarrow> lookup (rep_list p') s = lookup (rep_list p) s"
by blast+
have c: "lookup (rep_list p') (s + punit.lt (rep_list b)) = 0"
if "b \<in> set bs" and "rep_list b \<noteq> 0" and "t \<preceq> s + punit.lt (rep_list b)" and "s \<oplus> lt b \<prec>\<^sub>t lt p'" for b s
proof (cases "t \<prec> s + punit.lt (rep_list b)")
case True
hence "lookup (rep_list p') (s + punit.lt (rep_list b)) =
lookup (rep_list p) (s + punit.lt (rep_list b))" by (rule b)
also from that(1, 2) True that(4) have "... = 0" unfolding lt_p' by (rule a)
finally show ?thesis .
next
case False
with that(3) have t: "t = s + punit.lt (rep_list b)" by simp
show ?thesis
proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
case None
from that(1) have "spp_of b \<in> set (map spp_of bs)" by fastforce
with None show ?thesis
proof (rule find_sig_reducer_NoneE)
assume "snd (spp_of b) = 0"
with that(2) show ?thesis by (simp add: snd_spp_of)
next
assume "\<not> punit.lt (snd (spp_of b)) adds t"
thus ?thesis by (simp add: snd_spp_of t)
next
assume "\<not> (t - punit.lt (snd (spp_of b))) \<oplus> fst (spp_of b) \<prec>\<^sub>t lt p"
with that(4) show ?thesis by (simp add: fst_spp_of snd_spp_of t lt_p')
qed
next
case (Some i)
hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
have "sig_red_single (\<prec>\<^sub>t) (\<preceq>) p p' (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
unfolding p' using \<open>t \<in> keys (rep_list p)\<close> Some by (rule find_sig_reducer_SomeD_red_single)
hence r: "punit.red_single (rep_list p) (rep_list p') (rep_list (bs ! i)) (t - punit.lt (rep_list (bs ! i)))"
by (rule sig_red_single_red_single)
from Some have "punit.lt (rep_list (bs ! i)) adds t" by (rule find_sig_reducer_SomeD)
hence eq0: "(t - punit.lt (rep_list (bs ! i))) + punit.lt (rep_list (bs ! i)) = t" (is "?t = t")
by (rule adds_minus)
from r have "lookup (rep_list p') ((t - punit.lt (rep_list (bs ! i))) + punit.lt (rep_list (bs ! i))) = 0"
by (rule punit.red_single_lookup[simplified])
thus ?thesis by (simp only: eq0 t[symmetric])
qed
qed
show ?case
proof (simp add: sig_trd_aux.psimps[OF 1(1)] Let_def p'_def[symmetric] p''_def[symmetric], intro conjI impI)
assume "p'' = 0"
show "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p'"
proof
assume "is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p'"
then obtain b s where "b \<in> set bs" and "s \<in> keys (rep_list p')" and "rep_list b \<noteq> 0"
and adds: "punit.lt (rep_list b) adds s" and "s \<oplus> lt b \<prec>\<^sub>t punit.lt (rep_list b) \<oplus> lt p'"
by (rule is_sig_red_addsE)
let ?s = "s - punit.lt (rep_list b)"
from adds have eq0: "?s + punit.lt (rep_list b) = s" by (simp add: adds_minus)
show False
proof (cases "t \<preceq> s")
case True
note \<open>b \<in> set bs\<close> \<open>rep_list b \<noteq> 0\<close>
moreover from True have "t \<preceq> ?s + punit.lt (rep_list b)" by (simp only: eq0)
moreover from adds \<open>s \<oplus> lt b \<prec>\<^sub>t punit.lt (rep_list b) \<oplus> lt p'\<close> have "?s \<oplus> lt b \<prec>\<^sub>t lt p'"
by (simp add: term_is_le_rel_minus)
ultimately have "lookup (rep_list p') (?s + punit.lt (rep_list b)) = 0" by (rule c)
hence "s \<notin> keys (rep_list p')" by (simp add: eq0 in_keys_iff)
thus ?thesis using \<open>s \<in> keys (rep_list p')\<close> ..
next
case False
hence "s \<prec> t" by simp
hence "lookup (rep_list p') s = lookup (punit.lower (rep_list p') t) s"
by (simp add: punit.lookup_lower_when)
also from \<open>p'' = 0\<close> have "... = 0" by (simp add: p''_def)
finally have "s \<notin> keys (rep_list p')" by (simp add: in_keys_iff)
thus ?thesis using \<open>s \<in> keys (rep_list p')\<close> ..
qed
qed
next
assume "p'' \<noteq> 0"
with p'_def p''_def show "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) (sig_trd_aux (punit.lt p'', p'))"
proof (rule 1(2))
from \<open>p'' \<noteq> 0\<close> have "punit.lt p'' \<in> keys p''" by (rule punit.lt_in_keys)
also have "... \<subseteq> keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
finally show "fst (punit.lt p'', p') \<in> keys (rep_list (snd (punit.lt p'', p')))" by simp
next
fix b s
assume "b \<in> set bs" and "rep_list b \<noteq> 0"
assume "fst (punit.lt p'', p') \<prec> s + punit.lt (rep_list b)"
and "s \<oplus> lt b \<prec>\<^sub>t lt (snd (punit.lt p'', p'))"
hence "punit.lt p'' \<prec> s + punit.lt (rep_list b)" and "s \<oplus> lt b \<prec>\<^sub>t lt p'" by simp_all
have "lookup (rep_list p') (s + punit.lt (rep_list b)) = 0"
proof (cases "t \<preceq> s + punit.lt (rep_list b)")
case True
with \<open>b \<in> set bs\<close> \<open>rep_list b \<noteq> 0\<close> show ?thesis using \<open>s \<oplus> lt b \<prec>\<^sub>t lt p'\<close> by (rule c)
next
case False
hence "s + punit.lt (rep_list b) \<prec> t" by simp
hence "lookup (rep_list p') (s + punit.lt (rep_list b)) =
lookup (punit.lower (rep_list p') t) (s + punit.lt (rep_list b))"
by (simp add: punit.lookup_lower_when)
also have "... = 0"
proof (rule ccontr)
assume "lookup (punit.lower (rep_list p') t) (s + punit.lt (rep_list b)) \<noteq> 0"
hence "s + punit.lt (rep_list b) \<preceq> punit.lt (punit.lower (rep_list p') t)"
by (rule punit.lt_max)
also have "... = punit.lt p''" by (simp only: p''_def)
finally show False using \<open>punit.lt p'' \<prec> s + punit.lt (rep_list b)\<close> by simp
qed
finally show ?thesis .
qed
thus "lookup (rep_list (snd (punit.lt p'', p'))) (s + punit.lt (rep_list b)) = 0"
by (simp only: snd_conv)
qed
qed
qed
qed
corollary sig_trd_irred: "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) (sig_trd p)"
unfolding sig_trd_def
proof (split if_split, intro conjI impI)
assume "rep_list p = 0"
show "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p"
proof
assume "is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p"
then obtain t where "t \<in> keys (rep_list p)" by (rule is_sig_red_addsE)
thus False by (simp add: \<open>rep_list p = 0\<close>)
qed
next
assume "rep_list p \<noteq> 0"
show "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) (sig_trd_aux (punit.lt (rep_list p), p))"
proof (rule sig_trd_aux_irred)
from \<open>rep_list p \<noteq> 0\<close> have "punit.lt (rep_list p) \<in> keys (rep_list p)" by (rule punit.lt_in_keys)
thus "fst (punit.lt (rep_list p), p) \<in> keys (rep_list (snd (punit.lt (rep_list p), p)))" by simp
next
fix b s
assume "fst (punit.lt (rep_list p), p) \<prec> s + punit.lt (rep_list b)"
thus "lookup (rep_list (snd (punit.lt (rep_list p), p))) (s + punit.lt (rep_list b)) = 0"
using punit.lt_max by force
qed
qed
end
context
fixes bs :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list"
begin
context
fixes v :: 't
begin
fun sig_trd_spp_body :: "(('a \<Rightarrow>\<^sub>0 'b) \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> (('a \<Rightarrow>\<^sub>0 'b) \<times> ('a \<Rightarrow>\<^sub>0 'b))" where
"sig_trd_spp_body (p, r) =
(case find_sig_reducer bs v (punit.lt p) 0 of
None \<Rightarrow> (punit.tail p, r + monomial (punit.lc p) (punit.lt p))
| Some i \<Rightarrow> let b = snd (bs ! i) in
(punit.tail p - punit.monom_mult (punit.lc p / punit.lc b) (punit.lt p - punit.lt b) (punit.tail b), r))"
definition sig_trd_spp_aux :: "(('a \<Rightarrow>\<^sub>0 'b) \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b)"
where sig_trd_spp_aux_def [code del]: "sig_trd_spp_aux = tailrec.fun (\<lambda>x. fst x = 0) snd sig_trd_spp_body"
lemma sig_trd_spp_aux_simps [code]:
"sig_trd_spp_aux (p, r) = (if p = 0 then r else sig_trd_spp_aux (sig_trd_spp_body (p, r)))"
by (simp add: sig_trd_spp_aux_def tailrec.simps)
end
fun sig_trd_spp :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))" where
"sig_trd_spp (v, p) = (v, sig_trd_spp_aux v (p, 0))"
text \<open>We define function @{const sig_trd_spp}, operating on sig-poly-pairs, already here, to have
its definition in the right context. Lemmas are proved about it below in Section \<open>Sig-Poly-Pairs\<close>.\<close>
end
subsubsection \<open>Koszul Syzygies\<close>
text \<open>A @{emph \<open>Koszul syzygy\<close>} of the list @{term fs} of scalar polynomials is a syzygy of the form
@{term "(fs ! i) \<odot> (monomial 1 (term_of_pair (0, j))) - (fs ! j) \<odot> (monomial 1 (term_of_pair (0, i)))"},
for @{prop "i < j"} and @{prop "j < length fs"}.\<close>
primrec Koszul_syz_sigs_aux :: "('a \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> nat \<Rightarrow> 't list" where
"Koszul_syz_sigs_aux [] i = []" |
"Koszul_syz_sigs_aux (b # bs) i =
map_idx (\<lambda>b' j. ord_term_lin.max (term_of_pair (punit.lt b, j)) (term_of_pair (punit.lt b', i))) bs (Suc i) @
Koszul_syz_sigs_aux bs (Suc i)"
definition Koszul_syz_sigs :: "('a \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> 't list"
where "Koszul_syz_sigs bs = filter_min (adds\<^sub>t) (Koszul_syz_sigs_aux bs 0)"
fun new_syz_sigs :: "'t list \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> (('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat \<Rightarrow> 't list"
where
"new_syz_sigs ss bs (Inl (a, b)) = ss" |
"new_syz_sigs ss bs (Inr j) =
(if is_pot_ord then
filter_min_append (adds\<^sub>t) ss (filter_min (adds\<^sub>t) (map (\<lambda>b. term_of_pair (punit.lt (rep_list b), j)) bs))
else ss)"
fun new_syz_sigs_spp :: "'t list \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<Rightarrow> (('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat \<Rightarrow> 't list"
where
"new_syz_sigs_spp ss bs (Inl (a, b)) = ss" |
"new_syz_sigs_spp ss bs (Inr j) =
(if is_pot_ord then
filter_min_append (adds\<^sub>t) ss (filter_min (adds\<^sub>t) (map (\<lambda>b. term_of_pair (punit.lt (snd b), j)) bs))
else ss)"
lemma Koszul_syz_sigs_auxI:
assumes "i < j" and "j < length bs"
shows "ord_term_lin.max (term_of_pair (punit.lt (bs ! i), k + j)) (term_of_pair (punit.lt (bs ! j), k + i)) \<in>
set (Koszul_syz_sigs_aux bs k)"
using assms
proof (induct bs arbitrary: i j k)
case Nil
from Nil(2) show ?case by simp
next
case (Cons b bs)
from Cons(2) obtain j0 where j: "j = Suc j0" by (meson lessE)
from Cons(3) have "j0 < length bs" by (simp add: j)
let ?A = "(\<lambda>j. ord_term_lin.max (term_of_pair (punit.lt b, Suc (j + k))) (term_of_pair (punit.lt (bs ! j), k))) `
{0..<length bs}"
let ?B = "set (Koszul_syz_sigs_aux bs (Suc k))"
show ?case
proof (cases i)
case 0
from \<open>j0 < length bs\<close> have "j0 \<in> {0..<length bs}" by simp
hence "ord_term_lin.max (term_of_pair (punit.lt b, Suc (j0 + k)))
(term_of_pair (punit.lt (bs ! j0), k)) \<in> ?A" by (rule imageI)
thus ?thesis by (simp add: \<open>i = 0\<close> j set_map_idx ac_simps)
next
case (Suc i0)
from Cons(2) have "i0 < j0" by (simp add: \<open>i = Suc i0\<close> j)
hence "ord_term_lin.max (term_of_pair (punit.lt (bs ! i0), Suc k + j0))
(term_of_pair (punit.lt (bs ! j0), Suc k + i0)) \<in> ?B"
using \<open>j0 < length bs\<close> by (rule Cons(1))
thus ?thesis by (simp add: \<open>i = Suc i0\<close> j set_map_idx ac_simps)
qed
qed
lemma Koszul_syz_sigs_auxE:
assumes "v \<in> set (Koszul_syz_sigs_aux bs k)"
obtains i j where "i < j" and "j < length bs"
and "v = ord_term_lin.max (term_of_pair (punit.lt (bs ! i), k + j)) (term_of_pair (punit.lt (bs ! j), k + i))"
using assms
proof (induct bs arbitrary: k thesis)
case Nil
from Nil(2) show ?case by simp
next
case (Cons b bs)
have "v \<in> (\<lambda>j. ord_term_lin.max (term_of_pair (punit.lt b, Suc (j + k))) (term_of_pair (punit.lt (bs ! j), k))) `
{0..<length bs} \<union> set (Koszul_syz_sigs_aux bs (Suc k))" (is "v \<in> ?A \<union> ?B")
using Cons(3) by (simp add: set_map_idx)
thus ?case
proof
assume "v \<in> ?A"
then obtain j where "j \<in> {0..<length bs}"
and v: "v = ord_term_lin.max (term_of_pair (punit.lt b, Suc (j + k)))
(term_of_pair (punit.lt (bs ! j), k))" ..
from this(1) have "j < length bs" by simp
show ?thesis
proof (rule Cons(2))
show "0 < Suc j" by simp
next
from \<open>j < length bs\<close> show "Suc j < length (b # bs)" by simp
next
show "v = ord_term_lin.max (term_of_pair (punit.lt ((b # bs) ! 0), k + Suc j))
(term_of_pair (punit.lt ((b # bs) ! Suc j), k + 0))"
by (simp add: v ac_simps)
qed
next
assume "v \<in> ?B"
obtain i j where "i < j" and "j < length bs"
and v: "v = ord_term_lin.max (term_of_pair (punit.lt (bs ! i), Suc k + j))
(term_of_pair (punit.lt (bs ! j), Suc k + i))"
by (rule Cons(1), assumption, rule \<open>v \<in> ?B\<close>)
show ?thesis
proof (rule Cons(2))
from \<open>i < j\<close> show "Suc i < Suc j" by simp
next
from \<open>j < length bs\<close> show "Suc j < length (b # bs)" by simp
next
show "v = ord_term_lin.max (term_of_pair (punit.lt ((b # bs) ! Suc i), k + Suc j))
(term_of_pair (punit.lt ((b # bs) ! Suc j), k + Suc i))"
by (simp add: v)
qed
qed
qed
lemma lt_Koszul_syz_comp:
assumes "0 \<notin> set fs" and "i < length fs"
shows "lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j))) = term_of_pair (punit.lt (fs ! i), j)"
proof -
from assms(2) have "fs ! i \<in> set fs" by (rule nth_mem)
with assms(1) have "fs ! i \<noteq> 0" by auto
thus ?thesis by (simp add: lt_mult_scalar_monomial_right splus_def term_simps)
qed
lemma Koszul_syz_nonzero_lt:
assumes "rep_list a \<noteq> 0" and "rep_list b \<noteq> 0" and "component_of_term (lt a) < component_of_term (lt b)"
shows "rep_list a \<odot> b - rep_list b \<odot> a \<noteq> 0" (is "?p - ?q \<noteq> 0")
and "lt (rep_list a \<odot> b - rep_list b \<odot> a) =
ord_term_lin.max (punit.lt (rep_list a) \<oplus> lt b) (punit.lt (rep_list b) \<oplus> lt a)" (is "_ = ?r")
proof -
from assms(2) have "b \<noteq> 0" by (auto simp: rep_list_zero)
with assms(1) have lt_p: "lt ?p = punit.lt (rep_list a) \<oplus> lt b" by (rule lt_mult_scalar)
from assms(1) have "a \<noteq> 0" by (auto simp: rep_list_zero)
with assms(2) have lt_q: "lt ?q = punit.lt (rep_list b) \<oplus> lt a" by (rule lt_mult_scalar)
from assms(3) have "component_of_term (lt ?p) \<noteq> component_of_term (lt ?q)"
by (simp add: lt_p lt_q component_of_term_splus)
hence "lt ?p \<noteq> lt ?q" by auto
hence "lt (?p - ?q) = ord_term_lin.max (lt ?p) (lt ?q)" by (rule lt_minus_distinct_eq_max)
also have "... = ?r" by (simp only: lt_p lt_q)
finally show "lt (?p - ?q) = ?r" .
from \<open>lt ?p \<noteq> lt ?q\<close> show "?p - ?q \<noteq> 0" by auto
qed
lemma Koszul_syz_is_syz: "rep_list (rep_list a \<odot> b - rep_list b \<odot> a) = 0"
by (simp add: rep_list_minus rep_list_mult_scalar)
lemma dgrad_sig_set_closed_Koszul_syz:
assumes "dickson_grading dgrad" and "a \<in> dgrad_sig_set dgrad" and "b \<in> dgrad_sig_set dgrad"
shows "rep_list a \<odot> b - rep_list b \<odot> a \<in> dgrad_sig_set dgrad"
proof -
from assms(2, 3) have 1: "a \<in> dgrad_max_set dgrad" and 2: "b \<in> dgrad_max_set dgrad"
by (simp_all add: dgrad_sig_set'_def)
show ?thesis
by (intro dgrad_sig_set_closed_minus dgrad_sig_set_closed_mult_scalar dgrad_max_2 assms 1 2)
qed
corollary Koszul_syz_is_syz_sig:
assumes "dickson_grading dgrad" and "a \<in> dgrad_sig_set dgrad" and "b \<in> dgrad_sig_set dgrad"
and "rep_list a \<noteq> 0" and "rep_list b \<noteq> 0" and "component_of_term (lt a) < component_of_term (lt b)"
shows "is_syz_sig dgrad (ord_term_lin.max (punit.lt (rep_list a) \<oplus> lt b) (punit.lt (rep_list b) \<oplus> lt a))"
proof (rule is_syz_sigI)
from assms(4-6) show "rep_list a \<odot> b - rep_list b \<odot> a \<noteq> 0"
and "lt (rep_list a \<odot> b - rep_list b \<odot> a) =
ord_term_lin.max (punit.lt (rep_list a) \<oplus> lt b) (punit.lt (rep_list b) \<oplus> lt a)"
by (rule Koszul_syz_nonzero_lt)+
next
from assms(1-3) show "rep_list a \<odot> b - rep_list b \<odot> a \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_Koszul_syz)
qed (fact Koszul_syz_is_syz)
corollary lt_Koszul_syz_in_Koszul_syz_sigs_aux:
assumes "distinct fs" and "0 \<notin> set fs" and "i < j" and "j < length fs"
shows "lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j)) - (fs ! j) \<odot> monomial 1 (term_of_pair (0, i))) \<in>
set (Koszul_syz_sigs_aux fs 0)" (is "?l \<in> ?K")
proof -
let ?a = "monomial (1::'b) (term_of_pair (0, i))"
let ?b = "monomial (1::'b) (term_of_pair (0, j))"
from assms(3, 4) have "i < length fs" by simp
with assms(1) have a: "rep_list ?a = fs ! i" by (simp add: rep_list_monomial term_simps)
from assms(1, 4) have b: "rep_list ?b = fs ! j" by (simp add: rep_list_monomial term_simps)
have "?l = lt (rep_list ?a \<odot> ?b - rep_list ?b \<odot> ?a)" by (simp only: a b)
also have "... = ord_term_lin.max (punit.lt (rep_list ?a) \<oplus> lt ?b) (punit.lt (rep_list ?b) \<oplus> lt ?a)"
proof (rule Koszul_syz_nonzero_lt)
from \<open>i < length fs\<close> have "fs ! i \<in> set fs" by (rule nth_mem)
with assms(2) show "rep_list ?a \<noteq> 0" by (auto simp: a)
next
from assms(4) have "fs ! j \<in> set fs" by (rule nth_mem)
with assms(2) show "rep_list ?b \<noteq> 0" by (auto simp: b)
next
from assms(3) show "component_of_term (lt ?a) < component_of_term (lt ?b)"
by (simp add: lt_monomial component_of_term_of_pair)
qed
also have "... = ord_term_lin.max (term_of_pair (punit.lt (fs ! i), 0 + j)) (term_of_pair (punit.lt (fs ! j), 0 + i))"
by (simp add: a b lt_monomial splus_def term_simps)
also from assms(3, 4) have "... \<in> ?K" by (rule Koszul_syz_sigs_auxI)
thm Koszul_syz_sigs_auxI[OF assms(3, 4)]
finally show ?thesis .
qed
corollary lt_Koszul_syz_in_Koszul_syz_sigs:
assumes "\<not> is_pot_ord" and "distinct fs" and "0 \<notin> set fs" and "i < j" and "j < length fs"
obtains v where "v \<in> set (Koszul_syz_sigs fs)"
and "v adds\<^sub>t lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j)) - (fs ! j) \<odot> monomial 1 (term_of_pair (0, i)))"
proof -
have "transp (adds\<^sub>t)" by (rule transpI, drule adds_term_trans)
moreover have "lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j)) - (fs ! j) \<odot> monomial 1 (term_of_pair (0, i))) \<in>
set (Koszul_syz_sigs_aux fs 0)" (is "?l \<in> set ?ks")
using assms(2-5) by (rule lt_Koszul_syz_in_Koszul_syz_sigs_aux)
ultimately show ?thesis
proof (rule filter_min_cases)
assume "?l \<in> set (filter_min (adds\<^sub>t) ?ks)"
hence "?l \<in> set (Koszul_syz_sigs fs)" by (simp add: Koszul_syz_sigs_def assms(1))
thus ?thesis using adds_term_refl ..
next
fix v
assume "v \<in> set (filter_min (adds\<^sub>t) ?ks)"
hence "v \<in> set (Koszul_syz_sigs fs)" by (simp add: Koszul_syz_sigs_def assms(1))
moreover assume "v adds\<^sub>t ?l"
ultimately show ?thesis ..
qed
qed
lemma lt_Koszul_syz_init:
assumes "0 \<notin> set fs" and "i < j" and "j < length fs"
shows "lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j)) - (fs ! j) \<odot> monomial 1 (term_of_pair (0, i))) =
ord_term_lin.max (term_of_pair (punit.lt (fs ! i), j)) (term_of_pair (punit.lt (fs ! j), i))"
(is "lt (?p - ?q) = ?r")
proof -
from assms(2, 3) have "i < length fs" by simp
with assms(1) have lt_i: "lt ?p = term_of_pair (punit.lt (fs ! i), j)" by (rule lt_Koszul_syz_comp)
from assms(1, 3) have lt_j: "lt ?q = term_of_pair (punit.lt (fs ! j), i)" by (rule lt_Koszul_syz_comp)
from assms(2) have "component_of_term (lt ?p) \<noteq> component_of_term (lt ?q)"
by (simp add: lt_i lt_j component_of_term_of_pair)
hence "lt ?p \<noteq> lt ?q" by auto
hence "lt (?p - ?q) = ord_term_lin.max (lt ?p) (lt ?q)" by (rule lt_minus_distinct_eq_max)
also have "... = ?r" by (simp only: lt_i lt_j)
finally show ?thesis .
qed
corollary Koszul_syz_sigs_auxE_lt_Koszul_syz:
assumes "0 \<notin> set fs" and "v \<in> set (Koszul_syz_sigs_aux fs 0)"
obtains i j where "i < j" and "j < length fs"
and "v = lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j)) - (fs ! j) \<odot> monomial 1 (term_of_pair (0, i)))"
proof -
from assms(2) obtain i j where "i < j" and "j < length fs"
and "v = ord_term_lin.max (term_of_pair (punit.lt (fs ! i), 0 + j))
(term_of_pair (punit.lt (fs ! j), 0 + i))"
by (rule Koszul_syz_sigs_auxE)
with assms(1) have "v = lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j)) -
(fs ! j) \<odot> monomial 1 (term_of_pair (0, i)))"
by (simp add: lt_Koszul_syz_init)
with \<open>i < j\<close> \<open>j < length fs\<close> show ?thesis ..
qed
corollary Koszul_syz_sigs_is_syz_sig:
assumes "dickson_grading dgrad" and "distinct fs" and "0 \<notin> set fs" and "v \<in> set (Koszul_syz_sigs fs)"
shows "is_syz_sig dgrad v"
proof -
from assms(4) have "v \<in> set (Koszul_syz_sigs_aux fs 0)"
using filter_min_subset by (fastforce simp: Koszul_syz_sigs_def)
with assms(3) obtain i j where "i < j" and "j < length fs"
and v': "v = lt ((fs ! i) \<odot> monomial 1 (term_of_pair (0, j)) - (fs ! j) \<odot> monomial 1 (term_of_pair (0, i)))"
(is "v = lt (?p - ?q)")
by (rule Koszul_syz_sigs_auxE_lt_Koszul_syz)
let ?a = "monomial (1::'b) (term_of_pair (0, i))"
let ?b = "monomial (1::'b) (term_of_pair (0, j))"
from \<open>i < j\<close> \<open>j < length fs\<close> have "i < length fs" by simp
with assms(2) have a: "rep_list ?a = fs ! i" by (simp add: rep_list_monomial term_simps)
from assms(2) \<open>j < length fs\<close> have b: "rep_list ?b = fs ! j" by (simp add: rep_list_monomial term_simps)
note v'
also have "lt (?p - ?q) = ord_term_lin.max (term_of_pair (punit.lt (fs ! i), j)) (term_of_pair (punit.lt (fs ! j), i))"
using assms(3) \<open>i < j\<close> \<open>j < length fs\<close> by (rule lt_Koszul_syz_init)
also have "... = ord_term_lin.max (punit.lt (rep_list ?a) \<oplus> lt ?b) (punit.lt (rep_list ?b) \<oplus> lt ?a)"
by (simp add: a b lt_monomial splus_def term_simps)
finally have v: "v = ord_term_lin.max (punit.lt (rep_list ?a) \<oplus> lt ?b) (punit.lt (rep_list ?b) \<oplus> lt ?a)" .
show ?thesis unfolding v using assms(1)
proof (rule Koszul_syz_is_syz_sig)
show "?a \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 \<open>i < length fs\<close>)
next
show "?b \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 \<open>j < length fs\<close>)
next
from \<open>i < length fs\<close> have "fs ! i \<in> set fs" by (rule nth_mem)
with assms(3) show "rep_list ?a \<noteq> 0" by (fastforce simp: a)
next
from \<open>j < length fs\<close> have "fs ! j \<in> set fs" by (rule nth_mem)
with assms(3) show "rep_list ?b \<noteq> 0" by (fastforce simp: b)
next
from \<open>i < j\<close> show "component_of_term (lt ?a) < component_of_term (lt ?b)"
by (simp add: lt_monomial component_of_term_of_pair)
qed
qed
lemma Koszul_syz_sigs_minimal:
assumes "u \<in> set (Koszul_syz_sigs fs)" and "v \<in> set (Koszul_syz_sigs fs)" and "u adds\<^sub>t v"
shows "u = v"
proof -
from assms(1, 2) have "u \<in> set (filter_min (adds\<^sub>t) (Koszul_syz_sigs_aux fs 0))"
and "v \<in> set (filter_min (adds\<^sub>t) (Koszul_syz_sigs_aux fs 0))" by (simp_all add: Koszul_syz_sigs_def)
with _ show ?thesis using assms(3)
proof (rule filter_min_minimal)
show "transp (adds\<^sub>t)" by (rule transpI, drule adds_term_trans)
qed
qed
lemma Koszul_syz_sigs_distinct: "distinct (Koszul_syz_sigs fs)"
proof -
from adds_term_refl have "reflp (adds\<^sub>t)" by (rule reflpI)
thus ?thesis by (simp add: Koszul_syz_sigs_def filter_min_distinct)
qed
subsubsection \<open>Algorithms\<close>
definition spair_spp :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))"
where "spair_spp p q = (let t1 = punit.lt (snd p); t2 = punit.lt (snd q); l = lcs t1 t2 in
(ord_term_lin.max ((l - t1) \<oplus> fst p) ((l - t2) \<oplus> fst q),
punit.monom_mult (1 / punit.lc (snd p)) (l - t1) (snd p) -
punit.monom_mult (1 / punit.lc (snd q)) (l - t2) (snd q)))"
definition is_regular_spair_spp :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool"
where "is_regular_spair_spp p q \<longleftrightarrow>
(snd p \<noteq> 0 \<and> snd q \<noteq> 0 \<and> punit.lt (snd q) \<oplus> fst p \<noteq> punit.lt (snd p) \<oplus> fst q)"
definition spair_sigs :: "('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('t \<times> 't)"
where "spair_sigs p q =
(let t1 = punit.lt (rep_list p); t2 = punit.lt (rep_list q); l = lcs t1 t2 in
((l - t1) \<oplus> lt p, (l - t2) \<oplus> lt q))"
definition spair_sigs_spp :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> 't)"
where "spair_sigs_spp p q =
(let t1 = punit.lt (snd p); t2 = punit.lt (snd q); l = lcs t1 t2 in
((l - t1) \<oplus> fst p, (l - t2) \<oplus> fst q))"
fun poly_of_pair :: "((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b)"
where
"poly_of_pair (Inl (p, q)) = spair p q" |
"poly_of_pair (Inr j) = monomial 1 (term_of_pair (0, j))"
fun spp_of_pair :: "((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))"
where
"spp_of_pair (Inl (p, q)) = spair_spp p q" |
"spp_of_pair (Inr j) = (term_of_pair (0, j), fs ! j)"
fun sig_of_pair :: "((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) \<Rightarrow> 't"
where
"sig_of_pair (Inl (p, q)) = (let (u, v) = spair_sigs p q in ord_term_lin.max u v)" |
"sig_of_pair (Inr j) = term_of_pair (0, j)"
fun sig_of_pair_spp :: "((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) \<Rightarrow> 't"
where
"sig_of_pair_spp (Inl (p, q)) = (let (u, v) = spair_sigs_spp p q in ord_term_lin.max u v)" |
"sig_of_pair_spp (Inr j) = term_of_pair (0, j)"
definition pair_ord :: "((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) \<Rightarrow> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) \<Rightarrow> bool"
where "pair_ord x y \<longleftrightarrow> (sig_of_pair x \<preceq>\<^sub>t sig_of_pair y)"
definition pair_ord_spp :: "((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) \<Rightarrow>
((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) \<Rightarrow> bool"
where "pair_ord_spp x y \<longleftrightarrow> (sig_of_pair_spp x \<preceq>\<^sub>t sig_of_pair_spp y)"
primrec new_spairs :: "('t \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list" where
"new_spairs [] p = []" |
"new_spairs (b # bs) p =
(if is_regular_spair p b then insort_wrt pair_ord (Inl (p, b)) (new_spairs bs p) else new_spairs bs p)"
primrec new_spairs_spp :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow>
((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) list" where
"new_spairs_spp [] p = []" |
"new_spairs_spp (b # bs) p =
(if is_regular_spair_spp p b then
insort_wrt pair_ord_spp (Inl (p, b)) (new_spairs_spp bs p)
else new_spairs_spp bs p)"
definition add_spairs :: "((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow>
((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list"
where "add_spairs ps bs p = merge_wrt pair_ord (new_spairs bs p) ps"
definition add_spairs_spp :: "((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) list \<Rightarrow>
('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow>
((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) list"
where "add_spairs_spp ps bs p = merge_wrt pair_ord_spp (new_spairs_spp bs p) ps"
lemma spair_alt_spair_sigs:
"spair p q = monom_mult (1 / punit.lc (rep_list p)) (pp_of_term (fst (spair_sigs p q)) - lp p) p -
monom_mult (1 / punit.lc (rep_list q)) (pp_of_term (snd (spair_sigs p q)) - lp q) q"
by (simp add: spair_def spair_sigs_def Let_def term_simps)
lemma sig_of_spair:
assumes "is_regular_spair p q"
shows "sig_of_pair (Inl (p, q)) = lt (spair p q)"
proof -
from assms have "rep_list p \<noteq> 0" by (rule is_regular_spairD1)
hence 1: "punit.lc (rep_list p) \<noteq> 0" and "p \<noteq> 0" by (rule punit.lc_not_0, auto simp: rep_list_zero)
from assms have "rep_list q \<noteq> 0" by (rule is_regular_spairD2)
hence 2: "punit.lc (rep_list q) \<noteq> 0" and "q \<noteq> 0" by (rule punit.lc_not_0, auto simp: rep_list_zero)
let ?t1 = "punit.lt (rep_list p)"
let ?t2 = "punit.lt (rep_list q)"
let ?l = "lcs ?t1 ?t2"
from assms have "lt (monom_mult (1 / punit.lc (rep_list p)) (?l - ?t1) p) \<noteq>
lt (monom_mult (1 / punit.lc (rep_list q)) (?l - ?t2) q)"
by (rule is_regular_spairD3)
hence *: "lt (monom_mult (1 / punit.lc (rep_list p)) (pp_of_term (fst (spair_sigs p q)) - lp p) p) \<noteq>
lt (monom_mult (1 / punit.lc (rep_list q)) (pp_of_term (snd (spair_sigs p q)) - lp q) q)"
by (simp add: spair_sigs_def Let_def term_simps)
from 1 2 \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis
by (simp add: spair_alt_spair_sigs lt_monom_mult lt_minus_distinct_eq_max[OF *],
simp add: spair_sigs_def Let_def term_simps)
qed
lemma sig_of_spair_commute: "sig_of_pair (Inl (p, q)) = sig_of_pair (Inl (q, p))"
by (simp add: spair_sigs_def Let_def lcs_comm ord_term_lin.max.commute)
lemma in_new_spairsI:
assumes "b \<in> set bs" and "is_regular_spair p b"
shows "Inl (p, b) \<in> set (new_spairs bs p)"
using assms(1)
proof (induct bs)
case Nil
thus ?case by simp
next
case (Cons a bs)
from Cons(2) have "b = a \<or> b \<in> set bs" by simp
thus ?case
proof
assume "b = a"
from assms(2) show ?thesis by (simp add: \<open>b = a\<close>)
next
assume "b \<in> set bs"
hence "Inl (p, b) \<in> set (new_spairs bs p)" by (rule Cons(1))
thus ?thesis by simp
qed
qed
lemma in_new_spairsD:
assumes "Inl (a, b) \<in> set (new_spairs bs p)"
shows "a = p" and "b \<in> set bs" and "is_regular_spair p b"
proof -
from assms have "a = p \<and> b \<in> set bs \<and> is_regular_spair p b"
proof (induct bs)
case Nil
thus ?case by simp
next
case (Cons c bs)
from Cons(2) have "(is_regular_spair p c \<and> Inl (a, b) = Inl (p, c)) \<or> Inl (a, b) \<in> set (new_spairs bs p)"
by (simp split: if_split_asm)
thus ?case
proof
assume "is_regular_spair p c \<and> Inl (a, b) = Inl (p, c)"
hence "is_regular_spair p c" and "a = p" and "b = c" by simp_all
thus ?thesis by simp
next
assume "Inl (a, b) \<in> set (new_spairs bs p)"
hence "a = p \<and> b \<in> set bs \<and> is_regular_spair p b" by (rule Cons(1))
thus ?thesis by simp
qed
qed
thus "a = p" and "b \<in> set bs" and "is_regular_spair p b" by simp_all
qed
corollary in_new_spairs_iff:
"Inl (p, b) \<in> set (new_spairs bs p) \<longleftrightarrow> (b \<in> set bs \<and> is_regular_spair p b)"
by (auto intro: in_new_spairsI dest: in_new_spairsD)
lemma Inr_not_in_new_spairs: "Inr j \<notin> set (new_spairs bs p)"
by (induct bs, simp_all)
lemma sum_prodE:
assumes "\<And>a b. p = Inl (a, b) \<Longrightarrow> thesis" and "\<And>j. p = Inr j \<Longrightarrow> thesis"
shows thesis
using _ assms(2)
proof (rule sumE)
fix x
assume "p = Inl x"
moreover obtain a b where "x = (a, b)" by fastforce
ultimately have "p = Inl (a, b)" by simp
thus ?thesis by (rule assms(1))
qed
corollary in_new_spairsE:
assumes "q \<in> set (new_spairs bs p)"
obtains b where "b \<in> set bs" and "is_regular_spair p b" and "q = Inl (p, b)"
proof (rule sum_prodE)
fix a b
assume q: "q = Inl (a, b)"
from assms have "a = p" and "b \<in> set bs" and "is_regular_spair p b"
unfolding q by (rule in_new_spairsD)+
note this(2, 3)
moreover have "q = Inl (p, b)" by (simp only: q \<open>a = p\<close>)
ultimately show ?thesis ..
next
fix j
assume "q = Inr j"
with assms show ?thesis by (simp add: Inr_not_in_new_spairs)
qed
lemma new_spairs_sorted: "sorted_wrt pair_ord (new_spairs bs p)"
proof (induct bs)
case Nil
show ?case by simp
next
case (Cons a bs)
moreover have "transp pair_ord" by (rule transpI, simp add: pair_ord_def)
moreover have "pair_ord x y \<or> pair_ord y x" for x y by (simp add: pair_ord_def ord_term_lin.linear)
ultimately show ?case by (simp add: sorted_wrt_insort_wrt)
qed
lemma sorted_add_spairs:
assumes "sorted_wrt pair_ord ps"
shows "sorted_wrt pair_ord (add_spairs ps bs p)"
unfolding add_spairs_def using _ _ new_spairs_sorted assms
proof (rule sorted_merge_wrt)
show "transp pair_ord" by (rule transpI, simp add: pair_ord_def)
next
fix x y
show "pair_ord x y \<or> pair_ord y x" by (simp add: pair_ord_def ord_term_lin.linear)
qed
context
fixes rword_strict :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool" \<comment>\<open>Must be a @{emph \<open>strict\<close>} rewrite order.\<close>
begin
qualified definition rword :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool"
where "rword x y \<longleftrightarrow> \<not> rword_strict y x"
definition is_pred_syz :: "'t list \<Rightarrow> 't \<Rightarrow> bool"
where "is_pred_syz ss u = (\<exists>v\<in>set ss. v adds\<^sub>t u)"
definition is_rewritable :: "('t \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> ('t \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 't \<Rightarrow> bool"
where "is_rewritable bs p u = (\<exists>b\<in>set bs. b \<noteq> 0 \<and> lt b adds\<^sub>t u \<and> rword_strict (spp_of p) (spp_of b))"
definition is_rewritable_spp :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> 't \<Rightarrow> bool"
where "is_rewritable_spp bs p u = (\<exists>b\<in>set bs. fst b adds\<^sub>t u \<and> rword_strict p b)"
fun sig_crit :: "('t \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> 't list \<Rightarrow> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) \<Rightarrow> bool"
where
"sig_crit bs ss (Inl (p, q)) =
(let (u, v) = spair_sigs p q in
is_pred_syz ss u \<or> is_pred_syz ss v \<or> is_rewritable bs p u \<or> is_rewritable bs q v)" |
"sig_crit bs ss (Inr j) = is_pred_syz ss (term_of_pair (0, j))"
fun sig_crit' :: "('t \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) \<Rightarrow> bool"
where
"sig_crit' bs (Inl (p, q)) =
(let (u, v) = spair_sigs p q in
is_syz_sig dgrad u \<or> is_syz_sig dgrad v \<or> is_rewritable bs p u \<or> is_rewritable bs q v)" |
"sig_crit' bs (Inr j) = is_syz_sig dgrad (term_of_pair (0, j))"
fun sig_crit_spp :: "('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<Rightarrow> 't list \<Rightarrow> ((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) \<Rightarrow> bool"
where
"sig_crit_spp bs ss (Inl (p, q)) =
(let (u, v) = spair_sigs_spp p q in
is_pred_syz ss u \<or> is_pred_syz ss v \<or> is_rewritable_spp bs p u \<or> is_rewritable_spp bs q v)" |
"sig_crit_spp bs ss (Inr j) = is_pred_syz ss (term_of_pair (0, j))"
text \<open>@{const sig_crit} is used in algorithms, @{const sig_crit'} is only needed for proving.\<close>
fun rb_spp_body ::
"((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<times> 't list \<times> ((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) list) \<times> nat) \<Rightarrow>
((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<times> 't list \<times> ((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) list) \<times> nat)"
where
"rb_spp_body ((bs, ss, []), z) = ((bs, ss, []), z)" |
"rb_spp_body ((bs, ss, p # ps), z) =
(let ss' = new_syz_sigs_spp ss bs p in
if sig_crit_spp bs ss' p then
((bs, ss', ps), z)
else
let p' = sig_trd_spp bs (spp_of_pair p) in
if snd p' = 0 then
((bs, fst p' # ss', ps), Suc z)
else
((p' # bs, ss', add_spairs_spp ps bs p'), z))"
definition rb_spp_aux ::
"((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<times> 't list \<times> ((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) list) \<times> nat) \<Rightarrow>
((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) list \<times> 't list \<times> ((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) list) \<times> nat)"
where rb_spp_aux_def [code del]: "rb_spp_aux = tailrec.fun (\<lambda>x. snd (snd (fst x)) = []) (\<lambda>x. x) rb_spp_body"
lemma rb_spp_aux_Nil [code]: "rb_spp_aux ((bs, ss, []), z) = ((bs, ss, []), z)"
by (simp add: rb_spp_aux_def tailrec.simps)
lemma rb_spp_aux_Cons [code]:
"rb_spp_aux ((bs, ss, p # ps), z) = rb_spp_aux (rb_spp_body ((bs, ss, p # ps), z))"
by (simp add: rb_spp_aux_def tailrec.simps)
text \<open>The last parameter / return value of @{const rb_spp_aux}, @{term z}, counts the number of
zero-reductions. Below we will prove that this number remains $0$ under certain conditions.\<close>
context
assumes rword_is_strict_rewrite_ord: "is_strict_rewrite_ord rword_strict"
assumes dgrad: "dickson_grading dgrad"
begin
lemma rword: "is_rewrite_ord rword"
unfolding rword_def using rword_is_strict_rewrite_ord by (rule is_strict_rewrite_ordD)
lemma sig_crit'_sym: "sig_crit' bs (Inl (p, q)) \<Longrightarrow> sig_crit' bs (Inl (q, p))"
by (auto simp: spair_sigs_def Let_def lcs_comm)
lemma is_rewritable_ConsD:
assumes "is_rewritable (b # bs) p u" and "u \<prec>\<^sub>t lt b"
shows "is_rewritable bs p u"
proof -
from assms(1) obtain b' where "b' \<in> set (b # bs)" and "b' \<noteq> 0" and "lt b' adds\<^sub>t u"
and "rword_strict (spp_of p) (spp_of b')" unfolding is_rewritable_def by blast
from this(3) have "lt b' \<preceq>\<^sub>t u" by (rule ord_adds_term)
with assms(2) have "b' \<noteq> b" by auto
with \<open>b' \<in> set (b # bs)\<close> have "b' \<in> set bs" by simp
with \<open>b' \<noteq> 0\<close> \<open>lt b' adds\<^sub>t u\<close> \<open>rword_strict (spp_of p) (spp_of b')\<close> show ?thesis
by (auto simp: is_rewritable_def)
qed
lemma sig_crit'_ConsD:
assumes "sig_crit' (b # bs) p" and "sig_of_pair p \<prec>\<^sub>t lt b"
shows "sig_crit' bs p"
proof (rule sum_prodE)
fix x y
assume p: "p = Inl (x, y)"
define u where "u = fst (spair_sigs x y)"
define v where "v = snd (spair_sigs x y)"
have sigs: "spair_sigs x y = (u, v)" by (simp add: u_def v_def)
have "u \<preceq>\<^sub>t sig_of_pair p" and "v \<preceq>\<^sub>t sig_of_pair p" by (simp_all add: p sigs)
hence "u \<prec>\<^sub>t lt b" and "v \<prec>\<^sub>t lt b" using assms(2) by simp_all
with assms(1) show ?thesis by (auto simp: p sigs dest: is_rewritable_ConsD)
next
fix j
assume p: "p = Inr j"
from assms show ?thesis by (simp add: p)
qed
definition rb_aux_inv1 :: "('t \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> bool"
where "rb_aux_inv1 bs =
(set bs \<subseteq> dgrad_sig_set dgrad \<and> 0 \<notin> rep_list ` set bs \<and>
sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) bs \<and>
(\<forall>i<length bs. \<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i) bs)) (bs ! i)) \<and>
(\<forall>i<length bs.
((\<exists>j<length fs. lt (bs ! i) = lt (monomial (1::'b) (term_of_pair (0, j))) \<and>
punit.lt (rep_list (bs ! i)) \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) \<or>
(\<exists>p\<in>set bs. \<exists>q\<in>set bs. is_regular_spair p q \<and> rep_list (spair p q) \<noteq> 0 \<and>
lt (bs ! i) = lt (spair p q) \<and> punit.lt (rep_list (bs ! i)) \<preceq> punit.lt (rep_list (spair p q))))) \<and>
(\<forall>i<length bs. is_RB_upt dgrad rword (set (drop (Suc i) bs)) (lt (bs ! i))))"
fun rb_aux_inv :: "(('t \<Rightarrow>\<^sub>0 'b) list \<times> 't list \<times> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list) \<Rightarrow> bool"
where "rb_aux_inv (bs, ss, ps) =
(rb_aux_inv1 bs \<and>
(\<forall>u\<in>set ss. is_syz_sig dgrad u) \<and>
(\<forall>p q. Inl (p, q) \<in> set ps \<longrightarrow> (is_regular_spair p q \<and> p \<in> set bs \<and> q \<in> set bs)) \<and>
(\<forall>j. Inr j \<in> set ps \<longrightarrow> (j < length fs \<and> (\<forall>b\<in>set bs. lt b \<prec>\<^sub>t term_of_pair (0, j))) \<and>
length (filter (\<lambda>q. sig_of_pair q = term_of_pair (0, j)) ps) \<le> 1) \<and>
(sorted_wrt pair_ord ps) \<and>
(\<forall>p\<in>set ps. (\<forall>b1\<in>set bs. \<forall>b2\<in>set bs. is_regular_spair b1 b2 \<longrightarrow>
sig_of_pair p \<prec>\<^sub>t lt (spair b1 b2) \<longrightarrow> (Inl (b1, b2) \<in> set ps \<or> Inl (b2, b1) \<in> set ps)) \<and>
(\<forall>j<length fs. sig_of_pair p \<prec>\<^sub>t term_of_pair (0, j) \<longrightarrow> Inr j \<in> set ps)) \<and>
(\<forall>b\<in>set bs. \<forall>p\<in>set ps. lt b \<preceq>\<^sub>t sig_of_pair p) \<and>
(\<forall>a\<in>set bs. \<forall>b\<in>set bs. is_regular_spair a b \<longrightarrow> Inl (a, b) \<notin> set ps \<longrightarrow> Inl (b, a) \<notin> set ps \<longrightarrow>
\<not> is_RB_in dgrad rword (set bs) (lt (spair a b)) \<longrightarrow>
(\<exists>p\<in>set ps. sig_of_pair p = lt (spair a b) \<and> \<not> sig_crit' bs p)) \<and>
(\<forall>j<length fs. Inr j \<notin> set ps \<longrightarrow> (is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) \<and>
rep_list (monomial (1::'b) (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs))))"
lemmas [simp del] = rb_aux_inv.simps
lemma rb_aux_inv1_D1: "rb_aux_inv1 bs \<Longrightarrow> set bs \<subseteq> dgrad_sig_set dgrad"
by (simp add: rb_aux_inv1_def)
lemma rb_aux_inv1_D2: "rb_aux_inv1 bs \<Longrightarrow> 0 \<notin> rep_list ` set bs"
by (simp add: rb_aux_inv1_def)
lemma rb_aux_inv1_D3: "rb_aux_inv1 bs \<Longrightarrow> sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) bs"
by (simp add: rb_aux_inv1_def)
lemma rb_aux_inv1_D4:
"rb_aux_inv1 bs \<Longrightarrow> i < length bs \<Longrightarrow> \<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i) bs)) (bs ! i)"
by (simp add: rb_aux_inv1_def)
lemma rb_aux_inv1_D5:
"rb_aux_inv1 bs \<Longrightarrow> i < length bs \<Longrightarrow> is_RB_upt dgrad rword (set (drop (Suc i) bs)) (lt (bs ! i))"
by (simp add: rb_aux_inv1_def)
lemma rb_aux_inv1_E:
assumes "rb_aux_inv1 bs" and "i < length bs"
and "\<And>j. j < length fs \<Longrightarrow> lt (bs ! i) = lt (monomial (1::'b) (term_of_pair (0, j))) \<Longrightarrow>
punit.lt (rep_list (bs ! i)) \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j)))) \<Longrightarrow> thesis"
and "\<And>p q. p \<in> set bs \<Longrightarrow> q \<in> set bs \<Longrightarrow> is_regular_spair p q \<Longrightarrow> rep_list (spair p q) \<noteq> 0 \<Longrightarrow>
lt (bs ! i) = lt (spair p q) \<Longrightarrow> punit.lt (rep_list (bs ! i)) \<preceq> punit.lt (rep_list (spair p q)) \<Longrightarrow> thesis"
shows thesis
using assms unfolding rb_aux_inv1_def by blast
lemmas rb_aux_inv1_D = rb_aux_inv1_D1 rb_aux_inv1_D2 rb_aux_inv1_D3 rb_aux_inv1_D4
rb_aux_inv1_D5
lemma rb_aux_inv1_distinct_lt:
assumes "rb_aux_inv1 bs"
shows "distinct (map lt bs)"
proof (rule distinct_sorted_wrt_irrefl)
show "irreflp (\<succ>\<^sub>t)" by (simp add: irreflp_def)
next
show "transp (\<succ>\<^sub>t)" by (auto simp: transp_def)
next
from assms show "sorted_wrt (\<succ>\<^sub>t) (map lt bs)"
unfolding sorted_wrt_map conversep_iff by (rule rb_aux_inv1_D3)
qed
corollary rb_aux_inv1_lt_inj_on:
assumes "rb_aux_inv1 bs"
shows "inj_on lt (set bs)"
proof
fix a b
assume "a \<in> set bs"
then obtain i where i: "i < length bs" and a: "a = bs ! i" by (metis in_set_conv_nth)
assume "b \<in> set bs"
then obtain j where j: "j < length bs" and b: "b = bs ! j" by (metis in_set_conv_nth)
assume "lt a = lt b"
with i j have "(map lt bs) ! i = (map lt bs) ! j" by (simp add: a b)
moreover from assms have "distinct (map lt bs)" by (rule rb_aux_inv1_distinct_lt)
moreover from i have "i < length (map lt bs)" by simp
moreover from j have "j < length (map lt bs)" by simp
ultimately have "i = j" by (simp only: nth_eq_iff_index_eq)
thus "a = b" by (simp add: a b)
qed
lemma canon_rewriter_unique:
assumes "rb_aux_inv1 bs" and "is_canon_rewriter rword (set bs) u a"
and "is_canon_rewriter rword (set bs) u b"
shows "a = b"
proof -
from assms(1) have "inj_on lt (set bs)" by (rule rb_aux_inv1_lt_inj_on)
moreover from rword(1) assms(2, 3) have "lt a = lt b" by (rule is_rewrite_ord_canon_rewriterD2)
moreover from assms(2) have "a \<in> set bs" by (rule is_canon_rewriterD1)
moreover from assms(3) have "b \<in> set bs" by (rule is_canon_rewriterD1)
ultimately show ?thesis by (rule inj_onD)
qed
lemma rb_aux_inv_D1: "rb_aux_inv (bs, ss, ps) \<Longrightarrow> rb_aux_inv1 bs"
by (simp add: rb_aux_inv.simps)
lemma rb_aux_inv_D2: "rb_aux_inv (bs, ss, ps) \<Longrightarrow> u \<in> set ss \<Longrightarrow> is_syz_sig dgrad u"
by (simp add: rb_aux_inv.simps)
lemma rb_aux_inv_D3:
assumes "rb_aux_inv (bs, ss, ps)" and "Inl (p, q) \<in> set ps"
shows "p \<in> set bs" and "q \<in> set bs" and "is_regular_spair p q"
using assms by (simp_all add: rb_aux_inv.simps)
lemma rb_aux_inv_D4:
assumes "rb_aux_inv (bs, ss, ps)" and "Inr j \<in> set ps"
shows "j < length fs" and "\<And>b. b \<in> set bs \<Longrightarrow> lt b \<prec>\<^sub>t term_of_pair (0, j)"
and "length (filter (\<lambda>q. sig_of_pair q = term_of_pair (0, j)) ps) \<le> 1"
using assms by (simp_all add: rb_aux_inv.simps)
lemma rb_aux_inv_D5: "rb_aux_inv (bs, ss, ps) \<Longrightarrow> sorted_wrt pair_ord ps"
by (simp add: rb_aux_inv.simps)
lemma rb_aux_inv_D6_1:
assumes "rb_aux_inv (bs, ss, ps)" and "p \<in> set ps" and "b1 \<in> set bs" and "b2 \<in> set bs"
and "is_regular_spair b1 b2" and "sig_of_pair p \<prec>\<^sub>t lt (spair b1 b2)"
obtains "Inl (b1, b2) \<in> set ps" | "Inl (b2, b1) \<in> set ps"
using assms unfolding rb_aux_inv.simps by blast
lemma rb_aux_inv_D6_2:
"rb_aux_inv (bs, ss, ps) \<Longrightarrow> p \<in> set ps \<Longrightarrow> j < length fs \<Longrightarrow> sig_of_pair p \<prec>\<^sub>t term_of_pair (0, j) \<Longrightarrow>
Inr j \<in> set ps"
by (simp add: rb_aux_inv.simps)
lemma rb_aux_inv_D7: "rb_aux_inv (bs, ss, ps) \<Longrightarrow> b \<in> set bs \<Longrightarrow> p \<in> set ps \<Longrightarrow> lt b \<preceq>\<^sub>t sig_of_pair p"
by (simp add: rb_aux_inv.simps)
lemma rb_aux_inv_D8:
assumes "rb_aux_inv (bs, ss, ps)" and "a \<in> set bs" and "b \<in> set bs" and "is_regular_spair a b"
and "Inl (a, b) \<notin> set ps" and "Inl (b, a) \<notin> set ps" and "\<not> is_RB_in dgrad rword (set bs) (lt (spair a b))"
obtains p where "p \<in> set ps" and "sig_of_pair p = lt (spair a b)" and "\<not> sig_crit' bs p"
using assms unfolding rb_aux_inv.simps by meson
lemma rb_aux_inv_D9:
assumes "rb_aux_inv (bs, ss, ps)" and "j < length fs" and "Inr j \<notin> set ps"
shows "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))"
and "rep_list (monomial (1::'b) (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)"
using assms by (simp_all add: rb_aux_inv.simps)
lemma rb_aux_inv_is_RB_upt:
assumes "rb_aux_inv (bs, ss, ps)" and "\<And>p. p \<in> set ps \<Longrightarrow> u \<preceq>\<^sub>t sig_of_pair p"
shows "is_RB_upt dgrad rword (set bs) u"
proof -
from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
from dgrad rword(1) show ?thesis
proof (rule is_RB_upt_finite)
from inv1 show "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
next
from inv1 show "inj_on lt (set bs)" by (rule rb_aux_inv1_lt_inj_on)
next
show "finite (set bs)" by (fact finite_set)
next
fix g1 g2
assume 1: "g1 \<in> set bs" and 2: "g2 \<in> set bs" and 3: "is_regular_spair g1 g2"
and 4: "lt (spair g1 g2) \<prec>\<^sub>t u"
have 5: "p \<notin> set ps" if "sig_of_pair p = lt (spair g1 g2)" for p
proof
assume "p \<in> set ps"
hence "u \<preceq>\<^sub>t sig_of_pair p" by (rule assms(2))
also have "... \<prec>\<^sub>t u" unfolding that by (fact 4)
finally show False ..
qed
show "is_RB_in dgrad rword (set bs) (lt (spair g1 g2))"
proof (rule ccontr)
note assms(1) 1 2 3
moreover have "Inl (g1, g2) \<notin> set ps" by (rule 5, rule sig_of_spair, fact 3)
moreover have "Inl (g2, g1) \<notin> set ps"
by (rule 5, simp only: sig_of_spair_commute, rule sig_of_spair, fact 3)
moreover assume "\<not> is_RB_in dgrad rword (set bs) (lt (spair g1 g2))"
ultimately obtain p where "p \<in> set ps" and "sig_of_pair p = lt (spair g1 g2)"
by (rule rb_aux_inv_D8)
from this(2) have "p \<notin> set ps" by (rule 5)
thus False using \<open>p \<in> set ps\<close> ..
qed
next
fix j
assume 1: "term_of_pair (0, j) \<prec>\<^sub>t u"
note assms(1)
moreover assume "j < length fs"
moreover have "Inr j \<notin> set ps"
proof
assume "Inr j \<in> set ps"
hence "u \<preceq>\<^sub>t sig_of_pair (Inr j)" by (rule assms(2))
also have "... \<prec>\<^sub>t u" by (simp add: 1)
finally show False ..
qed
ultimately show "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))" by (rule rb_aux_inv_D9)
qed
qed
lemma rb_aux_inv_is_RB_upt_Cons:
assumes "rb_aux_inv (bs, ss, p # ps)"
shows "is_RB_upt dgrad rword (set bs) (sig_of_pair p)"
using assms
proof (rule rb_aux_inv_is_RB_upt)
fix q
assume "q \<in> set (p # ps)"
hence "q = p \<or> q \<in> set ps" by simp
thus "sig_of_pair p \<preceq>\<^sub>t sig_of_pair q"
proof
assume "q = p"
thus ?thesis by simp
next
assume "q \<in> set ps"
moreover from assms have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
ultimately show ?thesis by (simp add: pair_ord_def)
qed
qed
lemma Inr_in_tailD:
assumes "rb_aux_inv (bs, ss, p # ps)" and "Inr j \<in> set ps"
shows "sig_of_pair p \<noteq> term_of_pair (0, j)"
proof
assume eq: "sig_of_pair p = term_of_pair (0, j)"
from assms(2) have "Inr j \<in> set (p # ps)" by simp
let ?P = "\<lambda>q. sig_of_pair q = term_of_pair (0, j)"
from assms(2) obtain i1 where "i1 < length ps" and Inrj: "Inr j = ps ! i1"
by (metis in_set_conv_nth)
from assms(1) \<open>Inr j \<in> set (p # ps)\<close> have "length (filter ?P (p # ps)) \<le> 1"
by (rule rb_aux_inv_D4)
moreover from \<open>i1 < length ps\<close> have "Suc i1 < length (p # ps)" by simp
moreover have "0 < length (p # ps)" by simp
moreover have "?P ((p # ps) ! Suc i1)" by (simp add: Inrj[symmetric])
moreover have "?P ((p # ps) ! 0)" by (simp add: eq)
ultimately have "Suc i1 = 0" by (rule length_filter_le_1)
thus False ..
qed
lemma pair_list_aux:
assumes "rb_aux_inv (bs, ss, ps)" and "p \<in> set ps"
shows "sig_of_pair p = lt (poly_of_pair p) \<and> poly_of_pair p \<noteq> 0 \<and> poly_of_pair p \<in> dgrad_sig_set dgrad"
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence bs_sub: "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
from assms have "is_regular_spair a b" unfolding p by (rule rb_aux_inv_D3)
hence "sig_of_pair p = lt (poly_of_pair p)" and "poly_of_pair p \<noteq> 0"
unfolding p poly_of_pair.simps by (rule sig_of_spair, rule is_regular_spair_nonzero)
moreover from dgrad have "poly_of_pair p \<in> dgrad_sig_set dgrad" unfolding p poly_of_pair.simps
proof (rule dgrad_sig_set_closed_spair)
from assms have "a \<in> set bs" unfolding p by (rule rb_aux_inv_D3)
thus "a \<in> dgrad_sig_set dgrad" using bs_sub ..
next
from assms have "b \<in> set bs" unfolding p by (rule rb_aux_inv_D3)
thus "b \<in> dgrad_sig_set dgrad" using bs_sub ..
qed
ultimately show ?thesis by simp
next
fix j
assume "p = Inr j"
from assms have "j < length fs" unfolding \<open>p = Inr j\<close> by (rule rb_aux_inv_D4)
have "monomial 1 (term_of_pair (0, j)) \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_monomial, simp add: pp_of_term_of_pair dgrad_max_0,
simp add: component_of_term_of_pair \<open>j < length fs\<close>)
thus ?thesis by (simp add: \<open>p = Inr j\<close> lt_monomial monomial_0_iff)
qed
corollary pair_list_sig_of_pair:
"rb_aux_inv (bs, ss, ps) \<Longrightarrow> p \<in> set ps \<Longrightarrow> sig_of_pair p = lt (poly_of_pair p)"
by (simp add: pair_list_aux)
corollary pair_list_nonzero: "rb_aux_inv (bs, ss, ps) \<Longrightarrow> p \<in> set ps \<Longrightarrow> poly_of_pair p \<noteq> 0"
by (simp add: pair_list_aux)
corollary pair_list_dgrad_sig_set:
"rb_aux_inv (bs, ss, ps) \<Longrightarrow> p \<in> set ps \<Longrightarrow> poly_of_pair p \<in> dgrad_sig_set dgrad"
by (simp add: pair_list_aux)
lemma is_rewritableI_is_canon_rewriter:
assumes "rb_aux_inv1 bs" and "b \<in> set bs" and "b \<noteq> 0" and "lt b adds\<^sub>t u"
and "\<not> is_canon_rewriter rword (set bs) u b"
shows "is_rewritable bs b u"
proof -
from assms(2-5) obtain b' where "b' \<in> set bs" and "b' \<noteq> 0" and "lt b' adds\<^sub>t u"
and 1: "\<not> rword (spp_of b') (spp_of b)" by (auto simp: is_canon_rewriter_def)
show ?thesis unfolding is_rewritable_def
proof (intro bexI conjI)
from rword(1) have 2: "rword (spp_of b) (spp_of b')"
proof (rule is_rewrite_ordD3)
assume "rword (spp_of b') (spp_of b)"
with 1 show ?thesis ..
qed
from rword(1) 1 have "b \<noteq> b'" by (auto dest: is_rewrite_ordD1)
have "lt b \<noteq> lt b'"
proof
assume "lt b = lt b'"
with rb_aux_inv1_lt_inj_on[OF assms(1)] have "b = b'" using assms(2) \<open>b' \<in> set bs\<close>
by (rule inj_onD)
with \<open>b \<noteq> b'\<close> show False ..
qed
hence "fst (spp_of b) \<noteq> fst (spp_of b')" by (simp add: spp_of_def)
with rword_is_strict_rewrite_ord 2 show "rword_strict (spp_of b) (spp_of b')"
by (auto simp: rword_def dest: is_strict_rewrite_ord_antisym)
qed fact+
qed
lemma is_rewritableD_is_canon_rewriter:
assumes "rb_aux_inv1 bs" and "is_rewritable bs b u"
shows "\<not> is_canon_rewriter rword (set bs) u b"
proof
assume "is_canon_rewriter rword (set bs) u b"
hence "b \<in> set bs" and "b \<noteq> 0" and "lt b adds\<^sub>t u"
and 1: "\<And>a. a \<in> set bs \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> lt a adds\<^sub>t u \<Longrightarrow> rword (spp_of a) (spp_of b)"
by (rule is_canon_rewriterD)+
from assms(2) obtain b' where "b' \<in> set bs" and "b' \<noteq> 0" and "lt b' adds\<^sub>t u"
and 2: "rword_strict (spp_of b) (spp_of b')" unfolding is_rewritable_def by blast
from this(1, 2, 3) have "rword (spp_of b') (spp_of b)" by (rule 1)
moreover from rword_is_strict_rewrite_ord 2 have "rword (spp_of b) (spp_of b')"
unfolding rword_def by (rule is_strict_rewrite_ord_asym)
ultimately have "fst (spp_of b') = fst (spp_of b)" by (rule is_rewrite_ordD4[OF rword])
hence "lt b' = lt b" by (simp add: spp_of_def)
with rb_aux_inv1_lt_inj_on[OF assms(1)] have "b' = b" using \<open>b' \<in> set bs\<close> \<open>b \<in> set bs\<close>
by (rule inj_onD)
from rword_is_strict_rewrite_ord have "\<not> rword_strict (spp_of b) (spp_of b')"
unfolding \<open>b' = b\<close> by (rule is_strict_rewrite_ord_irrefl)
thus False using 2 ..
qed
lemma lemma_12:
assumes "rb_aux_inv (bs, ss, ps)" and "is_RB_upt dgrad rword (set bs) u"
and "dgrad (pp_of_term u) \<le> dgrad_max dgrad" and "is_canon_rewriter rword (set bs) u a"
and "\<not> is_syz_sig dgrad u" and "is_sig_red (\<prec>\<^sub>t) (=) (set bs) (monom_mult 1 (pp_of_term u - lp a) a)"
obtains p q where "p \<in> set bs" and "q \<in> set bs" and "is_regular_spair p q" and "lt (spair p q) = u"
and "\<not> sig_crit' bs (Inl (p, q))"
proof -
from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence inj: "inj_on lt (set bs)" by (rule rb_aux_inv1_lt_inj_on)
from assms(4) have "lt a adds\<^sub>t u" by (rule is_canon_rewriterD3)
hence "lp a adds pp_of_term u" and comp_a: "component_of_term (lt a) = component_of_term u"
by (simp_all add: adds_term_def)
let ?s = "pp_of_term u - lp a"
let ?a = "monom_mult 1 ?s a"
from assms(4) have "a \<in> set bs" by (rule is_canon_rewriterD1)
from assms(6) have "rep_list ?a \<noteq> 0" using is_sig_red_top_addsE by blast
hence "rep_list a \<noteq> 0" by (auto simp: rep_list_monom_mult)
hence "a \<noteq> 0" by (auto simp: rep_list_zero)
hence "lt ?a = ?s \<oplus> lt a" by (simp add: lt_monom_mult)
also from \<open>lp a adds pp_of_term u\<close> have eq0: "... = u"
by (simp add: splus_def comp_a adds_minus term_simps)
finally have "lt ?a = u" .
note dgrad rword(1)
moreover from assms(2) have "is_RB_upt dgrad rword (set bs) (lt ?a)" by (simp only: \<open>lt ?a = u\<close>)
moreover from dgrad have "?a \<in> dgrad_sig_set dgrad"
proof (rule dgrad_sig_set_closed_monom_mult)
from dgrad \<open>lp a adds pp_of_term u\<close> have "dgrad (pp_of_term u - lp a) \<le> dgrad (pp_of_term u)"
by (rule dickson_grading_minus)
thus "dgrad (pp_of_term u - lp a) \<le> dgrad_max dgrad" using assms(3) by (rule le_trans)
next
from inv1 have "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
with \<open>a \<in> set bs\<close> show "a \<in> dgrad_sig_set dgrad" ..
qed
ultimately obtain v b where "v \<prec>\<^sub>t lt ?a" and "dgrad (pp_of_term v) \<le> dgrad_max dgrad"
and "component_of_term v < length fs" and ns: "\<not> is_syz_sig dgrad v"
and v: "v = (punit.lt (rep_list ?a) - punit.lt (rep_list b)) \<oplus> lt b"
and cr: "is_canon_rewriter rword (set bs) v b" and "is_sig_red (\<prec>\<^sub>t) (=) {b} ?a"
using assms(6) by (rule lemma_11)
from this(6) have "b \<in> set bs" by (rule is_canon_rewriterD1)
with \<open>a \<in> set bs\<close> show ?thesis
proof
from dgrad rword(1) assms(2) inj assms(5, 4) \<open>b \<in> set bs\<close> \<open>is_sig_red (\<prec>\<^sub>t) (=) {b} ?a\<close> assms(3)
show "is_regular_spair a b" by (rule lemma_9(3))
next
from dgrad rword(1) assms(2) inj assms(5, 4) \<open>b \<in> set bs\<close> \<open>is_sig_red (\<prec>\<^sub>t) (=) {b} ?a\<close> assms(3)
show "lt (spair a b) = u" by (rule lemma_9(4))
next
from \<open>rep_list a \<noteq> 0\<close> have v': "v = (?s + punit.lt (rep_list a) - punit.lt (rep_list b)) \<oplus> lt b"
by (simp add: v rep_list_monom_mult punit.lt_monom_mult)
moreover from dgrad rword(1) assms(2) inj assms(5, 4) \<open>b \<in> set bs\<close> \<open>is_sig_red (\<prec>\<^sub>t) (=) {b} ?a\<close> assms(3)
have "lcs (punit.lt (rep_list a)) (punit.lt (rep_list b)) - punit.lt (rep_list a) = ?s"
and "lcs (punit.lt (rep_list a)) (punit.lt (rep_list b)) - punit.lt (rep_list b) =
?s + punit.lt (rep_list a) - punit.lt (rep_list b)"
by (rule lemma_9)+
ultimately have eq1: "spair_sigs a b = (u, v)" by (simp add: spair_sigs_def eq0)
show "\<not> sig_crit' bs (Inl (a, b))"
proof (simp add: eq1 assms(5) ns, intro conjI notI)
assume "is_rewritable bs a u"
with inv1 have "\<not> is_canon_rewriter rword (set bs) u a" by (rule is_rewritableD_is_canon_rewriter)
thus False using assms(4) ..
next
assume "is_rewritable bs b v"
with inv1 have "\<not> is_canon_rewriter rword (set bs) v b" by (rule is_rewritableD_is_canon_rewriter)
thus False using cr ..
qed
qed
qed
lemma is_canon_rewriterI_eq_sig:
assumes "rb_aux_inv1 bs" and "b \<in> set bs"
shows "is_canon_rewriter rword (set bs) (lt b) b"
proof -
from assms(2) have "rep_list b \<in> rep_list ` set bs" by (rule imageI)
moreover from assms(1) have "0 \<notin> rep_list ` set bs" by (rule rb_aux_inv1_D2)
ultimately have "b \<noteq> 0" by (auto simp: rep_list_zero)
with assms(2) show ?thesis
proof (rule is_canon_rewriterI)
fix a
assume "a \<in> set bs" and "a \<noteq> 0" and "lt a adds\<^sub>t lt b"
from assms(2) obtain i where "i < length bs" and b: "b = bs ! i" by (metis in_set_conv_nth)
from assms(1) this(1) have "is_RB_upt dgrad rword (set (drop (Suc i) bs)) (lt (bs ! i))"
by (rule rb_aux_inv1_D5)
with dgrad have "is_sig_GB_upt dgrad (set (drop (Suc i) bs)) (lt (bs ! i))"
by (rule is_RB_upt_is_sig_GB_upt)
hence "is_sig_GB_upt dgrad (set (drop (Suc i) bs)) (lt b)" by (simp only: b)
moreover have "set (drop (Suc i) bs) \<subseteq> set bs" by (rule set_drop_subset)
moreover from assms(1) have "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
ultimately have "is_sig_GB_upt dgrad (set bs) (lt b)" by (rule is_sig_GB_upt_mono)
with rword(1) dgrad show "rword (spp_of a) (spp_of b)"
proof (rule is_rewrite_ordD5)
from assms(1) \<open>i < length bs\<close> have "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i) bs)) (bs ! i)"
by (rule rb_aux_inv1_D4)
hence "\<not> is_sig_red (\<prec>\<^sub>t) (=) (set (drop (Suc i) bs)) b" by (simp add: b is_sig_red_top_tail_cases)
moreover have "\<not> is_sig_red (\<prec>\<^sub>t) (=) (set (take (Suc i) bs)) b"
proof
assume "is_sig_red (\<prec>\<^sub>t) (=) (set (take (Suc i) bs)) b"
then obtain f where f_in: "f \<in> set (take (Suc i) bs)" and "is_sig_red (\<prec>\<^sub>t) (=) {f} b"
by (rule is_sig_red_singletonI)
from this(2) have "lt f \<prec>\<^sub>t lt b" by (rule is_sig_red_regularD_lt)
from \<open>i < length bs\<close> have take_eq: "take (Suc i) bs = (take i bs) @ [b]"
unfolding b by (rule take_Suc_conv_app_nth)
from assms(1) have "sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) ((take (Suc i) bs) @ (drop (Suc i) bs))"
unfolding append_take_drop_id by (rule rb_aux_inv1_D3)
hence 1: "\<And>y. y \<in> set (take i bs) \<Longrightarrow> lt b \<prec>\<^sub>t lt y"
by (simp add: sorted_wrt_append take_eq del: append_take_drop_id)
from f_in have "f = b \<or> f \<in> set (take i bs)" by (simp add: take_eq)
hence "lt b \<preceq>\<^sub>t lt f"
proof
assume "f \<in> set (take i bs)"
hence "lt b \<prec>\<^sub>t lt f" by (rule 1)
thus ?thesis by simp
qed simp
with \<open>lt f \<prec>\<^sub>t lt b\<close> show False by simp
qed
ultimately have "\<not> is_sig_red (\<prec>\<^sub>t) (=) (set (take (Suc i) bs) \<union> set (drop (Suc i) bs)) b"
by (simp add: is_sig_red_Un)
thus "\<not> is_sig_red (\<prec>\<^sub>t) (=) (set bs) b" by (metis append_take_drop_id set_append)
qed fact+
qed (simp add: term_simps)
qed
lemma not_sig_crit:
assumes "rb_aux_inv (bs, ss, p # ps)" and "\<not> sig_crit bs (new_syz_sigs ss bs p) p" and "b \<in> set bs"
shows "lt b \<prec>\<^sub>t sig_of_pair p"
proof (rule sum_prodE)
fix x y
assume p: "p = Inl (x, y)"
have "p \<in> set (p # ps)" by simp
hence "Inl (x, y) \<in> set (p # ps)" by (simp only: p)
define t1 where "t1 = punit.lt (rep_list x)"
define t2 where "t2 = punit.lt (rep_list y)"
define u where "u = fst (spair_sigs x y)"
define v where "v = snd (spair_sigs x y)"
have u: "u = (lcs t1 t2 - t1) \<oplus> lt x" by (simp add: u_def spair_sigs_def t1_def t2_def Let_def)
have v: "v = (lcs t1 t2 - t2) \<oplus> lt y" by (simp add: v_def spair_sigs_def t1_def t2_def Let_def)
have spair_sigs: "spair_sigs x y = (u, v)" by (simp add: u_def v_def)
with assms(2) have "\<not> is_rewritable bs x u" and "\<not> is_rewritable bs y v"
by (simp_all add: p)
from assms(1) \<open>Inl (x, y) \<in> set (p # ps)\<close> have x_in: "x \<in> set bs" and y_in: "y \<in> set bs"
and "is_regular_spair x y" by (rule rb_aux_inv_D3)+
from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
from inv1 have "0 \<notin> rep_list ` set bs" by (rule rb_aux_inv1_D2)
with x_in y_in have "rep_list x \<noteq> 0" and "rep_list y \<noteq> 0" by auto
hence "x \<noteq> 0" and "y \<noteq> 0" by (auto simp: rep_list_zero)
from inv1 have sorted: "sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) bs" by (rule rb_aux_inv1_D3)
from x_in obtain i1 where "i1 < length bs" and x: "x = bs ! i1" by (metis in_set_conv_nth)
from y_in obtain i2 where "i2 < length bs" and y: "y = bs ! i2" by (metis in_set_conv_nth)
have "lt b \<noteq> sig_of_pair p"
proof
assume lt_b: "lt b = sig_of_pair p"
from inv1 have crw: "is_canon_rewriter rword (set bs) (lt b) b" using assms(3)
by (rule is_canon_rewriterI_eq_sig)
show False
proof (rule ord_term_lin.linorder_cases)
assume "u \<prec>\<^sub>t v"
hence "lt b = v" by (auto simp: lt_b p spair_sigs ord_term_lin.max_def)
with crw have crw_b: "is_canon_rewriter rword (set bs) v b" by simp
from v have "lt y adds\<^sub>t v" by (rule adds_termI)
hence "is_canon_rewriter rword (set bs) v y"
using inv1 y_in \<open>y \<noteq> 0\<close> \<open>\<not> is_rewritable bs y v\<close> is_rewritableI_is_canon_rewriter by blast
with inv1 crw_b have "b = y" by (rule canon_rewriter_unique)
with \<open>lt b = v\<close> have "lt y = v" by simp
from inv1 \<open>i2 < length bs\<close> have "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i2) bs)) (bs ! i2)"
by (rule rb_aux_inv1_D4)
moreover have "is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i2) bs)) (bs ! i2)"
proof (rule is_sig_red_singletonD)
have "is_sig_red (\<prec>\<^sub>t) (=) {x} y"
proof (rule is_sig_red_top_addsI)
from \<open>lt y = v\<close> have "(lcs t1 t2 - t2) \<oplus> lt y = lt y" by (simp only: v)
also have "... = 0 \<oplus> lt y" by (simp only: term_simps)
finally have "lcs t1 t2 - t2 = 0" by (simp only: splus_right_canc)
hence "lcs t1 t2 = t2" by (metis (full_types) add.left_neutral adds_minus adds_lcs_2)
with adds_lcs[of t1 t2] show "punit.lt (rep_list x) adds punit.lt (rep_list y)"
by (simp only: t1_def t2_def)
next
from \<open>u \<prec>\<^sub>t v\<close> show "punit.lt (rep_list y) \<oplus> lt x \<prec>\<^sub>t punit.lt (rep_list x) \<oplus> lt y"
by (simp add: t1_def t2_def u v term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
qed (simp|fact)+
thus "is_sig_red (\<prec>\<^sub>t) (\<preceq>) {x} (bs ! i2)" by (simp add: y is_sig_red_top_tail_cases)
next
have "lt x \<preceq>\<^sub>t 0 \<oplus> lt x" by (simp only: term_simps)
also have "... \<preceq>\<^sub>t u" unfolding u using zero_min by (rule splus_mono_left)
also have "... \<prec>\<^sub>t v" by fact
finally have *: "lt (bs ! i1) \<prec>\<^sub>t lt (bs ! i2)" by (simp only: \<open>lt y = v\<close> x y[symmetric])
have "i2 < i1"
proof (rule linorder_cases)
assume "i1 < i2"
with sorted have "lt (bs ! i2) \<prec>\<^sub>t lt (bs ! i1)" using \<open>i2 < length bs\<close>
by (rule sorted_wrt_nth_less)
with * show ?thesis by simp
next
assume "i1 = i2"
with * show ?thesis by simp
qed
hence "Suc i2 \<le> i1" by simp
thus "x \<in> set (drop (Suc i2) bs)" unfolding x using \<open>i1 < length bs\<close> by (rule nth_in_set_dropI)
qed
ultimately show ?thesis ..
next
assume "v \<prec>\<^sub>t u"
hence "lt b = u" by (auto simp: lt_b p spair_sigs ord_term_lin.max_def)
with crw have crw_b: "is_canon_rewriter rword (set bs) u b" by simp
from u have "lt x adds\<^sub>t u" by (rule adds_termI)
hence "is_canon_rewriter rword (set bs) u x"
using inv1 x_in \<open>x \<noteq> 0\<close> \<open>\<not> is_rewritable bs x u\<close> is_rewritableI_is_canon_rewriter by blast
with inv1 crw_b have "b = x" by (rule canon_rewriter_unique)
with \<open>lt b = u\<close> have "lt x = u" by simp
from inv1 \<open>i1 < length bs\<close> have "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i1) bs)) (bs ! i1)"
by (rule rb_aux_inv1_D4)
moreover have "is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i1) bs)) (bs ! i1)"
proof (rule is_sig_red_singletonD)
have "is_sig_red (\<prec>\<^sub>t) (=) {y} x"
proof (rule is_sig_red_top_addsI)
from \<open>lt x = u\<close> have "(lcs t1 t2 - t1) \<oplus> lt x = lt x" by (simp only: u)
also have "... = 0 \<oplus> lt x" by (simp only: term_simps)
finally have "lcs t1 t2 - t1 = 0" by (simp only: splus_right_canc)
hence "lcs t1 t2 = t1" by (metis (full_types) add.left_neutral adds_minus adds_lcs)
with adds_lcs_2[of t2 t1] show "punit.lt (rep_list y) adds punit.lt (rep_list x)"
by (simp only: t1_def t2_def)
next
from \<open>v \<prec>\<^sub>t u\<close> show "punit.lt (rep_list x) \<oplus> lt y \<prec>\<^sub>t punit.lt (rep_list y) \<oplus> lt x"
by (simp add: t1_def t2_def u v term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
qed (simp|fact)+
thus "is_sig_red (\<prec>\<^sub>t) (\<preceq>) {y} (bs ! i1)" by (simp add: x is_sig_red_top_tail_cases)
next
have "lt y \<preceq>\<^sub>t 0 \<oplus> lt y" by (simp only: term_simps)
also have "... \<preceq>\<^sub>t v" unfolding v using zero_min by (rule splus_mono_left)
also have "... \<prec>\<^sub>t u" by fact
finally have *: "lt (bs ! i2) \<prec>\<^sub>t lt (bs ! i1)" by (simp only: \<open>lt x = u\<close> y x[symmetric])
have "i1 < i2"
proof (rule linorder_cases)
assume "i2 < i1"
with sorted have "lt (bs ! i1) \<prec>\<^sub>t lt (bs ! i2)" using \<open>i1 < length bs\<close>
by (rule sorted_wrt_nth_less)
with * show ?thesis by simp
next
assume "i2 = i1"
with * show ?thesis by simp
qed
hence "Suc i1 \<le> i2" by simp
thus "y \<in> set (drop (Suc i1) bs)" unfolding y using \<open>i2 < length bs\<close> by (rule nth_in_set_dropI)
qed
ultimately show ?thesis ..
next
assume "u = v"
hence "punit.lt (rep_list x) \<oplus> lt y = punit.lt (rep_list y) \<oplus> lt x"
by (simp add: t1_def t2_def u v term_is_le_rel_minus_minus adds_lcs adds_lcs_2)
moreover from \<open>is_regular_spair x y\<close>
have "punit.lt (rep_list y) \<oplus> lt x \<noteq> punit.lt (rep_list x) \<oplus> lt y" by (rule is_regular_spairD3)
ultimately show ?thesis by simp
qed
qed
moreover from assms(1, 3) \<open>p \<in> set (p # ps)\<close> have "lt b \<preceq>\<^sub>t sig_of_pair p" by (rule rb_aux_inv_D7)
ultimately show ?thesis by simp
next
fix j
assume p: "p = Inr j"
have "Inr j \<in> set (p # ps)" by (simp add: p)
with assms(1) have "lt b \<prec>\<^sub>t term_of_pair (0, j)" using assms(3) by (rule rb_aux_inv_D4)
thus ?thesis by (simp add: p)
qed
context
assumes fs_distinct: "distinct fs"
assumes fs_nonzero: "0 \<notin> set fs"
begin
lemma rep_list_monomial': "rep_list (monomial 1 (term_of_pair (0, j))) = ((fs ! j) when j < length fs)"
by (simp add: rep_list_monomial fs_distinct term_simps)
lemma new_syz_sigs_is_syz_sig:
assumes "rb_aux_inv (bs, ss, p # ps)" and "v \<in> set (new_syz_sigs ss bs p)"
shows "is_syz_sig dgrad v"
proof (rule sum_prodE)
fix a b
assume "p = Inl (a, b)"
with assms(2) have "v \<in> set ss" by simp
with assms(1) show ?thesis by (rule rb_aux_inv_D2)
next
fix j
assume p: "p = Inr j"
let ?f = "\<lambda>b. term_of_pair (punit.lt (rep_list b), j)"
let ?a = "monomial (1::'b) (term_of_pair (0, j))"
from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
have "Inr j \<in> set (p # ps)" by (simp add: p)
with assms(1) have "j < length fs" by (rule rb_aux_inv_D4)
hence a: "rep_list ?a = fs ! j" by (simp add: rep_list_monomial')
show ?thesis
proof (cases "is_pot_ord")
case True
with assms(2) have "v \<in> set (filter_min_append (adds\<^sub>t) ss (filter_min (adds\<^sub>t) (map ?f bs)))"
by (simp add: p)
hence "v \<in> set ss \<union> ?f ` set bs" using filter_min_append_subset filter_min_subset by fastforce
thus ?thesis
proof
assume "v \<in> set ss"
with assms(1) show ?thesis by (rule rb_aux_inv_D2)
next
assume "v \<in> ?f ` set bs"
then obtain b where "b \<in> set bs" and "v = ?f b" ..
have comp_b: "component_of_term (lt b) < component_of_term (lt ?a)"
proof (rule ccontr)
have *: "pp_of_term (term_of_pair (0, j)) \<preceq> pp_of_term (lt b)"
by (simp add: pp_of_term_of_pair zero_min)
assume "\<not> component_of_term (lt b) < component_of_term (lt ?a)"
hence "component_of_term (term_of_pair (0, j)) \<le> component_of_term (lt b)"
by (simp add: lt_monomial)
with * have "term_of_pair (0, j) \<preceq>\<^sub>t lt b" by (rule ord_termI)
moreover from assms(1) \<open>Inr j \<in> set (p # ps)\<close> \<open>b \<in> set bs\<close> have "lt b \<prec>\<^sub>t term_of_pair (0, j)"
by (rule rb_aux_inv_D4)
ultimately show False by simp
qed
have "v = punit.lt (rep_list b) \<oplus> lt ?a"
by (simp add: \<open>v = ?f b\<close> a lt_monomial splus_def term_simps)
also have "... = ord_term_lin.max (punit.lt (rep_list b) \<oplus> lt ?a) (punit.lt (rep_list ?a) \<oplus> lt b)"
proof -
have "component_of_term (punit.lt (rep_list ?a) \<oplus> lt b) = component_of_term (lt b)"
by (simp only: term_simps)
also have "... < component_of_term (lt ?a)" by (fact comp_b)
also have "... = component_of_term (punit.lt (rep_list b) \<oplus> lt ?a)"
by (simp only: term_simps)
finally have "component_of_term (punit.lt (rep_list ?a) \<oplus> lt b) <
component_of_term (punit.lt (rep_list b) \<oplus> lt ?a)" .
with True have "punit.lt (rep_list ?a) \<oplus> lt b \<prec>\<^sub>t punit.lt (rep_list b) \<oplus> lt ?a"
by (rule is_pot_ordD)
thus ?thesis by (auto simp: ord_term_lin.max_def)
qed
finally have v: "v = ord_term_lin.max (punit.lt (rep_list b) \<oplus> lt ?a) (punit.lt (rep_list ?a) \<oplus> lt b)" .
show ?thesis unfolding v using dgrad
proof (rule Koszul_syz_is_syz_sig)
from inv1 have "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
with \<open>b \<in> set bs\<close> show "b \<in> dgrad_sig_set dgrad" ..
next
show "?a \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 \<open>j < length fs\<close>)
next
from inv1 have "0 \<notin> rep_list ` set bs" by (rule rb_aux_inv1_D2)
with \<open>b \<in> set bs\<close> show "rep_list b \<noteq> 0" by fastforce
next
from \<open>j < length fs\<close> have "fs ! j \<in> set fs" by (rule nth_mem)
with fs_nonzero show "rep_list ?a \<noteq> 0" by (auto simp: a)
qed (fact comp_b)
qed
next
case False
with assms(2) have "v \<in> set ss" by (simp add: p)
with assms(1) show ?thesis by (rule rb_aux_inv_D2)
qed
qed
lemma new_syz_sigs_minimal:
assumes "\<And>u' v'. u' \<in> set ss \<Longrightarrow> v' \<in> set ss \<Longrightarrow> u' adds\<^sub>t v' \<Longrightarrow> u' = v'"
assumes "u \<in> set (new_syz_sigs ss bs p)" and "v \<in> set (new_syz_sigs ss bs p)" and "u adds\<^sub>t v"
shows "u = v"
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
from assms(2, 3) have "u \<in> set ss" and "v \<in> set ss" by (simp_all add: p)
thus ?thesis using assms(4) by (rule assms(1))
next
fix j
assume p: "p = Inr j"
show ?thesis
proof (cases is_pot_ord)
case True
have "transp (adds\<^sub>t)" by (rule transpI, drule adds_term_trans)
define ss' where "ss' = filter_min (adds\<^sub>t) (map (\<lambda>b. term_of_pair (punit.lt (rep_list b), j)) bs)"
note assms(1)
moreover have "u' = v'" if "u' \<in> set ss'" and "v' \<in> set ss'" and "u' adds\<^sub>t v'" for u' v'
using \<open>transp (adds\<^sub>t)\<close> that unfolding ss'_def by (rule filter_min_minimal)
moreover from True assms(2, 3) have "u \<in> set (filter_min_append (adds\<^sub>t) ss ss')"
and "v \<in> set (filter_min_append (adds\<^sub>t) ss ss')" by (simp_all add: p ss'_def)
ultimately show ?thesis using assms(4) by (rule filter_min_append_minimal)
next
case False
with assms(2, 3) have "u \<in> set ss" and "v \<in> set ss" by (simp_all add: p)
thus ?thesis using assms(4) by (rule assms(1))
qed
qed
lemma new_syz_sigs_distinct:
assumes "distinct ss"
shows "distinct (new_syz_sigs ss bs p)"
proof (rule sum_prodE)
fix a b
assume "p = Inl (a, b)"
with assms show ?thesis by simp
next
fix j
assume p: "p = Inr j"
show ?thesis
proof (cases is_pot_ord)
case True
define ss' where "ss' = filter_min (adds\<^sub>t) (map (\<lambda>b. term_of_pair (punit.lt (rep_list b), j)) bs)"
from adds_term_refl have "reflp (adds\<^sub>t)" by (rule reflpI)
moreover note assms
moreover have "distinct ss'" unfolding ss'_def using \<open>reflp (adds\<^sub>t)\<close> by (rule filter_min_distinct)
ultimately have "distinct (filter_min_append (adds\<^sub>t) ss ss')" by (rule filter_min_append_distinct)
thus ?thesis by (simp add: p ss'_def True)
next
case False
with assms show ?thesis by (simp add: p)
qed
qed
lemma sig_crit'I_sig_crit:
assumes "rb_aux_inv (bs, ss, p # ps)" and "sig_crit bs (new_syz_sigs ss bs p) p"
shows "sig_crit' bs p"
proof -
have rl: "is_syz_sig dgrad u"
if "is_pred_syz (new_syz_sigs ss bs p) u" and "dgrad (pp_of_term u) \<le> dgrad_max dgrad" for u
proof -
from that(1) obtain s where "s \<in> set (new_syz_sigs ss bs p)" and adds: "s adds\<^sub>t u"
unfolding is_pred_syz_def ..
from assms(1) this(1) have "is_syz_sig dgrad s" by (rule new_syz_sigs_is_syz_sig)
with dgrad show ?thesis using adds that(2) by (rule is_syz_sig_adds)
qed
from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence bs_sub: "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
show ?thesis
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
hence "Inl (a, b) \<in> set (p # ps)" by simp
with assms(1) have "a \<in> set bs" and "b \<in> set bs" by (rule rb_aux_inv_D3)+
with bs_sub have a_in: "a \<in> dgrad_sig_set dgrad" and b_in: "b \<in> dgrad_sig_set dgrad" by fastforce+
define t1 where "t1 = punit.lt (rep_list a)"
define t2 where "t2 = punit.lt (rep_list b)"
define u where "u = fst (spair_sigs a b)"
define v where "v = snd (spair_sigs a b)"
from dgrad a_in have "dgrad t1 \<le> dgrad_max dgrad" unfolding t1_def by (rule dgrad_sig_setD_rep_list_lt)
moreover from dgrad b_in have "dgrad t2 \<le> dgrad_max dgrad"
unfolding t2_def by (rule dgrad_sig_setD_rep_list_lt)
ultimately have "ord_class.max (dgrad t1) (dgrad t2) \<le> dgrad_max dgrad" by simp
with dickson_grading_lcs[OF dgrad] have "dgrad (lcs t1 t2) \<le> dgrad_max dgrad" by (rule le_trans)
have u: "u = (lcs t1 t2 - t1) \<oplus> lt a" by (simp add: u_def spair_sigs_def t1_def t2_def Let_def)
have v: "v = (lcs t1 t2 - t2) \<oplus> lt b" by (simp add: v_def spair_sigs_def t1_def t2_def Let_def)
have 1: "spair_sigs a b = (u, v)" by (simp add: u_def v_def)
from assms(2) have "(is_pred_syz (new_syz_sigs ss bs p) u \<or> is_pred_syz (new_syz_sigs ss bs p) v) \<or>
(is_rewritable bs a u \<or> is_rewritable bs b v)" by (simp add: p 1)
thus ?thesis
proof
assume "is_pred_syz (new_syz_sigs ss bs p) u \<or> is_pred_syz (new_syz_sigs ss bs p) v"
thus ?thesis
proof
assume "is_pred_syz (new_syz_sigs ss bs p) u"
moreover have "dgrad (pp_of_term u) \<le> dgrad_max dgrad"
proof (simp add: u term_simps dickson_gradingD1[OF dgrad], rule)
from dgrad adds_lcs have "dgrad (lcs t1 t2 - t1) \<le> dgrad (lcs t1 t2)"
by (rule dickson_grading_minus)
also have "... \<le> dgrad_max dgrad" by fact
finally show "dgrad (lcs t1 t2 - t1) \<le> dgrad_max dgrad" .
next
from a_in show "dgrad (lp a) \<le> dgrad_max dgrad" by (rule dgrad_sig_setD_lp)
qed
ultimately have "is_syz_sig dgrad u" by (rule rl)
thus ?thesis by (simp add: p 1)
next
assume "is_pred_syz (new_syz_sigs ss bs p) v"
moreover have "dgrad (pp_of_term v) \<le> dgrad_max dgrad"
proof (simp add: v term_simps dickson_gradingD1[OF dgrad], rule)
from dgrad adds_lcs_2 have "dgrad (lcs t1 t2 - t2) \<le> dgrad (lcs t1 t2)"
by (rule dickson_grading_minus)
also have "... \<le> dgrad_max dgrad" by fact
finally show "dgrad (lcs t1 t2 - t2) \<le> dgrad_max dgrad" .
next
from b_in show "dgrad (lp b) \<le> dgrad_max dgrad" by (rule dgrad_sig_setD_lp)
qed
ultimately have "is_syz_sig dgrad v" by (rule rl)
thus ?thesis by (simp add: p 1)
qed
next
assume "is_rewritable bs a u \<or> is_rewritable bs b v"
thus ?thesis by (simp add: p 1)
qed
next
fix j
assume "p = Inr j"
with assms(2) have "is_pred_syz (new_syz_sigs ss bs p) (term_of_pair (0, j))" by simp
moreover have "dgrad (pp_of_term (term_of_pair (0, j))) \<le> dgrad_max dgrad"
by (simp add: pp_of_term_of_pair dgrad_max_0)
ultimately have "is_syz_sig dgrad (term_of_pair (0, j))" by (rule rl)
thus ?thesis by (simp add: \<open>p = Inr j\<close>)
qed
qed
lemma rb_aux_inv_preserved_0:
assumes "rb_aux_inv (bs, ss, p # ps)"
and "\<And>s. s \<in> set ss' \<Longrightarrow> is_syz_sig dgrad s"
and "\<And>a b. a \<in> set bs \<Longrightarrow> b \<in> set bs \<Longrightarrow> is_regular_spair a b \<Longrightarrow> Inl (a, b) \<notin> set ps \<Longrightarrow>
Inl (b, a) \<notin> set ps \<Longrightarrow> \<not> is_RB_in dgrad rword (set bs) (lt (spair a b)) \<Longrightarrow>
\<exists>q\<in>set ps. sig_of_pair q = lt (spair a b) \<and> \<not> sig_crit' bs q"
and "\<And>j. j < length fs \<Longrightarrow> p = Inr j \<Longrightarrow> Inr j \<notin> set ps \<Longrightarrow> is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) \<and>
rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)"
shows "rb_aux_inv (bs, ss', ps)"
proof -
from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
show ?thesis unfolding rb_aux_inv.simps
proof (intro conjI ballI allI impI)
fix s
assume "s \<in> set ss'"
thus "is_syz_sig dgrad s" by (rule assms(2))
next
fix q1 q2
assume "Inl (q1, q2) \<in> set ps"
hence "Inl (q1, q2) \<in> set (p # ps)" by simp
with assms(1) show "is_regular_spair q1 q2" and "q1 \<in> set bs" and "q2 \<in> set bs"
by (rule rb_aux_inv_D3)+
next
fix j
assume "Inr j \<in> set ps"
hence "Inr j \<in> set (p # ps)" by simp
with assms(1) have "j < length fs" and "length (filter (\<lambda>q. sig_of_pair q = term_of_pair (0, j)) (p # ps)) \<le> 1"
by (rule rb_aux_inv_D4)+
have "length (filter (\<lambda>q. sig_of_pair q = term_of_pair (0, j)) ps) \<le>
length (filter (\<lambda>q. sig_of_pair q = term_of_pair (0, j)) (p # ps))" by simp
also have "... \<le> 1" by fact
finally show "length (filter (\<lambda>q. sig_of_pair q = term_of_pair (0, j)) ps) \<le> 1" .
show "j < length fs" by fact
fix b
assume "b \<in> set bs"
with assms(1) \<open>Inr j \<in> set (p # ps)\<close> show "lt b \<prec>\<^sub>t term_of_pair (0, j)" by (rule rb_aux_inv_D4)
next
from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
thus "sorted_wrt pair_ord ps" by simp
next
fix q
assume "q \<in> set ps"
from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
hence "\<And>p'. p' \<in> set ps \<Longrightarrow> sig_of_pair p \<preceq>\<^sub>t sig_of_pair p'" by (simp add: pair_ord_def)
with \<open>q \<in> set ps\<close> have 1: "sig_of_pair p \<preceq>\<^sub>t sig_of_pair q" by blast
{
fix b1 b2
note assms(1)
moreover from \<open>q \<in> set ps\<close> have "q \<in> set (p # ps)" by simp
moreover assume "b1 \<in> set bs" and "b2 \<in> set bs" and "is_regular_spair b1 b2"
and 2: "sig_of_pair q \<prec>\<^sub>t lt (spair b1 b2)"
ultimately show "Inl (b1, b2) \<in> set ps \<or> Inl (b2, b1) \<in> set ps"
proof (rule rb_aux_inv_D6_1)
assume "Inl (b1, b2) \<in> set (p # ps)"
moreover from 1 2 have "sig_of_pair p \<prec>\<^sub>t lt (spair b1 b2)" by simp
ultimately have "Inl (b1, b2) \<in> set ps"
by (auto simp: sig_of_spair \<open>is_regular_spair b1 b2\<close> simp del: sig_of_pair.simps)
thus ?thesis ..
next
assume "Inl (b2, b1) \<in> set (p # ps)"
moreover from 1 2 have "sig_of_pair p \<prec>\<^sub>t lt (spair b1 b2)" by simp
ultimately have "Inl (b2, b1) \<in> set ps"
by (auto simp: sig_of_spair \<open>is_regular_spair b1 b2\<close> sig_of_spair_commute simp del: sig_of_pair.simps)
thus ?thesis ..
qed
}
{
fix j
note assms(1)
moreover from \<open>q \<in> set ps\<close> have "q \<in> set (p # ps)" by simp
moreover assume "j < length fs" and 2: "sig_of_pair q \<prec>\<^sub>t term_of_pair (0, j)"
ultimately have "Inr j \<in> set (p # ps)" by (rule rb_aux_inv_D6_2)
moreover from 1 2 have "sig_of_pair p \<prec>\<^sub>t sig_of_pair (Inr j)" by simp
ultimately show "Inr j \<in> set ps" by auto
}
next
fix b q
assume "b \<in> set bs" and "q \<in> set ps"
hence "b \<in> set bs" and "q \<in> set (p # ps)" by simp_all
with assms(1) show "lt b \<preceq>\<^sub>t sig_of_pair q" by (rule rb_aux_inv_D7)
next
fix j
assume "j < length fs" and "Inr j \<notin> set ps"
have "is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) \<and>
rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)"
proof (cases "p = Inr j")
case True
with \<open>j < length fs\<close> show ?thesis using \<open>Inr j \<notin> set ps\<close> by (rule assms(4))
next
case False
with \<open>Inr j \<notin> set ps\<close> have "Inr j \<notin> set (p # ps)" by simp
with assms(1) \<open>j < length fs\<close> rb_aux_inv_D9 show ?thesis by blast
qed
thus "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))"
and "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)" by simp_all
qed (fact, rule assms(3))
qed
lemma rb_aux_inv_preserved_1:
assumes "rb_aux_inv (bs, ss, p # ps)" and "sig_crit bs (new_syz_sigs ss bs p) p"
shows "rb_aux_inv (bs, new_syz_sigs ss bs p, ps)"
proof -
from assms(1) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence bs_sub: "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
from assms(1, 2) have "sig_crit' bs p" by (rule sig_crit'I_sig_crit)
from assms(1) show ?thesis
proof (rule rb_aux_inv_preserved_0)
fix s
assume "s \<in> set (new_syz_sigs ss bs p)"
with assms(1) show "is_syz_sig dgrad s" by (rule new_syz_sigs_is_syz_sig)
next
fix a b
assume 1: "a \<in> set bs" and 2: "b \<in> set bs" and 3: "is_regular_spair a b" and 4: "Inl (a, b) \<notin> set ps"
and 5: "Inl (b, a) \<notin> set ps" and 6: "\<not> is_RB_in dgrad rword (set bs) (lt (spair a b))"
from assms(1, 2) have "sig_crit' bs p" by (rule sig_crit'I_sig_crit)
show "\<exists>q\<in>set ps. sig_of_pair q = lt (spair a b) \<and> \<not> sig_crit' bs q"
proof (cases "p = Inl (a, b) \<or> p = Inl (b, a)")
case True
hence sig_of_p: "lt (spair a b) = sig_of_pair p"
proof
assume p: "p = Inl (a, b)"
from 3 show ?thesis by (simp only: p sig_of_spair)
next
assume p: "p = Inl (b, a)"
from 3 have "is_regular_spair b a" by (rule is_regular_spair_sym)
thus ?thesis by (simp only: p sig_of_spair spair_comm[of a] lt_uminus)
qed
note assms(1)
moreover have "is_RB_upt dgrad rword (set bs) (lt (spair a b))" unfolding sig_of_p
using assms(1) by (rule rb_aux_inv_is_RB_upt_Cons)
moreover have "dgrad (lp (spair a b)) \<le> dgrad_max dgrad"
proof (rule dgrad_sig_setD_lp, rule dgrad_sig_set_closed_spair, fact dgrad)
from \<open>a \<in> set bs\<close> bs_sub show "a \<in> dgrad_sig_set dgrad" ..
next
from \<open>b \<in> set bs\<close> bs_sub show "b \<in> dgrad_sig_set dgrad" ..
qed
moreover obtain c where crw: "is_canon_rewriter rword (set bs) (lt (spair a b)) c"
proof (rule ord_term_lin.linorder_cases)
from 3 have "rep_list b \<noteq> 0" by (rule is_regular_spairD2)
moreover assume "punit.lt (rep_list b) \<oplus> lt a \<prec>\<^sub>t punit.lt (rep_list a) \<oplus> lt b"
ultimately have "lt (spair b a) = (lcs (punit.lt (rep_list b)) (punit.lt (rep_list a)) - punit.lt (rep_list b)) \<oplus> lt b"
by (rule lt_spair)
hence "lt (spair a b) = (lcs (punit.lt (rep_list b)) (punit.lt (rep_list a)) - punit.lt (rep_list b)) \<oplus> lt b"
by (simp add: spair_comm[of a])
hence "lt b adds\<^sub>t lt (spair a b)" by (rule adds_termI)
from \<open>rep_list b \<noteq> 0\<close> have "b \<noteq> 0" by (auto simp: rep_list_zero)
show ?thesis by (rule is_rewrite_ord_finite_canon_rewriterE, fact rword, fact finite_set, fact+)
next
from 3 have "rep_list a \<noteq> 0" by (rule is_regular_spairD1)
moreover assume "punit.lt (rep_list a) \<oplus> lt b \<prec>\<^sub>t punit.lt (rep_list b) \<oplus> lt a"
ultimately have "lt (spair a b) = (lcs (punit.lt (rep_list a)) (punit.lt (rep_list b)) - punit.lt (rep_list a)) \<oplus> lt a"
by (rule lt_spair)
hence "lt a adds\<^sub>t lt (spair a b)" by (rule adds_termI)
from \<open>rep_list a \<noteq> 0\<close> have "a \<noteq> 0" by (auto simp: rep_list_zero)
show ?thesis by (rule is_rewrite_ord_finite_canon_rewriterE, fact rword, fact finite_set, fact+)
next
from 3 have "punit.lt (rep_list b) \<oplus> lt a \<noteq> punit.lt (rep_list a) \<oplus> lt b"
by (rule is_regular_spairD3)
moreover assume "punit.lt (rep_list b) \<oplus> lt a = punit.lt (rep_list a) \<oplus> lt b"
ultimately show ?thesis ..
qed
moreover from 6 have "\<not> is_syz_sig dgrad (lt (spair a b))" by (simp add: is_RB_in_def)
moreover from 6 crw have "is_sig_red (\<prec>\<^sub>t) (=) (set bs) (monom_mult 1 (lp (spair a b) - lp c) c)"
by (simp add: is_RB_in_def)
ultimately obtain x y where 7: "x \<in> set bs" and 8: "y \<in> set bs" and 9: "is_regular_spair x y"
and 10: "lt (spair x y) = lt (spair a b)" and 11: "\<not> sig_crit' bs (Inl (x, y))"
by (rule lemma_12)
from this(5) \<open>sig_crit' bs p\<close> have "Inl (x, y) \<noteq> p" and "Inl (y, x) \<noteq> p"
by (auto simp only: sig_crit'_sym)
show ?thesis
proof (cases "Inl (x, y) \<in> set ps \<or> Inl (y, x) \<in> set ps")
case True
thus ?thesis
proof
assume "Inl (x, y) \<in> set ps"
show ?thesis
proof (intro bexI conjI)
show "sig_of_pair (Inl (x, y)) = lt (spair a b)" by (simp only: sig_of_spair 9 10)
qed fact+
next
assume "Inl (y, x) \<in> set ps"
show ?thesis
proof (intro bexI conjI)
from 9 have "is_regular_spair y x" by (rule is_regular_spair_sym)
thus "sig_of_pair (Inl (y, x)) = lt (spair a b)"
by (simp only: sig_of_spair spair_comm[of y] lt_uminus 10)
next
from 11 show "\<not> sig_crit' bs (Inl (y, x))" by (auto simp only: sig_crit'_sym)
qed fact
qed
next
case False
note assms(1) 7 8 9
moreover from False \<open>Inl (x, y) \<noteq> p\<close> \<open>Inl (y, x) \<noteq> p\<close> have "Inl (x, y) \<notin> set (p # ps)"
and "Inl (y, x) \<notin> set (p # ps)" by auto
moreover from 6 have "\<not> is_RB_in dgrad rword (set bs) (lt (spair x y))" by (simp add: 10)
ultimately obtain q where 12: "q \<in> set (p # ps)" and 13: "sig_of_pair q = lt (spair x y)"
and 14: "\<not> sig_crit' bs q" by (rule rb_aux_inv_D8)
from 12 14 \<open>sig_crit' bs p\<close> have "q \<in> set ps" by auto
with 13 14 show ?thesis unfolding 10 by blast
qed
next
case False
with 4 5 have "Inl (a, b) \<notin> set (p # ps)" and "Inl (b, a) \<notin> set (p # ps)" by auto
with assms(1) 1 2 3 obtain q where 7: "q \<in> set (p # ps)" and 8: "sig_of_pair q = lt (spair a b)"
and 9: "\<not> sig_crit' bs q" using 6 by (rule rb_aux_inv_D8)
from 7 9 \<open>sig_crit' bs p\<close> have "q \<in> set ps" by auto
with 8 9 show ?thesis by blast
qed
next
fix j
assume "j < length fs"
assume p: "p = Inr j"
with \<open>sig_crit' bs p\<close> have "is_syz_sig dgrad (term_of_pair (0, j))" by simp
hence "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))" by (rule is_RB_inI2)
moreover have "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)"
proof (rule sig_red_zero_idealI, rule syzygy_crit)
from assms(1) have "is_RB_upt dgrad rword (set bs) (sig_of_pair p)"
by (rule rb_aux_inv_is_RB_upt_Cons)
with dgrad have "is_sig_GB_upt dgrad (set bs) (sig_of_pair p)"
by (rule is_RB_upt_is_sig_GB_upt)
thus "is_sig_GB_upt dgrad (set bs) (term_of_pair (0, j))" by (simp add: p)
next
show "monomial 1 (term_of_pair (0, j)) \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 \<open>j < length fs\<close>)
next
show "lt (monomial (1::'b) (term_of_pair (0, j))) = term_of_pair (0, j)" by (simp add: lt_monomial)
qed (fact dgrad, fact)
ultimately show "is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) \<and>
rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)" ..
qed
qed
lemma rb_aux_inv_preserved_2:
assumes "rb_aux_inv (bs, ss, p # ps)" and "rep_list (sig_trd bs (poly_of_pair p)) = 0"
shows "rb_aux_inv (bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps)"
proof -
let ?p = "sig_trd bs (poly_of_pair p)"
have 0: "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (poly_of_pair p) ?p"
by (rule sig_trd_red_rtrancl)
hence eq: "lt ?p = lt (poly_of_pair p)" by (rule sig_red_regular_rtrancl_lt)
from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
have *: "is_syz_sig dgrad (lt (poly_of_pair p))"
proof (rule is_syz_sigI)
have "poly_of_pair p \<noteq> 0" by (rule pair_list_nonzero, fact, simp)
hence "lc (poly_of_pair p) \<noteq> 0" by (rule lc_not_0)
moreover from 0 have "lc ?p = lc (poly_of_pair p)" by (rule sig_red_regular_rtrancl_lc)
ultimately have "lc ?p \<noteq> 0" by simp
thus "?p \<noteq> 0" by (simp add: lc_eq_zero_iff)
next
note dgrad(1)
moreover from inv1 have "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
moreover have "poly_of_pair p \<in> dgrad_sig_set dgrad" by (rule pair_list_dgrad_sig_set, fact, simp)
ultimately show "?p \<in> dgrad_sig_set dgrad" using 0 by (rule dgrad_sig_set_closed_sig_red_rtrancl)
qed (fact eq, fact assms(2))
hence rb: "is_RB_in dgrad rword (set bs) (lt (poly_of_pair p))" by (rule is_RB_inI2)
from assms(1) show ?thesis
proof (rule rb_aux_inv_preserved_0)
fix s
assume "s \<in> set (lt ?p # new_syz_sigs ss bs p)"
hence "s = lt (poly_of_pair p) \<or> s \<in> set (new_syz_sigs ss bs p)" by (simp add: eq)
thus "is_syz_sig dgrad s"
proof
assume "s = lt (poly_of_pair p)"
with * show ?thesis by simp
next
assume "s \<in> set (new_syz_sigs ss bs p)"
with assms(1) show ?thesis by (rule new_syz_sigs_is_syz_sig)
qed
next
fix a b
assume 1: "a \<in> set bs" and 2: "b \<in> set bs" and 3: "is_regular_spair a b" and 4: "Inl (a, b) \<notin> set ps"
and 5: "Inl (b, a) \<notin> set ps" and 6: "\<not> is_RB_in dgrad rword (set bs) (lt (spair a b))"
have "p \<in> set (p # ps)" by simp
with assms(1) have sig_of_p: "sig_of_pair p = lt (poly_of_pair p)" by (rule pair_list_sig_of_pair)
from rb 6 have neq: "lt (poly_of_pair p) \<noteq> lt (spair a b)" by auto
hence "p \<noteq> Inl (a, b)" and "p \<noteq> Inl (b, a)" by (auto simp: spair_comm[of a])
with 4 5 have "Inl (a, b) \<notin> set (p # ps)" and "Inl (b, a) \<notin> set (p # ps)" by auto
with assms(1) 1 2 3 obtain q where 7: "q \<in> set (p # ps)" and 8: "sig_of_pair q = lt (spair a b)"
and 9: "\<not> sig_crit' bs q" using 6 by (rule rb_aux_inv_D8)
from this(1, 2) neq have "q \<in> set ps" by (auto simp: sig_of_p)
thus "\<exists>q\<in>set ps. sig_of_pair q = lt (spair a b) \<and> \<not> sig_crit' bs q" using 8 9 by blast
next
fix j
assume "j < length fs"
assume p: "p = Inr j"
from rb have "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))" by (simp add: p lt_monomial)
moreover have "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)"
proof (rule sig_red_zero_idealI, rule sig_red_zeroI)
from 0 show "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) ?p" by (simp add: p)
qed fact
ultimately show "is_RB_in dgrad rword (set bs) (term_of_pair (0, j)) \<and>
rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)" ..
qed
qed
lemma rb_aux_inv_preserved_3:
assumes "rb_aux_inv (bs, ss, p # ps)" and "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
and "rep_list (sig_trd bs (poly_of_pair p)) \<noteq> 0"
shows "rb_aux_inv ((sig_trd bs (poly_of_pair p)) # bs, new_syz_sigs ss bs p,
add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
and "lt (sig_trd bs (poly_of_pair p)) \<notin> lt ` set bs"
proof -
have "p \<in> set (p # ps)" by simp
with assms(1) have sig_of_p: "sig_of_pair p = lt (poly_of_pair p)"
and p_in: "poly_of_pair p \<in> dgrad_sig_set dgrad"
by (rule pair_list_sig_of_pair, rule pair_list_dgrad_sig_set)
define p' where "p' = sig_trd bs (poly_of_pair p)"
from assms(1) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence bs_sub: "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
have p_red: "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (poly_of_pair p) p'"
and p'_irred: "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p'"
unfolding p'_def by (rule sig_trd_red_rtrancl, rule sig_trd_irred)
from dgrad bs_sub p_in p_red have p'_in: "p' \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_sig_red_rtrancl)
from p_red have lt_p': "lt p' = lt (poly_of_pair p)" by (rule sig_red_regular_rtrancl_lt)
have sig_merge: "sig_of_pair p \<preceq>\<^sub>t sig_of_pair q" if "q \<in> set (add_spairs ps bs p')" for q
using that unfolding add_spairs_def set_merge_wrt
proof
assume "q \<in> set (new_spairs bs p')"
then obtain b0 where "is_regular_spair p' b0" and "q = Inl (p', b0)" by (rule in_new_spairsE)
hence sig_of_q: "sig_of_pair q = lt (spair p' b0)" by (simp only: sig_of_spair)
show ?thesis unfolding sig_of_q sig_of_p lt_p'[symmetric] by (rule is_regular_spair_lt_ge_1, fact)
next
assume "q \<in> set ps"
moreover from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
ultimately show ?thesis by (simp add: pair_ord_def)
qed
have sig_of_p_less: "sig_of_pair p \<prec>\<^sub>t term_of_pair (0, j)" if "Inr j \<in> set ps" for j
proof (intro ord_term_lin.le_neq_trans)
from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
with \<open>Inr j \<in> set ps\<close> show "sig_of_pair p \<preceq>\<^sub>t term_of_pair (0, j)"
by (auto simp: pair_ord_def)
next
from assms(1) that show "sig_of_pair p \<noteq> term_of_pair (0, j)" by (rule Inr_in_tailD)
qed
have lt_p_gr: "lt b \<prec>\<^sub>t lt (poly_of_pair p)" if "b \<in> set bs" for b unfolding sig_of_p[symmetric]
using assms(1, 2) that by (rule not_sig_crit)
have inv1: "rb_aux_inv1 (p' # bs)" unfolding rb_aux_inv1_def
proof (intro conjI impI allI)
from bs_sub p'_in show "set (p' # bs) \<subseteq> dgrad_sig_set dgrad" by simp
next
from inv1 have "0 \<notin> rep_list ` set bs" by (rule rb_aux_inv1_D2)
with assms(3) show "0 \<notin> rep_list ` set (p' # bs)" by (simp add: p'_def)
next
from inv1 have "sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) bs" by (rule rb_aux_inv1_D3)
with lt_p_gr show "sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) (p' # bs)" by (simp add: lt_p')
next
fix i
assume "i < length (p' # bs)"
have "(\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i) (p' # bs))) ((p' # bs) ! i)) \<and>
((\<exists>j<length fs. lt ((p' # bs) ! i) = lt (monomial (1::'b) (term_of_pair (0, j))) \<and>
punit.lt (rep_list ((p' # bs) ! i)) \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) \<or>
(\<exists>p\<in>set (p' # bs). \<exists>q\<in>set (p' # bs). is_regular_spair p q \<and> rep_list (spair p q) \<noteq> 0 \<and>
lt ((p' # bs) ! i) = lt (spair p q) \<and>
punit.lt (rep_list ((p' # bs) ! i)) \<preceq> punit.lt (rep_list (spair p q)))) \<and>
is_RB_upt dgrad rword (set (drop (Suc i) (p' # bs))) (lt ((p' # bs) ! i))"
(is "?thesis1 \<and> ?thesis2 \<and> ?thesis3")
proof (cases i)
case 0
show ?thesis
proof (simp add: \<open>i = 0\<close> p'_irred del: bex_simps, rule conjI)
show "(\<exists>j<length fs. lt p' = lt (monomial (1::'b) (term_of_pair (0, j))) \<and>
punit.lt (rep_list p') \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) \<or>
(\<exists>p\<in>insert p' (set bs). \<exists>q\<in>insert p' (set bs). is_regular_spair p q \<and> rep_list (spair p q) \<noteq> 0 \<and>
lt p' = lt (spair p q) \<and> punit.lt (rep_list p') \<preceq> punit.lt (rep_list (spair p q)))"
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
have "Inl (a, b) \<in> set (p # ps)" by (simp add: p)
with assms(1) have "a \<in> set bs" and "b \<in> set bs" and "is_regular_spair a b"
by (rule rb_aux_inv_D3)+
from p_red have p'_red: "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (spair a b) p'" by (simp add: p)
hence "(punit.red (rep_list ` set bs))\<^sup>*\<^sup>* (rep_list (spair a b)) (rep_list p')"
by (rule sig_red_red_rtrancl)
moreover from assms(3) have "rep_list p' \<noteq> 0" by (simp add: p'_def)
ultimately have "rep_list (spair a b) \<noteq> 0" by (auto dest: punit.rtrancl_0)
moreover from p'_red have "lt p' = lt (spair a b)"
and "punit.lt (rep_list p') \<preceq> punit.lt (rep_list (spair a b))"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_rtrancl_lt_rep_list)
ultimately show ?thesis using \<open>a \<in> set bs\<close> \<open>b \<in> set bs\<close> \<open>is_regular_spair a b\<close> by blast
next
fix j
assume "p = Inr j"
hence "Inr j \<in> set (p # ps)" by simp
with assms(1) have "j < length fs" by (rule rb_aux_inv_D4)
from p_red have "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) p'"
by (simp add: \<open>p = Inr j\<close>)
hence "lt p' = lt (monomial (1::'b) (term_of_pair (0, j)))"
and "punit.lt (rep_list p') \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))"
by (rule sig_red_regular_rtrancl_lt, rule sig_red_rtrancl_lt_rep_list)
with \<open>j < length fs\<close> show ?thesis by blast
qed
next
from assms(1) show "is_RB_upt dgrad rword (set bs) (lt p')" unfolding lt_p' sig_of_p[symmetric]
by (rule rb_aux_inv_is_RB_upt_Cons)
qed
next
case (Suc i')
with \<open>i < length (p' # bs)\<close> have i': "i' < length bs" by simp
show ?thesis
proof (simp add: \<open>i = Suc i'\<close> del: bex_simps, intro conjI)
from inv1 i' show "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc i') bs)) (bs ! i')"
by (rule rb_aux_inv1_D4)
next
from inv1 i'
show "(\<exists>j<length fs. lt (bs ! i') = lt (monomial (1::'b) (term_of_pair (0, j))) \<and>
punit.lt (rep_list (bs ! i')) \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) \<or>
(\<exists>p\<in>insert p' (set bs). \<exists>q\<in>insert p' (set bs). is_regular_spair p q \<and> rep_list (spair p q) \<noteq> 0 \<and>
lt (bs ! i') = lt (spair p q) \<and> punit.lt (rep_list (bs ! i')) \<preceq> punit.lt (rep_list (spair p q)))"
by (auto elim!: rb_aux_inv1_E)
next
from inv1 i' show "is_RB_upt dgrad rword (set (drop (Suc i') bs)) (lt (bs ! i'))"
by (rule rb_aux_inv1_D5)
qed
qed
thus ?thesis1 and ?thesis2 and ?thesis3 by simp_all
qed
have rb: "is_RB_in dgrad rword (set (p' # bs)) (sig_of_pair p)"
proof (rule is_RB_inI1)
have "p' \<in> set (p' # bs)" by simp
with inv1 have "is_canon_rewriter rword (set (p' # bs)) (lt p') p'"
by (rule is_canon_rewriterI_eq_sig)
thus "is_canon_rewriter rword (set (p' # bs)) (sig_of_pair p) p'" by (simp add: lt_p' sig_of_p)
next
from p'_irred have "\<not> is_sig_red (\<prec>\<^sub>t) (=) (set bs) p'"
by (simp add: is_sig_red_top_tail_cases)
with sig_irred_regular_self have "\<not> is_sig_red (\<prec>\<^sub>t) (=) ({p'} \<union> set bs) p'"
by (simp add: is_sig_red_Un del: Un_insert_left)
thus "\<not> is_sig_red (\<prec>\<^sub>t) (=) (set (p' # bs)) (monom_mult 1 (pp_of_term (sig_of_pair p) - lp p') p')"
by (simp add: lt_p' sig_of_p)
qed
show "rb_aux_inv (p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p')"
unfolding rb_aux_inv.simps
proof (intro conjI ballI allI impI)
show "rb_aux_inv1 (p' # bs)" by (fact inv1)
next
fix s
assume "s \<in> set (new_syz_sigs ss bs p)"
with assms(1) show "is_syz_sig dgrad s" by (rule new_syz_sigs_is_syz_sig)
next
fix q1 q2
assume "Inl (q1, q2) \<in> set (add_spairs ps bs p')"
hence "Inl (q1, q2) \<in> set (new_spairs bs p') \<or> Inl (q1, q2) \<in> set (p # ps)"
by (auto simp: add_spairs_def set_merge_wrt)
hence "is_regular_spair q1 q2 \<and> q1 \<in> set (p' # bs) \<and> q2 \<in> set (p' # bs)"
proof
assume "Inl (q1, q2) \<in> set (new_spairs bs p')"
hence "q1 = p'" and "q2 \<in> set bs" and "is_regular_spair p' q2" by (rule in_new_spairsD)+
thus ?thesis by simp
next
assume "Inl (q1, q2) \<in> set (p # ps)"
with assms(1) have "is_regular_spair q1 q2" and "q1 \<in> set bs" and "q2 \<in> set bs"
by (rule rb_aux_inv_D3)+
thus ?thesis by simp
qed
thus "is_regular_spair q1 q2" and "q1 \<in> set (p' # bs)" and "q2 \<in> set (p' # bs)" by simp_all
next
fix j
assume "Inr j \<in> set (add_spairs ps bs p')"
hence "Inr j \<in> set ps" by (simp add: add_spairs_def set_merge_wrt Inr_not_in_new_spairs)
hence "Inr j \<in> set (p # ps)" by simp
with assms(1) show "j < length fs" by (rule rb_aux_inv_D4)
fix b
assume "b \<in> set (p' # bs)"
hence "b = p' \<or> b \<in> set bs" by simp
thus "lt b \<prec>\<^sub>t term_of_pair (0, j)"
proof
assume "b = p'"
hence "lt b = sig_of_pair p" by (simp only: lt_p' sig_of_p)
also from \<open>Inr j \<in> set ps\<close> have "... \<prec>\<^sub>t term_of_pair (0, j)" by (rule sig_of_p_less)
finally show ?thesis .
next
assume "b \<in> set bs"
with assms(1) \<open>Inr j \<in> set (p # ps)\<close> show ?thesis by (rule rb_aux_inv_D4)
qed
next
fix j
assume "Inr j \<in> set (add_spairs ps bs p')"
hence "Inr j \<in> set ps" by (simp add: add_spairs_def set_merge_wrt Inr_not_in_new_spairs)
hence "Inr j \<in> set (p # ps)" by simp
let ?P = "\<lambda>q. sig_of_pair q = term_of_pair (0, j)"
have "filter ?P (add_spairs ps bs p') = filter ?P ps" unfolding add_spairs_def
proof (rule filter_merge_wrt_2)
fix q
assume "q \<in> set (new_spairs bs p')"
then obtain b where "b \<in> set bs" and "is_regular_spair p' b" and "q = Inl (p', b)"
by (rule in_new_spairsE)
moreover assume "sig_of_pair q = term_of_pair (0, j)"
ultimately have "lt (spair p' b) = term_of_pair (0, j)"
by (simp add: sig_of_spair del: sig_of_pair.simps)
hence eq: "component_of_term (lt (spair p' b)) = j" by (simp add: component_of_term_of_pair)
have "component_of_term (lt p') < j"
proof (rule ccontr)
assume "\<not> component_of_term (lt p') < j"
hence "component_of_term (term_of_pair (0, j)) \<le> component_of_term (lt p')"
by (simp add: component_of_term_of_pair)
moreover have "pp_of_term (term_of_pair (0, j)) \<preceq> pp_of_term (lt p')"
by (simp add: pp_of_term_of_pair zero_min)
ultimately have "term_of_pair (0, j) \<preceq>\<^sub>t lt p'" using ord_termI by blast
moreover have "lt p' \<prec>\<^sub>t term_of_pair (0, j)" unfolding lt_p' sig_of_p[symmetric]
using \<open>Inr j \<in> set ps\<close> by (rule sig_of_p_less)
ultimately show False by simp
qed
moreover have "component_of_term (lt b) < j"
proof (rule ccontr)
assume "\<not> component_of_term (lt b) < j"
hence "component_of_term (term_of_pair (0, j)) \<le> component_of_term (lt b)"
by (simp add: component_of_term_of_pair)
moreover have "pp_of_term (term_of_pair (0, j)) \<preceq> pp_of_term (lt b)"
by (simp add: pp_of_term_of_pair zero_min)
ultimately have "term_of_pair (0, j) \<preceq>\<^sub>t lt b" using ord_termI by blast
moreover from assms(1) \<open>Inr j \<in> set (p # ps)\<close> \<open>b \<in> set bs\<close>
have "lt b \<prec>\<^sub>t term_of_pair (0, j)" by (rule rb_aux_inv_D4)
ultimately show False by simp
qed
ultimately have "component_of_term (lt (spair p' b)) < j"
using is_regular_spair_component_lt_cases[OF \<open>is_regular_spair p' b\<close>] by auto
thus False by (simp add: eq)
qed
hence "length (filter ?P (add_spairs ps bs p')) \<le> length (filter ?P (p # ps))"
by simp
also from assms(1) \<open>Inr j \<in> set (p # ps)\<close> have "... \<le> 1" by (rule rb_aux_inv_D4)
finally show "length (filter ?P (add_spairs ps bs p')) \<le> 1" .
next
from assms(1) have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
hence "sorted_wrt pair_ord ps" by simp
thus "sorted_wrt pair_ord (add_spairs ps bs p')" by (rule sorted_add_spairs)
next
fix q b1 b2
assume 1: "q \<in> set (add_spairs ps bs p')" and 2: "is_regular_spair b1 b2"
and 3: "sig_of_pair q \<prec>\<^sub>t lt (spair b1 b2)"
assume "b1 \<in> set (p' # bs)" and "b2 \<in> set (p' # bs)"
hence "b1 = p' \<or> b1 \<in> set bs" and "b2 = p' \<or> b2 \<in> set bs" by simp_all
thus "Inl (b1, b2) \<in> set (add_spairs ps bs p') \<or> Inl (b2, b1) \<in> set (add_spairs ps bs p')"
proof (elim disjE)
assume "b1 = p'" and "b2 = p'"
with 2 show ?thesis by (simp add: is_regular_spair_def)
next
assume "b1 = p'" and "b2 \<in> set bs"
from this(2) 2 have "Inl (b1, b2) \<in> set (new_spairs bs p')" unfolding \<open>b1 = p'\<close>
by (rule in_new_spairsI)
with 2 show ?thesis by (simp add: sig_of_spair add_spairs_def set_merge_wrt image_Un del: sig_of_pair.simps)
next
assume "b2 = p'" and "b1 \<in> set bs"
note this(2)
moreover from 2 have "is_regular_spair b2 b1" by (rule is_regular_spair_sym)
ultimately have "Inl (b2, b1) \<in> set (new_spairs bs p')" unfolding \<open>b2 = p'\<close>
by (rule in_new_spairsI)
with 2 show ?thesis
by (simp add: sig_of_spair_commute sig_of_spair add_spairs_def set_merge_wrt image_Un del: sig_of_pair.simps)
next
note assms(1) \<open>p \<in> set (p # ps)\<close>
moreover assume "b1 \<in> set bs" and "b2 \<in> set bs"
moreover note 2
moreover have 4: "sig_of_pair p \<prec>\<^sub>t lt (spair b1 b2)"
by (rule ord_term_lin.le_less_trans, rule sig_merge, fact 1, fact 3)
ultimately show ?thesis
proof (rule rb_aux_inv_D6_1)
assume "Inl (b1, b2) \<in> set (p # ps)"
with 4 have "Inl (b1, b2) \<in> set ps"
by (auto simp: sig_of_spair \<open>is_regular_spair b1 b2\<close> simp del: sig_of_pair.simps)
thus ?thesis by (simp add: add_spairs_def set_merge_wrt)
next
assume "Inl (b2, b1) \<in> set (p # ps)"
with 4 have "Inl (b2, b1) \<in> set ps"
by (auto simp: sig_of_spair sig_of_spair_commute \<open>is_regular_spair b1 b2\<close> simp del: sig_of_pair.simps)
thus ?thesis by (simp add: add_spairs_def set_merge_wrt)
qed
qed
next
fix q j
assume "j < length fs"
assume "q \<in> set (add_spairs ps bs p')"
hence "sig_of_pair p \<preceq>\<^sub>t sig_of_pair q" by (rule sig_merge)
also assume "sig_of_pair q \<prec>\<^sub>t term_of_pair (0, j)"
finally have 1: "sig_of_pair p \<prec>\<^sub>t term_of_pair (0, j)" .
with assms(1) \<open>p \<in> set (p # ps)\<close> \<open>j < length fs\<close> have "Inr j \<in> set (p # ps)"
by (rule rb_aux_inv_D6_2)
with 1 show "Inr j \<in> set (add_spairs ps bs p')" by (auto simp: add_spairs_def set_merge_wrt)
next
fix b q
assume "b \<in> set (p' # bs)" and q_in: "q \<in> set (add_spairs ps bs p')"
from this(1) have "b = p' \<or> b \<in> set bs" by simp
hence "lt b \<preceq>\<^sub>t lt p'"
proof
note assms(1)
moreover assume "b \<in> set bs"
moreover have "p \<in> set (p # ps)" by simp
ultimately have "lt b \<preceq>\<^sub>t sig_of_pair p" by (rule rb_aux_inv_D7)
thus ?thesis by (simp only: lt_p' sig_of_p)
qed simp
also have "... = sig_of_pair p" by (simp only: sig_of_p lt_p')
also from q_in have "... \<preceq>\<^sub>t sig_of_pair q" by (rule sig_merge)
finally show "lt b \<preceq>\<^sub>t sig_of_pair q" .
next
fix a b
assume 1: "a \<in> set (p' # bs)" and 2: "b \<in> set (p' # bs)" and 3: "is_regular_spair a b"
assume 6: "\<not> is_RB_in dgrad rword (set (p' # bs)) (lt (spair a b))"
with rb have neq: "lt (spair a b) \<noteq> lt (poly_of_pair p)" by (auto simp: sig_of_p)
assume "Inl (a, b) \<notin> set (add_spairs ps bs p')"
hence 40: "Inl (a, b) \<notin> set (new_spairs bs p')" and "Inl (a, b) \<notin> set ps"
by (simp_all add: add_spairs_def set_merge_wrt)
from this(2) neq have 4: "Inl (a, b) \<notin> set (p # ps)" by auto
assume "Inl (b, a) \<notin> set (add_spairs ps bs p')"
hence 50: "Inl (b, a) \<notin> set (new_spairs bs p')" and "Inl (b, a) \<notin> set ps"
by (simp_all add: add_spairs_def set_merge_wrt)
from this(2) neq have 5: "Inl (b, a) \<notin> set (p # ps)" by (auto simp: spair_comm[of a])
have "a \<noteq> p'"
proof
assume "a = p'"
with 3 have "b \<noteq> p'" by (auto simp: is_regular_spair_def)
with 2 have "b \<in> set bs" by simp
hence "Inl (a, b) \<in> set (new_spairs bs p')" using 3 unfolding \<open>a = p'\<close> by (rule in_new_spairsI)
with 40 show False ..
qed
with 1 have "a \<in> set bs" by simp
have "b \<noteq> p'"
proof
assume "b = p'"
with 3 have "a \<noteq> p'" by (auto simp: is_regular_spair_def)
with 1 have "a \<in> set bs" by simp
moreover from 3 have "is_regular_spair b a" by (rule is_regular_spair_sym)
ultimately have "Inl (b, a) \<in> set (new_spairs bs p')" unfolding \<open>b = p'\<close> by (rule in_new_spairsI)
with 50 show False ..
qed
with 2 have "b \<in> set bs" by simp
have lt_sp: "lt (spair a b) \<prec>\<^sub>t lt p'"
proof (rule ord_term_lin.linorder_cases)
assume "lt (spair a b) = lt p'"
with neq show ?thesis by (simp add: lt_p')
next
assume "lt p' \<prec>\<^sub>t lt (spair a b)"
hence "sig_of_pair p \<prec>\<^sub>t lt (spair a b)" by (simp only: lt_p' sig_of_p)
with assms(1) \<open>p \<in> set (p # ps)\<close> \<open>a \<in> set bs\<close> \<open>b \<in> set bs\<close> 3 show ?thesis
proof (rule rb_aux_inv_D6_1)
assume "Inl (a, b) \<in> set (p # ps)"
with 4 show ?thesis ..
next
assume "Inl (b, a) \<in> set (p # ps)"
with 5 show ?thesis ..
qed
qed
have "\<not> is_RB_in dgrad rword (set bs) (lt (spair a b))"
proof
assume "is_RB_in dgrad rword (set bs) (lt (spair a b))"
hence "is_RB_in dgrad rword (set (p' # bs)) (lt (spair a b))" unfolding set_simps using lt_sp
by (rule is_RB_in_insertI)
with 6 show False ..
qed
with assms(1) \<open>a \<in> set bs\<close> \<open>b \<in> set bs\<close> 3 4 5
obtain q where "q \<in> set (p # ps)" and 8: "sig_of_pair q = lt (spair a b)" and 9: "\<not> sig_crit' bs q"
by (rule rb_aux_inv_D8)
from this(1, 2) lt_sp have "q \<in> set ps" by (auto simp: lt_p' sig_of_p)
show "\<exists>q\<in>set (add_spairs ps bs p'). sig_of_pair q = lt (spair a b) \<and> \<not> sig_crit' (p' # bs) q"
proof (intro bexI conjI)
show "\<not> sig_crit' (p' # bs) q"
proof
assume "sig_crit' (p' # bs) q"
moreover from lt_sp have "sig_of_pair q \<prec>\<^sub>t lt p'" by (simp only: 8)
ultimately have "sig_crit' bs q" by (rule sig_crit'_ConsD)
with 9 show False ..
qed
next
from \<open>q \<in> set ps\<close> show "q \<in> set (add_spairs ps bs p')" by (simp add: add_spairs_def set_merge_wrt)
qed fact
next
fix j
assume "j < length fs"
assume "Inr j \<notin> set (add_spairs ps bs p')"
hence "Inr j \<notin> set ps" by (simp add: add_spairs_def set_merge_wrt)
show "is_RB_in dgrad rword (set (p' # bs)) (term_of_pair (0, j))"
proof (cases "term_of_pair (0, j) = sig_of_pair p")
case True
with rb show ?thesis by simp
next
case False
with \<open>Inr j \<notin> set ps\<close> have "Inr j \<notin> set (p # ps)" by auto
with assms(1) \<open>j < length fs\<close> have rb': "is_RB_in dgrad rword (set bs) (term_of_pair (0, j))"
by (rule rb_aux_inv_D9)
have "term_of_pair (0, j) \<prec>\<^sub>t lt p'"
proof (rule ord_term_lin.linorder_cases)
assume "term_of_pair (0, j) = lt p'"
with False show ?thesis by (simp add: lt_p' sig_of_p)
next
assume "lt p' \<prec>\<^sub>t term_of_pair (0, j)"
hence "sig_of_pair p \<prec>\<^sub>t term_of_pair (0, j)" by (simp only: lt_p' sig_of_p)
with assms(1) \<open>p \<in> set (p # ps)\<close> \<open>j < length fs\<close> have "Inr j \<in> set (p # ps)"
by (rule rb_aux_inv_D6_2)
with \<open>Inr j \<notin> set (p # ps)\<close> show ?thesis ..
qed
with rb' show ?thesis unfolding set_simps by (rule is_RB_in_insertI)
qed
show "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set (p' # bs))"
proof (cases "p = Inr j")
case True
show ?thesis
proof (rule sig_red_zero_idealI, rule sig_red_zeroI)
from p_red have "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) p'"
by (simp add: True)
moreover have "set bs \<subseteq> set (p' # bs)" by fastforce
ultimately have "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set (p' # bs)))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) p'"
by (rule sig_red_rtrancl_mono)
hence "(sig_red (\<preceq>\<^sub>t) (\<preceq>) (set (p' # bs)))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) p'"
by (rule sig_red_rtrancl_sing_regI)
also have "sig_red (\<preceq>\<^sub>t) (\<preceq>) (set (p' # bs)) p' 0" unfolding sig_red_def
proof (intro exI bexI)
from assms(3) have "rep_list p' \<noteq> 0" by (simp add: p'_def)
show "sig_red_single (\<preceq>\<^sub>t) (\<preceq>) p' 0 p' 0"
proof (rule sig_red_singleI)
show "rep_list p' \<noteq> 0" by fact
next
from \<open>rep_list p' \<noteq> 0\<close> have "punit.lt (rep_list p') \<in> keys (rep_list p')"
by (rule punit.lt_in_keys)
thus "0 + punit.lt (rep_list p') \<in> keys (rep_list p')" by simp
next
from \<open>rep_list p' \<noteq> 0\<close> have "punit.lc (rep_list p') \<noteq> 0" by (rule punit.lc_not_0)
thus "0 = p' - monom_mult (lookup (rep_list p') (0 + punit.lt (rep_list p')) / punit.lc (rep_list p')) 0 p'"
by (simp add: punit.lc_def[symmetric])
qed (simp_all add: term_simps)
qed simp
finally show "(sig_red (\<preceq>\<^sub>t) (\<preceq>) (set (p' # bs)))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) 0" .
qed (fact rep_list_zero)
next
case False
with \<open>Inr j \<notin> set ps\<close> have "Inr j \<notin> set (p # ps)" by simp
with assms(1) \<open>j < length fs\<close>
have "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` set bs)"
by (rule rb_aux_inv_D9)
also have "... \<subseteq> ideal (rep_list ` set (p' # bs))" by (rule ideal.span_mono, fastforce)
finally show ?thesis .
qed
qed
show "lt p' \<notin> lt ` set bs" unfolding lt_p'
proof
assume "lt (poly_of_pair p) \<in> lt ` set bs"
then obtain b where "b \<in> set bs" and "lt (poly_of_pair p) = lt b" ..
note this(2)
also from \<open>b \<in> set bs\<close> have "lt b \<prec>\<^sub>t lt (poly_of_pair p)" by (rule lt_p_gr)
finally show False ..
qed
qed
lemma rb_aux_inv_init: "rb_aux_inv ([], Koszul_syz_sigs fs, map Inr [0..<length fs])"
proof (simp add: rb_aux_inv.simps rb_aux_inv1_def o_def, intro conjI ballI allI impI)
fix v
assume "v \<in> set (Koszul_syz_sigs fs)"
with dgrad fs_distinct fs_nonzero show "is_syz_sig dgrad v" by (rule Koszul_syz_sigs_is_syz_sig)
next
fix p q :: "'t \<Rightarrow>\<^sub>0 'b"
show "Inl (p, q) \<notin> Inr ` {0..<length fs}" by blast
next
fix j
assume "Inr j \<in> Inr ` {0..<length fs}"
thus "j < length fs" by fastforce
next
fix j
have eq: "(term_of_pair (0, i) = term_of_pair (0, j)) \<longleftrightarrow> (j = i)" for i
by (auto dest: term_of_pair_injective)
show "length (filter (\<lambda>i. term_of_pair (0, i) = term_of_pair (0, j)) [0..<length fs]) \<le> Suc 0"
by (simp add: eq)
next
show "sorted_wrt pair_ord (map Inr [0..<length fs])"
proof (simp add: sorted_wrt_map pair_ord_def sorted_wrt_upt_iff, intro allI impI)
fix i j :: nat
assume "i < j"
hence "i \<le> j" by simp
show "term_of_pair (0, i) \<preceq>\<^sub>t term_of_pair (0, j)" by (rule ord_termI, simp_all add: term_simps \<open>i \<le> j\<close>)
qed
qed
corollary rb_aux_inv_init_fst:
"rb_aux_inv (fst (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))"
using rb_aux_inv_init by simp
function (domintros) rb_aux :: "((('t \<Rightarrow>\<^sub>0 'b) list \<times> 't list \<times> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list) \<times> nat) \<Rightarrow>
((('t \<Rightarrow>\<^sub>0 'b) list \<times> 't list \<times> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list) \<times> nat)"
where
"rb_aux ((bs, ss, []), z) = ((bs, ss, []), z)" |
"rb_aux ((bs, ss, p # ps), z) =
(let ss' = new_syz_sigs ss bs p in
if sig_crit bs ss' p then
rb_aux ((bs, ss', ps), z)
else
let p' = sig_trd bs (poly_of_pair p) in
if rep_list p' = 0 then
rb_aux ((bs, lt p' # ss', ps), Suc z)
else
rb_aux ((p' # bs, ss', add_spairs ps bs p'), z))"
by pat_completeness auto
definition rb :: "('t \<Rightarrow>\<^sub>0 'b) list \<times> nat"
where "rb = (let ((bs, _, _), z) = rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), 0) in (bs, z))"
text \<open>@{const rb} is only an auxiliary function used for stating some theorems about rewrite bases
and their computation in a readable way. Actual computations (of Gr\"obner bases) are performed
by function \<open>sig_gb\<close>, defined below.
The second return value of @{const rb} is the number of zero-reductions. It is only needed for
proving that under certain assumptions, there are no such zero-reductions.\<close>
text \<open>Termination\<close>
qualified definition "rb_aux_term1 \<equiv> {(x, y). \<exists>z. x = z # y}"
qualified definition "rb_aux_term2 \<equiv> {(x, y). (fst x, fst y) \<in> rb_aux_term1 \<or>
(fst x = fst y \<and> length (snd (snd x)) < length (snd (snd y)))}"
qualified definition "rb_aux_term \<equiv> rb_aux_term2 \<inter> {(x, y). rb_aux_inv x \<and> rb_aux_inv y}"
lemma wfp_on_rb_aux_term1: "wfp_on (\<lambda>x y. (x, y) \<in> rb_aux_term1) (Collect rb_aux_inv1)"
proof (rule wfp_onI_chain, rule, elim exE)
fix seq'
assume "\<forall>i. seq' i \<in> Collect rb_aux_inv1 \<and> (seq' (Suc i), seq' i) \<in> rb_aux_term1"
hence inv: "rb_aux_inv1 (seq' j)" and cons: "\<exists>b. seq' (Suc j) = b # seq' j" for j
by (simp_all add: rb_aux_term1_def)
from this(2) have 1: thesis0 if "\<And>j. i < length (seq' j) \<Longrightarrow> thesis0" for i thesis0
using that by (rule list_seq_indexE_length)
define seq where "seq = (\<lambda>i. let j = (SOME k. i < length (seq' k)) in rev (seq' j) ! i)"
have 2: "seq i = rev (seq' j) ! i" if "i < length (seq' j)" for i j
proof -
define k where "k = (SOME k. i < length (seq' k))"
from that have "i < length (seq' k)" unfolding k_def by (rule someI)
with cons that have "rev (seq' k) ! i = rev (seq' j) ! i" by (rule list_seq_nth')
thus ?thesis by (simp add: seq_def k_def[symmetric])
qed
have 3: "seq i \<in> set (seq' j)" if "i < length (seq' j)" for i j
proof -
from that have "i < length (rev (seq' j))" by simp
moreover from that have "seq i = rev (seq' j) ! i" by (rule 2)
ultimately have "seq i \<in> set (rev (seq' j))" by (metis nth_mem)
thus ?thesis by simp
qed
have 4: "seq ` {0..<i} = set (take i (rev (seq' j)))" if "i < length (seq' j)" for i j
proof -
from refl have "seq ` {0..<i} = (!) (rev (seq' j)) ` {0..<i}"
proof (rule image_cong)
fix i'
assume "i' \<in> {0..<i}"
hence "i' < i" by simp
hence "i' < length (seq' j)" using that by simp
thus "seq i' = rev (seq' j) ! i'" by (rule 2)
qed
also have "... = set (take i (rev (seq' j)))" by (rule nth_image, simp add: that less_imp_le_nat)
finally show ?thesis .
qed
from dgrad show False
proof (rule rb_termination)
have "seq i \<in> dgrad_sig_set dgrad" for i
proof -
obtain j where "i < length (seq' j)" by (rule 1)
hence "seq i \<in> set (seq' j)" by (rule 3)
moreover from inv have "set (seq' j) \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
ultimately show ?thesis ..
qed
thus "range seq \<subseteq> dgrad_sig_set dgrad" by blast
next
have "rep_list (seq i) \<noteq> 0" for i
proof -
obtain j where "i < length (seq' j)" by (rule 1)
hence "seq i \<in> set (seq' j)" by (rule 3)
moreover from inv have "0 \<notin> rep_list ` set (seq' j)" by (rule rb_aux_inv1_D2)
ultimately show ?thesis by auto
qed
thus "0 \<notin> rep_list ` range seq" by fastforce
next
fix i1 i2 :: nat
assume "i1 < i2"
also obtain j where i2: "i2 < length (seq' j)" by (rule 1)
finally have i1: "i1 < length (seq' j)" .
from i1 have s1: "seq i1 = rev (seq' j) ! i1" by (rule 2)
from i2 have s2: "seq i2 = rev (seq' j) ! i2" by (rule 2)
from inv have "sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) (seq' j)" by (rule rb_aux_inv1_D3)
hence "sorted_wrt (\<lambda>x y. lt x \<prec>\<^sub>t lt y) (rev (seq' j))" by (simp add: sorted_wrt_rev)
moreover note \<open>i1 < i2\<close>
moreover from i2 have "i2 < length (rev (seq' j))" by simp
ultimately have "lt (rev (seq' j) ! i1) \<prec>\<^sub>t lt (rev (seq' j) ! i2)" by (rule sorted_wrt_nth_less)
thus "lt (seq i1) \<prec>\<^sub>t lt (seq i2)" by (simp only: s1 s2)
next
fix i
obtain j where i: "i < length (seq' j)" by (rule 1)
hence eq1: "seq i = rev (seq' j) ! i" and eq2: "seq ` {0..<i} = set (take i (rev (seq' j)))"
by (rule 2, rule 4)
let ?i = "length (seq' j) - Suc i"
from i have "?i < length (seq' j)" by simp
with inv have "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set (drop (Suc ?i) (seq' j))) ((seq' j) ! ?i)"
by (rule rb_aux_inv1_D4)
thus "\<not> is_sig_red (\<prec>\<^sub>t) (\<preceq>) (seq ` {0..<i}) (seq i)"
using i by (simp add: eq1 eq2 rev_nth take_rev Suc_diff_Suc)
from inv \<open>?i < length (seq' j)\<close>
show "(\<exists>j<length fs. lt (seq i) = lt (monomial (1::'b) (term_of_pair (0, j))) \<and>
punit.lt (rep_list (seq i)) \<preceq> punit.lt (rep_list (monomial 1 (term_of_pair (0, j))))) \<or>
(\<exists>j k. is_regular_spair (seq j) (seq k) \<and> rep_list (spair (seq j) (seq k)) \<noteq> 0 \<and>
lt (seq i) = lt (spair (seq j) (seq k)) \<and>
punit.lt (rep_list (seq i)) \<preceq> punit.lt (rep_list (spair (seq j) (seq k))))" (is "?l \<or> ?r")
proof (rule rb_aux_inv1_E)
fix j0
assume "j0 < length fs"
and "lt (seq' j ! (length (seq' j) - Suc i)) = lt (monomial (1::'b) (term_of_pair (0, j0)))"
and "punit.lt (rep_list (seq' j ! (length (seq' j) - Suc i))) \<preceq>
punit.lt (rep_list (monomial 1 (term_of_pair (0, j0))))"
hence ?l using i by (auto simp: eq1 eq2 rev_nth take_rev Suc_diff_Suc)
thus ?thesis ..
next
fix p q
assume "p \<in> set (seq' j)"
then obtain pi where "pi < length (seq' j)" and "p = (seq' j) ! pi" by (metis in_set_conv_nth)
hence p: "p = seq (length (seq' j) - Suc pi)"
by (metis "2" \<open>p \<in> set (seq' j)\<close> diff_Suc_less length_pos_if_in_set length_rev rev_nth rev_rev_ident)
assume "q \<in> set (seq' j)"
then obtain qi where "qi < length (seq' j)" and "q = (seq' j) ! qi" by (metis in_set_conv_nth)
hence q: "q = seq (length (seq' j) - Suc qi)"
by (metis "2" \<open>q \<in> set (seq' j)\<close> diff_Suc_less length_pos_if_in_set length_rev rev_nth rev_rev_ident)
assume "is_regular_spair p q" and "rep_list (spair p q) \<noteq> 0"
and "lt (seq' j ! (length (seq' j) - Suc i)) = lt (spair p q)"
and "punit.lt (rep_list (seq' j ! (length (seq' j) - Suc i))) \<preceq> punit.lt (rep_list (spair p q))"
hence ?r using i by (auto simp: eq1 eq2 p q rev_nth take_rev Suc_diff_Suc)
thus ?thesis ..
qed
from inv \<open>?i < length (seq' j)\<close>
have "is_RB_upt dgrad rword (set (drop (Suc ?i) (seq' j))) (lt ((seq' j) ! ?i))"
by (rule rb_aux_inv1_D5)
with dgrad have "is_sig_GB_upt dgrad (set (drop (Suc ?i) (seq' j))) (lt ((seq' j) ! ?i))"
by (rule is_RB_upt_is_sig_GB_upt)
thus "is_sig_GB_upt dgrad (seq ` {0..<i}) (lt (seq i))"
using i by (simp add: eq1 eq2 rev_nth take_rev Suc_diff_Suc)
qed
qed
lemma wfp_on_rb_aux_term2: "wfp_on (\<lambda>x y. (x, y) \<in> rb_aux_term2) (Collect rb_aux_inv)"
proof (rule wfp_onI_min)
fix x Q
assume "x \<in> Q" and Q_sub: "Q \<subseteq> Collect rb_aux_inv"
from this(1) have "fst x \<in> fst ` Q" by (rule imageI)
have "fst ` Q \<subseteq> Collect rb_aux_inv1"
proof
fix y
assume "y \<in> fst ` Q"
then obtain z where "z \<in> Q" and y: "y = fst z" by fastforce
obtain bs ss ps where z: "z = (bs, ss, ps)" by (rule rb_aux_inv.cases)
from \<open>z \<in> Q\<close> Q_sub have "rb_aux_inv z" by blast
thus "y \<in> Collect rb_aux_inv1" by (simp add: y z rb_aux_inv.simps)
qed
with wfp_on_rb_aux_term1 \<open>fst x \<in> fst ` Q\<close> obtain z' where "z' \<in> fst ` Q"
and z'_min: "\<And>y. (y, z') \<in> rb_aux_term1 \<Longrightarrow> y \<notin> fst ` Q" by (rule wfp_onE_min) blast
from this(1) obtain z0 where "z0 \<in> Q" and z': "z' = fst z0" by fastforce
define Q0 where "Q0 = {z. z \<in> Q \<and> fst z = fst z0}"
from \<open>z0 \<in> Q\<close> have "z0 \<in> Q0" by (simp add: Q0_def)
hence "length (snd (snd z0)) \<in> length ` snd ` snd ` Q0" by (intro imageI)
with wf_less obtain n where n1: "n \<in> length ` snd ` snd ` Q0"
and n2: "\<And>n'. n' < n \<Longrightarrow> n' \<notin> length ` snd ` snd ` Q0" by (rule wfE_min, blast)
from n1 obtain z where "z \<in> Q0" and n3: "n = length (snd (snd z))" by fastforce
have z_min: "y \<notin> Q0" if "length (snd (snd y)) < length (snd (snd z))" for y
proof
assume "y \<in> Q0"
hence "length (snd (snd y)) \<in> length ` snd ` snd ` Q0" by (intro imageI)
with n2 have "\<not> length (snd (snd y)) < length (snd (snd z))" unfolding n3[symmetric] by blast
thus False using that ..
qed
show "\<exists>z\<in>Q. \<forall>y\<in>Collect rb_aux_inv. (y, z) \<in> rb_aux_term2 \<longrightarrow> y \<notin> Q"
proof (intro bexI ballI impI)
fix y
assume "y \<in> Collect rb_aux_inv"
assume "(y, z) \<in> rb_aux_term2"
hence "(fst y, fst z) \<in> rb_aux_term1 \<or> (fst y = fst z \<and> length (snd (snd y)) < length (snd (snd z)))"
by (simp add: rb_aux_term2_def)
thus "y \<notin> Q"
proof
assume "(fst y, fst z) \<in> rb_aux_term1"
moreover from \<open>z \<in> Q0\<close> have "fst z = fst z0" by (simp add: Q0_def)
ultimately have "(fst y, z') \<in> rb_aux_term1" by (simp add: rb_aux_term1_def z')
hence "fst y \<notin> fst ` Q" by (rule z'_min)
thus ?thesis by blast
next
assume "fst y = fst z \<and> length (snd (snd y)) < length (snd (snd z))"
hence "fst y = fst z" and "length (snd (snd y)) < length (snd (snd z))" by simp_all
from this(2) have "y \<notin> Q0" by (rule z_min)
moreover from \<open>z \<in> Q0\<close> have "fst y = fst z0" by (simp add: Q0_def \<open>fst y = fst z\<close>)
ultimately show ?thesis by (simp add: Q0_def)
qed
next
from \<open>z \<in> Q0\<close> show "z \<in> Q" by (simp add: Q0_def)
qed
qed
corollary wf_rb_aux_term: "wf rb_aux_term"
proof (rule wfI_min)
fix x::"('t \<Rightarrow>\<^sub>0 'b) list \<times> 't list \<times> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list" and Q
assume "x \<in> Q"
show "\<exists>z\<in>Q. \<forall>y. (y, z) \<in> rb_aux_term \<longrightarrow> y \<notin> Q"
proof (cases "rb_aux_inv x")
case True
let ?Q = "Q \<inter> Collect rb_aux_inv"
note wfp_on_rb_aux_term2
moreover from \<open>x \<in> Q\<close> True have "x \<in> ?Q" by simp
moreover have "?Q \<subseteq> Collect rb_aux_inv" by simp
ultimately obtain z where "z \<in> ?Q" and z_min: "\<And>y. (y, z) \<in> rb_aux_term2 \<Longrightarrow> y \<notin> ?Q"
by (rule wfp_onE_min) blast
show ?thesis
proof (intro bexI allI impI)
fix y
assume "(y, z) \<in> rb_aux_term"
hence "(y, z) \<in> rb_aux_term2" and "rb_aux_inv y" by (simp_all add: rb_aux_term_def)
from this(1) have "y \<notin> ?Q" by (rule z_min)
with \<open>rb_aux_inv y\<close> show "y \<notin> Q" by simp
next
from \<open>z \<in> ?Q\<close> show "z \<in> Q" by simp
qed
next
case False
show ?thesis
proof (intro bexI allI impI)
fix y
assume "(y, x) \<in> rb_aux_term"
hence "rb_aux_inv x" by (simp add: rb_aux_term_def)
with False show "y \<notin> Q" ..
qed fact
qed
qed
lemma rb_aux_domI:
assumes "rb_aux_inv (fst args)"
shows "rb_aux_dom args"
proof -
let ?rel = "rb_aux_term <*lex*> ({}::(nat \<times> nat) set)"
from wf_rb_aux_term wf_empty have "wf ?rel" ..
thus ?thesis using assms
proof (induct args)
case (less args)
obtain bs ss ps0 z where args: "args = ((bs, ss, ps0), z)" using prod.exhaust by metis
show ?case
proof (cases ps0)
case Nil
show ?thesis unfolding args Nil by (rule rb_aux.domintros)
next
case (Cons p ps)
from less(1) have 1: "\<And>y. (y, ((bs, ss, p # ps), z)) \<in> ?rel \<Longrightarrow> rb_aux_inv (fst y) \<Longrightarrow> rb_aux_dom y"
by (simp only: args Cons)
from less(2) have 2: "rb_aux_inv (bs, ss, p # ps)" by (simp only: args Cons fst_conv)
show ?thesis unfolding args Cons
proof (rule rb_aux.domintros)
assume "sig_crit bs (new_syz_sigs ss bs p) p"
with 2 have a: "rb_aux_inv (bs, (new_syz_sigs ss bs p), ps)" by (rule rb_aux_inv_preserved_1)
with 2 have "((bs, (new_syz_sigs ss bs p), ps), bs, ss, p # ps) \<in> rb_aux_term"
by (simp add: rb_aux_term_def rb_aux_term2_def)
hence "(((bs, (new_syz_sigs ss bs p), ps), z), (bs, ss, p # ps), z) \<in> ?rel" by simp
moreover from a have "rb_aux_inv (fst ((bs, (new_syz_sigs ss bs p), ps), z))"
by (simp only: fst_conv)
ultimately show "rb_aux_dom ((bs, (new_syz_sigs ss bs p), ps), z)" by (rule 1)
next
assume "rep_list (sig_trd bs (poly_of_pair p)) = 0"
with 2 have a: "rb_aux_inv (bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps)"
by (rule rb_aux_inv_preserved_2)
with 2 have "((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), bs, ss, p # ps) \<in>
rb_aux_term"
by (simp add: rb_aux_term_def rb_aux_term2_def)
hence "(((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z), (bs, ss, p # ps), z) \<in>
?rel" by simp
moreover from a have "rb_aux_inv (fst ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z))"
by (simp only: fst_conv)
ultimately show "rb_aux_dom ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z)"
by (rule 1)
next
let ?args = "(sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p, add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
assume "\<not> sig_crit bs (new_syz_sigs ss bs p) p" and "rep_list (sig_trd bs (poly_of_pair p)) \<noteq> 0"
with 2 have a: "rb_aux_inv ?args" by (rule rb_aux_inv_preserved_3)
with 2 have "(?args, bs, ss, p # ps) \<in> rb_aux_term"
by (simp add: rb_aux_term_def rb_aux_term2_def rb_aux_term1_def)
hence "((?args, z), (bs, ss, p # ps), z) \<in> ?rel" by simp
moreover from a have "rb_aux_inv (fst (?args, z))" by (simp only: fst_conv)
ultimately show "rb_aux_dom (?args, z)" by (rule 1)
qed
qed
qed
qed
text \<open>Invariant\<close>
lemma rb_aux_inv_invariant:
assumes "rb_aux_inv (fst args)"
shows "rb_aux_inv (fst (rb_aux args))"
proof -
from assms have "rb_aux_dom args" by (rule rb_aux_domI)
thus ?thesis using assms
proof (induct args rule: rb_aux.pinduct)
case (1 bs ss z)
thus ?case by (simp only: rb_aux.psimps(1))
next
case (2 bs ss p ps z)
from 2(5) have *: "rb_aux_inv (bs, ss, p # ps)" by (simp only: fst_conv)
show ?case
proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
assume a: "sig_crit bs (new_syz_sigs ss bs p) p"
with * have "rb_aux_inv (bs, new_syz_sigs ss bs p, ps)"
by (rule rb_aux_inv_preserved_1)
hence "rb_aux_inv (fst ((bs, new_syz_sigs ss bs p, ps), z))" by (simp only: fst_conv)
with refl a show "rb_aux_inv (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))" by (rule 2(2))
thus "rb_aux_inv (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))" .
next
assume a: "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
assume b: "rep_list (sig_trd bs (poly_of_pair p)) = 0"
with * have "rb_aux_inv (bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps)"
by (rule rb_aux_inv_preserved_2)
hence "rb_aux_inv (fst ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z))"
by (simp only: fst_conv)
with refl a refl b
show "rb_aux_inv (fst (rb_aux ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z)))"
by (rule 2(3))
next
let ?args = "(sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p,
add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
assume a: "\<not> sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list (sig_trd bs (poly_of_pair p)) \<noteq> 0"
with * have "rb_aux_inv ?args" by (rule rb_aux_inv_preserved_3)
hence "rb_aux_inv (fst (?args, z))" by (simp only: fst_conv)
with refl a refl b
show "rb_aux_inv (fst (rb_aux (?args, z)))"
by (rule 2(4))
qed
qed
qed
lemma rb_aux_inv_last_Nil:
assumes "rb_aux_dom args"
shows "snd (snd (fst (rb_aux args))) = []"
using assms
proof (induct args rule: rb_aux.pinduct)
case (1 bs ss z)
thus ?case by (simp add: rb_aux.psimps(1))
next
case (2 bs ss p ps z)
show ?case
proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
assume "sig_crit bs (new_syz_sigs ss bs p) p"
with refl show "snd (snd (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))) = []"
and "snd (snd (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))) = []"
by (rule 2(2))+
next
assume a: "\<not> sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list (sig_trd bs (poly_of_pair p)) = 0"
from refl a refl b
show "snd (snd (fst (rb_aux ((bs, lt (sig_trd bs (poly_of_pair p)) # new_syz_sigs ss bs p, ps), Suc z)))) = []"
by (rule 2(3))
next
assume a: "\<not> sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list (sig_trd bs (poly_of_pair p)) \<noteq> 0"
from refl a refl b
show "snd (snd (fst (rb_aux ((sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p,
add_spairs ps bs (sig_trd bs (poly_of_pair p))), z)))) = []"
by (rule 2(4))
qed
qed
corollary rb_aux_shape:
assumes "rb_aux_dom args"
obtains bs ss z where "rb_aux args = ((bs, ss, []), z)"
proof -
obtain bs ss ps z where "rb_aux args = ((bs, ss, ps), z)" using prod.exhaust by metis
moreover from assms have "snd (snd (fst (rb_aux args))) = []" by (rule rb_aux_inv_last_Nil)
ultimately have "rb_aux args = ((bs, ss, []), z)" by simp
thus ?thesis ..
qed
lemma rb_aux_is_RB_upt:
"is_RB_upt dgrad rword (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) u"
proof -
let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
by (rule rb_aux_inv_invariant)
ultimately have "rb_aux_inv (bs, ss, [])" by simp
have "is_RB_upt dgrad rword (set bs) u" by (rule rb_aux_inv_is_RB_upt, fact, simp)
thus ?thesis by (simp add: eq)
qed
corollary rb_is_RB_upt: "is_RB_upt dgrad rword (set (fst rb)) u"
using rb_aux_is_RB_upt[of 0 u] by (auto simp add: rb_def split: prod.split)
corollary rb_aux_is_sig_GB_upt:
"is_sig_GB_upt dgrad (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) u"
using dgrad rb_aux_is_RB_upt by (rule is_RB_upt_is_sig_GB_upt)
corollary rb_aux_is_sig_GB_in:
"is_sig_GB_in dgrad (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) u"
proof -
let ?u = "term_of_pair (pp_of_term u, Suc (component_of_term u))"
have "u \<prec>\<^sub>t ?u"
proof (rule ord_term_lin.le_neq_trans)
show "u \<preceq>\<^sub>t ?u" by (rule ord_termI, simp_all add: term_simps)
next
show "u \<noteq> ?u"
proof
assume "u = ?u"
hence "component_of_term u = component_of_term ?u" by simp
thus False by (simp add: term_simps)
qed
qed
with rb_aux_is_sig_GB_upt show ?thesis by (rule is_sig_GB_uptD2)
qed
corollary rb_aux_is_Groebner_basis:
assumes "hom_grading dgrad"
shows "punit.is_Groebner_basis (set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))))"
proof -
let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
by (rule rb_aux_inv_invariant)
ultimately have "rb_aux_inv (bs, ss, [])" by simp
hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
hence "set (fst (fst (rb_aux ?args))) \<subseteq> dgrad_max_set dgrad" by (simp add: eq dgrad_sig_set'_def)
with dgrad assms have "punit.is_Groebner_basis (rep_list ` set (fst (fst (rb_aux ?args))))"
using rb_aux_is_sig_GB_in by (rule is_sig_GB_is_Groebner_basis)
thus ?thesis by simp
qed
lemma ideal_rb_aux:
"ideal (set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))) =
ideal (set fs)" (is "ideal ?l = ideal ?r")
proof
show "ideal ?l \<subseteq> ideal ?r" by (rule ideal.span_subset_spanI, auto simp: rep_list_in_ideal)
next
show "ideal ?r \<subseteq> ideal ?l"
proof (rule ideal.span_subset_spanI, rule subsetI)
fix f
assume "f \<in> set fs"
then obtain j where "j < length fs" and f: "f = fs ! j" by (metis in_set_conv_nth)
let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
by (rule rb_aux_inv_invariant)
ultimately have "rb_aux_inv (bs, ss, [])" by simp
moreover note \<open>j < length fs\<close>
moreover have "Inr j \<notin> set []" by simp
ultimately have "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal ?l"
unfolding eq set_map fst_conv by (rule rb_aux_inv_D9)
thus "f \<in> ideal ?l" by (simp add: rep_list_monomial' \<open>j < length fs\<close> f)
qed
qed
corollary ideal_rb: "ideal (rep_list ` set (fst rb)) = ideal (set fs)"
proof -
have "ideal (rep_list ` set (fst rb)) =
ideal (set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), 0))))))"
by (auto simp: rb_def split: prod.splits)
also have "... = ideal (set fs)" by (fact ideal_rb_aux)
finally show ?thesis .
qed
lemma
shows dgrad_max_set_closed_rb_aux:
"set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) \<subseteq>
punit_dgrad_max_set dgrad" (is ?thesis1)
and rb_aux_nonzero:
"0 \<notin> set (map rep_list (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))"
(is ?thesis2)
proof -
let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
from rb_aux_inv_init_fst have "rb_aux_dom ?args" by (rule rb_aux_domI)
then obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')" by (rule rb_aux_shape)
moreover from rb_aux_inv_init_fst have "rb_aux_inv (fst (rb_aux ?args))"
by (rule rb_aux_inv_invariant)
ultimately have "rb_aux_inv (bs, ss, [])" by simp
hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence "set bs \<subseteq> dgrad_sig_set dgrad" and *: "0 \<notin> rep_list ` set bs"
by (rule rb_aux_inv1_D1, rule rb_aux_inv1_D2)
from this(1) have "set bs \<subseteq> dgrad_max_set dgrad" by (simp add: dgrad_sig_set'_def)
with dgrad show ?thesis1 by (simp add: eq dgrad_max_3)
from * show ?thesis2 by (simp add: eq)
qed
subsubsection \<open>Minimality of the Computed Basis\<close>
lemma rb_aux_top_irred':
assumes "rword_strict = rw_rat_strict" and "rb_aux_inv (bs, ss, p # ps)"
and "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
shows "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set bs) (sig_trd bs (poly_of_pair p))"
proof -
have "rword = rw_rat" by (intro ext, simp only: rword_def rw_rat_alt, simp add: assms(1))
have lt_p: "sig_of_pair p = lt (poly_of_pair p)" by (rule pair_list_sig_of_pair, fact, simp)
define p' where "p' = sig_trd bs (poly_of_pair p)"
have red_p: "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (poly_of_pair p) p'"
unfolding p'_def by (rule sig_trd_red_rtrancl)
hence lt_p': "lt p' = sig_of_pair p"
and lt_p'': "punit.lt (rep_list p') \<preceq> punit.lt (rep_list (poly_of_pair p))"
unfolding lt_p by (rule sig_red_regular_rtrancl_lt, rule sig_red_rtrancl_lt_rep_list)
have "\<not> is_sig_red (=) (=) (set bs) p'"
proof
assume "is_sig_red (=) (=) (set bs) p'"
then obtain b where "b \<in> set bs" and "rep_list b \<noteq> 0" and "rep_list p' \<noteq> 0"
and 1: "punit.lt (rep_list b) adds punit.lt (rep_list p')"
and 2: "punit.lt (rep_list p') \<oplus> lt b = punit.lt (rep_list b) \<oplus> lt p'"
by (rule is_sig_red_top_addsE)
note this(3)
moreover from red_p have "(punit.red (rep_list ` set bs))\<^sup>*\<^sup>* (rep_list (poly_of_pair p)) (rep_list p')"
by (rule sig_red_red_rtrancl)
ultimately have "rep_list (poly_of_pair p) \<noteq> 0" by (auto simp: punit.rtrancl_0)
define x where "x = punit.lt (rep_list p') - punit.lt (rep_list b)"
from 1 2 have x1: "x \<oplus> lt b = lt p'" by (simp add: term_is_le_rel_minus x_def)
from this[symmetric] have "lt b adds\<^sub>t sig_of_pair p" unfolding lt_p' by (rule adds_termI)
from 1 have x2: "x + punit.lt (rep_list b) = punit.lt (rep_list p')" by (simp add: x_def adds_minus)
from \<open>rep_list b \<noteq> 0\<close> have "b \<noteq> 0" by (auto simp: rep_list_zero)
show False
proof (rule sum_prodE)
fix a0 b0
assume p: "p = Inl (a0, b0)"
hence "Inl (a0, b0) \<in> set (p # ps)" by simp
with assms(2) have reg: "is_regular_spair a0 b0" and "a0 \<in> set bs" and "b0 \<in> set bs"
by (rule rb_aux_inv_D3)+
from assms(2) have inv1: "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence "0 \<notin> rep_list ` set bs" by (rule rb_aux_inv1_D2)
with \<open>a0 \<in> set bs\<close> \<open>b0 \<in> set bs\<close> have "rep_list a0 \<noteq> 0" and "rep_list b0 \<noteq> 0" by fastforce+
hence "a0 \<noteq> 0" and "b0 \<noteq> 0" by (auto simp: rep_list_zero)
let ?t1 = "punit.lt (rep_list a0)"
let ?t2 = "punit.lt (rep_list b0)"
let ?l = "lcs ?t1 ?t2"
from \<open>rep_list (poly_of_pair p) \<noteq> 0\<close> have "punit.spoly (rep_list a0) (rep_list b0) \<noteq> 0"
by (simp add: p rep_list_spair)
with \<open>rep_list a0 \<noteq> 0\<close> \<open>rep_list b0 \<noteq> 0\<close>
have "punit.lt (punit.spoly (rep_list a0) (rep_list b0)) \<prec> ?l"
by (rule punit.lt_spoly_less_lcs[simplified])
obtain b' where 3: "is_canon_rewriter rword (set bs) (sig_of_pair p) b'"
and 4: "punit.lt (rep_list (poly_of_pair p)) \<prec>
(pp_of_term (sig_of_pair p) - lp b') + punit.lt (rep_list b')"
proof (cases "(?l - ?t1) \<oplus> lt a0 \<preceq>\<^sub>t (?l - ?t2) \<oplus> lt b0")
case True
have "sig_of_pair p = lt (spair a0 b0)" unfolding lt_p by (simp add: p)
also from reg have "... = (?l - ?t2) \<oplus> lt b0"
by (simp add: True is_regular_spair_lt ord_term_lin.max_def)
finally have eq1: "sig_of_pair p = (?l - ?t2) \<oplus> lt b0" .
hence "lt b0 adds\<^sub>t sig_of_pair p" by (rule adds_termI)
moreover from assms(3) have "\<not> is_rewritable bs b0 ((?l - ?t2) \<oplus> lt b0)"
by (simp add: p spair_sigs_def Let_def)
ultimately have "is_canon_rewriter rword (set bs) (sig_of_pair p) b0"
unfolding eq1[symmetric] using inv1 \<open>b0 \<in> set bs\<close> \<open>b0 \<noteq> 0\<close> is_rewritableI_is_canon_rewriter
by blast
thus ?thesis
proof
have "punit.lt (rep_list (poly_of_pair p)) = punit.lt (punit.spoly (rep_list a0) (rep_list b0))"
by (simp add: p rep_list_spair)
also have "... \<prec> ?l" by fact
also have "... = (?l - ?t2) + ?t2" by (simp only: adds_minus adds_lcs_2)
also have "... = (pp_of_term (sig_of_pair p) - lp b0) + ?t2"
by (simp only: eq1 pp_of_term_splus add_diff_cancel_right')
finally show "punit.lt (rep_list (poly_of_pair p)) \<prec> pp_of_term (sig_of_pair p) - lp b0 + ?t2" .
qed
next
case False
have "sig_of_pair p = lt (spair a0 b0)" unfolding lt_p by (simp add: p)
also from reg have "... = (?l - ?t1) \<oplus> lt a0"
by (simp add: False is_regular_spair_lt ord_term_lin.max_def)
finally have eq1: "sig_of_pair p = (?l - ?t1) \<oplus> lt a0" .
hence "lt a0 adds\<^sub>t sig_of_pair p" by (rule adds_termI)
moreover from assms(3) have "\<not> is_rewritable bs a0 ((?l - ?t1) \<oplus> lt a0)"
by (simp add: p spair_sigs_def Let_def)
ultimately have "is_canon_rewriter rword (set bs) (sig_of_pair p) a0"
unfolding eq1[symmetric] using inv1 \<open>a0 \<in> set bs\<close> \<open>a0 \<noteq> 0\<close> is_rewritableI_is_canon_rewriter
by blast
thus ?thesis
proof
have "punit.lt (rep_list (poly_of_pair p)) = punit.lt (punit.spoly (rep_list a0) (rep_list b0))"
by (simp add: p rep_list_spair)
also have "... \<prec> ?l" by fact
also have "... = (?l - ?t1) + ?t1" by (simp only: adds_minus adds_lcs)
also have "... = (pp_of_term (sig_of_pair p) - lp a0) + ?t1"
by (simp only: eq1 pp_of_term_splus add_diff_cancel_right')
finally show "punit.lt (rep_list (poly_of_pair p)) \<prec> pp_of_term (sig_of_pair p) - lp a0 + ?t1" .
qed
qed
define y where "y = pp_of_term (sig_of_pair p) - lp b'"
from lt_p'' 4 have y2: "punit.lt (rep_list p') \<prec> y + punit.lt (rep_list b')"
unfolding y_def by (rule ordered_powerprod_lin.le_less_trans)
from 3 have "lt b' adds\<^sub>t sig_of_pair p" by (rule is_canon_rewriterD3)
hence "lp b' adds lp p'" and "component_of_term (lt b') = component_of_term (lt p')"
by (simp_all add: adds_term_def lt_p')
hence y1: "y \<oplus> lt b' = lt p'" by (simp add: y_def splus_def lt_p' adds_minus term_simps)
from 3 \<open>b \<in> set bs\<close> \<open>b \<noteq> 0\<close> \<open>lt b adds\<^sub>t sig_of_pair p\<close>
have "rword (spp_of b) (spp_of b')" by (rule is_canon_rewriterD)
hence "punit.lt (rep_list b') \<oplus> lt b \<preceq>\<^sub>t punit.lt (rep_list b) \<oplus> lt b'"
by (auto simp: \<open>rword = rw_rat\<close> rw_rat_def Let_def spp_of_def)
hence "(x + y) \<oplus> (punit.lt (rep_list b') \<oplus> lt b) \<preceq>\<^sub>t (x + y) \<oplus> (punit.lt (rep_list b) \<oplus> lt b')"
by (rule splus_mono)
hence "(y + punit.lt (rep_list b')) \<oplus> (x \<oplus> lt b) \<preceq>\<^sub>t (x + punit.lt (rep_list b)) \<oplus> (y \<oplus> lt b')"
by (simp add: ac_simps)
hence "(y + punit.lt (rep_list b')) \<oplus> lt p' \<preceq>\<^sub>t punit.lt (rep_list p') \<oplus> lt p'"
by (simp only: x1 x2 y1)
hence "y + punit.lt (rep_list b') \<preceq> punit.lt (rep_list p')" by (rule ord_term_canc_left)
with y2 show ?thesis by simp
next
fix j
assume p: "p = Inr j"
hence "lt p' = term_of_pair (0, j)" by (simp add: lt_p')
with x1 term_of_pair_pair[of "lt b"] have "lt b = term_of_pair (0, j)"
by (auto simp: splus_def dest!: term_of_pair_injective plus_eq_zero_2)
moreover have "lt b \<prec>\<^sub>t term_of_pair (0, j)" by (rule rb_aux_inv_D4, fact, simp add: p, fact)
ultimately show ?thesis by simp
qed
qed
moreover have "\<not> is_sig_red (\<prec>\<^sub>t) (=) (set bs) p'"
proof
assume "is_sig_red (\<prec>\<^sub>t) (=) (set bs) p'"
hence "is_sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs) p'" by (simp add: is_sig_red_top_tail_cases)
with sig_trd_irred show False unfolding p'_def ..
qed
ultimately show ?thesis by (simp add: p'_def is_sig_red_sing_reg_cases)
qed
lemma rb_aux_top_irred:
assumes "rword_strict = rw_rat_strict" and "rb_aux_inv (fst args)" and "b \<in> set (fst (fst (rb_aux args)))"
and "\<And>b0. b0 \<in> set (fst (fst args)) \<Longrightarrow> \<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst args)) - {b0}) b0"
shows "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst (rb_aux args))) - {b}) b"
proof -
from assms(2) have "rb_aux_dom args" by (rule rb_aux_domI)
thus ?thesis using assms(2, 3, 4)
proof (induct args rule: rb_aux.pinduct)
case (1 bs ss z)
let ?nil = "[]::((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list"
from 1(3) have "b \<in> set (fst (fst ((bs, ss, ?nil), z)))" by (simp add: rb_aux.psimps(1)[OF 1(1)])
hence "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst ((bs, ss, ?nil), z))) - {b}) b" by (rule 1(4))
thus ?case by (simp add: rb_aux.psimps(1)[OF 1(1)])
next
case (2 bs ss p ps z)
from 2(5) have *: "rb_aux_inv (bs, ss, p # ps)" by (simp only: fst_conv)
define p' where "p' = sig_trd bs (poly_of_pair p)"
from 2(6) show ?case
proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def p'_def[symmetric] split: if_splits)
note refl
moreover assume "sig_crit bs (new_syz_sigs ss bs p) p"
moreover from * this have "rb_aux_inv (fst ((bs, new_syz_sigs ss bs p, ps), z))"
unfolding fst_conv by (rule rb_aux_inv_preserved_1)
moreover assume "b \<in> set (fst (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z))))"
ultimately show "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)))) - {b}) b"
proof (rule 2(2))
fix b0
assume "b0 \<in> set (fst (fst ((bs, new_syz_sigs ss bs p, ps), z)))"
hence "b0 \<in> set (fst (fst ((bs, ss, p # ps), z)))" by simp
hence "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst ((bs, ss, p # ps), z))) - {b0}) b0" by (rule 2(7))
thus "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst ((bs, new_syz_sigs ss bs p, ps), z))) - {b0}) b0"
by simp
qed
next
note refl
moreover assume "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
moreover note refl
moreover assume "rep_list p' = 0"
moreover from * this have "rb_aux_inv (fst ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z))"
unfolding p'_def fst_conv by (rule rb_aux_inv_preserved_2)
moreover assume "b \<in> set (fst (fst (rb_aux ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z))))"
ultimately show "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst (rb_aux ((bs,
lt p' # new_syz_sigs ss bs p, ps), Suc z)))) - {b}) b"
proof (rule 2(3)[simplified p'_def[symmetric]])
fix b0
assume "b0 \<in> set (fst (fst ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z)))"
hence "b0 \<in> set (fst (fst ((bs, ss, p # ps), z)))" by simp
hence "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst ((bs, ss, p # ps), z))) - {b0}) b0" by (rule 2(7))
thus "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst ((bs, lt p' # new_syz_sigs ss bs p, ps), Suc z))) - {b0}) b0"
by simp
qed
next
note refl
moreover assume "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
moreover note refl
moreover assume "rep_list p' \<noteq> 0"
moreover from * \<open>\<not> sig_crit bs (new_syz_sigs ss bs p) p\<close> this
have inv: "rb_aux_inv (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))"
unfolding p'_def fst_conv by (rule rb_aux_inv_preserved_3)
moreover assume "b \<in> set (fst (fst (rb_aux ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))))"
ultimately show "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst (rb_aux ((p' # bs,
new_syz_sigs ss bs p, add_spairs ps bs p'), z)))) - {b}) b"
proof (rule 2(4)[simplified p'_def[symmetric]])
fix b0
assume "b0 \<in> set (fst (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)))"
hence "b0 = p' \<or> b0 \<in> set bs" by simp
hence "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (({p'} - {b0}) \<union> (set bs - {b0})) b0"
proof
assume "b0 = p'"
have "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set bs - {b0}) p'"
proof
assume "is_sig_red (\<preceq>\<^sub>t) (=) (set bs - {b0}) p'"
moreover have "set bs - {b0} \<subseteq> set bs" by fastforce
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) (set bs) p'" by (rule is_sig_red_mono)
moreover have "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set bs) p'" unfolding p'_def
using assms(1) * \<open>\<not> sig_crit bs (new_syz_sigs ss bs p) p\<close> by (rule rb_aux_top_irred')
ultimately show False by simp
qed
thus ?thesis by (simp add: \<open>b0 = p'\<close>)
next
assume "b0 \<in> set bs"
hence "b0 \<in> set (fst (fst ((bs, ss, p # ps), z)))" by simp
hence "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst ((bs, ss, p # ps), z))) - {b0}) b0" by (rule 2(7))
hence "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set bs - {b0}) b0" by simp
moreover have "\<not> is_sig_red (\<preceq>\<^sub>t) (=) ({p'} - {b0}) b0"
proof
assume "is_sig_red (\<preceq>\<^sub>t) (=) ({p'} - {b0}) b0"
moreover have "{p'} - {b0} \<subseteq> {p'}" by fastforce
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) {p'} b0" by (rule is_sig_red_mono)
hence "lt p' \<preceq>\<^sub>t lt b0" by (rule is_sig_redD_lt)
from inv have "rb_aux_inv (p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p')"
by (simp only: fst_conv)
hence "rb_aux_inv1 (p' # bs)" by (rule rb_aux_inv_D1)
hence "sorted_wrt (\<lambda>x y. lt y \<prec>\<^sub>t lt x) (p' # bs)" by (rule rb_aux_inv1_D3)
with \<open>b0 \<in> set bs\<close> have "lt b0 \<prec>\<^sub>t lt p'" by simp
with \<open>lt p' \<preceq>\<^sub>t lt b0\<close> show False by simp
qed
ultimately show ?thesis by (simp add: is_sig_red_Un)
qed
thus "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))) - {b0}) b0"
by (simp add: Un_Diff[symmetric])
qed
qed
qed
qed
corollary rb_aux_is_min_sig_GB:
assumes "rword_strict = rw_rat_strict"
shows "is_min_sig_GB dgrad (set (fst (fst (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))"
(is "is_min_sig_GB _ (set (fst (fst (rb_aux ?args))))")
unfolding is_min_sig_GB_def
proof (intro conjI allI ballI impI)
from rb_aux_inv_init_fst have inv: "rb_aux_inv (fst (rb_aux ?args))"
and "rb_aux_dom ?args"
by (rule rb_aux_inv_invariant, rule rb_aux_domI)
from this(2) obtain bs ss z' where eq: "rb_aux ?args = ((bs, ss, []), z')"
by (rule rb_aux_shape)
from inv have "rb_aux_inv (bs, ss, [])" by (simp only: eq fst_conv)
hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
thus "set (fst (fst (rb_aux ?args))) \<subseteq> dgrad_sig_set dgrad" by (simp add: eq)
next
fix u
show "is_sig_GB_in dgrad (set (fst (fst (rb_aux ?args)))) u" by (fact rb_aux_is_sig_GB_in)
next
fix g
assume "g \<in> set (fst (fst (rb_aux ?args)))"
with assms(1) rb_aux_inv_init_fst
show "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (set (fst (fst (rb_aux ?args))) - {g}) g"
by (rule rb_aux_top_irred) simp
qed
corollary rb_is_min_sig_GB:
assumes "rword_strict = rw_rat_strict"
shows "is_min_sig_GB dgrad (set (fst rb))"
using rb_aux_is_min_sig_GB[OF assms, of 0] by (auto simp: rb_def split: prod.split)
subsubsection \<open>No Zero-Reductions\<close>
fun rb_aux_inv2 :: "(('t \<Rightarrow>\<^sub>0 'b) list \<times> 't list \<times> ((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list) \<Rightarrow> bool"
where "rb_aux_inv2 (bs, ss, ps) =
(rb_aux_inv (bs, ss, ps) \<and>
(\<forall>j<length fs. Inr j \<notin> set ps \<longrightarrow>
(fs ! j \<in> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < Suc j) bs)) \<and>
(\<forall>b\<in>set bs. component_of_term (lt b) < j \<longrightarrow>
(\<exists>s\<in>set ss. s adds\<^sub>t term_of_pair (punit.lt (rep_list b), j))))))"
lemma rb_aux_inv2_D1: "rb_aux_inv2 args \<Longrightarrow> rb_aux_inv args"
by (metis prod.exhaust rb_aux_inv2.simps)
lemma rb_aux_inv2_D2:
"rb_aux_inv2 (bs, ss, ps) \<Longrightarrow> j < length fs \<Longrightarrow> Inr j \<notin> set ps \<Longrightarrow>
fs ! j \<in> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < Suc j) bs))"
by simp
lemma rb_aux_inv2_E:
assumes "rb_aux_inv2 (bs, ss, ps)" and "j < length fs" and "Inr j \<notin> set ps" and "b \<in> set bs"
and "component_of_term (lt b) < j"
obtains s where "s \<in> set ss" and "s adds\<^sub>t term_of_pair (punit.lt (rep_list b), j)"
using assms by auto
context
assumes pot: "is_pot_ord"
begin
lemma sig_red_zero_filter:
assumes "sig_red_zero (\<preceq>\<^sub>t) (set bs) r" and "component_of_term (lt r) < j"
shows "sig_red_zero (\<preceq>\<^sub>t) (set (filter (\<lambda>b. component_of_term (lt b) < j) bs)) r"
proof -
have "(\<preceq>\<^sub>t) = (\<preceq>\<^sub>t) \<or> (\<preceq>\<^sub>t) = (\<prec>\<^sub>t)" by simp
with assms(1) have "sig_red_zero (\<preceq>\<^sub>t) {b\<in>set bs. lt b \<preceq>\<^sub>t lt r} r" by (rule sig_red_zero_subset)
moreover have "{b\<in>set bs. lt b \<preceq>\<^sub>t lt r} \<subseteq> set (filter (\<lambda>b. component_of_term (lt b) < j) bs)"
proof
fix b
assume "b \<in> {b\<in>set bs. lt b \<preceq>\<^sub>t lt r}"
hence "b \<in> set bs" and "lt b \<preceq>\<^sub>t lt r" by simp_all
from pot this(2) have "component_of_term (lt b) \<le> component_of_term (lt r)" by (rule is_pot_ordD2)
also have "... < j" by (fact assms(2))
finally have "component_of_term (lt b) < j" .
with \<open>b \<in> set bs\<close> show "b \<in> set (filter (\<lambda>b. component_of_term (lt b) < j) bs)" by simp
qed
ultimately show ?thesis by (rule sig_red_zero_mono)
qed
lemma rb_aux_inv2_preserved_0:
assumes "rb_aux_inv2 (bs, ss, p # ps)" and "j < length fs" and "Inr j \<notin> set ps"
and "b \<in> set bs" and "component_of_term (lt b) < j"
shows "\<exists>s\<in>set (new_syz_sigs ss bs p). s adds\<^sub>t term_of_pair (punit.lt (rep_list b), j)"
proof (rule sum_prodE)
fix x y
assume p: "p = Inl (x, y)"
with assms(3) have "Inr j \<notin> set (p # ps)" by simp
with assms(1, 2) obtain s where "s \<in> set ss" and *: "s adds\<^sub>t term_of_pair (punit.lt (rep_list b), j)"
using assms(4, 5) by (rule rb_aux_inv2_E)
from this(1) have "s \<in> set (new_syz_sigs ss bs p)" by (simp add: p)
with * show ?thesis ..
next
fix i
assume p: "p = Inr i"
have trans: "transp (adds\<^sub>t)" by (rule transpI, drule adds_term_trans)
from adds_term_refl have refl: "reflp (adds\<^sub>t)" by (rule reflpI)
let ?v = "term_of_pair (punit.lt (rep_list b), j)"
let ?f = "\<lambda>b. term_of_pair (punit.lt (rep_list b), i)"
define ss' where "ss' = filter_min (adds\<^sub>t) (map ?f bs)"
have eq: "new_syz_sigs ss bs p = filter_min_append (adds\<^sub>t) ss ss'" by (simp add: p ss'_def pot)
show ?thesis
proof (cases "i = j")
case True
from \<open>b \<in> set bs\<close> have "?v \<in> ?f ` set bs" unfolding \<open>i = j\<close> by (rule imageI)
hence "?v \<in> set ss \<union> set (map ?f bs)" by simp
thus ?thesis
proof
assume "?v \<in> set ss"
hence "?v \<in> set ss \<union> set ss'" by simp
with trans refl obtain s where "s \<in> set (new_syz_sigs ss bs p)" and "s adds\<^sub>t ?v"
unfolding eq by (rule filter_min_append_relE)
thus ?thesis ..
next
assume "?v \<in> set (map ?f bs)"
with trans refl obtain s where "s \<in> set ss'" and "s adds\<^sub>t ?v"
unfolding ss'_def by (rule filter_min_relE)
from this(1) have "s \<in> set ss \<union> set ss'" by simp
with trans refl obtain s' where s': "s' \<in> set (new_syz_sigs ss bs p)" and "s' adds\<^sub>t s"
unfolding eq by (rule filter_min_append_relE)
from this(2) \<open>s adds\<^sub>t ?v\<close> have "s' adds\<^sub>t ?v" by (rule adds_term_trans)
with s' show ?thesis ..
qed
next
case False
with assms(3) have "Inr j \<notin> set (p # ps)" by (simp add: p)
with assms(1, 2) obtain s where "s \<in> set ss" and "s adds\<^sub>t ?v"
using assms(4, 5) by (rule rb_aux_inv2_E)
from this(1) have "s \<in> set ss \<union> set (map ?f bs)" by simp
thus ?thesis
proof
assume "s \<in> set ss"
hence "s \<in> set ss \<union> set ss'" by simp
with trans refl obtain s' where s': "s' \<in> set (new_syz_sigs ss bs p)" and "s' adds\<^sub>t s"
unfolding eq by (rule filter_min_append_relE)
from this(2) \<open>s adds\<^sub>t ?v\<close> have "s' adds\<^sub>t ?v" by (rule adds_term_trans)
with s' show ?thesis ..
next
assume "s \<in> set (map ?f bs)"
with trans refl obtain s' where "s' \<in> set ss'" and "s' adds\<^sub>t s"
unfolding ss'_def by (rule filter_min_relE)
from this(1) have "s' \<in> set ss \<union> set ss'" by simp
with trans refl obtain s'' where s'': "s'' \<in> set (new_syz_sigs ss bs p)" and "s'' adds\<^sub>t s'"
unfolding eq by (rule filter_min_append_relE)
from this(2) \<open>s' adds\<^sub>t s\<close> have "s'' adds\<^sub>t s" by (rule adds_term_trans)
hence "s'' adds\<^sub>t ?v" using \<open>s adds\<^sub>t ?v\<close> by (rule adds_term_trans)
with s'' show ?thesis ..
qed
qed
qed
lemma rb_aux_inv2_preserved_1:
assumes "rb_aux_inv2 (bs, ss, p # ps)" and "sig_crit bs (new_syz_sigs ss bs p) p"
shows "rb_aux_inv2 (bs, new_syz_sigs ss bs p, ps)"
unfolding rb_aux_inv2.simps
proof (intro allI conjI impI ballI)
from assms(1) have inv: "rb_aux_inv (bs, ss, p # ps)" by (rule rb_aux_inv2_D1)
thus "rb_aux_inv (bs, new_syz_sigs ss bs p, ps)"
using assms(2) by (rule rb_aux_inv_preserved_1)
fix j
assume "j < length fs" and "Inr j \<notin> set ps"
show "fs ! j \<in> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < Suc j) bs))"
proof (cases "p = Inr j")
case True
with assms(2) have "is_pred_syz (new_syz_sigs ss bs p) (term_of_pair (0, j))" by simp
let ?X = "set (filter (\<lambda>b. component_of_term (lt b) < Suc j) bs)"
have "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` ?X)"
proof (rule sig_red_zero_idealI)
have "sig_red_zero (\<prec>\<^sub>t) (set bs) (monomial 1 (term_of_pair (0, j)))"
proof (rule syzygy_crit)
from inv have "is_RB_upt dgrad rword (set bs) (sig_of_pair p)"
by (rule rb_aux_inv_is_RB_upt_Cons)
with dgrad have "is_sig_GB_upt dgrad (set bs) (sig_of_pair p)"
by (rule is_RB_upt_is_sig_GB_upt)
thus "is_sig_GB_upt dgrad (set bs) (term_of_pair (0, j))" by (simp add: \<open>p = Inr j\<close>)
next
show "monomial 1 (term_of_pair (0, j)) \<in> dgrad_sig_set dgrad"
by (rule dgrad_sig_set_closed_monomial, simp_all add: term_simps dgrad_max_0 \<open>j < length fs\<close>)
next
show "lt (monomial (1::'b) (term_of_pair (0, j))) = term_of_pair (0, j)" by (simp add: lt_monomial)
next
from inv assms(2) have "sig_crit' bs p" by (rule sig_crit'I_sig_crit)
thus "is_syz_sig dgrad (term_of_pair (0, j))" by (simp add: \<open>p = Inr j\<close>)
qed (fact dgrad)
hence "sig_red_zero (\<preceq>\<^sub>t) (set bs) (monomial 1 (term_of_pair (0, j)))"
by (rule sig_red_zero_sing_regI)
moreover have "component_of_term (lt (monomial (1::'b) (term_of_pair (0, j)))) < Suc j"
by (simp add: lt_monomial component_of_term_of_pair)
ultimately show "sig_red_zero (\<preceq>\<^sub>t) ?X (monomial 1 (term_of_pair (0, j)))"
by (rule sig_red_zero_filter)
qed
thus ?thesis by (simp add: rep_list_monomial' \<open>j < length fs\<close>)
next
case False
with \<open>Inr j \<notin> set ps\<close> have "Inr j \<notin> set (p # ps)" by simp
with assms(1) \<open>j < length fs\<close> show ?thesis by (rule rb_aux_inv2_D2)
qed
next
fix j b
assume "j < length fs" and "Inr j \<notin> set ps" and "b \<in> set bs" and "component_of_term (lt b) < j"
with assms(1) show "\<exists>s\<in>set (new_syz_sigs ss bs p). s adds\<^sub>t term_of_pair (punit.lt (rep_list b), j)"
by (rule rb_aux_inv2_preserved_0)
qed
lemma rb_aux_inv2_preserved_3:
assumes "rb_aux_inv2 (bs, ss, p # ps)" and "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
and "rep_list (sig_trd bs (poly_of_pair p)) \<noteq> 0"
shows "rb_aux_inv2 (sig_trd bs (poly_of_pair p) # bs, new_syz_sigs ss bs p,
add_spairs ps bs (sig_trd bs (poly_of_pair p)))"
proof -
from assms(1) have inv: "rb_aux_inv (bs, ss, p # ps)" by (rule rb_aux_inv2_D1)
define p' where "p' = sig_trd bs (poly_of_pair p)"
from sig_trd_red_rtrancl[of bs "poly_of_pair p"] have "lt p' = lt (poly_of_pair p)"
unfolding p'_def by (rule sig_red_regular_rtrancl_lt)
also have "... = sig_of_pair p" by (rule sym, rule pair_list_sig_of_pair, fact inv, simp)
finally have lt_p': "lt p' = sig_of_pair p" .
show ?thesis unfolding rb_aux_inv2.simps p'_def[symmetric]
proof (intro allI conjI impI ballI)
show "rb_aux_inv (p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p')"
unfolding p'_def using inv assms(2, 3) by (rule rb_aux_inv_preserved_3)
next
fix j
assume "j < length fs" and *: "Inr j \<notin> set (add_spairs ps bs p')"
show "fs ! j \<in> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < Suc j) (p' # bs)))"
proof (cases "p = Inr j")
case True
let ?X = "set (filter (\<lambda>b. component_of_term (lt b) < Suc j) (p' # bs))"
have "rep_list (monomial 1 (term_of_pair (0, j))) \<in> ideal (rep_list ` ?X)"
proof (rule sig_red_zero_idealI)
have "sig_red_zero (\<preceq>\<^sub>t) (set (p' # bs)) (monomial 1 (term_of_pair (0, j)))"
proof (rule sig_red_zeroI)
have "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) p'"
using sig_trd_red_rtrancl[of bs "poly_of_pair p"] by (simp add: True p'_def)
moreover have "set bs \<subseteq> set (p' # bs)" by fastforce
ultimately have "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set (p' # bs)))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) p'"
by (rule sig_red_rtrancl_mono)
hence "(sig_red (\<preceq>\<^sub>t) (\<preceq>) (set (p' # bs)))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) p'"
by (rule sig_red_rtrancl_sing_regI)
also have "sig_red (\<preceq>\<^sub>t) (\<preceq>) (set (p' # bs)) p' 0" unfolding sig_red_def
proof (intro exI bexI)
from assms(3) have "rep_list p' \<noteq> 0" by (simp add: p'_def)
show "sig_red_single (\<preceq>\<^sub>t) (\<preceq>) p' 0 p' 0"
proof (rule sig_red_singleI)
show "rep_list p' \<noteq> 0" by fact
next
from \<open>rep_list p' \<noteq> 0\<close> have "punit.lt (rep_list p') \<in> keys (rep_list p')"
by (rule punit.lt_in_keys)
thus "0 + punit.lt (rep_list p') \<in> keys (rep_list p')" by simp
next
from \<open>rep_list p' \<noteq> 0\<close> have "punit.lc (rep_list p') \<noteq> 0" by (rule punit.lc_not_0)
thus "0 = p' - monom_mult (lookup (rep_list p') (0 + punit.lt (rep_list p')) / punit.lc (rep_list p')) 0 p'"
by (simp add: punit.lc_def[symmetric])
qed (simp_all add: term_simps)
qed simp
finally show "(sig_red (\<preceq>\<^sub>t) (\<preceq>) (set (p' # bs)))\<^sup>*\<^sup>* (monomial 1 (term_of_pair (0, j))) 0" .
qed (fact rep_list_zero)
moreover have "component_of_term (lt (monomial (1::'b) (term_of_pair (0, j)))) < Suc j"
by (simp add: lt_monomial component_of_term_of_pair)
ultimately show "sig_red_zero (\<preceq>\<^sub>t) ?X (monomial 1 (term_of_pair (0, j)))"
by (rule sig_red_zero_filter)
qed
thus ?thesis by (simp add: rep_list_monomial' \<open>j < length fs\<close>)
next
case False
from * have "Inr j \<notin> set ps" by (simp add: add_spairs_def set_merge_wrt)
hence "Inr j \<notin> set (p # ps)" using False by simp
with assms(1) \<open>j < length fs\<close>
have "fs ! j \<in> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < Suc j) bs))"
by (rule rb_aux_inv2_D2)
also have "... \<subseteq> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < Suc j) (p' # bs)))"
by (intro ideal.span_mono image_mono, fastforce)
finally show ?thesis .
qed
next
fix j and b::"'t \<Rightarrow>\<^sub>0 'b"
assume "j < length fs" and *: "component_of_term (lt b) < j"
assume "Inr j \<notin> set (add_spairs ps bs p')"
hence "Inr j \<notin> set ps" by (simp add: add_spairs_def set_merge_wrt)
assume "b \<in> set (p' # bs)"
hence "b = p' \<or> b \<in> set bs" by simp
thus "\<exists>s\<in>set (new_syz_sigs ss bs p). s adds\<^sub>t term_of_pair (punit.lt (rep_list b), j)"
proof
assume "b = p'"
with * have "component_of_term (sig_of_pair p) < component_of_term (term_of_pair (0, j))"
by (simp only: lt_p' component_of_term_of_pair)
with pot have **: "sig_of_pair p \<prec>\<^sub>t term_of_pair (0, j)" by (rule is_pot_ordD)
have "p \<in> set (p # ps)" by simp
with inv have "Inr j \<in> set (p # ps)" using \<open>j < length fs\<close> ** by (rule rb_aux_inv_D6_2)
with \<open>Inr j \<notin> set ps\<close> have "p = Inr j" by simp
with ** show ?thesis by simp
next
assume "b \<in> set bs"
with assms(1) \<open>j < length fs\<close> \<open>Inr j \<notin> set ps\<close> show ?thesis
using * by (rule rb_aux_inv2_preserved_0)
qed
qed
qed
lemma rb_aux_inv2_ideal_subset:
assumes "rb_aux_inv2 (bs, ss, ps)" and "\<And>p0. p0 \<in> set ps \<Longrightarrow> j \<le> component_of_term (sig_of_pair p0)"
shows "ideal (set (take j fs)) \<subseteq> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < j) bs))"
(is "ideal ?B \<subseteq> ideal ?A")
proof (intro ideal.span_subset_spanI subsetI)
fix f
assume "f \<in> ?B"
then obtain i where "i < length (take j fs)" and "f = (take j fs) ! i"
by (metis in_set_conv_nth)
hence "i < length fs" and "i < j" and f: "f = fs ! i" by auto
from this(2) have "Suc i \<le> j" by simp
have "f \<in> ideal (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < Suc i) bs))"
unfolding f using assms(1) \<open>i < length fs\<close>
proof (rule rb_aux_inv2_D2)
show "Inr i \<notin> set ps"
proof
assume "Inr i \<in> set ps"
hence "j \<le> component_of_term (sig_of_pair (Inr i))" by (rule assms(2))
hence "j \<le> i" by (simp add: component_of_term_of_pair)
with \<open>i < j\<close> show False by simp
qed
qed
also have "... \<subseteq> ideal ?A"
by (intro ideal.span_mono image_mono, auto dest: order_less_le_trans[OF _ \<open>Suc i \<le> j\<close>])
finally show "f \<in> ideal ?A" .
qed
lemma rb_aux_inv_is_Groebner_basis:
assumes "hom_grading dgrad" and "rb_aux_inv (bs, ss, ps)"
and "\<And>p0. p0 \<in> set ps \<Longrightarrow> j \<le> component_of_term (sig_of_pair p0)"
shows "punit.is_Groebner_basis (rep_list ` set (filter (\<lambda>b. component_of_term (lt b) < j) bs))"
(is "punit.is_Groebner_basis (rep_list ` set ?bs)")
using dgrad assms(1)
proof (rule is_sig_GB_upt_is_Groebner_basis)
show "set ?bs \<subseteq> dgrad_sig_set' j dgrad"
proof
fix b
assume "b \<in> set ?bs"
hence "b \<in> set bs" and "component_of_term (lt b) < j" by simp_all
show "b \<in> dgrad_sig_set' j dgrad" unfolding dgrad_sig_set'_def
proof
from assms(2) have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
with \<open>b \<in> set bs\<close> have "b \<in> dgrad_sig_set dgrad" ..
thus "b \<in> dgrad_max_set dgrad" by (simp add: dgrad_sig_set'_def)
next
show "b \<in> sig_inv_set' j"
proof (rule sig_inv_setI')
fix v
assume "v \<in> keys b"
hence "v \<preceq>\<^sub>t lt b" by (rule lt_max_keys)
with pot have "component_of_term v \<le> component_of_term (lt b)" by (rule is_pot_ordD2)
also have "... < j" by fact
finally show "component_of_term v < j" .
qed
qed
qed
next
fix u
assume u: "component_of_term u < j"
from dgrad have "is_sig_GB_upt dgrad (set bs) (term_of_pair (0, j))"
proof (rule is_RB_upt_is_sig_GB_upt)
from assms(2) show "is_RB_upt dgrad rword (set bs) (term_of_pair (0, j))"
proof (rule rb_aux_inv_is_RB_upt)
fix p
assume "p \<in> set ps"
hence "j \<le> component_of_term (sig_of_pair p)" by (rule assms(3))
with pot show "term_of_pair (0, j) \<preceq>\<^sub>t sig_of_pair p"
by (auto simp: is_pot_ord term_simps zero_min)
qed
qed
moreover from pot have "u \<prec>\<^sub>t term_of_pair (0, j)"
by (rule is_pot_ordD) (simp only: u component_of_term_of_pair)
ultimately have 1: "is_sig_GB_in dgrad (set bs) u" by (rule is_sig_GB_uptD2)
show "is_sig_GB_in dgrad (set ?bs) u"
proof (rule is_sig_GB_inI)
fix r :: "'t \<Rightarrow>\<^sub>0 'b"
assume "lt r = u"
assume "r \<in> dgrad_sig_set dgrad"
with 1 have "sig_red_zero (\<preceq>\<^sub>t) (set bs) r" using \<open>lt r = u\<close> by (rule is_sig_GB_inD)
moreover from u have "component_of_term (lt r) < j" by (simp only: \<open>lt r = u\<close>)
ultimately show "sig_red_zero (\<preceq>\<^sub>t) (set ?bs) r" by (rule sig_red_zero_filter)
qed
qed
lemma rb_aux_inv2_no_zero_red:
assumes "hom_grading dgrad" and "is_regular_sequence fs" and "rb_aux_inv2 (bs, ss, p # ps)"
and "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
shows "rep_list (sig_trd bs (poly_of_pair p)) \<noteq> 0"
proof
from assms(3) have inv: "rb_aux_inv (bs, ss, p # ps)" by (rule rb_aux_inv2_D1)
moreover have "p \<in> set (p # ps)" by simp
ultimately have sig_p: "sig_of_pair p = lt (poly_of_pair p)" and "poly_of_pair p \<noteq> 0"
and p_in: "poly_of_pair p \<in> dgrad_sig_set dgrad"
by (rule pair_list_sig_of_pair, rule pair_list_nonzero, rule pair_list_dgrad_sig_set)
from this(2) have "lc (poly_of_pair p) \<noteq> 0" by (rule lc_not_0)
from inv have "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence bs_sub: "set bs \<subseteq> dgrad_sig_set dgrad" by (rule rb_aux_inv1_D1)
define p' where "p' = sig_trd bs (poly_of_pair p)"
define j where "j = component_of_term (lt p')"
define q where "q = lookup (vectorize_poly p') j"
let ?bs = "filter (\<lambda>b. component_of_term (lt b) < j) bs"
let ?fs = "take (Suc j) fs"
have "p' \<in> dgrad_sig_set dgrad" unfolding p'_def using dgrad bs_sub p_in sig_trd_red_rtrancl
by (rule dgrad_sig_set_closed_sig_red_rtrancl)
hence "p' \<in> sig_inv_set" by (simp add: dgrad_sig_set'_def)
have lt_p': "lt p' = lt (poly_of_pair p)" and "lc p' = lc (poly_of_pair p)"
unfolding p'_def using sig_trd_red_rtrancl
by (rule sig_red_regular_rtrancl_lt, rule sig_red_regular_rtrancl_lc)
from this(2) \<open>lc (poly_of_pair p) \<noteq> 0\<close> have "p' \<noteq> 0" by (simp add: lc_eq_zero_iff[symmetric])
hence "lt p' \<in> keys p'" by (rule lt_in_keys)
hence "j \<in> keys (vectorize_poly p')" by (simp add: keys_vectorize_poly j_def)
hence "q \<noteq> 0" by (simp add: q_def in_keys_iff)
from \<open>p' \<in> sig_inv_set\<close> \<open>lt p' \<in> keys p'\<close> have "j < length fs"
unfolding j_def by (rule sig_inv_setD')
with le_refl have "fs ! j \<in> set (drop j fs)" by (rule nth_in_set_dropI)
with fs_distinct le_refl have 0: "fs ! j \<notin> set (take j fs)"
by (auto dest: set_take_disj_set_drop_if_distinct)
have 1: "j \<le> component_of_term (sig_of_pair p0)" if "p0 \<in> set (p # ps)" for p0
proof -
from that have "p0 = p \<or> p0 \<in> set ps" by simp
thus ?thesis
proof
assume "p0 = p"
thus ?thesis by (simp add: j_def lt_p' sig_p)
next
assume "p0 \<in> set ps"
from inv have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
hence "Ball (set ps) (pair_ord p)" by simp
hence "pair_ord p p0" using \<open>p0 \<in> set ps\<close> ..
hence "lt p' \<preceq>\<^sub>t sig_of_pair p0" by (simp add: pair_ord_def lt_p' sig_p)
thus ?thesis using pot by (auto simp add: is_pot_ord j_def term_simps)
qed
qed
with assms(1) inv have gb: "punit.is_Groebner_basis (rep_list ` set ?bs)"
by (rule rb_aux_inv_is_Groebner_basis)
have "p' \<in> sig_inv_set' (Suc j)"
proof (rule sig_inv_setI')
fix v
assume "v \<in> keys p'"
hence "v \<preceq>\<^sub>t lt p'" by (rule lt_max_keys)
with pot have "component_of_term v \<le> j" unfolding j_def by (rule is_pot_ordD2)
thus "component_of_term v < Suc j" by simp
qed
hence 2: "keys (vectorize_poly p') \<subseteq> {0..<Suc j}" by (rule sig_inv_setD)
moreover assume "rep_list p' = 0"
ultimately have "0 = (\<Sum>k\<in>keys (pm_of_idx_pm ?fs (vectorize_poly p')).
lookup (pm_of_idx_pm ?fs (vectorize_poly p')) k * k)"
by (simp add: rep_list_def ideal.rep_def pm_of_idx_pm_take)
also have "... = (\<Sum>k\<in>set ?fs. lookup (pm_of_idx_pm ?fs (vectorize_poly p')) k * k)"
using finite_set keys_pm_of_idx_pm_subset by (rule sum.mono_neutral_left) (simp add: in_keys_iff)
also from 2 have "... = (\<Sum>k\<in>set ?fs. lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
by (simp only: pm_of_idx_pm_take)
also have "... = lookup (pm_of_idx_pm fs (vectorize_poly p')) (fs ! j) * fs ! j +
(\<Sum>k\<in>set (take j fs). lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
using \<open>j < length fs\<close> by (simp add: take_Suc_conv_app_nth q_def sum.insert[OF finite_set 0])
also have "... = q * fs ! j + (\<Sum>k\<in>set (take j fs). lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
using fs_distinct \<open>j < length fs\<close> by (simp only: lookup_pm_of_idx_pm_distinct q_def)
finally have "- (q * fs ! j) =
(\<Sum>k\<in>set (take j fs). lookup (pm_of_idx_pm fs (vectorize_poly p')) k * k)"
by (simp add: add_eq_0_iff)
hence "- (q * fs ! j) \<in> ideal (set (take j fs))" by (simp add: ideal.sum_in_spanI)
hence "- (- (q * fs ! j)) \<in> ideal (set (take j fs))" by (rule ideal.span_neg)
hence "q * fs ! j \<in> ideal (set (take j fs))" by simp
with assms(2) \<open>j < length fs\<close> have "q \<in> ideal (set (take j fs))" by (rule is_regular_sequenceD)
also from assms(3) 1 have "... \<subseteq> ideal (rep_list ` set ?bs)"
by (rule rb_aux_inv2_ideal_subset)
finally have "q \<in> ideal (rep_list ` set ?bs)" .
with gb obtain g where "g \<in> rep_list ` set ?bs" and "g \<noteq> 0" and "punit.lt g adds punit.lt q"
using \<open>q \<noteq> 0\<close> by (rule punit.GB_adds_lt[simplified])
from this(1) obtain b where "b \<in> set bs" and "component_of_term (lt b) < j" and g: "g = rep_list b"
by auto
from assms(3) \<open>j < length fs\<close> _ this(1, 2)
have "\<exists>s\<in>set (new_syz_sigs ss bs p). s adds\<^sub>t term_of_pair (punit.lt (rep_list b), j)"
proof (rule rb_aux_inv2_preserved_0)
show "Inr j \<notin> set ps"
proof
assume "Inr j \<in> set ps"
with inv have "sig_of_pair p \<noteq> term_of_pair (0, j)" by (rule Inr_in_tailD)
hence "lt p' \<noteq> term_of_pair (0, j)" by (simp add: lt_p' sig_p)
from inv have "sorted_wrt pair_ord (p # ps)" by (rule rb_aux_inv_D5)
hence "Ball (set ps) (pair_ord p)" by simp
hence "pair_ord p (Inr j)" using \<open>Inr j \<in> set ps\<close> ..
hence "lt p' \<preceq>\<^sub>t term_of_pair (0, j)" by (simp add: pair_ord_def lt_p' sig_p)
hence "lp p' \<preceq> 0" using pot by (simp add: is_pot_ord j_def term_simps)
- hence "lp p' = 0" using zero_min by (rule ordered_powerprod_lin.antisym)
+ hence "lp p' = 0" using zero_min by (rule ordered_powerprod_lin.order_antisym)
hence "lt p' = term_of_pair (0, j)" by (metis j_def term_of_pair_pair)
with \<open>lt p' \<noteq> term_of_pair (0, j)\<close> show False ..
qed
qed
then obtain s where s_in: "s \<in> set (new_syz_sigs ss bs p)" and "s adds\<^sub>t term_of_pair (punit.lt g, j)"
unfolding g ..
from this(2) \<open>punit.lt g adds punit.lt q\<close> have "s adds\<^sub>t term_of_pair (punit.lt q, j)"
by (metis adds_minus_splus adds_term_splus component_of_term_of_pair pp_of_term_of_pair)
also have "... = lt p'" by (simp only: q_def j_def lt_lookup_vectorize term_simps)
finally have "s adds\<^sub>t sig_of_pair p" by (simp only: lt_p' sig_p)
with s_in have pred: "is_pred_syz (new_syz_sigs ss bs p) (sig_of_pair p)"
by (auto simp: is_pred_syz_def)
have "sig_crit bs (new_syz_sigs ss bs p) p"
proof (rule sum_prodE)
fix x y
assume "p = Inl (x, y)"
thus ?thesis using pred by (auto simp: ord_term_lin.max_def split: if_splits)
next
fix i
assume "p = Inr i"
thus ?thesis using pred by simp
qed
with assms(4) show False ..
qed
corollary rb_aux_no_zero_red':
assumes "hom_grading dgrad" and "is_regular_sequence fs" and "rb_aux_inv2 (fst args)"
shows "snd (rb_aux args) = snd args"
proof -
from assms(3) have "rb_aux_inv (fst args)" by (rule rb_aux_inv2_D1)
hence "rb_aux_dom args" by (rule rb_aux_domI)
thus ?thesis using assms(3)
proof (induct args rule: rb_aux.pinduct)
case (1 bs ss z)
show ?case by (simp only: rb_aux.psimps(1)[OF 1(1)])
next
case (2 bs ss p ps z)
from 2(5) have *: "rb_aux_inv2 (bs, ss, p # ps)" by (simp only: fst_conv)
show ?case
proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
note refl
moreover assume "sig_crit bs (new_syz_sigs ss bs p) p"
moreover from * this have "rb_aux_inv2 (fst ((bs, new_syz_sigs ss bs p, ps), z))"
unfolding fst_conv by (rule rb_aux_inv2_preserved_1)
ultimately have "snd (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
snd ((bs, new_syz_sigs ss bs p, ps), z)" by (rule 2(2))
thus "snd (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) = z" by (simp only: snd_conv)
thus "snd (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) = z" .
next
assume "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
with assms(1, 2) * have "rep_list (sig_trd bs (poly_of_pair p)) \<noteq> 0"
by (rule rb_aux_inv2_no_zero_red)
moreover assume "rep_list (sig_trd bs (poly_of_pair p)) = 0"
ultimately show "snd (rb_aux ((bs, lt (sig_trd bs (poly_of_pair p)) #
new_syz_sigs ss bs p, ps), Suc z)) = z" ..
next
define p' where "p' = sig_trd bs (poly_of_pair p)"
note refl
moreover assume a: "\<not> sig_crit bs (new_syz_sigs ss bs p) p"
moreover note p'_def
moreover assume b: "rep_list p' \<noteq> 0"
moreover have "rb_aux_inv2 (fst ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z))"
using * a b unfolding fst_conv p'_def by (rule rb_aux_inv2_preserved_3)
ultimately have "snd (rb_aux ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)) =
snd ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)"
by (rule 2(4))
thus "snd (rb_aux ((p' # bs, new_syz_sigs ss bs p, add_spairs ps bs p'), z)) = z"
by (simp only: snd_conv)
qed
qed
qed
corollary rb_aux_no_zero_red:
assumes "hom_grading dgrad" and "is_regular_sequence fs"
shows "snd (rb_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)) = z"
proof -
let ?args = "(([]::('t \<Rightarrow>\<^sub>0 'b) list, Koszul_syz_sigs fs,
(map Inr [0..<length fs])::((('t \<Rightarrow>\<^sub>0 'b) \<times> ('t \<Rightarrow>\<^sub>0 'b)) + nat) list), z)"
from rb_aux_inv_init have "rb_aux_inv2 (fst ?args)" by simp
with assms have "snd (rb_aux ?args) = snd ?args" by (rule rb_aux_no_zero_red')
thus ?thesis by (simp only: snd_conv)
qed
corollary rb_no_zero_red:
assumes "hom_grading dgrad" and "is_regular_sequence fs"
shows "snd rb = 0"
using rb_aux_no_zero_red[OF assms, of 0] by (auto simp: rb_def split: prod.split)
end
subsection \<open>Sig-Poly-Pairs\<close>
text \<open>We now prove that the algorithms defined for sig-poly-pairs (i.\,e. those whose names end with
\<open>_spp\<close>) behave exactly as those defined for module elements. More precisely, if \<open>A\<close> is some
algorithm defined for module elements, we prove something like
@{prop "spp_of (A x) = A_spp (spp_of x)"}.\<close>
fun spp_inv_pair :: "((('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<times> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b))) + nat) \<Rightarrow> bool" where
"spp_inv_pair (Inl (p, q)) = (spp_inv p \<and> spp_inv q)" |
"spp_inv_pair (Inr j) = True"
fun app_pair :: "('x \<Rightarrow> 'y) \<Rightarrow> (('x \<times> 'x) + nat) \<Rightarrow> (('y \<times> 'y) + nat)" where
"app_pair f (Inl (p, q)) = Inl (f p, f q)" |
"app_pair f (Inr j) = Inr j"
fun app_args :: "('x \<Rightarrow> 'y) \<Rightarrow> (('x list \<times> 'z \<times> ((('x \<times> 'x) + nat) list)) \<times> nat) \<Rightarrow>
(('y list \<times> 'z \<times> ((('y \<times> 'y) + nat) list)) \<times> nat)" where
"app_args f ((as, bs, cs), n) = ((map f as, bs, map (app_pair f) cs), n)"
lemma app_pair_spp_of_vec_of:
assumes "spp_inv_pair p"
shows "app_pair spp_of (app_pair vec_of p) = p"
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
from assms have "spp_inv a" and "spp_inv b" by (simp_all add: p)
thus ?thesis by (simp add: p spp_of_vec_of)
qed simp
lemma map_app_pair_spp_of_vec_of:
assumes "list_all spp_inv_pair ps"
shows "map (app_pair spp_of \<circ> app_pair vec_of) ps = ps"
proof (rule map_idI)
fix p
assume "p \<in> set ps"
with assms have "spp_inv_pair p" by (simp add: list_all_def)
hence "app_pair spp_of (app_pair vec_of p) = p" by (rule app_pair_spp_of_vec_of)
thus "(app_pair spp_of \<circ> app_pair vec_of) p = p" by simp
qed
lemma snd_app_args: "snd (app_args f args) = snd args"
by (metis prod.exhaust app_args.simps snd_conv)
lemma new_syz_sigs_alt_spp:
"new_syz_sigs ss bs p = new_syz_sigs_spp ss (map spp_of bs) (app_pair spp_of p)"
proof (rule sum_prodE)
fix a b
assume "p = Inl (a, b)"
thus ?thesis by simp
next
fix j
assume "p = Inr j"
thus ?thesis by (simp add: comp_def spp_of_def)
qed
lemma is_rewritable_alt_spp:
assumes "0 \<notin> set bs"
shows "is_rewritable bs p u = is_rewritable_spp (map spp_of bs) (spp_of p) u"
proof -
from assms have "b \<in> set bs \<Longrightarrow> b \<noteq> 0" for b by blast
thus ?thesis by (auto simp: is_rewritable_def is_rewritable_spp_def fst_spp_of)
qed
lemma spair_sigs_alt_spp: "spair_sigs p q = spair_sigs_spp (spp_of p) (spp_of q)"
by (simp add: spair_sigs_def spair_sigs_spp_def Let_def fst_spp_of snd_spp_of)
lemma sig_crit_alt_spp:
assumes "0 \<notin> set bs"
shows "sig_crit bs ss p = sig_crit_spp (map spp_of bs) ss (app_pair spp_of p)"
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
from assms show ?thesis by (simp add: p spair_sigs_alt_spp is_rewritable_alt_spp)
qed simp
lemma spair_alt_spp:
assumes "is_regular_spair p q"
shows "spp_of (spair p q) = spair_spp (spp_of p) (spp_of q)"
proof -
let ?t1 = "punit.lt (rep_list p)"
let ?t2 = "punit.lt (rep_list q)"
let ?l = "lcs ?t1 ?t2"
from assms have p: "rep_list p \<noteq> 0" and q: "rep_list q \<noteq> 0"
by (rule is_regular_spairD1, rule is_regular_spairD2)
hence "p \<noteq> 0" and "q \<noteq> 0" and 1: "punit.lc (rep_list p) \<noteq> 0" and 2: "punit.lc (rep_list q) \<noteq> 0"
by (auto simp: rep_list_zero punit.lc_eq_zero_iff)
from assms have "lt (monom_mult (1 / punit.lc (rep_list p)) (?l - ?t1) p) \<noteq>
lt (monom_mult (1 / punit.lc (rep_list q)) (?l - ?t2) q)" (is "?u \<noteq> ?v")
by (rule is_regular_spairD3)
hence "lt (monom_mult (1 / punit.lc (rep_list p)) (?l - ?t1) p - monom_mult (1 / punit.lc (rep_list q)) (?l - ?t2) q) =
ord_term_lin.max ?u ?v" by (rule lt_minus_distinct_eq_max)
moreover from \<open>p \<noteq> 0\<close> 1 have "?u = (?l - ?t1) \<oplus> fst (spp_of p)" by (simp add: lt_monom_mult fst_spp_of)
moreover from \<open>q \<noteq> 0\<close> 2 have "?v = (?l - ?t2) \<oplus> fst (spp_of q)" by (simp add: lt_monom_mult fst_spp_of)
ultimately show ?thesis
by (simp add: spair_spp_def spair_def Let_def spp_of_def rep_list_minus rep_list_monom_mult)
qed
lemma sig_trd_spp_body_alt_Some:
assumes "find_sig_reducer (map spp_of bs) v (punit.lt p) 0 = Some i"
shows "sig_trd_spp_body (map spp_of bs) v (p, r) =
(punit.lower (p - local.punit.monom_mult (punit.lc p / punit.lc (rep_list (bs ! i)))
(punit.lt p - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i))) (punit.lt p), r)"
(is ?thesis1)
and "sig_trd_spp_body (map spp_of bs) v (p, r) =
(p - local.punit.monom_mult (punit.lc p / punit.lc (rep_list (bs ! i)))
(punit.lt p - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i)), r)"
(is ?thesis2)
proof -
have "?thesis1 \<and> ?thesis2"
proof (cases "p = 0")
case True
show ?thesis by (simp add: assms, simp add: True)
next
case False
from assms have "i < length bs" by (rule find_sig_reducer_SomeD)
hence eq1: "snd (map spp_of bs ! i) = rep_list (bs ! i)" by (simp add: snd_spp_of)
from assms have "rep_list (bs ! i) \<noteq> 0" and "punit.lt (rep_list (bs ! i)) adds punit.lt p"
by (rule find_sig_reducer_SomeD)+
hence nz: "rep_list (bs ! i) \<noteq> 0" and adds: "punit.lt (rep_list (bs ! i)) adds punit.lt p"
by (simp_all add: snd_spp_of)
from nz have "punit.lc (rep_list (bs ! i)) \<noteq> 0" by (rule punit.lc_not_0)
moreover from False have "punit.lc p \<noteq> 0" by (rule punit.lc_not_0)
ultimately have eq2: "punit.lt (punit.monom_mult (punit.lc p / punit.lc (rep_list (bs ! i)))
(punit.lt p - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i))) = punit.lt p"
(is "punit.lt ?p = _") using nz adds by (simp add: lp_monom_mult adds_minus)
have ?thesis1 by (simp add: assms Let_def eq1 punit.lower_minus punit.tail_monom_mult[symmetric],
simp add: punit.tail_def eq2)
moreover have ?thesis2
proof (simp add: \<open>?thesis1\<close> punit.lower_id_iff disj_commute[of "p = ?p"] del: sig_trd_spp_body.simps)
show "punit.lt (p - ?p) \<prec> punit.lt p \<or> p = ?p"
proof (rule disjCI)
assume "p \<noteq> ?p"
hence "p - ?p \<noteq> 0" by simp
moreover note eq2
moreover from \<open>punit.lc (rep_list (bs ! i)) \<noteq> 0\<close> have "punit.lc ?p = punit.lc p" by simp
ultimately show "punit.lt (p - ?p) \<prec> punit.lt p" by (rule punit.lt_minus_lessI)
qed
qed
ultimately show ?thesis ..
qed
thus ?thesis1 and ?thesis2 by blast+
qed
lemma sig_trd_aux_alt_spp:
assumes "fst args \<in> keys (rep_list (snd args))"
shows "rep_list (sig_trd_aux bs args) =
sig_trd_spp_aux (map spp_of bs) (lt (snd args))
(rep_list (snd args) - punit.higher (rep_list (snd args)) (fst args),
punit.higher (rep_list (snd args)) (fst args))"
proof -
from assms have "sig_trd_aux_dom bs args" by (rule sig_trd_aux_domI)
thus ?thesis using assms
proof (induct args rule: sig_trd_aux.pinduct)
case (1 t p)
define p' where "p' = (case find_sig_reducer (map spp_of bs) (lt p) t 0 of
None \<Rightarrow> p
| Some i \<Rightarrow> p -
monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i))"
define p'' where "p'' = punit.lower (rep_list p') t"
from 1(3) have t_in: "t \<in> keys (rep_list p)" by simp
hence "t \<in> keys (rep_list p - punit.higher (rep_list p) t)" (is "_ \<in> keys ?p")
by (simp add: punit.keys_minus_higher)
hence "?p \<noteq> 0" by auto
hence eq1: "sig_trd_spp_aux bs0 v0 (?p, r0) = sig_trd_spp_aux bs0 v0 (sig_trd_spp_body bs0 v0 (?p, r0))"
for bs0 v0 r0 by (simp add: sig_trd_spp_aux_simps del: sig_trd_spp_body.simps)
from t_in have lt_p: "punit.lt ?p = t" and lc_p: "punit.lc ?p = lookup (rep_list p) t"
and tail_p: "punit.tail ?p = punit.lower (rep_list p) t"
by (rule punit.lt_minus_higher, rule punit.lc_minus_higher, rule punit.tail_minus_higher)
have "lt p' = lt p \<and> punit.higher (rep_list p') t = punit.higher (rep_list p) t \<and>
(\<forall>i. find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i \<longrightarrow> lookup (rep_list p') t = 0)"
(is "?A \<and> ?B \<and> ?C")
proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
case None
thus ?thesis by (simp add: p'_def)
next
case (Some i)
hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
from Some have "punit.lt (rep_list (bs ! i)) adds t" by (rule find_sig_reducer_SomeD)
hence eq: "t - punit.lt (rep_list (bs ! i)) + punit.lt (rep_list (bs ! i)) = t" by (rule adds_minus)
from t_in Some have *: "sig_red_single (\<prec>\<^sub>t) (\<preceq>) p p' (bs ! i) (t - punit.lt (rep_list (bs ! i)))"
unfolding p' by (rule find_sig_reducer_SomeD_red_single)
hence **: "punit.red_single (rep_list p) (rep_list p') (rep_list (bs ! i)) (t - punit.lt (rep_list (bs ! i)))"
by (rule sig_red_single_red_single)
from * have ?A by (rule sig_red_single_regular_lt)
moreover from punit.red_single_higher[OF **] have ?B by (simp add: eq)
moreover have ?C
proof (intro allI impI)
from punit.red_single_lookup[OF **] show "lookup (rep_list p') t = 0" by (simp add: eq)
qed
ultimately show ?thesis by (intro conjI)
qed
hence lt_p': "lt p' = lt p" and higher_p': "punit.higher (rep_list p') t = punit.higher (rep_list p) t"
and lookup_p': "\<And>i. find_sig_reducer (map spp_of bs) (lt p) t 0 = Some i \<Longrightarrow> lookup (rep_list p') t = 0"
by blast+
show ?case
proof (simp add: sig_trd_aux.psimps[OF 1(1)] Let_def p'_def[symmetric] p''_def[symmetric], intro conjI impI)
assume "p'' = 0"
hence p'_decomp: "punit.higher (rep_list p) t + monomial (lookup (rep_list p') t) t = rep_list p'"
using punit.higher_lower_decomp[of "rep_list p'" t] by (simp add: p''_def higher_p')
show "rep_list p' = sig_trd_spp_aux (map spp_of bs) (lt p) (?p, punit.higher (rep_list p) t)"
proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
case None
hence p': "p' = p" by (simp add: p'_def)
from \<open>p'' = 0\<close> have eq2: "punit.tail ?p = 0" by (simp add: tail_p p''_def p')
from p'_decomp show ?thesis by (simp add: p' eq1 lt_p lc_p None eq2 sig_trd_spp_aux_simps)
next
case (Some i)
hence p': "p' = p - monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (bs ! i)" by (simp add: p'_def)
from \<open>p'' = 0\<close> have eq2: "punit.lower (rep_list p - punit.higher (rep_list p) t -
punit.monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i)))
t = 0"
by (simp add: p''_def p' rep_list_minus rep_list_monom_mult punit.lower_minus punit.lower_higher_zeroI)
from Some have "lookup (rep_list p') t = 0" by (rule lookup_p')
with p'_decomp have eq3: "rep_list p' = punit.higher (rep_list p) t" by simp
show ?thesis by (simp add: sig_trd_spp_body_alt_Some(1) eq1 eq2 lt_p lc_p Some del: sig_trd_spp_body.simps,
simp add: sig_trd_spp_aux_simps eq3)
qed
next
assume "p'' \<noteq> 0"
hence "punit.lt p'' \<prec> t" unfolding p''_def by (rule punit.lt_lower_less)
have higher_p'_2: "punit.higher (rep_list p') (punit.lt p'') =
punit.higher (rep_list p) t + monomial (lookup (rep_list p') t) t"
proof (simp add: higher_p'[symmetric], rule poly_mapping_eqI)
fix s
show "lookup (punit.higher (rep_list p') (punit.lt p'')) s =
lookup (punit.higher (rep_list p') t + monomial (lookup (rep_list p') t) t) s"
proof (rule ordered_powerprod_lin.linorder_cases)
assume "t \<prec> s"
moreover from \<open>punit.lt p'' \<prec> t\<close> this have "punit.lt p'' \<prec> s"
by (rule ordered_powerprod_lin.less_trans)
ultimately show ?thesis by (simp add: lookup_add punit.lookup_higher_when lookup_single)
next
assume "t = s"
with \<open>punit.lt p'' \<prec> t\<close> show ?thesis by (simp add: lookup_add punit.lookup_higher_when)
next
assume "s \<prec> t"
show ?thesis
proof (cases "punit.lt p'' \<prec> s")
case True
hence "lookup (punit.higher (rep_list p') (punit.lt p'')) s = lookup (rep_list p') s"
by (simp add: punit.lookup_higher_when)
also from \<open>s \<prec> t\<close> have "... = lookup p'' s" by (simp add: p''_def punit.lookup_lower_when)
also from True have "... = 0" using punit.lt_le_iff by auto
finally show ?thesis using \<open>s \<prec> t\<close>
by (simp add: lookup_add lookup_single punit.lookup_higher_when)
next
case False
with \<open>s \<prec> t\<close> show ?thesis by (simp add: lookup_add punit.lookup_higher_when lookup_single)
qed
qed
qed
have "rep_list (sig_trd_aux bs (punit.lt p'', p')) =
sig_trd_spp_aux (map spp_of bs) (lt (snd (punit.lt p'', p')))
(rep_list (snd (punit.lt p'', p')) -
punit.higher (rep_list (snd (punit.lt p'', p'))) (fst (punit.lt p'', p')),
punit.higher (rep_list (snd (punit.lt p'', p'))) (fst (punit.lt p'', p')))"
using p'_def p''_def \<open>p'' \<noteq> 0\<close>
proof (rule 1(2))
from \<open>p'' \<noteq> 0\<close> have "punit.lt p'' \<in> keys p''" by (rule punit.lt_in_keys)
also have "... \<subseteq> keys (rep_list p')" by (auto simp: p''_def punit.keys_lower)
finally show "fst (punit.lt p'', p') \<in> keys (rep_list (snd (punit.lt p'', p')))" by simp
qed
also have "... = sig_trd_spp_aux (map spp_of bs) (lt p)
(rep_list p' - punit.higher (rep_list p') (punit.lt p''),
punit.higher (rep_list p') (punit.lt p''))"
by (simp only: lt_p' fst_conv snd_conv)
also have "... = sig_trd_spp_aux (map spp_of bs) (lt p) (?p, punit.higher (rep_list p) t)"
proof (cases "find_sig_reducer (map spp_of bs) (lt p) t 0")
case None
hence p': "p' = p" by (simp add: p'_def)
have "rep_list p - (punit.higher (rep_list p) t + monomial (lookup (rep_list p) t) t) =
punit.lower (rep_list p) t"
using punit.higher_lower_decomp[of "rep_list p" t] by (simp add: diff_eq_eq ac_simps)
with higher_p'_2 show ?thesis by (simp add: eq1 lt_p lc_p tail_p p' None)
next
case (Some i)
hence p': "rep_list p - punit.monom_mult (lookup (rep_list p) t / punit.lc (rep_list (bs ! i)))
(t - punit.lt (rep_list (bs ! i))) (rep_list (bs ! i)) = rep_list p'"
by (simp add: p'_def rep_list_minus rep_list_monom_mult)
from Some have "lookup (rep_list p') t = 0" by (rule lookup_p')
with higher_p'_2 show ?thesis
by (simp add: sig_trd_spp_body_alt_Some(2) eq1 lt_p lc_p tail_p Some
diff_right_commute[of "rep_list p" "punit.higher (rep_list p) t"] p' del: sig_trd_spp_body.simps)
qed
finally show "rep_list (sig_trd_aux bs (punit.lt p'', p')) =
sig_trd_spp_aux (map spp_of bs) (lt p) (?p, punit.higher (rep_list p) t)" .
qed
qed
qed
lemma sig_trd_alt_spp: "spp_of (sig_trd bs p) = sig_trd_spp (map spp_of bs) (spp_of p)"
unfolding sig_trd_def
proof (split if_split, intro conjI impI)
assume "rep_list p = 0"
thus "spp_of p = sig_trd_spp (map spp_of bs) (spp_of p)" by (simp add: spp_of_def sig_trd_spp_aux_simps)
next
let ?args = "(punit.lt (rep_list p), p)"
assume "rep_list p \<noteq> 0"
hence a: "fst ?args \<in> keys (rep_list (snd ?args))" by (simp add: punit.lt_in_keys)
hence "(sig_red (\<prec>\<^sub>t) (\<preceq>) (set bs))\<^sup>*\<^sup>* (snd ?args) (sig_trd_aux bs ?args)"
by (rule sig_trd_aux_red_rtrancl)
hence eq1: "lt (sig_trd_aux bs ?args) = lt (snd ?args)" by (rule sig_red_regular_rtrancl_lt)
have eq2: "punit.higher (rep_list p) (punit.lt (rep_list p)) = 0"
by (auto simp: punit.higher_eq_zero_iff punit.lt_max simp flip: not_in_keys_iff_lookup_eq_zero
dest: punit.lt_max_keys)
show "spp_of (sig_trd_aux bs (punit.lt (rep_list p), p)) = sig_trd_spp (map spp_of bs) (spp_of p)"
by (simp add: spp_of_def eq1 eq2 sig_trd_aux_alt_spp[OF a])
qed
lemma is_regular_spair_alt_spp: "is_regular_spair p q \<longleftrightarrow> is_regular_spair_spp (spp_of p) (spp_of q)"
by (auto simp: is_regular_spair_spp_def fst_spp_of snd_spp_of intro: is_regular_spairI
dest: is_regular_spairD1 is_regular_spairD2 is_regular_spairD3)
lemma sig_of_spair_alt_spp: "sig_of_pair p = sig_of_pair_spp (app_pair spp_of p)"
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
show ?thesis by (simp add: p spair_sigs_def spair_sigs_spp_def spp_of_def)
qed simp
lemma pair_ord_alt_spp: "pair_ord x y \<longleftrightarrow> pair_ord_spp (app_pair spp_of x) (app_pair spp_of y)"
by (simp add: pair_ord_spp_def pair_ord_def sig_of_spair_alt_spp)
lemma new_spairs_alt_spp:
"map (app_pair spp_of) (new_spairs bs p) = new_spairs_spp (map spp_of bs) (spp_of p)"
proof (induct bs)
case Nil
show ?case by simp
next
case (Cons b bs)
have "map (app_pair spp_of) (insort_wrt pair_ord (Inl (p, b)) (new_spairs bs p)) =
insort_wrt pair_ord_spp (app_pair spp_of (Inl (p, b))) (map (app_pair spp_of) (new_spairs bs p))"
by (rule map_insort_wrt, rule pair_ord_alt_spp[symmetric])
thus ?case by (simp add: is_regular_spair_alt_spp Cons)
qed
lemma add_spairs_alt_spp:
assumes "\<And>x. x \<in> set bs \<Longrightarrow> Inl (spp_of p, spp_of x) \<notin> app_pair spp_of ` set ps"
shows "map (app_pair spp_of) (add_spairs ps bs p) =
add_spairs_spp (map (app_pair spp_of) ps) (map spp_of bs) (spp_of p)"
proof -
have "map (app_pair spp_of) (merge_wrt pair_ord (new_spairs bs p) ps) =
merge_wrt pair_ord_spp (map (app_pair spp_of) (new_spairs bs p)) (map (app_pair spp_of) ps)"
proof (rule map_merge_wrt, rule ccontr)
assume "app_pair spp_of ` set (new_spairs bs p) \<inter> app_pair spp_of ` set ps \<noteq> {}"
then obtain q' where "q' \<in> app_pair spp_of ` set (new_spairs bs p)"
and q'_in: "q' \<in> app_pair spp_of ` set ps" by blast
from this(1) obtain q where "q \<in> set (new_spairs bs p)" and q': "q' = app_pair spp_of q" ..
from this(1) obtain x where x_in: "x \<in> set bs" and q: "q = Inl (p, x)"
by (rule in_new_spairsE)
have q': "q' = Inl (spp_of p, spp_of x)" by (simp add: q q')
have "q' \<notin> app_pair spp_of ` set ps" unfolding q' using x_in by (rule assms)
thus False using q'_in ..
qed (simp only: pair_ord_alt_spp)
thus ?thesis by (simp add: add_spairs_def add_spairs_spp_def new_spairs_alt_spp)
qed
lemma rb_aux_invD_app_args:
assumes "rb_aux_inv (fst (app_args vec_of ((bs, ss, ps), z)))"
shows "list_all spp_inv bs" and "list_all spp_inv_pair ps"
proof -
from assms(1) have inv: "rb_aux_inv (map vec_of bs, ss, map (app_pair vec_of) ps)" by simp
hence "rb_aux_inv1 (map vec_of bs)" by (rule rb_aux_inv_D1)
hence "0 \<notin> rep_list ` set (map vec_of bs)" by (rule rb_aux_inv1_D2)
hence "0 \<notin> vec_of ` set bs" using rep_list_zero by fastforce
hence 1: "b \<in> set bs \<Longrightarrow> spp_inv b" for b by (auto simp: spp_inv_alt)
thus "list_all spp_inv bs" by (simp add: list_all_def)
have 2: "x \<in> set bs" if "vec_of x \<in> set (map vec_of bs)" for x
proof -
from that have "vec_of x \<in> vec_of ` set bs" by simp
then obtain y where "y \<in> set bs" and eq: "vec_of x = vec_of y" ..
from this(1) have "spp_inv y" by (rule 1)
moreover have "vec_of y = vec_of x" by (simp only: eq)
ultimately have "y = x" by (rule vec_of_inj)
with \<open>y \<in> set bs\<close> show ?thesis by simp
qed
show "list_all spp_inv_pair ps" unfolding list_all_def
proof (rule ballI)
fix p
assume "p \<in> set ps"
show "spp_inv_pair p"
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
from \<open>p \<in> set ps\<close> have "Inl (a, b) \<in> set ps" by (simp only: p)
hence "app_pair vec_of (Inl (a, b)) \<in> app_pair vec_of ` set ps" by (rule imageI)
hence "Inl (vec_of a, vec_of b) \<in> set (map (app_pair vec_of) ps)" by simp
with inv have "vec_of a \<in> set (map vec_of bs)" and "vec_of b \<in> set (map vec_of bs)"
by (rule rb_aux_inv_D3)+
have "spp_inv a" by (rule 1, rule 2, fact)
moreover have "spp_inv b" by (rule 1, rule 2, fact)
ultimately show ?thesis by (simp add: p)
qed simp
qed
qed
lemma app_args_spp_of_vec_of:
assumes "rb_aux_inv (fst (app_args vec_of args))"
shows "app_args spp_of (app_args vec_of args) = args"
proof -
obtain bs ss ps z where args: "args = ((bs, ss, ps), z)" using prod.exhaust by metis
from assms have "list_all spp_inv bs" and *: "list_all spp_inv_pair ps" unfolding args
by (rule rb_aux_invD_app_args)+
from this(1) have "map (spp_of \<circ> vec_of) bs = bs" by (rule map_spp_of_vec_of)
moreover from * have "map (app_pair spp_of \<circ> app_pair vec_of) ps = ps"
by (rule map_app_pair_spp_of_vec_of)
ultimately show ?thesis by (simp add: args)
qed
lemma poly_of_pair_alt_spp:
assumes "distinct fs" and "rb_aux_inv (bs, ss, p # ps)"
shows "spp_of (poly_of_pair p) = spp_of_pair (app_pair spp_of p)"
proof -
show ?thesis
proof (rule sum_prodE)
fix a b
assume p: "p = Inl (a, b)"
hence "Inl (a, b) \<in> set (p # ps)" by simp
with assms(2) have "is_regular_spair a b" by (rule rb_aux_inv_D3)
thus ?thesis by (simp add: p spair_alt_spp)
next
fix j
assume p: "p = Inr j"
hence "Inr j \<in> set (p # ps)" by simp
with assms(2) have "j < length fs" by (rule rb_aux_inv_D4)
thus ?thesis by (simp add: p spp_of_def lt_monomial rep_list_monomial[OF assms(1)] term_simps)
qed
qed
lemma rb_aux_alt_spp:
assumes "rb_aux_inv (fst args)"
shows "app_args spp_of (rb_aux args) = rb_spp_aux (app_args spp_of args)"
proof -
from assms have "rb_aux_dom args" by (rule rb_aux_domI)
thus ?thesis using assms
proof (induct args rule: rb_aux.pinduct)
case (1 bs ss z)
show ?case by (simp add: rb_aux.psimps(1)[OF 1(1)] rb_spp_aux_Nil)
next
case (2 bs ss p ps z)
let ?q = "sig_trd bs (poly_of_pair p)"
from 2(5) have *: "rb_aux_inv (bs, ss, p # ps)" by (simp only: fst_conv)
hence "rb_aux_inv1 bs" by (rule rb_aux_inv_D1)
hence "0 \<notin> rep_list ` set bs" by (rule rb_aux_inv1_D2)
hence "0 \<notin> set bs" by (force simp: rep_list_zero)
hence eq1: "sig_crit_spp (map spp_of bs) ss' (app_pair spp_of p) \<longleftrightarrow> sig_crit bs ss' p" for ss'
by (simp add: sig_crit_alt_spp)
from fs_distinct * have eq2: "sig_trd_spp (map spp_of bs) (spp_of_pair (app_pair spp_of p)) = spp_of ?q"
by (simp only: sig_trd_alt_spp poly_of_pair_alt_spp)
show ?case
proof (simp add: rb_aux.psimps(2)[OF 2(1)] Let_def, intro conjI impI)
note refl
moreover assume a: "sig_crit bs (new_syz_sigs ss bs p) p"
moreover from * this have "rb_aux_inv (fst ((bs, new_syz_sigs ss bs p, ps), z))"
unfolding fst_conv by (rule rb_aux_inv_preserved_1)
ultimately have "app_args spp_of (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
rb_spp_aux (app_args spp_of ((bs, new_syz_sigs ss bs p, ps), z))"
by (rule 2(2))
also have "... = rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)"
by (simp add: rb_spp_aux_Cons eq1 a new_syz_sigs_alt_spp[symmetric])
finally show "app_args spp_of (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
thus "app_args spp_of (rb_aux ((bs, new_syz_sigs ss bs p, ps), z)) =
rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
next
assume a: "\<not> sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list ?q = 0"
from * b have "rb_aux_inv (fst ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z))"
unfolding fst_conv by (rule rb_aux_inv_preserved_2)
with refl a refl b have "app_args spp_of (rb_aux ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z)) =
rb_spp_aux (app_args spp_of ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z))"
by (rule 2(3))
also have "... = rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)"
by (simp add: rb_spp_aux_Cons eq1 a Let_def eq2 snd_spp_of b fst_spp_of new_syz_sigs_alt_spp[symmetric])
finally show "app_args spp_of (rb_aux ((bs, lt ?q # new_syz_sigs ss bs p, ps), Suc z)) =
rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
next
assume a: "\<not> sig_crit bs (new_syz_sigs ss bs p) p" and b: "rep_list ?q \<noteq> 0"
have "Inl (spp_of ?q, spp_of x) \<notin> app_pair spp_of ` set ps" for x
proof
assume "Inl (spp_of ?q, spp_of x) \<in> app_pair spp_of ` set ps"
then obtain y where "y \<in> set ps" and eq0: "Inl (spp_of ?q, spp_of x) = app_pair spp_of y" ..
obtain a b where y: "y = Inl (a, b)" and "spp_of ?q = spp_of a"
proof (rule sum_prodE)
fix a b
assume "y = Inl (a, b)"
moreover from eq0 have "spp_of ?q = spp_of a" by (simp add: \<open>y = Inl (a, b)\<close>)
ultimately show ?thesis ..
next
fix j
assume "y = Inr j"
with eq0 show ?thesis by simp
qed
from this(2) have "lt ?q = lt a" by (simp add: spp_of_def)
from \<open>y \<in> set ps\<close> have "y \<in> set (p # ps)" by simp
with * have "a \<in> set bs" unfolding y by (rule rb_aux_inv_D3(1))
hence "lt ?q \<in> lt ` set bs" unfolding \<open>lt ?q = lt a\<close> by (rule imageI)
moreover from * a b have "lt ?q \<notin> lt ` set bs" by (rule rb_aux_inv_preserved_3)
ultimately show False by simp
qed
hence eq3: "add_spairs_spp (map (app_pair spp_of) ps) (map spp_of bs) (spp_of ?q) =
map (app_pair spp_of) (add_spairs ps bs ?q)" by (simp add: add_spairs_alt_spp)
from * a b have "rb_aux_inv (fst ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z))"
unfolding fst_conv by (rule rb_aux_inv_preserved_3)
with refl a refl b
have "app_args spp_of (rb_aux ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z)) =
rb_spp_aux (app_args spp_of ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z))"
by (rule 2(4))
also have "... = rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)"
by (simp add: rb_spp_aux_Cons eq1 a Let_def eq2 fst_spp_of snd_spp_of b eq3 new_syz_sigs_alt_spp[symmetric])
finally show "app_args spp_of (rb_aux ((?q # bs, new_syz_sigs ss bs p, add_spairs ps bs ?q), z)) =
rb_spp_aux ((map spp_of bs, ss, app_pair spp_of p # map (app_pair spp_of) ps), z)" .
qed
qed
qed
corollary rb_spp_aux_alt:
"rb_aux_inv (fst (app_args vec_of args)) \<Longrightarrow>
rb_spp_aux args = app_args spp_of (rb_aux (app_args vec_of args))"
by (simp only: rb_aux_alt_spp app_args_spp_of_vec_of)
corollary rb_spp_aux:
"hom_grading dgrad \<Longrightarrow>
punit.is_Groebner_basis (set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))))"
(is "_ \<Longrightarrow> ?thesis1")
"ideal (set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))) = ideal (set fs)"
(is "?thesis2")
"set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z))))) \<subseteq> punit_dgrad_max_set dgrad"
(is "?thesis3")
"0 \<notin> set (map snd (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))))"
(is "?thesis4")
"hom_grading dgrad \<Longrightarrow> is_pot_ord \<Longrightarrow> is_regular_sequence fs \<Longrightarrow>
snd (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)) = z"
(is "_ \<Longrightarrow> _ \<Longrightarrow> _ \<Longrightarrow> ?thesis5")
"rword_strict = rw_rat_strict \<Longrightarrow> p \<in> set (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))) \<Longrightarrow>
q \<in> set (fst (fst (rb_spp_aux (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)))) \<Longrightarrow> p \<noteq> q \<Longrightarrow>
punit.lt (snd p) adds punit.lt (snd q) \<Longrightarrow> punit.lt (snd p) \<oplus> fst q \<prec>\<^sub>t punit.lt (snd q) \<oplus> fst p"
proof -
let ?args = "(([], Koszul_syz_sigs fs, map Inr [0..<length fs]), z)"
have eq0: "app_pair vec_of \<circ> Inr = Inr" by (intro ext, simp)
have eq1: "fst (fst (app_args spp_of a)) = map spp_of (fst (fst a))" for a::"(_ \<times> ('t list) \<times> _) \<times> _"
proof -
obtain bs ss ps z where "a = ((bs, ss, ps), z)" using prod.exhaust by metis
thus ?thesis by simp
qed
have eq2: "snd \<circ> spp_of = rep_list" by (intro ext, simp add: snd_spp_of)
have "rb_aux_inv (fst (app_args vec_of ?args))" by (simp add: eq0 rb_aux_inv_init)
hence eq3: "rb_spp_aux ?args = app_args spp_of (rb_aux (app_args vec_of ?args))"
by (rule rb_spp_aux_alt)
{
assume "hom_grading dgrad"
with rb_aux_is_Groebner_basis show ?thesis1 by (simp add: eq0 eq1 eq2 eq3 del: set_map)
}
from ideal_rb_aux show ?thesis2 by (simp add: eq0 eq1 eq2 eq3 del: set_map)
from dgrad_max_set_closed_rb_aux show ?thesis3 by (simp add: eq0 eq1 eq2 eq3 del: set_map)
from rb_aux_nonzero show ?thesis4 by (simp add: eq0 eq1 eq2 eq3 del: set_map)
{
assume "is_pot_ord" and "hom_grading dgrad" and "is_regular_sequence fs"
hence "snd (rb_aux ?args) = z" by (rule rb_aux_no_zero_red)
thus ?thesis5 by (simp add: snd_app_args eq0 eq3)
}
{
from rb_aux_nonzero have "0 \<notin> rep_list ` set (fst (fst (rb_aux ?args)))"
(is "0 \<notin> rep_list ` ?G") by simp
assume "rword_strict = rw_rat_strict"
hence "is_min_sig_GB dgrad ?G" by (rule rb_aux_is_min_sig_GB)
hence rl: "\<And>g. g \<in> ?G \<Longrightarrow> \<not> is_sig_red (\<preceq>\<^sub>t) (=) (?G - {g}) g" by (simp add: is_min_sig_GB_def)
assume "p \<in> set (fst (fst (rb_spp_aux ?args)))"
also have "... = spp_of ` ?G" by (simp add: eq0 eq1 eq3)
finally obtain p' where "p' \<in> ?G" and p: "p = spp_of p'" ..
assume "q \<in> set (fst (fst (rb_spp_aux ?args)))"
also have "... = spp_of ` ?G" by (simp add: eq0 eq1 eq3)
finally obtain q' where "q' \<in> ?G" and q: "q = spp_of q'" ..
from this(1) have 1: "\<not> is_sig_red (\<preceq>\<^sub>t) (=) (?G - {q'}) q'" by (rule rl)
assume "p \<noteq> q" and "punit.lt (snd p) adds punit.lt (snd q)"
hence "p' \<noteq> q'" and adds: "punit.lt (rep_list p') adds punit.lt (rep_list q')"
by (auto simp: p q snd_spp_of)
show "punit.lt (snd p) \<oplus> fst q \<prec>\<^sub>t punit.lt (snd q) \<oplus> fst p"
proof (rule ccontr)
assume "\<not> punit.lt (snd p) \<oplus> fst q \<prec>\<^sub>t punit.lt (snd q) \<oplus> fst p"
hence le: "punit.lt (rep_list q') \<oplus> lt p' \<preceq>\<^sub>t punit.lt (rep_list p') \<oplus> lt q'"
by (simp add: p q spp_of_def)
from \<open>p' \<noteq> q'\<close> \<open>p' \<in> ?G\<close> have "p' \<in> ?G - {q'}" by simp
moreover from \<open>p' \<in> ?G\<close> \<open>0 \<notin> rep_list ` ?G\<close> have "rep_list p' \<noteq> 0" by fastforce
moreover from \<open>q' \<in> ?G\<close> \<open>0 \<notin> rep_list ` ?G\<close> have "rep_list q' \<noteq> 0" by fastforce
moreover note adds
moreover have "ord_term_lin.is_le_rel (\<preceq>\<^sub>t)" by simp
ultimately have "is_sig_red (\<preceq>\<^sub>t) (=) (?G - {q'}) q'" using le by (rule is_sig_red_top_addsI)
with 1 show False ..
qed
}
qed
end
end
end
end
end
definition gb_sig_z ::
"(('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> (('t \<times> ('a \<Rightarrow>\<^sub>0 'b::field)) list \<times> nat)"
where "gb_sig_z rword_strict fs0 =
(let fs = rev (remdups (rev (removeAll 0 fs0)));
res = rb_spp_aux fs rword_strict (([], Koszul_syz_sigs fs, map Inr [0..<length fs]), 0) in
(fst (fst res), snd res))"
text \<open>The second return value of @{const gb_sig_z} is the total number of zero-reductions.\<close>
definition gb_sig :: "(('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> ('t \<times> ('a \<Rightarrow>\<^sub>0 'b)) \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) list \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b::field) list"
where "gb_sig rword_strict fs0 = map snd (fst (gb_sig_z rword_strict fs0))"
theorem
assumes "\<And>fs. is_strict_rewrite_ord fs rword_strict"
shows gb_sig_isGB: "punit.is_Groebner_basis (set (gb_sig rword_strict fs))" (is ?thesis1)
and gb_sig_ideal: "ideal (set (gb_sig rword_strict fs)) = ideal (set fs)" (is ?thesis2)
and dgrad_p_set_closed_gb_sig:
"dickson_grading d \<Longrightarrow> set fs \<subseteq> punit.dgrad_p_set d m \<Longrightarrow> set (gb_sig rword_strict fs) \<subseteq> punit.dgrad_p_set d m"
(is "_ \<Longrightarrow> _ \<Longrightarrow> ?thesis3")
and gb_sig_nonzero: "0 \<notin> set (gb_sig rword_strict fs)" (is ?thesis4)
and gb_sig_no_zero_red: "is_pot_ord \<Longrightarrow> is_regular_sequence fs \<Longrightarrow> snd (gb_sig_z rword_strict fs) = 0"
proof -
from ex_hgrad obtain d0::"'a \<Rightarrow> nat" where "dickson_grading d0 \<and> hom_grading d0" ..
hence dg: "dickson_grading d0" and hg: "hom_grading d0" by simp_all
define fs1 where "fs1 = rev (remdups (rev (removeAll 0 fs)))"
note assms dg
moreover have "distinct fs1" and "0 \<notin> set fs1" by (simp_all add: fs1_def)
ultimately have "ideal (set (gb_sig rword_strict fs)) = ideal (set fs1)" and ?thesis4
unfolding gb_sig_def gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)+
thus ?thesis2 and ?thesis4 by (simp_all add: fs1_def ideal.span_Diff_zero)
from assms dg \<open>distinct fs1\<close> \<open>0 \<notin> set fs1\<close> hg show ?thesis1
unfolding gb_sig_def gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)
{
assume dg: "dickson_grading d" and *: "set fs \<subseteq> punit.dgrad_p_set d m"
show ?thesis3
proof (cases "set fs \<subseteq> {0}")
case True
hence "removeAll 0 fs = []"
by (metis (no_types, lifting) Diff_iff ex_in_conv set_empty2 set_removeAll subset_singleton_iff)
thus ?thesis by (simp add: gb_sig_def gb_sig_z_def Let_def rb_spp_aux_Nil)
next
case False
have "set fs1 \<subseteq> set fs" by (fastforce simp: fs1_def)
hence "Keys (set fs1) \<subseteq> Keys (set fs)" by (rule Keys_mono)
hence "d ` Keys (set fs1) \<subseteq> d ` Keys (set fs)" by (rule image_mono)
hence "insert (d 0) (d ` Keys (set fs1)) \<subseteq> insert (d 0) (d ` Keys (set fs))" by (rule insert_mono)
moreover have "insert (d 0) (d ` Keys (set fs1)) \<noteq> {}" by simp
moreover have "finite (insert (d 0) (d ` Keys (set fs)))"
by (simp add: finite_Keys)
ultimately have le: "Max (insert (d 0) (d ` Keys (set fs1))) \<le>
Max (insert (d 0) (d ` Keys (set fs)))" by (rule Max_mono)
from assms dg have "set (gb_sig rword_strict fs) \<subseteq> punit_dgrad_max_set (TYPE('b)) fs1 d"
using \<open>distinct fs1\<close> \<open>0 \<notin> set fs1\<close>
unfolding gb_sig_def gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)
also have "punit_dgrad_max_set (TYPE('b)) fs1 d \<subseteq> punit_dgrad_max_set (TYPE('b)) fs d"
by (rule punit.dgrad_p_set_subset, simp add: dgrad_max_def le)
also from dg * False have "... \<subseteq> punit.dgrad_p_set d m"
by (rule punit_dgrad_max_set_subset_dgrad_p_set)
finally show ?thesis .
qed
}
{
assume "is_regular_sequence fs"
note assms dg \<open>distinct fs1\<close> \<open>0 \<notin> set fs1\<close> hg
moreover assume "is_pot_ord"
moreover from \<open>is_regular_sequence fs\<close> have "is_regular_sequence fs1" unfolding fs1_def
by (intro is_regular_sequence_remdups is_regular_sequence_removeAll_zero)
ultimately show "snd (gb_sig_z rword_strict fs) = 0"
unfolding gb_sig_def gb_sig_z_def snd_conv fs1_def Let_def by (rule rb_spp_aux)
}
qed
theorem gb_sig_z_is_min_sig_GB:
assumes "p \<in> set (fst (gb_sig_z rw_rat_strict fs))" and "q \<in> set (fst (gb_sig_z rw_rat_strict fs))"
and "p \<noteq> q" and "punit.lt (snd p) adds punit.lt (snd q)"
shows "punit.lt (snd p) \<oplus> fst q \<prec>\<^sub>t punit.lt (snd q) \<oplus> fst p"
proof -
define fs1 where "fs1 = rev (remdups (rev (removeAll 0 fs)))"
from ex_hgrad obtain d0::"'a \<Rightarrow> nat" where "dickson_grading d0 \<and> hom_grading d0" ..
hence "dickson_grading d0" ..
note rw_rat_strict_is_strict_rewrite_ord this
moreover have "distinct fs1" and "0 \<notin> set fs1" by (simp_all add: fs1_def)
moreover note refl assms
ultimately show ?thesis unfolding gb_sig_z_def fst_conv fs1_def Let_def by (rule rb_spp_aux)
qed
text \<open>Summarizing, these are the four main results proved in this theory:
\<^item> @{thm gb_sig_isGB},
\<^item> @{thm gb_sig_ideal},
\<^item> @{thm gb_sig_no_zero_red}, and
\<^item> @{thm gb_sig_z_is_min_sig_GB}.\<close>
end (* qpm_inf_term *)
end (* theory *)
diff --git a/thys/Simplex/Simplex_Algebra.thy b/thys/Simplex/Simplex_Algebra.thy
--- a/thys/Simplex/Simplex_Algebra.thy
+++ b/thys/Simplex/Simplex_Algebra.thy
@@ -1,282 +1,282 @@
(* Authors: F. Maric, M. Spasic *)
section \<open>Linearly Ordered Rational Vectors\<close>
theory Simplex_Algebra
imports
HOL.Rat
HOL.Real_Vector_Spaces
begin
class scaleRat =
fixes scaleRat :: "rat \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "*R" 75)
begin
abbreviation
divideRat :: "'a \<Rightarrow> rat \<Rightarrow> 'a" (infixl "'/R" 70)
where
"x /R r == scaleRat (inverse r) x"
end
class rational_vector = scaleRat + ab_group_add +
assumes scaleRat_right_distrib: "scaleRat a (x + y) = scaleRat a x + scaleRat a y"
and scaleRat_left_distrib: "scaleRat (a + b) x = scaleRat a x + scaleRat b x"
and scaleRat_scaleRat: "scaleRat a (scaleRat b x) = scaleRat (a * b) x"
and scaleRat_one: "scaleRat 1 x = x"
interpretation rational_vector:
vector_space "scaleRat :: rat \<Rightarrow> 'a \<Rightarrow> 'a::rational_vector"
by (unfold_locales) (simp_all add: scaleRat_right_distrib scaleRat_left_distrib scaleRat_scaleRat scaleRat_one)
class ordered_rational_vector = rational_vector + order
class linordered_rational_vector = ordered_rational_vector + linorder +
assumes plus_less: "(a::'a) < b \<Longrightarrow> a + c < b + c" and
scaleRat_less1: "\<lbrakk>(a::'a) < b; k > 0\<rbrakk> \<Longrightarrow> (k *R a) < (k *R b)" and
scaleRat_less2: "\<lbrakk>(a::'a) < b; k < 0\<rbrakk> \<Longrightarrow> (k *R a) > (k *R b)"
begin
lemma scaleRat_leq1: "\<lbrakk> a \<le> b; k > 0\<rbrakk> \<Longrightarrow> k *R a \<le> k *R b"
unfolding le_less
using scaleRat_less1[of a b k]
by auto
lemma scaleRat_leq2: "\<lbrakk> a \<le> b; k < 0\<rbrakk> \<Longrightarrow> k *R a \<ge> k *R b"
unfolding le_less
using scaleRat_less2[of a b k]
by auto
lemma zero_scaleRat
[simp]: "0 *R v = zero"
using scaleRat_left_distrib[of 0 0 v]
by auto
lemma scaleRat_zero
[simp]: "a *R (0::'a) = 0"
using scaleRat_right_distrib[of a 0 0]
by auto
lemma scaleRat_uminus [simp]:
"-1 *R x = - (x :: 'a)"
proof-
have "0 = -1 *R x + x"
using scaleRat_left_distrib[of "-1" 1 x]
by (simp add: scaleRat_one)
have "-x = 0 - x"
by simp
then have "-x = -1 *R x + x - x"
using \<open>0 = -1 *R x + x\<close>
by simp
then show ?thesis
by (simp add: add_assoc)
qed
lemma minus_lt: "(a::'a) < b \<longleftrightarrow> a - b < 0"
using plus_less[of a b "-b"]
using plus_less[of "a - b" 0 b]
by (auto simp add: add_assoc)
lemma minus_gt: "(a::'a) < b \<longleftrightarrow> 0 < b - a"
using plus_less[of a b "-a"]
using plus_less[of 0 "b-a" a]
by (auto simp add: add_assoc)
lemma minus_leq:
"(a::'a) \<le> b \<longleftrightarrow> a - b \<le> 0"
proof-
have *: "a \<le> b \<Longrightarrow> a - b \<le> (0 :: 'a)"
using minus_gt[of a b]
using scaleRat_less2[of 0 "b-a" "-1"]
by (auto simp add: not_less_iff_gr_or_eq)
have **: "a - b \<le> 0 \<Longrightarrow> a \<le> b"
proof-
assume "a - b \<le> 0"
show ?thesis
proof(cases "a - b < 0")
case True
then show ?thesis
using plus_less[of "a - b" 0 b]
by (simp add: add_assoc )
next
case False
then show ?thesis
using \<open>a - b \<le> 0\<close>
by (simp add:antisym_conv1)
qed
qed
show ?thesis
using * **
by auto
qed
lemma minus_geq: "(a::'a) \<le> b \<longleftrightarrow> 0 \<le> b - a"
proof-
have *: "a \<le> b \<Longrightarrow> 0 \<le> b - a"
using minus_gt[of a b]
by (auto simp add: not_less_iff_gr_or_eq)
have **: "0 \<le> b - a \<Longrightarrow> a \<le> b"
proof-
assume "0 \<le> b - a"
show ?thesis
proof(cases "0 < b - a")
case True
then show ?thesis
using plus_less[of 0 "b - a" a]
by (simp add: add_assoc)
next
case False
then show ?thesis
using \<open>0 \<le> b - a\<close>
- using eq_iff[of "b - a" 0]
+ using order.eq_iff[of "b - a" 0]
by auto
qed
qed
show ?thesis
using * **
by auto
qed
lemma divide_lt:
"\<lbrakk>c *R (a::'a) < b; (c::rat) > 0 \<rbrakk> \<Longrightarrow> a < (1/c) *R b"
using scaleRat_less1[of "c *R a" b "1/c"]
by (simp add: scaleRat_one scaleRat_scaleRat)
lemma divide_gt:
"\<lbrakk>c *R (a::'a) > b;(c::rat) > 0\<rbrakk> \<Longrightarrow> a > (1/c) *R b"
using scaleRat_less1[of b "c *R a" "1/c"]
by (simp add: scaleRat_one scaleRat_scaleRat)
lemma divide_leq:
"\<lbrakk>c *R (a::'a) \<le> b; (c::rat) > 0\<rbrakk> \<Longrightarrow> a \<le> (1/c) *R b"
proof(cases "c *R a < b")
assume "c > 0"
case True
then show ?thesis
using divide_lt[of c a b]
using \<open>c > 0\<close>
by simp
next
assume "c *R a \<le> b" "c > 0"
case False
then have *: "c *R a = b"
using \<open>c *R a \<le> b\<close>
by simp
then show ?thesis
using \<open>c > 0\<close>
by (auto simp add: scaleRat_one scaleRat_scaleRat)
qed
lemma divide_geq:
"\<lbrakk>c *R (a::'a) \<ge> b; (c::rat) > 0\<rbrakk> \<Longrightarrow> a \<ge> (1/c) *R b"
proof(cases "c *R a > b")
assume "c > 0"
case True
then show ?thesis
using divide_gt[of b c a]
using \<open>c > 0\<close>
by simp
next
assume "c *R a \<ge> b" "c > 0"
case False
then have *: "c *R a = b"
using \<open>c *R a \<ge> b\<close>
by simp
then show ?thesis
using \<open>c > 0\<close>
by (auto simp add: scaleRat_one scaleRat_scaleRat)
qed
lemma divide_lt1:
"\<lbrakk>c *R (a::'a) < b; (c::rat) < 0\<rbrakk> \<Longrightarrow> a > (1/c) *R b"
using scaleRat_less2[of "c *R a" b "1/c"]
by (simp add: scaleRat_scaleRat scaleRat_one)
lemma divide_gt1:
"\<lbrakk>c *R (a::'a) > b; (c::rat) < 0 \<rbrakk> \<Longrightarrow> a < (1/c) *R b"
using scaleRat_less2[of b "c *R a" "1/c"]
by (simp add: scaleRat_scaleRat scaleRat_one)
lemma divide_leq1:
"\<lbrakk>c *R (a::'a) \<le> b; (c::rat) < 0\<rbrakk> \<Longrightarrow> a \<ge> (1/c) *R b"
proof(cases "c *R a < b")
assume "c < 0"
case True
then show ?thesis
using divide_lt1[of c a b]
using \<open>c < 0\<close>
by simp
next
assume "c *R a \<le> b" "c < 0"
case False
then have *: "c *R a = b"
using \<open>c *R a \<le> b\<close>
by simp
then show ?thesis
using \<open>c < 0\<close>
by (auto simp add: scaleRat_one scaleRat_scaleRat)
qed
lemma divide_geq1:
"\<lbrakk>c *R (a::'a) \<ge> b; (c::rat) < 0\<rbrakk> \<Longrightarrow> a \<le> (1/c) *R b"
proof(cases "c *R a > b")
assume "c < 0"
case True
then show ?thesis
using divide_gt1[of b c a]
using \<open>c < 0\<close>
by simp
next
assume "c *R a \<ge> b" "c < 0"
case False
then have *: "c *R a = b"
using \<open>c *R a \<ge> b\<close>
by simp
then show ?thesis
using \<open>c < 0\<close>
by (auto simp add: scaleRat_one scaleRat_scaleRat)
qed
end
class lrv = linordered_rational_vector + one +
assumes zero_neq_one: "0 \<noteq> 1"
subclass (in linordered_rational_vector) ordered_ab_semigroup_add
proof
fix a b c
assume "a \<le> b"
then show "c + a \<le> c + b"
using plus_less[of a b c]
by (auto simp add: add_ac le_less)
qed
instantiation rat :: rational_vector
begin
definition scaleRat_rat :: "rat \<Rightarrow> rat \<Rightarrow> rat" where
[simp]: "x *R y = x * y"
instance by standard (auto simp add: field_simps)
end
instantiation rat :: ordered_rational_vector
begin
instance ..
end
instantiation rat :: linordered_rational_vector
begin
instance by standard (auto simp add: field_simps)
end
instantiation rat :: lrv
begin
instance by standard (auto simp add: field_simps)
end
lemma uminus_less_lrv[simp]: fixes a b :: "'a :: lrv"
shows "- a < - b \<longleftrightarrow> b < a"
proof -
have "(- a < - b) = (-1 *R a < -1 *R b)" by simp
also have "\<dots> \<longleftrightarrow> (b < a)"
using scaleRat_less2[of _ _ "-1"] scaleRat_less2[of "-1 *R a" "-1 *R b" "-1"] by auto
finally show ?thesis .
qed
end
diff --git a/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy b/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy
--- a/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy
+++ b/thys/Sqrt_Babylonian/Sqrt_Babylonian_Auxiliary.thy
@@ -1,136 +1,137 @@
(* Title: Computing Square Roots using the Babylonian Method
Author: René Thiemann <rene.thiemann@uibk.ac.at>
Maintainer: René Thiemann
License: LGPL
*)
(*
Copyright 2009-2014 René Thiemann
This file is part of IsaFoR/CeTA.
IsaFoR/CeTA is free software: you can redistribute it and/or modify it under the
terms of the GNU Lesser General Public License as published by the Free Software
Foundation, either version 3 of the License, or (at your option) any later
version.
IsaFoR/CeTA is distributed in the hope that it will be useful, but WITHOUT ANY
WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A
PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details.
You should have received a copy of the GNU Lesser General Public License along
with IsaFoR/CeTA. If not, see <http://www.gnu.org/licenses/>.
*)
section \<open>Auxiliary lemmas which might be moved into the Isabelle distribution.\<close>
theory Sqrt_Babylonian_Auxiliary
imports
Complex_Main
begin
lemma mod_div_equality_int: "(n :: int) div x * x = n - n mod x"
using div_mult_mod_eq[of n x] by arith
lemma div_is_floor_divide_rat: "n div y = \<lfloor>rat_of_int n / rat_of_int y\<rfloor>"
unfolding Fract_of_int_quotient[symmetric] floor_Fract by simp
lemma div_is_floor_divide_real: "n div y = \<lfloor>real_of_int n / of_int y\<rfloor>"
unfolding div_is_floor_divide_rat[of n y]
by (metis Ratreal_def of_rat_divide of_rat_of_int_eq real_floor_code)
lemma floor_div_pos_int:
fixes r :: "'a :: floor_ceiling"
assumes n: "n > 0"
shows "\<lfloor>r / of_int n\<rfloor> = \<lfloor>r\<rfloor> div n" (is "?l = ?r")
proof -
let ?of_int = "of_int :: int \<Rightarrow> 'a"
define rhs where "rhs = \<lfloor>r\<rfloor> div n"
let ?n = "?of_int n"
define m where "m = \<lfloor>r\<rfloor> mod n"
let ?m = "?of_int m"
from div_mult_mod_eq[of "floor r" n] have dm: "rhs * n + m = \<lfloor>r\<rfloor>" unfolding rhs_def m_def by simp
have mn: "m < n" and m0: "m \<ge> 0" using n m_def by auto
define e where "e = r - ?of_int \<lfloor>r\<rfloor>"
have e0: "e \<ge> 0" unfolding e_def
by (metis diff_self eq_iff floor_diff_of_int zero_le_floor)
have e1: "e < 1" unfolding e_def
by (metis diff_self dual_order.refl floor_diff_of_int floor_le_zero)
have "r = ?of_int \<lfloor>r\<rfloor> + e" unfolding e_def by simp
also have "\<lfloor>r\<rfloor> = rhs * n + m" using dm by simp
finally have "r = ?of_int (rhs * n + m) + e" .
hence "r / ?n = ?of_int (rhs * n) / ?n + (e + ?m) / ?n" using n by (simp add: field_simps)
also have "?of_int (rhs * n) / ?n = ?of_int rhs" using n by auto
finally have *: "r / ?of_int n = (e + ?of_int m) / ?of_int n + ?of_int rhs" by simp
have "?l = rhs + floor ((e + ?m) / ?n)" unfolding * by simp
also have "floor ((e + ?m) / ?n) = 0"
proof (rule floor_unique)
show "?of_int 0 \<le> (e + ?m) / ?n" using e0 m0 n
by (metis add_increasing2 divide_nonneg_pos of_int_0 of_int_0_le_iff of_int_0_less_iff)
show "(e + ?m) / ?n < ?of_int 0 + 1"
proof (rule ccontr)
from n have n': "?n > 0" "?n \<ge> 0" by simp_all
assume "\<not> ?thesis"
hence "(e + ?m) / ?n \<ge> 1" by auto
from mult_right_mono[OF this n'(2)]
have "?n \<le> e + ?m" using n'(1) by simp
also have "?m \<le> ?n - 1" using mn
by (metis of_int_1 of_int_diff of_int_le_iff zle_diff1_eq)
finally have "?n \<le> e + ?n - 1" by auto
with e1 show False by arith
qed
qed
finally show ?thesis unfolding rhs_def by simp
qed
lemma floor_div_neg_int:
fixes r :: "'a :: floor_ceiling"
assumes n: "n < 0"
shows "\<lfloor>r / of_int n\<rfloor> = \<lceil>r\<rceil> div n"
proof -
from n have n': "- n > 0" by auto
have "\<lfloor>r / of_int n\<rfloor> = \<lfloor> - r / of_int (- n)\<rfloor>" using n
by (metis floor_of_int floor_zero less_int_code(1) minus_divide_left minus_minus nonzero_minus_divide_right of_int_minus)
also have "\<dots> = \<lfloor> - r \<rfloor> div (- n)" by (rule floor_div_pos_int[OF n'])
also have "\<dots> = \<lceil> r \<rceil> div n" using n
by (metis ceiling_def div_minus_right)
finally show ?thesis .
qed
lemma divide_less_floor1: "n / y < of_int (floor (n / y)) + 1"
by (metis floor_less_iff less_add_one of_int_1 of_int_add)
context linordered_idom
begin
lemma sgn_int_pow_if [simp]:
"sgn x ^ p = (if even p then 1 else sgn x)" if "x \<noteq> 0"
using that by (induct p) simp_all
lemma compare_pow_le_iff: "p > 0 \<Longrightarrow> (x :: 'a) \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> (x ^ p \<le> y ^ p) = (x \<le> y)"
- by (metis eq_iff linear power_eq_imp_eq_base power_mono)
+ by (rule power_mono_iff)
lemma compare_pow_less_iff: "p > 0 \<Longrightarrow> (x :: 'a) \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> (x ^ p < y ^ p) = (x < y)"
- by (metis power_less_imp_less_base power_strict_mono)
+ using compare_pow_le_iff [of p x y]
+ using local.dual_order.order_iff_strict local.power_strict_mono by blast
end
lemma quotient_of_int[simp]: "quotient_of (of_int i) = (i,1)"
by (metis Rat.of_int_def quotient_of_int)
lemma quotient_of_nat[simp]: "quotient_of (of_nat i) = (int i,1)"
by (metis Rat.of_int_def Rat.quotient_of_int of_int_of_nat_eq)
lemma square_lesseq_square: "\<And> x y. 0 \<le> (x :: 'a :: linordered_field) \<Longrightarrow> 0 \<le> y \<Longrightarrow> (x * x \<le> y * y) = (x \<le> y)"
by (metis mult_mono mult_strict_mono' not_less)
lemma square_less_square: "\<And> x y. 0 \<le> (x :: 'a :: linordered_field) \<Longrightarrow> 0 \<le> y \<Longrightarrow> (x * x < y * y) = (x < y)"
by (metis mult_mono mult_strict_mono' not_less)
lemma sqrt_sqrt[simp]: "x \<ge> 0 \<Longrightarrow> sqrt x * sqrt x = x"
by (metis real_sqrt_pow2 power2_eq_square)
lemma abs_lesseq_square: "abs (x :: real) \<le> abs y \<longleftrightarrow> x * x \<le> y * y"
using square_lesseq_square[of "abs x" "abs y"] by auto
end
diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Messages.thy b/thys/Stateful_Protocol_Composition_and_Typing/Messages.thy
--- a/thys/Stateful_Protocol_Composition_and_Typing/Messages.thy
+++ b/thys/Stateful_Protocol_Composition_and_Typing/Messages.thy
@@ -1,538 +1,540 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2015-2020
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 holder 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.
*)
(* Title: Messages.thy
Author: Andreas Viktor Hess, DTU
*)
section \<open>Protocol Messages as (First-Order) Terms\<close>
theory Messages
imports Miscellaneous "First_Order_Terms.Term"
begin
subsection \<open>Term-related definitions: subterms and free variables\<close>
abbreviation "the_Fun \<equiv> un_Fun1"
lemmas the_Fun_def = un_Fun1_def
fun subterms::"('a,'b) term \<Rightarrow> ('a,'b) terms" where
"subterms (Var x) = {Var x}"
| "subterms (Fun f T) = {Fun f T} \<union> (\<Union>t \<in> set T. subterms t)"
abbreviation subtermeq (infix "\<sqsubseteq>" 50) where "t' \<sqsubseteq> t \<equiv> (t' \<in> subterms t)"
abbreviation subterm (infix "\<sqsubset>" 50) where "t' \<sqsubset> t \<equiv> (t' \<sqsubseteq> t \<and> t' \<noteq> t)"
abbreviation "subterms\<^sub>s\<^sub>e\<^sub>t M \<equiv> \<Union>(subterms ` M)"
abbreviation subtermeqset (infix "\<sqsubseteq>\<^sub>s\<^sub>e\<^sub>t" 50) where "t \<sqsubseteq>\<^sub>s\<^sub>e\<^sub>t M \<equiv> (t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M)"
abbreviation fv where "fv \<equiv> vars_term"
lemmas fv_simps = term.simps(17,18)
fun fv\<^sub>s\<^sub>e\<^sub>t where "fv\<^sub>s\<^sub>e\<^sub>t M = \<Union>(fv ` M)"
abbreviation fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r where "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r p \<equiv> case p of (t,t') \<Rightarrow> fv t \<union> fv t'"
fun fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s where "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = \<Union>(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set F)"
abbreviation ground where "ground M \<equiv> fv\<^sub>s\<^sub>e\<^sub>t M = {}"
subsection \<open>Variants that return lists insteads of sets\<close>
fun fv_list where
"fv_list (Var x) = [x]"
| "fv_list (Fun f T) = concat (map fv_list T)"
definition fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s where
"fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \<equiv> concat (map (\<lambda>(t,t'). fv_list t@fv_list t') F)"
fun subterms_list::"('a,'b) term \<Rightarrow> ('a,'b) term list" where
"subterms_list (Var x) = [Var x]"
| "subterms_list (Fun f T) = remdups (Fun f T#concat (map subterms_list T))"
lemma fv_list_is_fv: "fv t = set (fv_list t)"
by (induct t) auto
lemma fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = set (fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)"
by (induct F) (auto simp add: fv_list_is_fv fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_def)
lemma subterms_list_is_subterms: "subterms t = set (subterms_list t)"
by (induct t) auto
subsection \<open>The subterm relation defined as a function\<close>
fun subterm_of where
"subterm_of t (Var y) = (t = Var y)"
| "subterm_of t (Fun f T) = (t = Fun f T \<or> list_ex (subterm_of t) T)"
lemma subterm_of_iff_subtermeq[code_unfold]: "t \<sqsubseteq> t' = subterm_of t t'"
proof (induction t')
case (Fun f T) thus ?case
proof (cases "t = Fun f T")
case False thus ?thesis
using Fun.IH subterm_of.simps(2)[of t f T]
unfolding list_ex_iff by fastforce
qed simp
qed simp
lemma subterm_of_ex_set_iff_subtermeqset[code_unfold]: "t \<sqsubseteq>\<^sub>s\<^sub>e\<^sub>t M = (\<exists>t' \<in> M. subterm_of t t')"
using subterm_of_iff_subtermeq by blast
subsection \<open>The subterm relation is a partial order on terms\<close>
-interpretation "term": order "(\<sqsubseteq>)" "(\<sqsubset>)"
+
+interpretation "term": ordering "(\<sqsubseteq>)" "(\<sqsubset>)"
proof
show "s \<sqsubseteq> s" for s :: "('a,'b) term"
by (induct s rule: subterms.induct) auto
show trans: "s \<sqsubseteq> t \<Longrightarrow> t \<sqsubseteq> u \<Longrightarrow> s \<sqsubseteq> u" for s t u :: "('a,'b) term"
by (induct u rule: subterms.induct) auto
show "s \<sqsubseteq> t \<Longrightarrow> t \<sqsubseteq> s \<Longrightarrow> s = t" for s t :: "('a,'b) term"
proof (induction s arbitrary: t rule: subterms.induct[case_names Var Fun])
case (Fun f T)
{ assume 0: "t \<noteq> Fun f T"
then obtain u::"('a,'b) term" where u: "u \<in> set T" "t \<sqsubseteq> u" using Fun.prems(2) by auto
hence 1: "Fun f T \<sqsubseteq> u" using trans[OF Fun.prems(1)] by simp
have 2: "u \<sqsubseteq> Fun f T"
by (cases u) (use u(1) in force, use u(1) subterms.simps(2)[of f T] in fastforce)
hence 3: "u = Fun f T" using Fun.IH[OF u(1) _ 1] by simp
have "u \<sqsubseteq> t" using trans[OF 2 Fun.prems(1)] by simp
hence 4: "u = t" using Fun.IH[OF u(1) _ u(2)] by simp
have "t = Fun f T" using 3 4 by simp
hence False using 0 by simp
}
thus ?case by auto
qed simp
- thus "(s \<sqsubset> t) = (s \<sqsubseteq> t \<and> \<not>(t \<sqsubseteq> s))" for s t :: "('a,'b) term"
- by blast
+ show \<open>s \<sqsubset> t \<longleftrightarrow> s \<sqsubset> t\<close> for s t :: "('a,'b) term" ..
qed
+interpretation "term": order "(\<sqsubseteq>)" "(\<sqsubset>)"
+ by (rule ordering_orderI) (fact term.ordering_axioms)
subsection \<open>Lemmata concerning subterms and free variables\<close>
lemma fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append: "fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) = fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F@fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G"
by (simp add: fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_def)
lemma distinct_fv_list_idx_fv_disjoint:
assumes t: "distinct (fv_list t)" "Fun f T \<sqsubseteq> t"
and ij: "i < length T" "j < length T" "i < j"
shows "fv (T ! i) \<inter> fv (T ! j) = {}"
using t
proof (induction t rule: fv_list.induct)
case (2 g S)
have "distinct (fv_list s)" when s: "s \<in> set S" for s
by (metis (no_types, lifting) s "2.prems"(1) concat_append distinct_append
map_append split_list fv_list.simps(2) concat.simps(2) list.simps(9))
hence IH: "fv (T ! i) \<inter> fv (T ! j) = {}"
when s: "s \<in> set S" "Fun f T \<sqsubseteq> s" for s
using "2.IH" s by blast
show ?case
proof (cases "Fun f T = Fun g S")
case True
define U where "U \<equiv> map fv_list T"
have a: "distinct (concat U)"
using "2.prems"(1) True unfolding U_def by auto
have b: "i < length U" "j < length U"
using ij(1,2) unfolding U_def by simp_all
show ?thesis
using b distinct_concat_idx_disjoint[OF a b ij(3)]
fv_list_is_fv[of "T ! i"] fv_list_is_fv[of "T ! j"]
unfolding U_def by force
qed (use IH "2.prems"(2) in auto)
qed force
lemmas subtermeqI'[intro] = term.eq_refl
lemma subtermeqI''[intro]: "t \<in> set T \<Longrightarrow> t \<sqsubseteq> Fun f T"
by force
lemma finite_fv_set[intro]: "finite M \<Longrightarrow> finite (fv\<^sub>s\<^sub>e\<^sub>t M)"
by auto
lemma finite_fun_symbols[simp]: "finite (funs_term t)"
by (induct t) simp_all
lemma fv_set_mono: "M \<subseteq> N \<Longrightarrow> fv\<^sub>s\<^sub>e\<^sub>t M \<subseteq> fv\<^sub>s\<^sub>e\<^sub>t N"
by auto
lemma subterms\<^sub>s\<^sub>e\<^sub>t_mono: "M \<subseteq> N \<Longrightarrow> subterms\<^sub>s\<^sub>e\<^sub>t M \<subseteq> subterms\<^sub>s\<^sub>e\<^sub>t N"
by auto
lemma ground_empty[simp]: "ground {}"
by simp
lemma ground_subset: "M \<subseteq> N \<Longrightarrow> ground N \<Longrightarrow> ground M"
by auto
lemma fv_map_fv_set: "\<Union>(set (map fv L)) = fv\<^sub>s\<^sub>e\<^sub>t (set L)"
by (induct L) auto
lemma fv\<^sub>s\<^sub>e\<^sub>t_union: "fv\<^sub>s\<^sub>e\<^sub>t (M \<union> N) = fv\<^sub>s\<^sub>e\<^sub>t M \<union> fv\<^sub>s\<^sub>e\<^sub>t N"
by auto
lemma finite_subset_Union:
fixes A::"'a set" and f::"'a \<Rightarrow> 'b set"
assumes "finite (\<Union>a \<in> A. f a)"
shows "\<exists>B. finite B \<and> B \<subseteq> A \<and> (\<Union>b \<in> B. f b) = (\<Union>a \<in> A. f a)"
by (metis assms eq_iff finite_subset_image finite_UnionD)
lemma inv_set_fv: "finite M \<Longrightarrow> \<Union>(set (map fv (inv set M))) = fv\<^sub>s\<^sub>e\<^sub>t M"
using fv_map_fv_set[of "inv set M"] inv_set_fset by auto
lemma ground_subterm: "fv t = {} \<Longrightarrow> t' \<sqsubseteq> t \<Longrightarrow> fv t' = {}" by (induct t) auto
lemma empty_fv_not_var: "fv t = {} \<Longrightarrow> t \<noteq> Var x" by auto
lemma empty_fv_exists_fun: "fv t = {} \<Longrightarrow> \<exists>f X. t = Fun f X" by (cases t) auto
lemma vars_iff_subtermeq: "x \<in> fv t \<longleftrightarrow> Var x \<sqsubseteq> t" by (induct t) auto
lemma vars_iff_subtermeq_set: "x \<in> fv\<^sub>s\<^sub>e\<^sub>t M \<longleftrightarrow> Var x \<in> subterms\<^sub>s\<^sub>e\<^sub>t M"
using vars_iff_subtermeq[of x] by auto
lemma vars_if_subtermeq_set: "Var x \<in> subterms\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> x \<in> fv\<^sub>s\<^sub>e\<^sub>t M"
by (metis vars_iff_subtermeq_set)
lemma subtermeq_set_if_vars: "x \<in> fv\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> Var x \<in> subterms\<^sub>s\<^sub>e\<^sub>t M"
by (metis vars_iff_subtermeq_set)
lemma vars_iff_subterm_or_eq: "x \<in> fv t \<longleftrightarrow> Var x \<sqsubset> t \<or> Var x = t"
by (induct t) (auto simp add: vars_iff_subtermeq)
lemma var_is_subterm: "x \<in> fv t \<Longrightarrow> Var x \<in> subterms t"
by (simp add: vars_iff_subtermeq)
lemma subterm_is_var: "Var x \<in> subterms t \<Longrightarrow> x \<in> fv t"
by (simp add: vars_iff_subtermeq)
lemma no_var_subterm: "\<not>t \<sqsubset> Var v" by auto
lemma fun_if_subterm: "t \<sqsubset> u \<Longrightarrow> \<exists>f X. u = Fun f X" by (induct u) simp_all
lemma subtermeq_vars_subset: "M \<sqsubseteq> N \<Longrightarrow> fv M \<subseteq> fv N" by (induct N) auto
lemma fv_subterms[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (subterms t) = fv t"
by (induct t) auto
lemma fv_subterms_set[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M) = fv\<^sub>s\<^sub>e\<^sub>t M"
using subtermeq_vars_subset by auto
lemma fv_subset: "t \<in> M \<Longrightarrow> fv t \<subseteq> fv\<^sub>s\<^sub>e\<^sub>t M"
by auto
lemma fv_subset_subterms: "t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> fv t \<subseteq> fv\<^sub>s\<^sub>e\<^sub>t M"
using fv_subset fv_subterms_set by metis
lemma subterms_finite[simp]: "finite (subterms t)" by (induction rule: subterms.induct) auto
lemma subterms_union_finite: "finite M \<Longrightarrow> finite (\<Union>t \<in> M. subterms t)"
by (induction rule: subterms.induct) auto
lemma subterms_subset: "t' \<sqsubseteq> t \<Longrightarrow> subterms t' \<subseteq> subterms t"
by (induction rule: subterms.induct) auto
lemma subterms_subset_set: "M \<subseteq> subterms t \<Longrightarrow> subterms\<^sub>s\<^sub>e\<^sub>t M \<subseteq> subterms t"
by (metis SUP_least contra_subsetD subterms_subset)
lemma subset_subterms_Union[simp]: "M \<subseteq> subterms\<^sub>s\<^sub>e\<^sub>t M" by auto
lemma in_subterms_Union: "t \<in> M \<Longrightarrow> t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M" using subset_subterms_Union by blast
lemma in_subterms_subset_Union: "t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> subterms t \<subseteq> subterms\<^sub>s\<^sub>e\<^sub>t M"
using subterms_subset by auto
lemma subterm_param_split:
assumes "t \<sqsubset> Fun f X"
shows "\<exists>pre x suf. t \<sqsubseteq> x \<and> X = pre@x#suf"
proof -
obtain x where "t \<sqsubseteq> x" "x \<in> set X" using assms by auto
then obtain pre suf where "X = pre@x#suf" "x \<notin> set pre \<or> x \<notin> set suf"
by (meson split_list_first split_list_last)
thus ?thesis using \<open>t \<sqsubseteq> x\<close> by auto
qed
lemma ground_iff_no_vars: "ground (M::('a,'b) terms) \<longleftrightarrow> (\<forall>v. Var v \<notin> (\<Union>m \<in> M. subterms m))"
proof
assume "ground M"
hence "\<forall>v. \<forall>m \<in> M. v \<notin> fv m" by auto
hence "\<forall>v. \<forall>m \<in> M. Var v \<notin> subterms m" by (simp add: vars_iff_subtermeq)
thus "(\<forall>v. Var v \<notin> (\<Union>m \<in> M. subterms m))" by simp
next
assume no_vars: "\<forall>v. Var v \<notin> (\<Union>m \<in> M. subterms m)"
moreover
{ assume "\<not>ground M"
then obtain v and m::"('a,'b) term" where "m \<in> M" "fv m \<noteq> {}" "v \<in> fv m" by auto
hence "Var v \<in> (subterms m)" by (simp add: vars_iff_subtermeq)
hence "\<exists>v. Var v \<in> (\<Union>t \<in> M. subterms t)" using \<open>m \<in> M\<close> by auto
hence False using no_vars by simp
}
ultimately show "ground M" by blast
qed
lemma index_Fun_subterms_subset[simp]: "i < length T \<Longrightarrow> subterms (T ! i) \<subseteq> subterms (Fun f T)"
by auto
lemma index_Fun_fv_subset[simp]: "i < length T \<Longrightarrow> fv (T ! i) \<subseteq> fv (Fun f T)"
using subtermeq_vars_subset by fastforce
lemma subterms_union_ground:
assumes "ground M"
shows "ground (subterms\<^sub>s\<^sub>e\<^sub>t M)"
proof -
{ fix t assume "t \<in> M"
hence "fv t = {}"
using ground_iff_no_vars[of M] assms
by auto
hence "\<forall>t' \<in> subterms t. fv t' = {}" using subtermeq_vars_subset[of _ t] by simp
hence "ground (subterms t)" by auto
}
thus ?thesis by auto
qed
lemma Var_subtermeq: "t \<sqsubseteq> Var v \<Longrightarrow> t = Var v" by simp
lemma subtermeq_imp_funs_term_subset: "s \<sqsubseteq> t \<Longrightarrow> funs_term s \<subseteq> funs_term t"
by (induct t arbitrary: s) auto
lemma subterms_const: "subterms (Fun f []) = {Fun f []}" by simp
lemma subterm_subtermeq_neq: "\<lbrakk>t \<sqsubset> u; u \<sqsubseteq> v\<rbrakk> \<Longrightarrow> t \<noteq> v"
by (metis term.eq_iff)
lemma subtermeq_subterm_neq: "\<lbrakk>t \<sqsubseteq> u; u \<sqsubset> v\<rbrakk> \<Longrightarrow> t \<noteq> v"
by (metis term.eq_iff)
lemma subterm_size_lt: "x \<sqsubset> y \<Longrightarrow> size x < size y"
using not_less_eq size_list_estimation by (induct y, simp, fastforce)
lemma in_subterms_eq: "\<lbrakk>x \<in> subterms y; y \<in> subterms x\<rbrakk> \<Longrightarrow> subterms x = subterms y"
using term.antisym by auto
lemma Fun_gt_params: "Fun f X \<notin> (\<Union>x \<in> set X. subterms x)"
proof -
have "size_list size X < size (Fun f X)" by simp
hence "Fun f X \<notin> set X" by (meson less_not_refl size_list_estimation)
hence "\<forall>x \<in> set X. Fun f X \<notin> subterms x \<or> x \<notin> subterms (Fun f X)"
by (metis term.antisym[of "Fun f X" _])
moreover have "\<forall>x \<in> set X. x \<in> subterms (Fun f X)" by fastforce
ultimately show ?thesis by auto
qed
lemma params_subterms[simp]: "set X \<subseteq> subterms (Fun f X)" by auto
lemma params_subterms_Union[simp]: "subterms\<^sub>s\<^sub>e\<^sub>t (set X) \<subseteq> subterms (Fun f X)" by auto
lemma Fun_subterm_inside_params: "t \<sqsubset> Fun f X \<longleftrightarrow> t \<in> (\<Union>x \<in> (set X). subterms x)"
using Fun_gt_params by fastforce
lemma Fun_param_is_subterm: "x \<in> set X \<Longrightarrow> x \<sqsubset> Fun f X"
using Fun_subterm_inside_params by fastforce
lemma Fun_param_in_subterms: "x \<in> set X \<Longrightarrow> x \<in> subterms (Fun f X)"
using Fun_subterm_inside_params by fastforce
lemma Fun_not_in_param: "x \<in> set X \<Longrightarrow> \<not>Fun f X \<sqsubset> x"
using term.antisym by fast
lemma Fun_ex_if_subterm: "t \<sqsubset> s \<Longrightarrow> \<exists>f T. Fun f T \<sqsubseteq> s \<and> t \<in> set T"
proof (induction s)
case (Fun f T)
then obtain s' where s': "s' \<in> set T" "t \<sqsubseteq> s'" by auto
show ?case
proof (cases "t = s'")
case True thus ?thesis using s' by blast
next
case False
thus ?thesis
using Fun.IH[OF s'(1)] s'(2) term.order_trans[OF _ Fun_param_in_subterms[OF s'(1), of f]]
by metis
qed
qed simp
lemma const_subterm_obtain:
assumes "fv t = {}"
obtains c where "Fun c [] \<sqsubseteq> t"
using assms
proof (induction t)
case (Fun f T) thus ?case by (cases "T = []") force+
qed simp
lemma const_subterm_obtain': "fv t = {} \<Longrightarrow> \<exists>c. Fun c [] \<sqsubseteq> t"
by (metis const_subterm_obtain)
lemma subterms_singleton:
assumes "(\<exists>v. t = Var v) \<or> (\<exists>f. t = Fun f [])"
shows "subterms t = {t}"
using assms by (cases t) auto
lemma subtermeq_Var_const:
assumes "s \<sqsubseteq> t"
shows "t = Var v \<Longrightarrow> s = Var v" "t = Fun f [] \<Longrightarrow> s = Fun f []"
using assms by fastforce+
lemma subterms_singleton':
assumes "subterms t = {t}"
shows "(\<exists>v. t = Var v) \<or> (\<exists>f. t = Fun f [])"
proof (cases t)
case (Fun f T)
{ fix s S assume "T = s#S"
hence "s \<in> subterms t" using Fun by auto
hence "s = t" using assms by auto
hence False
using Fun_param_is_subterm[of s "s#S" f] \<open>T = s#S\<close> Fun
by auto
}
hence "T = []" by (cases T) auto
thus ?thesis using Fun by simp
qed (simp add: assms)
lemma funs_term_subterms_eq[simp]:
"(\<Union>s \<in> subterms t. funs_term s) = funs_term t"
"(\<Union>s \<in> subterms\<^sub>s\<^sub>e\<^sub>t M. funs_term s) = \<Union>(funs_term ` M)"
proof -
show "\<And>t. \<Union>(funs_term ` subterms t) = funs_term t"
using term.order_refl subtermeq_imp_funs_term_subset by blast
thus "\<Union>(funs_term ` (subterms\<^sub>s\<^sub>e\<^sub>t M)) = \<Union>(funs_term ` M)" by force
qed
lemmas subtermI'[intro] = Fun_param_is_subterm
lemma funs_term_Fun_subterm: "f \<in> funs_term t \<Longrightarrow> \<exists>T. Fun f T \<in> subterms t"
proof (induction t)
case (Fun g T)
hence "f = g \<or> (\<exists>s \<in> set T. f \<in> funs_term s)" by simp
thus ?case
proof
assume "\<exists>s \<in> set T. f \<in> funs_term s"
then obtain s where "s \<in> set T" "\<exists>T. Fun f T \<in> subterms s" using Fun.IH by auto
thus ?thesis by auto
qed (auto simp add: Fun)
qed simp
lemma funs_term_Fun_subterm': "Fun f T \<in> subterms t \<Longrightarrow> f \<in> funs_term t"
by (induct t) auto
lemma zip_arg_subterm:
assumes "(s,t) \<in> set (zip X Y)"
shows "s \<sqsubset> Fun f X" "t \<sqsubset> Fun g Y"
proof -
from assms have *: "s \<in> set X" "t \<in> set Y" by (meson in_set_zipE)+
show "s \<sqsubset> Fun f X" by (metis Fun_param_is_subterm[OF *(1)])
show "t \<sqsubset> Fun g Y" by (metis Fun_param_is_subterm[OF *(2)])
qed
lemma fv_disj_Fun_subterm_param_cases:
assumes "fv t \<inter> X = {}" "Fun f T \<in> subterms t"
shows "T = [] \<or> (\<exists>s\<in>set T. s \<notin> Var ` X)"
proof (cases T)
case (Cons s S)
hence "s \<in> subterms t"
using assms(2) term.order_trans[of _ "Fun f T" t]
by auto
hence "fv s \<inter> X = {}" using assms(1) fv_subterms by force
thus ?thesis using Cons by auto
qed simp
lemma fv_eq_FunI:
assumes "length T = length S" "\<And>i. i < length T \<Longrightarrow> fv (T ! i) = fv (S ! i)"
shows "fv (Fun f T) = fv (Fun g S)"
using assms
proof (induction T arbitrary: S)
case (Cons t T S')
then obtain s S where S': "S' = s#S" by (cases S') simp_all
thus ?case using Cons by fastforce
qed simp
lemma fv_eq_FunI':
assumes "length T = length S" "\<And>i. i < length T \<Longrightarrow> x \<in> fv (T ! i) \<longleftrightarrow> x \<in> fv (S ! i)"
shows "x \<in> fv (Fun f T) \<longleftrightarrow> x \<in> fv (Fun g S)"
using assms
proof (induction T arbitrary: S)
case (Cons t T S')
then obtain s S where S': "S' = s#S" by (cases S') simp_all
thus ?case using Cons by fastforce
qed simp
lemma finite_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[simp]: "finite (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s x)" by auto
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_Nil[simp]: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s [] = {}" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_singleton[simp]: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s [(t,s)] = fv t \<union> fv s" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_Cons: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) = fv s \<union> fv t \<union> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \<union> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono: "set M \<subseteq> set N \<Longrightarrow> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s M \<subseteq> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s N" by auto
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI[intro]:
"f \<in> set F \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r f \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"f \<in> set F \<Longrightarrow> x \<in> fv (fst f) \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"f \<in> set F \<Longrightarrow> x \<in> fv (snd f) \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"(t,s) \<in> set F \<Longrightarrow> x \<in> fv t \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"(t,s) \<in> set F \<Longrightarrow> x \<in> fv s \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
using UN_I by fastforce+
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_cons_subset: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \<subseteq> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F)"
by auto
subsection \<open>Other lemmata\<close>
lemma nonvar_term_has_composed_shallow_term:
fixes t::"('f,'v) term"
assumes "\<not>(\<exists>x. t = Var x)"
shows "\<exists>f T. Fun f T \<sqsubseteq> t \<and> (\<forall>s \<in> set T. (\<exists>c. s = Fun c []) \<or> (\<exists>x. s = Var x))"
proof -
let ?Q = "\<lambda>S. \<forall>s \<in> set S. (\<exists>c. s = Fun c []) \<or> (\<exists>x. s = Var x)"
let ?P = "\<lambda>t. \<exists>g S. Fun g S \<sqsubseteq> t \<and> ?Q S"
{ fix t::"('f,'v) term"
have "(\<exists>x. t = Var x) \<or> ?P t"
proof (induction t)
case (Fun h R) show ?case
proof (cases "R = [] \<or> (\<forall>r \<in> set R. \<exists>x. r = Var x)")
case False
then obtain r g S where "r \<in> set R" "?P r" "Fun g S \<sqsubseteq> r" "?Q S" using Fun.IH by fast
thus ?thesis by auto
qed force
qed simp
} thus ?thesis using assms by blast
qed
end
diff --git a/thys/Stone_Algebras/Filters.thy b/thys/Stone_Algebras/Filters.thy
--- a/thys/Stone_Algebras/Filters.thy
+++ b/thys/Stone_Algebras/Filters.thy
@@ -1,943 +1,943 @@
(* Title: Filters
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Filters\<close>
text \<open>
This theory develops filters based on orders, semilattices, lattices and distributive lattices.
We prove the ultrafilter lemma for orders with a least element.
We show the following structure theorems:
\begin{itemize}
\item The set of filters over a directed semilattice forms a lattice with a greatest element.
\item The set of filters over a bounded semilattice forms a bounded lattice.
\item The set of filters over a distributive lattice with a greatest element forms a bounded distributive lattice.
\end{itemize}
Another result is that in a distributive lattice ultrafilters are prime filters.
We also prove a lemma of Gr\"atzer and Schmidt about principal filters.
We apply these results in proving the construction theorem for Stone algebras (described in a separate theory).
See, for example, \cite{BalbesDwinger1974,Birkhoff1967,Blyth2005,DaveyPriestley2002,Graetzer1971} for further results about filters.
\<close>
theory Filters
imports Lattice_Basics
begin
subsection \<open>Orders\<close>
text \<open>
This section gives the basic definitions related to filters in terms of orders.
The main result is the ultrafilter lemma.
\<close>
context ord
begin
abbreviation down :: "'a \<Rightarrow> 'a set" ("\<down>_" [81] 80)
where "\<down>x \<equiv> { y . y \<le> x }"
abbreviation down_set :: "'a set \<Rightarrow> 'a set" ("\<Down>_" [81] 80)
where "\<Down>X \<equiv> { y . \<exists>x\<in>X . y \<le> x }"
abbreviation is_down_set :: "'a set \<Rightarrow> bool"
where "is_down_set X \<equiv> \<forall>x\<in>X . \<forall>y . y \<le> x \<longrightarrow> y\<in>X"
abbreviation is_principal_down :: "'a set \<Rightarrow> bool"
where "is_principal_down X \<equiv> \<exists>x . X = \<down>x"
abbreviation up :: "'a \<Rightarrow> 'a set" ("\<up>_" [81] 80)
where "\<up>x \<equiv> { y . x \<le> y }"
abbreviation up_set :: "'a set \<Rightarrow> 'a set" ("\<Up>_" [81] 80)
where "\<Up>X \<equiv> { y . \<exists>x\<in>X . x \<le> y }"
abbreviation is_up_set :: "'a set \<Rightarrow> bool"
where "is_up_set X \<equiv> \<forall>x\<in>X . \<forall>y . x \<le> y \<longrightarrow> y\<in>X"
abbreviation is_principal_up :: "'a set \<Rightarrow> bool"
where "is_principal_up X \<equiv> \<exists>x . X = \<up>x"
text \<open>
A filter is a non-empty, downward directed, up-closed set.
\<close>
definition filter :: "'a set \<Rightarrow> bool"
where "filter F \<equiv> (F \<noteq> {}) \<and> (\<forall>x\<in>F . \<forall>y\<in>F . \<exists>z\<in>F . z \<le> x \<and> z \<le> y) \<and> is_up_set F"
abbreviation proper_filter :: "'a set \<Rightarrow> bool"
where "proper_filter F \<equiv> filter F \<and> F \<noteq> UNIV"
abbreviation ultra_filter :: "'a set \<Rightarrow> bool"
where "ultra_filter F \<equiv> proper_filter F \<and> (\<forall>G . proper_filter G \<and> F \<subseteq> G \<longrightarrow> F = G)"
abbreviation filters :: "'a set set"
where "filters \<equiv> { F::'a set . filter F }"
lemma filter_map_filter:
assumes "filter F"
and "mono f"
and "\<forall>x y . f x \<le> y \<longrightarrow> (\<exists>z . x \<le> z \<and> y = f z)"
shows "filter (f ` F)"
proof (unfold ord_class.filter_def, intro conjI)
show "f ` F \<noteq> {}"
using assms(1) ord_class.filter_def by auto
next
show "\<forall>x\<in>f ` F . \<forall>y\<in>f ` F . \<exists>z\<in>f ` F . z \<le> x \<and> z \<le> y"
proof (intro ballI)
fix x y
assume "x \<in> f ` F" and "y \<in> f ` F"
then obtain u v where 1: "x = f u \<and> u \<in> F \<and> y = f v \<and> v \<in> F"
by auto
then obtain w where "w \<le> u \<and> w \<le> v \<and> w \<in> F"
by (meson assms(1) ord_class.filter_def)
thus "\<exists>z\<in>f ` F . z \<le> x \<and> z \<le> y"
using 1 assms(2) mono_def image_eqI by blast
qed
next
show "is_up_set (f ` F)"
proof
fix x
assume "x \<in> f ` F"
then obtain u where 1: "x = f u \<and> u \<in> F"
by auto
show "\<forall>y . x \<le> y \<longrightarrow> y \<in> f ` F"
proof (rule allI, rule impI)
fix y
assume "x \<le> y"
hence "f u \<le> y"
using 1 by simp
then obtain z where "u \<le> z \<and> y = f z"
using assms(3) by auto
thus "y \<in> f ` F"
by (meson 1 assms(1) image_iff ord_class.filter_def)
qed
qed
qed
end
context order
begin
lemma self_in_downset [simp]:
"x \<in> \<down>x"
by simp
lemma self_in_upset [simp]:
"x \<in> \<up>x"
by simp
lemma up_filter [simp]:
"filter (\<up>x)"
using filter_def order_lesseq_imp by auto
lemma up_set_up_set [simp]:
"is_up_set (\<Up>X)"
using order.trans by fastforce
lemma up_injective:
"\<up>x = \<up>y \<Longrightarrow> x = y"
- using antisym by auto
+ using order.antisym by auto
lemma up_antitone:
"x \<le> y \<longleftrightarrow> \<up>y \<subseteq> \<up>x"
by auto
end
context order_bot
begin
lemma bot_in_downset [simp]:
"bot \<in> \<down>x"
by simp
lemma down_bot [simp]:
"\<down>bot = {bot}"
by (simp add: bot_unique)
lemma up_bot [simp]:
"\<up>bot = UNIV"
by simp
text \<open>
The following result is the ultrafilter lemma, generalised from \cite[10.17]{DaveyPriestley2002} to orders with a least element.
Its proof uses Isabelle/HOL's \<open>Zorn_Lemma\<close>, which requires closure under union of arbitrary (possibly empty) chains.
Actually, the proof does not use any of the underlying order properties except \<open>bot_least\<close>.
\<close>
lemma ultra_filter:
assumes "proper_filter F"
shows "\<exists>G . ultra_filter G \<and> F \<subseteq> G"
proof -
let ?A = "{ G . (proper_filter G \<and> F \<subseteq> G) \<or> G = {} }"
have "\<forall>C \<in> chains ?A . \<Union>C \<in> ?A"
proof
fix C :: "'a set set"
let ?D = "C - {{}}"
assume 1: "C \<in> chains ?A"
hence 2: "\<forall>x\<in>\<Union>?D . \<exists>H\<in>?D . x \<in> H \<and> proper_filter H"
using chainsD2 by fastforce
have 3: "\<Union>?D = \<Union>C"
by blast
have "\<Union>?D \<in> ?A"
proof (cases "?D = {}")
assume "?D = {}"
thus ?thesis
by auto
next
assume 4: "?D \<noteq> {}"
then obtain G where "G \<in> ?D"
by auto
hence 5: "F \<subseteq> \<Union>?D"
using 1 chainsD2 by blast
have 6: "is_up_set (\<Union>?D)"
proof
fix x
assume "x \<in> \<Union>?D"
then obtain H where "x \<in> H \<and> H \<in> ?D \<and> filter H"
using 2 by auto
thus "\<forall>y . x \<le> y \<longrightarrow> y\<in>\<Union>?D"
using filter_def UnionI by fastforce
qed
have 7: "\<Union>?D \<noteq> UNIV"
proof (rule ccontr)
assume "\<not> \<Union>?D \<noteq> UNIV"
then obtain H where "bot \<in> H \<and> proper_filter H"
using 2 by blast
thus False
by (meson UNIV_I bot_least filter_def subsetI subset_antisym)
qed
{
fix x y
assume "x\<in>\<Union>?D \<and> y\<in>\<Union>?D"
then obtain H I where 8: "x \<in> H \<and> H \<in> ?D \<and> filter H \<and> y \<in> I \<and> I \<in> ?D \<and> filter I"
using 2 by metis
have "\<exists>z\<in>\<Union>?D . z \<le> x \<and> z \<le> y"
proof (cases "H \<subseteq> I")
assume "H \<subseteq> I"
hence "\<exists>z\<in>I . z \<le> x \<and> z \<le> y"
using 8 by (metis subsetCE filter_def)
thus ?thesis
using 8 by (metis UnionI)
next
assume "\<not> (H \<subseteq> I)"
hence "I \<subseteq> H"
using 1 8 by (meson DiffE chainsD)
hence "\<exists>z\<in>H . z \<le> x \<and> z \<le> y"
using 8 by (metis subsetCE filter_def)
thus ?thesis
using 8 by (metis UnionI)
qed
}
thus ?thesis
using 4 5 6 7 filter_def by auto
qed
thus "\<Union>C \<in> ?A"
using 3 by simp
qed
hence "\<exists>M\<in>?A . \<forall>X\<in>?A . M \<subseteq> X \<longrightarrow> X = M"
by (rule Zorn_Lemma)
then obtain M where 9: "M \<in> ?A \<and> (\<forall>X\<in>?A . M \<subseteq> X \<longrightarrow> X = M)"
by auto
hence 10: "M \<noteq> {}"
using assms filter_def by auto
{
fix G
assume 11: "proper_filter G \<and> M \<subseteq> G"
hence "F \<subseteq> G"
using 9 10 by blast
hence "M = G"
using 9 11 by auto
}
thus ?thesis
using 9 10 by blast
qed
end
context order_top
begin
lemma down_top [simp]:
"\<down>top = UNIV"
by simp
lemma top_in_upset [simp]:
"top \<in> \<up>x"
by simp
lemma up_top [simp]:
"\<up>top = {top}"
by (simp add: top_unique)
lemma filter_top [simp]:
"filter {top}"
using filter_def top_unique by auto
lemma top_in_filter [simp]:
"filter F \<Longrightarrow> top \<in> F"
using filter_def by fastforce
end
text \<open>
The existence of proper filters and ultrafilters requires that the underlying order contains at least two elements.
\<close>
context non_trivial_order
begin
lemma proper_filter_exists:
"\<exists>F . proper_filter F"
proof -
from consistent obtain x y :: 'a where "x \<noteq> y"
by auto
hence "\<up>x \<noteq> UNIV \<or> \<up>y \<noteq> UNIV"
- using antisym by blast
+ using order.antisym by blast
hence "proper_filter (\<up>x) \<or> proper_filter (\<up>y)"
by simp
thus ?thesis
by blast
qed
end
context non_trivial_order_bot
begin
lemma ultra_filter_exists:
"\<exists>F . ultra_filter F"
using ultra_filter proper_filter_exists by blast
end
context non_trivial_bounded_order
begin
lemma proper_filter_top:
"proper_filter {top}"
using bot_not_top filter_top by blast
lemma ultra_filter_top:
"\<exists>G . ultra_filter G \<and> top \<in> G"
using ultra_filter proper_filter_top by fastforce
end
subsection \<open>Lattices\<close>
text \<open>
This section develops the lattice structure of filters based on a semilattice structure of the underlying order.
The main results are that filters over a directed semilattice form a lattice with a greatest element and that filters over a bounded semilattice form a bounded lattice.
\<close>
context semilattice_sup
begin
abbreviation prime_filter :: "'a set \<Rightarrow> bool"
where "prime_filter F \<equiv> proper_filter F \<and> (\<forall>x y . x \<squnion> y \<in> F \<longrightarrow> x \<in> F \<or> y \<in> F)"
end
context semilattice_inf
begin
lemma filter_inf_closed:
"filter F \<Longrightarrow> x \<in> F \<Longrightarrow> y \<in> F \<Longrightarrow> x \<sqinter> y \<in> F"
by (meson filter_def inf.boundedI)
lemma filter_univ:
"filter UNIV"
by (meson UNIV_I UNIV_not_empty filter_def inf.cobounded1 inf.cobounded2)
text \<open>
The operation \<open>filter_sup\<close> is the join operation in the lattice of filters.
\<close>
definition "filter_sup F G \<equiv> { z . \<exists>x\<in>F . \<exists>y\<in>G . x \<sqinter> y \<le> z }"
lemma filter_sup:
assumes "filter F"
and "filter G"
shows "filter (filter_sup F G)"
proof -
have "F \<noteq> {} \<and> G \<noteq> {}"
using assms filter_def by blast
hence 1: "filter_sup F G \<noteq> {}"
using filter_sup_def by blast
have 2: "\<forall>x\<in>filter_sup F G . \<forall>y\<in>filter_sup F G . \<exists>z\<in>filter_sup F G . z \<le> x \<and> z \<le> y"
proof
fix x
assume "x\<in>filter_sup F G"
then obtain t u where 3: "t \<in> F \<and> u \<in> G \<and> t \<sqinter> u \<le> x"
using filter_sup_def by auto
show "\<forall>y\<in>filter_sup F G . \<exists>z\<in>filter_sup F G . z \<le> x \<and> z \<le> y"
proof
fix y
assume "y\<in>filter_sup F G"
then obtain v w where 4: "v \<in> F \<and> w \<in> G \<and> v \<sqinter> w \<le> y"
using filter_sup_def by auto
let ?z = "(t \<sqinter> v) \<sqinter> (u \<sqinter> w)"
have 5: "?z \<le> x \<and> ?z \<le> y"
using 3 4 by (meson order.trans inf.cobounded1 inf.cobounded2 inf_mono)
have "?z \<in> filter_sup F G"
unfolding filter_sup_def using assms 3 4 filter_inf_closed by blast
thus "\<exists>z\<in>filter_sup F G . z \<le> x \<and> z \<le> y"
using 5 by blast
qed
qed
have "\<forall>x\<in>filter_sup F G . \<forall>y . x \<le> y \<longrightarrow> y \<in> filter_sup F G"
unfolding filter_sup_def using order_trans by blast
thus ?thesis
using 1 2 filter_def by presburger
qed
lemma filter_sup_left_upper_bound:
assumes "filter G"
shows "F \<subseteq> filter_sup F G"
proof -
from assms obtain y where "y\<in>G"
using all_not_in_conv filter_def by auto
thus ?thesis
unfolding filter_sup_def using inf.cobounded1 by blast
qed
lemma filter_sup_symmetric:
"filter_sup F G = filter_sup G F"
unfolding filter_sup_def using inf.commute by fastforce
lemma filter_sup_right_upper_bound:
"filter F \<Longrightarrow> G \<subseteq> filter_sup F G"
using filter_sup_symmetric filter_sup_left_upper_bound by simp
lemma filter_sup_least_upper_bound:
assumes "filter H"
and "F \<subseteq> H"
and "G \<subseteq> H"
shows "filter_sup F G \<subseteq> H"
proof
fix x
assume "x \<in> filter_sup F G"
then obtain y z where 1: "y \<in> F \<and> z \<in> G \<and> y \<sqinter> z \<le> x"
using filter_sup_def by auto
hence "y \<in> H \<and> z \<in> H"
using assms(2-3) by auto
hence "y \<sqinter> z \<in> H"
by (simp add: assms(1) filter_inf_closed)
thus "x \<in> H"
using 1 assms(1) filter_def by auto
qed
lemma filter_sup_left_isotone:
"G \<subseteq> H \<Longrightarrow> filter_sup G F \<subseteq> filter_sup H F"
unfolding filter_sup_def by blast
lemma filter_sup_right_isotone:
"G \<subseteq> H \<Longrightarrow> filter_sup F G \<subseteq> filter_sup F H"
unfolding filter_sup_def by blast
lemma filter_sup_right_isotone_var:
"filter_sup F (G \<inter> H) \<subseteq> filter_sup F H"
unfolding filter_sup_def by blast
lemma up_dist_inf:
"\<up>(x \<sqinter> y) = filter_sup (\<up>x) (\<up>y)"
proof
show "\<up>(x \<sqinter> y) \<subseteq> filter_sup (\<up>x) (\<up>y)"
unfolding filter_sup_def by blast
next
show "filter_sup (\<up>x) (\<up>y) \<subseteq> \<up>(x \<sqinter> y)"
proof
fix z
assume "z \<in> filter_sup (\<up>x) (\<up>y)"
then obtain u v where "u\<in>\<up>x \<and> v\<in>\<up>y \<and> u \<sqinter> v \<le> z"
using filter_sup_def by auto
hence "x \<sqinter> y \<le> z"
using order.trans inf_mono by blast
thus "z \<in> \<up>(x \<sqinter> y)"
by blast
qed
qed
text \<open>
The following result is part of \cite[Exercise 2.23]{DaveyPriestley2002}.
\<close>
lemma filter_inf_filter [simp]:
assumes "filter F"
shows "filter (\<Up>{ y . \<exists>z\<in>F . x \<sqinter> z = y})"
proof -
let ?G = "\<Up>{ y . \<exists>z\<in>F . x \<sqinter> z = y}"
have "F \<noteq> {}"
using assms filter_def by simp
hence 1: "?G \<noteq> {}"
by blast
have 2: "is_up_set ?G"
by auto
{
fix y z
assume "y \<in> ?G \<and> z \<in> ?G"
then obtain v w where "v \<in> F \<and> w \<in> F \<and> x \<sqinter> v \<le> y \<and> x \<sqinter> w \<le> z"
by auto
hence "v \<sqinter> w \<in> F \<and> x \<sqinter> (v \<sqinter> w) \<le> y \<sqinter> z"
by (meson assms filter_inf_closed order.trans inf.boundedI inf.cobounded1 inf.cobounded2)
hence "\<exists>u\<in>?G . u \<le> y \<and> u \<le> z"
by auto
}
hence "\<forall>x\<in>?G . \<forall>y\<in>?G . \<exists>z\<in>?G . z \<le> x \<and> z \<le> y"
by auto
thus ?thesis
using 1 2 filter_def by presburger
qed
end
context directed_semilattice_inf
begin
text \<open>
Set intersection is the meet operation in the lattice of filters.
\<close>
lemma filter_inf:
assumes "filter F"
and "filter G"
shows "filter (F \<inter> G)"
proof (unfold filter_def, intro conjI)
from assms obtain x y where 1: "x\<in>F \<and> y\<in>G"
using all_not_in_conv filter_def by auto
from ub obtain z where "x \<le> z \<and> y \<le> z"
by auto
hence "z \<in> F \<inter> G"
using 1 by (meson assms Int_iff filter_def)
thus "F \<inter> G \<noteq> {}"
by blast
next
show "is_up_set (F \<inter> G)"
by (meson assms Int_iff filter_def)
next
show "\<forall>x\<in>F \<inter> G . \<forall>y\<in>F \<inter> G . \<exists>z\<in>F \<inter> G . z \<le> x \<and> z \<le> y"
by (metis assms Int_iff filter_inf_closed inf.cobounded2 inf.commute)
qed
end
text \<open>
We introduce the following type of filters to instantiate the lattice classes and thereby inherit the results shown about lattices.
\<close>
typedef (overloaded) 'a filter = "{ F::'a::order set . filter F }"
by (meson mem_Collect_eq up_filter)
lemma simp_filter [simp]:
"filter (Rep_filter x)"
using Rep_filter by simp
setup_lifting type_definition_filter
text \<open>
The set of filters over a directed semilattice forms a lattice with a greatest element.
\<close>
instantiation filter :: (directed_semilattice_inf) bounded_lattice_top
begin
lift_definition top_filter :: "'a filter" is UNIV
by (simp add: filter_univ)
lift_definition sup_filter :: "'a filter \<Rightarrow> 'a filter \<Rightarrow> 'a filter" is filter_sup
by (simp add: filter_sup)
lift_definition inf_filter :: "'a filter \<Rightarrow> 'a filter \<Rightarrow> 'a filter" is inter
by (simp add: filter_inf)
lift_definition less_eq_filter :: "'a filter \<Rightarrow> 'a filter \<Rightarrow> bool" is subset_eq .
lift_definition less_filter :: "'a filter \<Rightarrow> 'a filter \<Rightarrow> bool" is subset .
instance
apply intro_classes
subgoal apply transfer by (simp add: less_le_not_le)
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by (simp add: filter_sup_left_upper_bound)
subgoal apply transfer by (simp add: filter_sup_right_upper_bound)
subgoal apply transfer by (simp add: filter_sup_least_upper_bound)
subgoal apply transfer by simp
done
end
context bounded_semilattice_inf_top
begin
abbreviation "filter_complements F G \<equiv> filter F \<and> filter G \<and> filter_sup F G = UNIV \<and> F \<inter> G = {top}"
end
text \<open>
The set of filters over a bounded semilattice forms a bounded lattice.
\<close>
instantiation filter :: (bounded_semilattice_inf_top) bounded_lattice
begin
lift_definition bot_filter :: "'a filter" is "{top}"
by simp
instance
apply intro_classes
apply transfer
by simp
end
context lattice
begin
lemma up_dist_sup:
"\<up>(x \<squnion> y) = \<up>x \<inter> \<up>y"
by auto
end
text \<open>
For convenience, the following function injects principal filters into the filter type.
We cannot define it in the \<open>order\<close> class since the type filter requires the sort constraint \<open>order\<close> that is not available in the class.
The result of the function is a filter by lemma \<open>up_filter\<close>.
\<close>
abbreviation up_filter :: "'a::order \<Rightarrow> 'a filter"
where "up_filter x \<equiv> Abs_filter (\<up>x)"
lemma up_filter_dist_inf:
"up_filter ((x::'a::lattice) \<sqinter> y) = up_filter x \<squnion> up_filter y"
by (simp add: eq_onp_def sup_filter.abs_eq up_dist_inf)
lemma up_filter_dist_sup:
"up_filter ((x::'a::lattice) \<squnion> y) = up_filter x \<sqinter> up_filter y"
by (metis eq_onp_def inf_filter.abs_eq up_dist_sup up_filter)
lemma up_filter_injective:
"up_filter x = up_filter y \<Longrightarrow> x = y"
by (metis Abs_filter_inject mem_Collect_eq up_filter up_injective)
lemma up_filter_antitone:
"x \<le> y \<longleftrightarrow> up_filter y \<le> up_filter x"
by (metis eq_onp_same_args less_eq_filter.abs_eq up_antitone up_filter)
text \<open>
The following definition applies a function to each element of a filter.
The subsequent lemma gives conditions under which the result of this application is a filter.
\<close>
abbreviation filter_map :: "('a::order \<Rightarrow> 'b::order) \<Rightarrow> 'a filter \<Rightarrow> 'b filter"
where "filter_map f F \<equiv> Abs_filter (f ` Rep_filter F)"
lemma filter_map_filter:
assumes "mono f"
and "\<forall>x y . f x \<le> y \<longrightarrow> (\<exists>z . x \<le> z \<and> y = f z)"
shows "filter (f ` Rep_filter F)"
by (simp add: assms inf.filter_map_filter)
subsection \<open>Distributive Lattices\<close>
text \<open>
In this section we additionally assume that the underlying order forms a distributive lattice.
Then filters form a bounded distributive lattice if the underlying order has a greatest element.
Moreover ultrafilters are prime filters.
We also prove a lemma of Gr\"atzer and Schmidt about principal filters.
\<close>
context distrib_lattice
begin
lemma filter_sup_left_dist_inf:
assumes "filter F"
and "filter G"
and "filter H"
shows "filter_sup F (G \<inter> H) = filter_sup F G \<inter> filter_sup F H"
proof
show "filter_sup F (G \<inter> H) \<subseteq> filter_sup F G \<inter> filter_sup F H"
unfolding filter_sup_def using filter_sup_right_isotone_var by blast
next
show "filter_sup F G \<inter> filter_sup F H \<subseteq> filter_sup F (G \<inter> H)"
proof
fix x
assume "x \<in> filter_sup F G \<inter> filter_sup F H"
then obtain t u v w where 1: "t \<in> F \<and> u \<in> G \<and> v \<in> F \<and> w \<in> H \<and> t \<sqinter> u \<le> x \<and> v \<sqinter> w \<le> x"
using filter_sup_def by auto
let ?y = "t \<sqinter> v"
let ?z = "u \<squnion> w"
have 2: "?y \<in> F"
using 1 by (simp add: assms(1) filter_inf_closed)
have 3: "?z \<in> G \<inter> H"
using 1 by (meson assms(2-3) Int_iff filter_def sup_ge1 sup_ge2)
have "?y \<sqinter> ?z = (t \<sqinter> v \<sqinter> u) \<squnion> (t \<sqinter> v \<sqinter> w)"
by (simp add: inf_sup_distrib1)
also have "... \<le> (t \<sqinter> u) \<squnion> (v \<sqinter> w)"
by (metis inf.cobounded1 inf.cobounded2 inf.left_idem inf_mono sup.mono)
also have "... \<le> x"
using 1 by (simp add: le_supI)
finally show "x \<in> filter_sup F (G \<inter> H)"
unfolding filter_sup_def using 2 3 by blast
qed
qed
lemma filter_inf_principal_rep:
"F \<inter> G = \<up>z \<Longrightarrow> (\<exists>x\<in>F . \<exists>y\<in>G . z = x \<squnion> y)"
by force
lemma filter_sup_principal_rep:
assumes "filter F"
and "filter G"
and "filter_sup F G = \<up>z"
shows "\<exists>x\<in>F . \<exists>y\<in>G . z = x \<sqinter> y"
proof -
from assms(3) obtain x y where 1: "x\<in>F \<and> y\<in>G \<and> x \<sqinter> y \<le> z"
unfolding filter_sup_def using order_refl by blast
hence 2: "x \<squnion> z \<in> F \<and> y \<squnion> z \<in> G"
by (meson assms(1-2) sup_ge1 filter_def)
have "(x \<squnion> z) \<sqinter> (y \<squnion> z) = z"
using 1 sup_absorb2 sup_inf_distrib2 by fastforce
thus ?thesis
using 2 by force
qed
lemma inf_sup_principal_aux:
assumes "filter F"
and "filter G"
and "is_principal_up (filter_sup F G)"
and "is_principal_up (F \<inter> G)"
shows "is_principal_up F"
proof -
from assms(3-4) obtain x y where 1: "filter_sup F G = \<up>x \<and> F \<inter> G = \<up>y"
by blast
from filter_inf_principal_rep obtain t u where 2: "t\<in>F \<and> u\<in>G \<and> y = t \<squnion> u"
using 1 by meson
from filter_sup_principal_rep obtain v w where 3: "v\<in>F \<and> w\<in>G \<and> x = v \<sqinter> w"
using 1 by (meson assms(1-2))
have "t \<in> filter_sup F G \<and> u \<in> filter_sup F G"
unfolding filter_sup_def using 2 inf.cobounded1 inf.cobounded2 by blast
hence "x \<le> t \<and> x \<le> u"
using 1 by blast
hence 4: "(t \<sqinter> v) \<sqinter> (u \<sqinter> w) = x"
using 3 by (simp add: inf.absorb2 inf.assoc inf.left_commute)
have "(t \<sqinter> v) \<squnion> (u \<sqinter> w) \<in> F \<and> (t \<sqinter> v) \<squnion> (u \<sqinter> w) \<in> G"
using 2 3 by (metis (no_types, lifting) assms(1-2) filter_inf_closed sup.cobounded1 sup.cobounded2 filter_def)
hence "y \<le> (t \<sqinter> v) \<squnion> (u \<sqinter> w)"
using 1 Int_iff by blast
hence 5: "(t \<sqinter> v) \<squnion> (u \<sqinter> w) = y"
- using 2 by (simp add: antisym inf.coboundedI1)
+ using 2 by (simp add: order.antisym inf.coboundedI1)
have "F = \<up>(t \<sqinter> v)"
proof
show "F \<subseteq> \<up>(t \<sqinter> v)"
proof
fix z
assume 6: "z \<in> F"
hence "z \<in> filter_sup F G"
unfolding filter_sup_def using 2 inf.cobounded1 by blast
hence "x \<le> z"
using 1 by simp
hence 7: "(t \<sqinter> v \<sqinter> z) \<sqinter> (u \<sqinter> w) = x"
using 4 by (metis inf.absorb1 inf.assoc inf.commute)
have "z \<squnion> u \<in> F \<and> z \<squnion> u \<in> G \<and> z \<squnion> w \<in> F \<and> z \<squnion> w \<in> G"
using 2 3 6 by (meson assms(1-2) filter_def sup_ge1 sup_ge2)
hence "y \<le> (z \<squnion> u) \<sqinter> (z \<squnion> w)"
using 1 Int_iff filter_inf_closed by auto
hence 8: "(t \<sqinter> v \<sqinter> z) \<squnion> (u \<sqinter> w) = y"
using 5 by (metis inf.absorb1 sup.commute sup_inf_distrib2)
have "t \<sqinter> v \<sqinter> z = t \<sqinter> v"
using 4 5 7 8 relative_equality by blast
thus "z \<in> \<up>(t \<sqinter> v)"
by (simp add: inf.orderI)
qed
next
show "\<up>(t \<sqinter> v) \<subseteq> F"
proof
fix z
have 9: "t \<sqinter> v \<in> F"
using 2 3 by (simp add: assms(1) filter_inf_closed)
assume "z \<in> \<up>(t \<sqinter> v)"
hence "t \<sqinter> v \<le> z" by simp
thus "z \<in> F"
using assms(1) 9 filter_def by auto
qed
qed
thus ?thesis
by blast
qed
text \<open>
The following result is \cite[Lemma II]{GraetzerSchmidt1958}.
If both join and meet of two filters are principal filters, both filters are principal filters.
\<close>
lemma inf_sup_principal:
assumes "filter F"
and "filter G"
and "is_principal_up (filter_sup F G)"
and "is_principal_up (F \<inter> G)"
shows "is_principal_up F \<and> is_principal_up G"
proof -
have "filter G \<and> filter F \<and> is_principal_up (filter_sup G F) \<and> is_principal_up (G \<inter> F)"
by (simp add: assms Int_commute filter_sup_symmetric)
thus ?thesis
using assms(3) inf_sup_principal_aux by blast
qed
lemma filter_sup_absorb_inf: "filter F \<Longrightarrow> filter G \<Longrightarrow> filter_sup (F \<inter> G) G = G"
by (simp add: filter_inf filter_sup_least_upper_bound filter_sup_left_upper_bound filter_sup_symmetric subset_antisym)
lemma filter_inf_absorb_sup: "filter F \<Longrightarrow> filter G \<Longrightarrow> filter_sup F G \<inter> G = G"
apply (rule subset_antisym)
apply simp
by (simp add: filter_sup_right_upper_bound)
lemma filter_inf_right_dist_sup:
assumes "filter F"
and "filter G"
and "filter H"
shows "filter_sup F G \<inter> H = filter_sup (F \<inter> H) (G \<inter> H)"
proof -
have "filter_sup (F \<inter> H) (G \<inter> H) = filter_sup (F \<inter> H) G \<inter> filter_sup (F \<inter> H) H"
by (simp add: assms filter_sup_left_dist_inf filter_inf)
also have "... = filter_sup (F \<inter> H) G \<inter> H"
using assms(1,3) filter_sup_absorb_inf by simp
also have "... = filter_sup F G \<inter> filter_sup G H \<inter> H"
using assms filter_sup_left_dist_inf filter_sup_symmetric by simp
also have "... = filter_sup F G \<inter> H"
by (simp add: assms(2-3) filter_inf_absorb_sup semilattice_inf_class.inf_assoc)
finally show ?thesis
by simp
qed
text \<open>
The following result generalises \cite[10.11]{DaveyPriestley2002} to distributive lattices as remarked after that section.
\<close>
lemma ultra_filter_prime:
assumes "ultra_filter F"
shows "prime_filter F"
proof -
{
fix x y
assume 1: "x \<squnion> y \<in> F \<and> x \<notin> F"
let ?G = "\<Up>{ z . \<exists>w\<in>F . x \<sqinter> w = z }"
have 2: "filter ?G"
using assms filter_inf_filter by simp
have "x \<in> ?G"
using 1 by auto
hence 3: "F \<noteq> ?G"
using 1 by auto
have "F \<subseteq> ?G"
using inf_le2 order_trans by blast
hence "?G = UNIV"
using 2 3 assms by blast
then obtain z where 4: "z \<in> F \<and> x \<sqinter> z \<le> y"
by blast
hence "y \<sqinter> z = (x \<squnion> y) \<sqinter> z"
by (simp add: inf_sup_distrib2 sup_absorb2)
also have "... \<in> F"
using 1 4 assms filter_inf_closed by auto
finally have "y \<in> F"
using assms by (simp add: filter_def)
}
thus ?thesis
using assms by blast
qed
lemma up_dist_inf_inter:
assumes "is_up_set S"
shows "\<up>(x \<sqinter> y) \<inter> S = filter_sup (\<up>x \<inter> S) (\<up>y \<inter> S) \<inter> S"
proof
show "\<up>(x \<sqinter> y) \<inter> S \<subseteq> filter_sup (\<up>x \<inter> S) (\<up>y \<inter> S) \<inter> S"
proof
fix z
let ?x = "x \<squnion> z"
let ?y = "y \<squnion> z"
assume "z \<in> \<up>(x \<sqinter> y) \<inter> S"
hence 1: "x \<sqinter> y \<le> z \<and> z \<in> S"
by auto
hence "?x \<in> (\<up>x \<inter> S) \<and> ?y \<in> (\<up>y \<inter> S) \<and> ?x \<sqinter> ?y \<le> z"
using assms sup_absorb2 sup_inf_distrib2 by fastforce
thus "z \<in> filter_sup (\<up>x \<inter> S) (\<up>y \<inter> S) \<inter> S"
using filter_sup_def 1 by fastforce
qed
next
show "filter_sup (\<up>x \<inter> S) (\<up>y \<inter> S) \<inter> S \<subseteq> \<up>(x \<sqinter> y) \<inter> S"
proof
fix z
assume "z \<in> filter_sup (\<up>x \<inter> S) (\<up>y \<inter> S) \<inter> S"
then obtain u v where 2: "u\<in>\<up>x \<and> v\<in>\<up>y \<and> u \<sqinter> v \<le> z \<and> z \<in> S"
using filter_sup_def by auto
hence "x \<sqinter> y \<le> z"
using order.trans inf_mono by blast
thus "z \<in> \<up>(x \<sqinter> y) \<inter> S"
using 2 by blast
qed
qed
end
context distrib_lattice_bot
begin
lemma prime_filter:
"proper_filter F \<Longrightarrow> \<exists>G . prime_filter G \<and> F \<subseteq> G"
by (metis ultra_filter ultra_filter_prime)
end
context distrib_lattice_top
begin
lemma complemented_filter_inf_principal:
assumes "filter_complements F G"
shows "is_principal_up (F \<inter> \<up>x)"
proof -
have 1: "filter F \<and> filter G"
by (simp add: assms)
hence 2: "filter (F \<inter> \<up>x) \<and> filter (G \<inter> \<up>x)"
by (simp add: filter_inf)
have "(F \<inter> \<up>x) \<inter> (G \<inter> \<up>x) = {top}"
using assms Int_assoc Int_insert_left_if1 inf_bot_left inf_sup_aci(3) top_in_upset inf.idem by auto
hence 3: "is_principal_up ((F \<inter> \<up>x) \<inter> (G \<inter> \<up>x))"
using up_top by blast
have "filter_sup (F \<inter> \<up>x) (G \<inter> \<up>x) = filter_sup F G \<inter> \<up>x"
using 1 filter_inf_right_dist_sup up_filter by auto
also have "... = \<up>x"
by (simp add: assms)
finally have "is_principal_up (filter_sup (F \<inter> \<up>x) (G \<inter> \<up>x))"
by auto
thus ?thesis
using 1 2 3 inf_sup_principal_aux by blast
qed
end
text \<open>
The set of filters over a distributive lattice with a greatest element forms a bounded distributive lattice.
\<close>
instantiation filter :: (distrib_lattice_top) bounded_distrib_lattice
begin
instance
apply intro_classes
apply transfer
by (simp add: filter_sup_left_dist_inf)
end
end
diff --git a/thys/Stone_Algebras/Lattice_Basics.thy b/thys/Stone_Algebras/Lattice_Basics.thy
--- a/thys/Stone_Algebras/Lattice_Basics.thy
+++ b/thys/Stone_Algebras/Lattice_Basics.thy
@@ -1,558 +1,561 @@
(* Title: Lattice Basics
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Lattice Basics\<close>
text \<open>
This theory provides notations, basic definitions and facts of lattice-related structures used throughout the subsequent development.
\<close>
theory Lattice_Basics
imports Main
begin
subsection \<open>General Facts and Notations\<close>
text \<open>
The following results extend basic Isabelle/HOL facts.
\<close>
lemma imp_as_conj:
assumes "P x \<Longrightarrow> Q x"
shows "P x \<and> Q x \<longleftrightarrow> P x"
using assms by auto
lemma if_distrib_2:
"f (if c then x else y) (if c then z else w) = (if c then f x z else f y w)"
by simp
lemma left_invertible_inj:
"(\<forall>x . g (f x) = x) \<Longrightarrow> inj f"
by (metis injI)
lemma invertible_bij:
assumes "\<forall>x . g (f x) = x"
and "\<forall>y . f (g y) = y"
shows "bij f"
by (metis assms bijI')
lemma finite_ne_subset_induct [consumes 3, case_names singleton insert]:
assumes "finite F"
and "F \<noteq> {}"
and "F \<subseteq> S"
and singleton: "\<And>x . P {x}"
and insert: "\<And>x F . finite F \<Longrightarrow> F \<noteq> {} \<Longrightarrow> F \<subseteq> S \<Longrightarrow> x \<in> S \<Longrightarrow> x \<notin> F \<Longrightarrow> P F \<Longrightarrow> P (insert x F)"
shows "P F"
using assms(1-3)
apply (induct rule: finite_ne_induct)
apply (simp add: singleton)
by (simp add: insert)
lemma finite_set_of_finite_funs_pred:
assumes "finite { x::'a . True }"
and "finite { y::'b . P y }"
shows "finite { f . (\<forall>x::'a . P (f x)) }"
using assms finite_set_of_finite_funs by force
text \<open>
We use the following notations for the join, meet and complement operations.
Changing the precedence of the unary complement allows us to write terms like \<open>--x\<close> instead of \<open>-(-x)\<close>.
\<close>
context sup
begin
notation sup (infixl "\<squnion>" 65)
definition additive :: "('a \<Rightarrow> 'a) \<Rightarrow> bool"
where "additive f \<equiv> \<forall>x y . f (x \<squnion> y) = f x \<squnion> f y"
end
context inf
begin
notation inf (infixl "\<sqinter>" 67)
end
context uminus
begin
no_notation uminus ("- _" [81] 80)
notation uminus ("- _" [80] 80)
end
subsection \<open>Orders\<close>
text \<open>
We use the following definition of monotonicity for operations defined in classes.
The standard \<open>mono\<close> places a sort constraint on the target type.
We also give basic properties of Galois connections and lift orders to functions.
\<close>
context ord
begin
definition isotone :: "('a \<Rightarrow> 'a) \<Rightarrow> bool"
where "isotone f \<equiv> \<forall>x y . x \<le> y \<longrightarrow> f x \<le> f y"
definition galois :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
where "galois l u \<equiv> \<forall>x y . l x \<le> y \<longleftrightarrow> x \<le> u y"
definition lifted_less_eq :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" ("(_ \<le>\<le> _)" [51, 51] 50)
where "f \<le>\<le> g \<equiv> \<forall>x . f x \<le> g x"
end
context order
begin
lemma order_lesseq_imp:
"(\<forall>z . x \<le> z \<longrightarrow> y \<le> z) \<longleftrightarrow> y \<le> x"
using order_trans by blast
lemma galois_char:
"galois l u \<longleftrightarrow> (\<forall>x . x \<le> u (l x)) \<and> (\<forall>x . l (u x) \<le> x) \<and> isotone l \<and> isotone u"
apply (rule iffI)
apply (metis (full_types) galois_def isotone_def order_refl order_trans)
using galois_def isotone_def order_trans by blast
lemma galois_closure:
"galois l u \<Longrightarrow> l x = l (u (l x)) \<and> u x = u (l (u x))"
- by (simp add: galois_char isotone_def antisym)
+ by (simp add: galois_char isotone_def order.antisym)
lemma lifted_reflexive:
"f = g \<Longrightarrow> f \<le>\<le> g"
by (simp add: lifted_less_eq_def)
lemma lifted_transitive:
"f \<le>\<le> g \<Longrightarrow> g \<le>\<le> h \<Longrightarrow> f \<le>\<le> h"
using lifted_less_eq_def order_trans by blast
lemma lifted_antisymmetric:
"f \<le>\<le> g \<Longrightarrow> g \<le>\<le> f \<Longrightarrow> f = g"
- by (metis (full_types) antisym ext lifted_less_eq_def)
+ by (rule ext, rule order.antisym) (simp_all add: lifted_less_eq_def)
text \<open>
If the image of a finite non-empty set under \<open>f\<close> is a totally ordered, there is an element that minimises the value of \<open>f\<close>.
\<close>
lemma finite_set_minimal:
assumes "finite s"
and "s \<noteq> {}"
and "\<forall>x\<in>s . \<forall>y\<in>s . f x \<le> f y \<or> f y \<le> f x"
shows "\<exists>m\<in>s . \<forall>z\<in>s . f m \<le> f z"
apply (rule finite_ne_subset_induct[where S=s])
apply (rule assms(1))
apply (rule assms(2))
apply simp
apply simp
by (metis assms(3) insert_iff order_trans subsetD)
end
subsection \<open>Semilattices\<close>
text \<open>
The following are basic facts in semilattices.
\<close>
context semilattice_sup
begin
lemma sup_left_isotone:
"x \<le> y \<Longrightarrow> x \<squnion> z \<le> y \<squnion> z"
using sup.mono by blast
lemma sup_right_isotone:
"x \<le> y \<Longrightarrow> z \<squnion> x \<le> z \<squnion> y"
using sup.mono by blast
lemma sup_left_divisibility:
"x \<le> y \<longleftrightarrow> (\<exists>z . x \<squnion> z = y)"
using sup.absorb2 sup.cobounded1 by blast
lemma sup_right_divisibility:
"x \<le> y \<longleftrightarrow> (\<exists>z . z \<squnion> x = y)"
by (metis sup.cobounded2 sup.orderE)
lemma sup_same_context:
"x \<le> y \<squnion> z \<Longrightarrow> y \<le> x \<squnion> z \<Longrightarrow> x \<squnion> z = y \<squnion> z"
by (simp add: le_iff_sup sup_left_commute)
lemma sup_relative_same_increasing:
"x \<le> y \<Longrightarrow> x \<squnion> z = x \<squnion> w \<Longrightarrow> y \<squnion> z = y \<squnion> w"
using sup.assoc sup_right_divisibility by auto
end
text \<open>
Every bounded semilattice is a commutative monoid.
Finite sums defined in commutative monoids are available via the following sublocale.
\<close>
context bounded_semilattice_sup_bot
begin
sublocale sup_monoid: comm_monoid_add where plus = sup and zero = bot
apply unfold_locales
apply (simp add: sup_assoc)
apply (simp add: sup_commute)
by simp
end
context semilattice_inf
begin
lemma inf_same_context:
"x \<le> y \<sqinter> z \<Longrightarrow> y \<le> x \<sqinter> z \<Longrightarrow> x \<sqinter> z = y \<sqinter> z"
- using antisym by auto
+ using order.antisym by auto
end
text \<open>
The following class requires only the existence of upper bounds, which is a property common to bounded semilattices and (not necessarily bounded) lattices.
We use it in our development of filters.
\<close>
class directed_semilattice_inf = semilattice_inf +
assumes ub: "\<exists>z . x \<le> z \<and> y \<le> z"
text \<open>
We extend the \<open>inf\<close> sublocale, which dualises the order in semilattices, to bounded semilattices.
\<close>
context bounded_semilattice_inf_top
begin
subclass directed_semilattice_inf
apply unfold_locales
using top_greatest by blast
sublocale inf: bounded_semilattice_sup_bot where sup = inf and less_eq = greater_eq and less = greater and bot = top
by unfold_locales (simp_all add: less_le_not_le)
end
subsection \<open>Lattices\<close>
context lattice
begin
subclass directed_semilattice_inf
apply unfold_locales
using sup_ge1 sup_ge2 by blast
definition dual_additive :: "('a \<Rightarrow> 'a) \<Rightarrow> bool"
where "dual_additive f \<equiv> \<forall>x y . f (x \<squnion> y) = f x \<sqinter> f y"
end
text \<open>
Not every bounded lattice has complements, but two elements might still be complements of each other as captured in the following definition.
In this situation we can apply, for example, the shunting property shown below.
We introduce most definitions using the \<open>abbreviation\<close> command.
\<close>
context bounded_lattice
begin
abbreviation "complement x y \<equiv> x \<squnion> y = top \<and> x \<sqinter> y = bot"
lemma complement_symmetric:
"complement x y \<Longrightarrow> complement y x"
by (simp add: inf.commute sup.commute)
definition conjugate :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
where "conjugate f g \<equiv> \<forall>x y . f x \<sqinter> y = bot \<longleftrightarrow> x \<sqinter> g y = bot"
end
class dense_lattice = bounded_lattice +
assumes bot_meet_irreducible: "x \<sqinter> y = bot \<longrightarrow> x = bot \<or> y = bot"
context distrib_lattice
begin
lemma relative_equality:
"x \<squnion> z = y \<squnion> z \<Longrightarrow> x \<sqinter> z = y \<sqinter> z \<Longrightarrow> x = y"
by (metis inf.commute inf_sup_absorb inf_sup_distrib2)
end
text \<open>
Distributive lattices with a greatest element are widely used in the construction theorem for Stone algebras.
\<close>
class distrib_lattice_bot = bounded_lattice_bot + distrib_lattice
class distrib_lattice_top = bounded_lattice_top + distrib_lattice
class bounded_distrib_lattice = bounded_lattice + distrib_lattice
begin
subclass distrib_lattice_bot ..
subclass distrib_lattice_top ..
lemma complement_shunting:
assumes "complement z w"
shows "z \<sqinter> x \<le> y \<longleftrightarrow> x \<le> w \<squnion> y"
proof
assume 1: "z \<sqinter> x \<le> y"
have "x = (z \<squnion> w) \<sqinter> x"
by (simp add: assms)
also have "... \<le> y \<squnion> (w \<sqinter> x)"
using 1 sup.commute sup.left_commute inf_sup_distrib2 sup_right_divisibility by fastforce
also have "... \<le> w \<squnion> y"
by (simp add: inf.coboundedI1)
finally show "x \<le> w \<squnion> y"
.
next
assume "x \<le> w \<squnion> y"
hence "z \<sqinter> x \<le> z \<sqinter> (w \<squnion> y)"
using inf.sup_right_isotone by auto
also have "... = z \<sqinter> y"
by (simp add: assms inf_sup_distrib1)
also have "... \<le> y"
by simp
finally show "z \<sqinter> x \<le> y"
.
qed
end
subsection \<open>Linear Orders\<close>
text \<open>
We next consider lattices with a linear order structure.
In such lattices, join and meet are selective operations, which give the maximum and the minimum of two elements, respectively.
Moreover, the lattice is automatically distributive.
\<close>
class bounded_linorder = linorder + order_bot + order_top
class linear_lattice = lattice + linorder
begin
lemma max_sup:
"max x y = x \<squnion> y"
by (metis max.boundedI max.cobounded1 max.cobounded2 sup_unique)
lemma min_inf:
"min x y = x \<sqinter> y"
by (simp add: inf.absorb1 inf.absorb2 min_def)
lemma sup_inf_selective:
"(x \<squnion> y = x \<and> x \<sqinter> y = y) \<or> (x \<squnion> y = y \<and> x \<sqinter> y = x)"
by (meson inf.absorb1 inf.absorb2 le_cases sup.absorb1 sup.absorb2)
lemma sup_selective:
"x \<squnion> y = x \<or> x \<squnion> y = y"
using sup_inf_selective by blast
lemma inf_selective:
"x \<sqinter> y = x \<or> x \<sqinter> y = y"
using sup_inf_selective by blast
subclass distrib_lattice
- apply unfold_locales
- by (metis inf_selective antisym distrib_sup_le inf.commute inf_le2)
+ apply standard
+ apply (rule order.antisym)
+ apply (auto simp add: le_supI2)
+ apply (metis inf_selective inf.coboundedI1 inf.coboundedI2 order.eq_iff)
+ done
lemma sup_less_eq:
"x \<le> y \<squnion> z \<longleftrightarrow> x \<le> y \<or> x \<le> z"
by (metis le_supI1 le_supI2 sup_selective)
lemma inf_less_eq:
"x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> z \<or> y \<le> z"
by (metis inf.coboundedI1 inf.coboundedI2 inf_selective)
lemma sup_inf_sup:
"x \<squnion> y = (x \<squnion> y) \<squnion> (x \<sqinter> y)"
by (metis sup_commute sup_inf_absorb sup_left_commute)
end
text \<open>
The following class derives additional properties if the linear order of the lattice has a least and a greatest element.
\<close>
class linear_bounded_lattice = bounded_lattice + linorder
begin
subclass linear_lattice ..
subclass bounded_linorder ..
subclass bounded_distrib_lattice ..
lemma sup_dense:
"x \<noteq> top \<Longrightarrow> y \<noteq> top \<Longrightarrow> x \<squnion> y \<noteq> top"
by (metis sup_selective)
lemma inf_dense:
"x \<noteq> bot \<Longrightarrow> y \<noteq> bot \<Longrightarrow> x \<sqinter> y \<noteq> bot"
by (metis inf_selective)
lemma sup_not_bot:
"x \<noteq> bot \<Longrightarrow> x \<squnion> y \<noteq> bot"
by simp
lemma inf_not_top:
"x \<noteq> top \<Longrightarrow> x \<sqinter> y \<noteq> top"
by simp
subclass dense_lattice
apply unfold_locales
using inf_dense by blast
end
text \<open>
Every bounded linear order can be expanded to a bounded lattice.
Join and meet are maximum and minimum, respectively.
\<close>
class linorder_lattice_expansion = bounded_linorder + sup + inf +
assumes sup_def [simp]: "x \<squnion> y = max x y"
assumes inf_def [simp]: "x \<sqinter> y = min x y"
begin
subclass linear_bounded_lattice
apply unfold_locales
by auto
end
subsection \<open>Non-trivial Algebras\<close>
text \<open>
Some results, such as the existence of certain filters, require that the algebras are not trivial.
This is not an assumption of the order and lattice classes that come with Isabelle/HOL; for example, \<open>bot = top\<close> may hold in bounded lattices.
\<close>
class non_trivial =
assumes consistent: "\<exists>x y . x \<noteq> y"
class non_trivial_order = non_trivial + order
class non_trivial_order_bot = non_trivial_order + order_bot
class non_trivial_bounded_order = non_trivial_order_bot + order_top
begin
lemma bot_not_top:
"bot \<noteq> top"
proof -
from consistent obtain x y :: 'a where "x \<noteq> y"
by auto
thus ?thesis
by (metis bot_less top.extremum_strict)
qed
end
subsection \<open>Homomorphisms\<close>
text \<open>
This section gives definitions of lattice homomorphisms and isomorphisms and basic properties.
\<close>
class sup_inf_top_bot_uminus = sup + inf + top + bot + uminus
class sup_inf_top_bot_uminus_ord = sup_inf_top_bot_uminus + ord
context boolean_algebra
begin
subclass sup_inf_top_bot_uminus_ord .
end
abbreviation sup_homomorphism :: "('a::sup \<Rightarrow> 'b::sup) \<Rightarrow> bool"
where "sup_homomorphism f \<equiv> \<forall>x y . f (x \<squnion> y) = f x \<squnion> f y"
abbreviation inf_homomorphism :: "('a::inf \<Rightarrow> 'b::inf) \<Rightarrow> bool"
where "inf_homomorphism f \<equiv> \<forall>x y . f (x \<sqinter> y) = f x \<sqinter> f y"
abbreviation bot_homomorphism :: "('a::bot \<Rightarrow> 'b::bot) \<Rightarrow> bool"
where "bot_homomorphism f \<equiv> f bot = bot"
abbreviation top_homomorphism :: "('a::top \<Rightarrow> 'b::top) \<Rightarrow> bool"
where "top_homomorphism f \<equiv> f top = top"
abbreviation minus_homomorphism :: "('a::minus \<Rightarrow> 'b::minus) \<Rightarrow> bool"
where "minus_homomorphism f \<equiv> \<forall>x y . f (x - y) = f x - f y"
abbreviation uminus_homomorphism :: "('a::uminus \<Rightarrow> 'b::uminus) \<Rightarrow> bool"
where "uminus_homomorphism f \<equiv> \<forall>x . f (-x) = -f x"
abbreviation sup_inf_homomorphism :: "('a::{sup,inf} \<Rightarrow> 'b::{sup,inf}) \<Rightarrow> bool"
where "sup_inf_homomorphism f \<equiv> sup_homomorphism f \<and> inf_homomorphism f"
abbreviation sup_inf_top_homomorphism :: "('a::{sup,inf,top} \<Rightarrow> 'b::{sup,inf,top}) \<Rightarrow> bool"
where "sup_inf_top_homomorphism f \<equiv> sup_inf_homomorphism f \<and> top_homomorphism f"
abbreviation sup_inf_top_bot_homomorphism :: "('a::{sup,inf,top,bot} \<Rightarrow> 'b::{sup,inf,top,bot}) \<Rightarrow> bool"
where "sup_inf_top_bot_homomorphism f \<equiv> sup_inf_top_homomorphism f \<and> bot_homomorphism f"
abbreviation bounded_lattice_homomorphism :: "('a::bounded_lattice \<Rightarrow> 'b::bounded_lattice) \<Rightarrow> bool"
where "bounded_lattice_homomorphism f \<equiv> sup_inf_top_bot_homomorphism f"
abbreviation sup_inf_top_bot_uminus_homomorphism :: "('a::sup_inf_top_bot_uminus \<Rightarrow> 'b::sup_inf_top_bot_uminus) \<Rightarrow> bool"
where "sup_inf_top_bot_uminus_homomorphism f \<equiv> sup_inf_top_bot_homomorphism f \<and> uminus_homomorphism f"
abbreviation sup_inf_top_bot_uminus_ord_homomorphism :: "('a::sup_inf_top_bot_uminus_ord \<Rightarrow> 'b::sup_inf_top_bot_uminus_ord) \<Rightarrow> bool"
where "sup_inf_top_bot_uminus_ord_homomorphism f \<equiv> sup_inf_top_bot_uminus_homomorphism f \<and> (\<forall>x y . x \<le> y \<longrightarrow> f x \<le> f y)"
abbreviation sup_inf_top_isomorphism :: "('a::{sup,inf,top} \<Rightarrow> 'b::{sup,inf,top}) \<Rightarrow> bool"
where "sup_inf_top_isomorphism f \<equiv> sup_inf_top_homomorphism f \<and> bij f"
abbreviation bounded_lattice_top_isomorphism :: "('a::bounded_lattice_top \<Rightarrow> 'b::bounded_lattice_top) \<Rightarrow> bool"
where "bounded_lattice_top_isomorphism f \<equiv> sup_inf_top_isomorphism f"
abbreviation sup_inf_top_bot_uminus_isomorphism :: "('a::sup_inf_top_bot_uminus \<Rightarrow> 'b::sup_inf_top_bot_uminus) \<Rightarrow> bool"
where "sup_inf_top_bot_uminus_isomorphism f \<equiv> sup_inf_top_bot_uminus_homomorphism f \<and> bij f"
abbreviation boolean_algebra_isomorphism :: "('a::boolean_algebra \<Rightarrow> 'b::boolean_algebra) \<Rightarrow> bool"
where "boolean_algebra_isomorphism f \<equiv> sup_inf_top_bot_uminus_isomorphism f \<and> minus_homomorphism f"
lemma sup_homomorphism_mono:
"sup_homomorphism (f::'a::semilattice_sup \<Rightarrow> 'b::semilattice_sup) \<Longrightarrow> mono f"
by (metis le_iff_sup monoI)
lemma sup_isomorphism_ord_isomorphism:
assumes "sup_homomorphism (f::'a::semilattice_sup \<Rightarrow> 'b::semilattice_sup)"
and "bij f"
shows "x \<le> y \<longleftrightarrow> f x \<le> f y"
proof
assume "x \<le> y"
thus "f x \<le> f y"
by (metis assms(1) le_iff_sup)
next
assume "f x \<le> f y"
hence "f (x \<squnion> y) = f y"
by (simp add: assms(1) le_iff_sup)
hence "x \<squnion> y = y"
by (metis injD bij_is_inj assms(2))
thus "x \<le> y"
by (simp add: le_iff_sup)
qed
lemma minus_homomorphism_default:
assumes "\<forall>x y::'a::{inf,minus,uminus} . x - y = x \<sqinter> -y"
and "\<forall>x y::'b::{inf,minus,uminus} . x - y = x \<sqinter> -y"
and "inf_homomorphism (f::'a \<Rightarrow> 'b)"
and "uminus_homomorphism f"
shows "minus_homomorphism f"
by (simp add: assms)
end
diff --git a/thys/Stone_Algebras/P_Algebras.thy b/thys/Stone_Algebras/P_Algebras.thy
--- a/thys/Stone_Algebras/P_Algebras.thy
+++ b/thys/Stone_Algebras/P_Algebras.thy
@@ -1,1324 +1,1324 @@
(* Title: Pseudocomplemented Algebras
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Pseudocomplemented Algebras\<close>
text \<open>
This theory expands lattices with a pseudocomplement operation.
In particular, we consider the following algebraic structures:
\begin{itemize}
\item pseudocomplemented lattices (p-algebras)
\item pseudocomplemented distributive lattices (distributive p-algebras)
\item Stone algebras
\item Heyting semilattices
\item Heyting lattices
\item Heyting algebras
\item Heyting-Stone algebras
\item Brouwer algebras
\item Boolean algebras
\end{itemize}
Most of these structures and many results in this theory are discussed in \cite{BalbesDwinger1974,Birkhoff1967,Blyth2005,Curry1977,Graetzer1971,Maddux1996}.
\<close>
theory P_Algebras
imports Lattice_Basics
begin
subsection \<open>P-Algebras\<close>
text \<open>
In this section we add a pseudocomplement operation to lattices and to distributive lattices.
\<close>
subsubsection \<open>Pseudocomplemented Lattices\<close>
text \<open>
The pseudocomplement of an element \<open>y\<close> is the greatest element whose meet with \<open>y\<close> is the least element of the lattice.
\<close>
class p_algebra = bounded_lattice + uminus +
assumes pseudo_complement: "x \<sqinter> y = bot \<longleftrightarrow> x \<le> -y"
begin
subclass sup_inf_top_bot_uminus_ord .
text \<open>
Regular elements and dense elements are frequently used in pseudocomplemented algebras.
\<close>
abbreviation "regular x \<equiv> x = --x"
abbreviation "dense x \<equiv> -x = bot"
abbreviation "complemented x \<equiv> \<exists>y . x \<sqinter> y = bot \<and> x \<squnion> y = top"
abbreviation "in_p_image x \<equiv> \<exists>y . x = -y"
abbreviation "selection s x \<equiv> s = --s \<sqinter> x"
abbreviation "dense_elements \<equiv> { x . dense x }"
abbreviation "regular_elements \<equiv> { x . in_p_image x }"
lemma p_bot [simp]:
"-bot = top"
using inf_top.left_neutral pseudo_complement top_unique by blast
lemma p_top [simp]:
"-top = bot"
by (metis eq_refl inf_top.comm_neutral pseudo_complement)
text \<open>
The pseudocomplement satisfies the following half of the requirements of a complement.
\<close>
lemma inf_p [simp]:
"x \<sqinter> -x = bot"
using inf.commute pseudo_complement by fastforce
lemma p_inf [simp]:
"-x \<sqinter> x = bot"
by (simp add: inf_commute)
lemma pp_inf_p:
"--x \<sqinter> -x = bot"
by simp
text \<open>
The double complement is a closure operation.
\<close>
lemma pp_increasing:
"x \<le> --x"
using inf_p pseudo_complement by blast
lemma ppp [simp]:
"---x = -x"
- by (metis antisym inf.commute order_trans pseudo_complement pp_increasing)
+ by (metis order.antisym inf.commute order_trans pseudo_complement pp_increasing)
lemma pp_idempotent:
"----x = --x"
by simp
lemma regular_in_p_image_iff:
"regular x \<longleftrightarrow> in_p_image x"
by auto
lemma pseudo_complement_pp:
"x \<sqinter> y = bot \<longleftrightarrow> --x \<le> -y"
by (metis inf_commute pseudo_complement ppp)
lemma p_antitone:
"x \<le> y \<Longrightarrow> -y \<le> -x"
by (metis inf_commute order_trans pseudo_complement pp_increasing)
lemma p_antitone_sup:
"-(x \<squnion> y) \<le> -x"
by (simp add: p_antitone)
lemma p_antitone_inf:
"-x \<le> -(x \<sqinter> y)"
by (simp add: p_antitone)
lemma p_antitone_iff:
"x \<le> -y \<longleftrightarrow> y \<le> -x"
using order_lesseq_imp p_antitone pp_increasing by blast
lemma pp_isotone:
"x \<le> y \<Longrightarrow> --x \<le> --y"
by (simp add: p_antitone)
lemma pp_isotone_sup:
"--x \<le> --(x \<squnion> y)"
by (simp add: p_antitone)
lemma pp_isotone_inf:
"--(x \<sqinter> y) \<le> --x"
by (simp add: p_antitone)
text \<open>
One of De Morgan's laws holds in pseudocomplemented lattices.
\<close>
lemma p_dist_sup [simp]:
"-(x \<squnion> y) = -x \<sqinter> -y"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: p_antitone)
using inf_le1 inf_le2 le_sup_iff p_antitone_iff by blast
lemma p_supdist_inf:
"-x \<squnion> -y \<le> -(x \<sqinter> y)"
by (simp add: p_antitone)
lemma pp_dist_pp_sup [simp]:
"--(--x \<squnion> --y) = --(x \<squnion> y)"
by simp
lemma p_sup_p [simp]:
"-(x \<squnion> -x) = bot"
by simp
lemma pp_sup_p [simp]:
"--(x \<squnion> -x) = top"
by simp
lemma dense_pp:
"dense x \<longleftrightarrow> --x = top"
by (metis p_bot p_top ppp)
lemma dense_sup_p:
"dense (x \<squnion> -x)"
by simp
lemma regular_char:
"regular x \<longleftrightarrow> (\<exists>y . x = -y)"
by auto
lemma pp_inf_bot_iff:
"x \<sqinter> y = bot \<longleftrightarrow> --x \<sqinter> y = bot"
by (simp add: pseudo_complement_pp)
text \<open>
Weak forms of the shunting property hold.
Most require a pseudocomplemented element on the right-hand side.
\<close>
lemma p_shunting_swap:
"x \<sqinter> y \<le> -z \<longleftrightarrow> x \<sqinter> z \<le> -y"
by (metis inf_assoc inf_commute pseudo_complement)
lemma pp_inf_below_iff:
"x \<sqinter> y \<le> -z \<longleftrightarrow> --x \<sqinter> y \<le> -z"
by (simp add: inf_commute p_shunting_swap)
lemma p_inf_pp [simp]:
"-(x \<sqinter> --y) = -(x \<sqinter> y)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: inf.coboundedI2 p_antitone pp_increasing)
using inf_commute p_antitone_iff pp_inf_below_iff by auto
lemma p_inf_pp_pp [simp]:
"-(--x \<sqinter> --y) = -(x \<sqinter> y)"
by (simp add: inf_commute)
lemma regular_closed_inf:
"regular x \<Longrightarrow> regular y \<Longrightarrow> regular (x \<sqinter> y)"
by (metis p_dist_sup ppp)
lemma regular_closed_p:
"regular (-x)"
by simp
lemma regular_closed_pp:
"regular (--x)"
by simp
lemma regular_closed_bot:
"regular bot"
by simp
lemma regular_closed_top:
"regular top"
by simp
lemma pp_dist_inf [simp]:
"--(x \<sqinter> y) = --x \<sqinter> --y"
by (metis p_dist_sup p_inf_pp_pp ppp)
lemma inf_import_p [simp]:
"x \<sqinter> -(x \<sqinter> y) = x \<sqinter> -y"
- apply (rule antisym)
+ apply (rule order.antisym)
using p_shunting_swap apply fastforce
using inf.sup_right_isotone p_antitone by auto
text \<open>
Pseudocomplements are unique.
\<close>
lemma p_unique:
"(\<forall>x . x \<sqinter> y = bot \<longleftrightarrow> x \<le> z) \<Longrightarrow> z = -y"
- using inf.eq_iff pseudo_complement by auto
+ using inf.order_eq_iff pseudo_complement by auto
lemma maddux_3_5:
"x \<squnion> x = x \<squnion> -(y \<squnion> -y)"
by simp
lemma shunting_1_pp:
"x \<le> --y \<longleftrightarrow> x \<sqinter> -y = bot"
by (simp add: pseudo_complement)
lemma pp_pp_inf_bot_iff:
"x \<sqinter> y = bot \<longleftrightarrow> --x \<sqinter> --y = bot"
by (simp add: pseudo_complement_pp)
lemma inf_pp_semi_commute:
"x \<sqinter> --y \<le> --(x \<sqinter> y)"
using inf.eq_refl p_antitone_iff p_inf_pp by presburger
lemma inf_pp_commute:
"--(--x \<sqinter> y) = --x \<sqinter> --y"
by simp
lemma sup_pp_semi_commute:
"x \<squnion> --y \<le> --(x \<squnion> y)"
by (simp add: p_antitone_iff)
lemma regular_sup:
"regular z \<Longrightarrow> (x \<le> z \<and> y \<le> z \<longleftrightarrow> --(x \<squnion> y) \<le> z)"
apply (rule iffI)
apply (metis le_supI pp_isotone)
using dual_order.trans sup_ge2 pp_increasing pp_isotone_sup by blast
lemma dense_closed_inf:
"dense x \<Longrightarrow> dense y \<Longrightarrow> dense (x \<sqinter> y)"
by (simp add: dense_pp)
lemma dense_closed_sup:
"dense x \<Longrightarrow> dense y \<Longrightarrow> dense (x \<squnion> y)"
by simp
lemma dense_closed_pp:
"dense x \<Longrightarrow> dense (--x)"
by simp
lemma dense_closed_top:
"dense top"
by simp
lemma dense_up_closed:
"dense x \<Longrightarrow> x \<le> y \<Longrightarrow> dense y"
using dense_pp top_le pp_isotone by auto
lemma regular_dense_top:
"regular x \<Longrightarrow> dense x \<Longrightarrow> x = top"
using p_bot by blast
lemma selection_char:
"selection s x \<longleftrightarrow> (\<exists>y . s = -y \<sqinter> x)"
by (metis inf_import_p inf_commute regular_closed_p)
lemma selection_closed_inf:
"selection s x \<Longrightarrow> selection t x \<Longrightarrow> selection (s \<sqinter> t) x"
by (metis inf_assoc inf_commute inf_idem pp_dist_inf)
lemma selection_closed_pp:
"regular x \<Longrightarrow> selection s x \<Longrightarrow> selection (--s) x"
by (metis pp_dist_inf)
lemma selection_closed_bot:
"selection bot x"
by simp
lemma selection_closed_id:
"selection x x"
using inf.le_iff_sup pp_increasing by auto
text \<open>
Conjugates are usually studied for Boolean algebras, however, some of their properties generalise to pseudocomplemented algebras.
\<close>
lemma conjugate_unique_p:
assumes "conjugate f g"
and "conjugate f h"
shows "uminus \<circ> g = uminus \<circ> h"
proof -
have "\<forall>x y . x \<sqinter> g y = bot \<longleftrightarrow> x \<sqinter> h y = bot"
using assms conjugate_def inf.commute by simp
hence "\<forall>x y . x \<le> -(g y) \<longleftrightarrow> x \<le> -(h y)"
using inf.commute pseudo_complement by simp
hence "\<forall>y . -(g y) = -(h y)"
- using eq_iff by blast
+ using order.eq_iff by blast
thus ?thesis
by auto
qed
lemma conjugate_symmetric:
"conjugate f g \<Longrightarrow> conjugate g f"
by (simp add: conjugate_def inf_commute)
lemma additive_isotone:
"additive f \<Longrightarrow> isotone f"
by (metis additive_def isotone_def le_iff_sup)
lemma dual_additive_antitone:
assumes "dual_additive f"
shows "isotone (uminus \<circ> f)"
proof -
have "\<forall>x y . f (x \<squnion> y) \<le> f x"
using assms dual_additive_def by simp
hence "\<forall>x y . x \<le> y \<longrightarrow> f y \<le> f x"
by (metis sup_absorb2)
hence "\<forall>x y . x \<le> y \<longrightarrow> -(f x) \<le> -(f y)"
by (simp add: p_antitone)
thus ?thesis
by (simp add: isotone_def)
qed
lemma conjugate_dual_additive:
assumes "conjugate f g"
shows "dual_additive (uminus \<circ> f)"
proof -
have 1: "\<forall>x y z . -z \<le> -(f (x \<squnion> y)) \<longleftrightarrow> -z \<le> -(f x) \<and> -z \<le> -(f y)"
proof (intro allI)
fix x y z
have "(-z \<le> -(f (x \<squnion> y))) = (f (x \<squnion> y) \<sqinter> -z = bot)"
by (simp add: p_antitone_iff pseudo_complement)
also have "... = ((x \<squnion> y) \<sqinter> g(-z) = bot)"
using assms conjugate_def by auto
also have "... = (x \<squnion> y \<le> -(g(-z)))"
by (simp add: pseudo_complement)
also have "... = (x \<le> -(g(-z)) \<and> y \<le> -(g(-z)))"
by (simp add: le_sup_iff)
also have "... = (x \<sqinter> g(-z) = bot \<and> y \<sqinter> g(-z) = bot)"
by (simp add: pseudo_complement)
also have "... = (f x \<sqinter> -z = bot \<and> f y \<sqinter> -z = bot)"
using assms conjugate_def by auto
also have "... = (-z \<le> -(f x) \<and> -z \<le> -(f y))"
by (simp add: p_antitone_iff pseudo_complement)
finally show "-z \<le> -(f (x \<squnion> y)) \<longleftrightarrow> -z \<le> -(f x) \<and> -z \<le> -(f y)"
by simp
qed
have "\<forall>x y . -(f (x \<squnion> y)) = -(f x) \<sqinter> -(f y)"
proof (intro allI)
fix x y
have "-(f x) \<sqinter> -(f y) = --(-(f x) \<sqinter> -(f y))"
by simp
hence "-(f x) \<sqinter> -(f y) \<le> -(f (x \<squnion> y))"
using 1 by (metis inf_le1 inf_le2)
thus "-(f (x \<squnion> y)) = -(f x) \<sqinter> -(f y)"
- using 1 antisym by fastforce
+ using 1 order.antisym by fastforce
qed
thus ?thesis
using dual_additive_def by simp
qed
lemma conjugate_isotone_pp:
"conjugate f g \<Longrightarrow> isotone (uminus \<circ> uminus \<circ> f)"
by (simp add: comp_assoc conjugate_dual_additive dual_additive_antitone)
lemma conjugate_char_1_pp:
"conjugate f g \<longleftrightarrow> (\<forall>x y . f(x \<sqinter> -(g y)) \<le> --f x \<sqinter> -y \<and> g(y \<sqinter> -(f x)) \<le> --g y \<sqinter> -x)"
proof
assume 1: "conjugate f g"
show "\<forall>x y . f(x \<sqinter> -(g y)) \<le> --f x \<sqinter> -y \<and> g(y \<sqinter> -(f x)) \<le> --g y \<sqinter> -x"
proof (intro allI)
fix x y
have 2: "f(x \<sqinter> -(g y)) \<le> -y"
using 1 by (simp add: conjugate_def pseudo_complement)
have "f(x \<sqinter> -(g y)) \<le> --f(x \<sqinter> -(g y))"
by (simp add: pp_increasing)
also have "... \<le> --f x"
using 1 conjugate_isotone_pp isotone_def by simp
finally have 3: "f(x \<sqinter> -(g y)) \<le> --f x \<sqinter> -y"
using 2 by simp
have 4: "isotone (uminus \<circ> uminus \<circ> g)"
using 1 conjugate_isotone_pp conjugate_symmetric by auto
have 5: "g(y \<sqinter> -(f x)) \<le> -x"
using 1 by (metis conjugate_def inf.cobounded2 inf_commute pseudo_complement)
have "g(y \<sqinter> -(f x)) \<le> --g(y \<sqinter> -(f x))"
by (simp add: pp_increasing)
also have "... \<le> --g y"
using 4 isotone_def by auto
finally have "g(y \<sqinter> -(f x)) \<le> --g y \<sqinter> -x"
using 5 by simp
thus "f(x \<sqinter> -(g y)) \<le> --f x \<sqinter> -y \<and> g(y \<sqinter> -(f x)) \<le> --g y \<sqinter> -x"
using 3 by simp
qed
next
assume 6: "\<forall>x y . f(x \<sqinter> -(g y)) \<le> --f x \<sqinter> -y \<and> g(y \<sqinter> -(f x)) \<le> --g y \<sqinter> -x"
hence 7: "\<forall>x y . f x \<sqinter> y = bot \<longrightarrow> x \<sqinter> g y = bot"
by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement)
have "\<forall>x y . x \<sqinter> g y = bot \<longrightarrow> f x \<sqinter> y = bot"
using 6 by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement)
thus "conjugate f g"
using 7 conjugate_def by auto
qed
lemma conjugate_char_1_isotone:
"conjugate f g \<Longrightarrow> isotone f \<Longrightarrow> isotone g \<Longrightarrow> f(x \<sqinter> -(g y)) \<le> f x \<sqinter> -y \<and> g(y \<sqinter> -(f x)) \<le> g y \<sqinter> -x"
by (simp add: conjugate_char_1_pp ord.isotone_def)
lemma dense_lattice_char_1:
"(\<forall>x y . x \<sqinter> y = bot \<longrightarrow> x = bot \<or> y = bot) \<longleftrightarrow> (\<forall>x . x \<noteq> bot \<longrightarrow> dense x)"
by (metis inf_top.left_neutral p_bot p_inf pp_inf_bot_iff)
lemma dense_lattice_char_2:
"(\<forall>x y . x \<sqinter> y = bot \<longrightarrow> x = bot \<or> y = bot) \<longleftrightarrow> (\<forall>x . regular x \<longrightarrow> x = bot \<or> x = top)"
by (metis dense_lattice_char_1 inf_top.left_neutral p_inf regular_closed_p regular_closed_top)
lemma restrict_below_Rep_eq:
"x \<sqinter> --y \<le> z \<Longrightarrow> x \<sqinter> y = x \<sqinter> z \<sqinter> y"
by (metis inf.absorb2 inf.commute inf.left_commute pp_increasing)
(*
lemma p_inf_sup_below: "-x \<sqinter> (x \<squnion> y) \<le> y" nitpick [expect=genuine] oops
lemma complement_p: "x \<sqinter> y = bot \<and> x \<squnion> y = top \<longrightarrow> -x = y" nitpick [expect=genuine] oops
lemma complemented_regular: "complemented x \<longrightarrow> regular x" nitpick [expect=genuine] oops
*)
end
text \<open>
The following class gives equational axioms for the pseudocomplement operation.
\<close>
class p_algebra_eq = bounded_lattice + uminus +
assumes p_bot_eq: "-bot = top"
and p_top_eq: "-top = bot"
and inf_import_p_eq: "x \<sqinter> -(x \<sqinter> y) = x \<sqinter> -y"
begin
lemma inf_p_eq:
"x \<sqinter> -x = bot"
by (metis inf_bot_right inf_import_p_eq inf_top_right p_top_eq)
subclass p_algebra
apply unfold_locales
apply (rule iffI)
apply (metis inf.orderI inf_import_p_eq inf_top.right_neutral p_bot_eq)
by (metis (full_types) inf.left_commute inf.orderE inf_bot_right inf_commute inf_p_eq)
end
subsubsection \<open>Pseudocomplemented Distributive Lattices\<close>
text \<open>
We obtain further properties if we assume that the lattice operations are distributive.
\<close>
class pd_algebra = p_algebra + bounded_distrib_lattice
begin
lemma p_inf_sup_below:
"-x \<sqinter> (x \<squnion> y) \<le> y"
by (simp add: inf_sup_distrib1)
lemma pp_inf_sup_p [simp]:
"--x \<sqinter> (x \<squnion> -x) = x"
using inf.absorb2 inf_sup_distrib1 pp_increasing by auto
lemma complement_p:
"x \<sqinter> y = bot \<Longrightarrow> x \<squnion> y = top \<Longrightarrow> -x = y"
by (metis pseudo_complement inf.commute inf_top.left_neutral sup.absorb_iff1 sup.commute sup_bot.right_neutral sup_inf_distrib2 p_inf)
lemma complemented_regular:
"complemented x \<Longrightarrow> regular x"
using complement_p inf.commute sup.commute by fastforce
lemma regular_inf_dense:
"\<exists>y z . regular y \<and> dense z \<and> x = y \<sqinter> z"
by (metis pp_inf_sup_p dense_sup_p ppp)
lemma maddux_3_12 [simp]:
"(x \<squnion> -y) \<sqinter> (x \<squnion> y) = x"
by (metis p_inf sup_bot_right sup_inf_distrib1)
lemma maddux_3_13 [simp]:
"(x \<squnion> y) \<sqinter> -x = y \<sqinter> -x"
by (simp add: inf_sup_distrib2)
lemma maddux_3_20:
"((v \<sqinter> w) \<squnion> (-v \<sqinter> x)) \<sqinter> -((v \<sqinter> y) \<squnion> (-v \<sqinter> z)) = (v \<sqinter> w \<sqinter> -y) \<squnion> (-v \<sqinter> x \<sqinter> -z)"
proof -
have "v \<sqinter> w \<sqinter> -(v \<sqinter> y) \<sqinter> -(-v \<sqinter> z) = v \<sqinter> w \<sqinter> -(v \<sqinter> y)"
by (meson inf.cobounded1 inf_absorb1 le_infI1 p_antitone_iff)
also have "... = v \<sqinter> w \<sqinter> -y"
using inf.sup_relative_same_increasing inf_import_p inf_le1 by blast
finally have 1: "v \<sqinter> w \<sqinter> -(v \<sqinter> y) \<sqinter> -(-v \<sqinter> z) = v \<sqinter> w \<sqinter> -y"
.
have "-v \<sqinter> x \<sqinter> -(v \<sqinter> y) \<sqinter> -(-v \<sqinter> z) = -v \<sqinter> x \<sqinter> -(-v \<sqinter> z)"
by (simp add: inf.absorb1 le_infI1 p_antitone_inf)
also have "... = -v \<sqinter> x \<sqinter> -z"
by (simp add: inf.assoc inf_left_commute)
finally have 2: "-v \<sqinter> x \<sqinter> -(v \<sqinter> y) \<sqinter> -(-v \<sqinter> z) = -v \<sqinter> x \<sqinter> -z"
.
have "((v \<sqinter> w) \<squnion> (-v \<sqinter> x)) \<sqinter> -((v \<sqinter> y) \<squnion> (-v \<sqinter> z)) = (v \<sqinter> w \<sqinter> -(v \<sqinter> y) \<sqinter> -(-v \<sqinter> z)) \<squnion> (-v \<sqinter> x \<sqinter> -(v \<sqinter> y) \<sqinter> -(-v \<sqinter> z))"
by (simp add: inf_assoc inf_sup_distrib2)
also have "... = (v \<sqinter> w \<sqinter> -y) \<squnion> (-v \<sqinter> x \<sqinter> -z)"
using 1 2 by simp
finally show ?thesis
.
qed
lemma order_char_1:
"x \<le> y \<longleftrightarrow> x \<le> y \<squnion> -x"
by (metis inf.sup_left_isotone inf_sup_absorb le_supI1 maddux_3_12 sup_commute)
lemma order_char_2:
"x \<le> y \<longleftrightarrow> x \<squnion> -x \<le> y \<squnion> -x"
using order_char_1 by auto
(*
lemma pp_dist_sup [simp]: "--(x \<squnion> y) = --x \<squnion> --y" nitpick [expect=genuine] oops
lemma regular_closed_sup: "regular x \<and> regular y \<longrightarrow> regular (x \<squnion> y)" nitpick [expect=genuine] oops
lemma regular_complemented_iff: "regular x \<longleftrightarrow> complemented x" nitpick [expect=genuine] oops
lemma selection_closed_sup: "selection s x \<and> selection t x \<longrightarrow> selection (s \<squnion> t) x" nitpick [expect=genuine] oops
lemma stone [simp]: "-x \<squnion> --x = top" nitpick [expect=genuine] oops
*)
end
subsection \<open>Stone Algebras\<close>
text \<open>
A Stone algebra is a distributive lattice with a pseudocomplement that satisfies the following equation.
We thus obtain the other half of the requirements of a complement at least for the regular elements.
\<close>
class stone_algebra = pd_algebra +
assumes stone [simp]: "-x \<squnion> --x = top"
begin
text \<open>
As a consequence, we obtain both De Morgan's laws for all elements.
\<close>
lemma p_dist_inf [simp]:
"-(x \<sqinter> y) = -x \<squnion> -y"
proof (rule p_unique[THEN sym], rule allI, rule iffI)
fix w
assume "w \<sqinter> (x \<sqinter> y) = bot"
hence "w \<sqinter> --x \<sqinter> y = bot"
using inf_commute inf_left_commute pseudo_complement by auto
hence 1: "w \<sqinter> --x \<le> -y"
by (simp add: pseudo_complement)
have "w = (w \<sqinter> -x) \<squnion> (w \<sqinter> --x)"
using distrib_imp2 sup_inf_distrib1 by auto
thus "w \<le> -x \<squnion> -y"
using 1 by (metis inf_le2 sup.mono)
next
fix w
assume "w \<le> -x \<squnion> -y"
thus "w \<sqinter> (x \<sqinter> y) = bot"
using order_trans p_supdist_inf pseudo_complement by blast
qed
lemma pp_dist_sup [simp]:
"--(x \<squnion> y) = --x \<squnion> --y"
by simp
lemma regular_closed_sup:
"regular x \<Longrightarrow> regular y \<Longrightarrow> regular (x \<squnion> y)"
by simp
text \<open>
The regular elements are precisely the ones having a complement.
\<close>
lemma regular_complemented_iff:
"regular x \<longleftrightarrow> complemented x"
by (metis inf_p stone complemented_regular)
lemma selection_closed_sup:
"selection s x \<Longrightarrow> selection t x \<Longrightarrow> selection (s \<squnion> t) x"
by (simp add: inf_sup_distrib2)
lemma huntington_3_pp [simp]:
"-(-x \<squnion> -y) \<squnion> -(-x \<squnion> y) = --x"
by (metis p_dist_inf p_inf sup.commute sup_bot_left sup_inf_distrib1)
lemma maddux_3_3 [simp]:
"-(x \<squnion> y) \<squnion> -(x \<squnion> -y) = -x"
by (simp add: sup_commute sup_inf_distrib1)
lemma maddux_3_11_pp:
"(x \<sqinter> -y) \<squnion> (x \<sqinter> --y) = x"
by (metis inf_sup_distrib1 inf_top_right stone)
lemma maddux_3_19_pp:
"(-x \<sqinter> y) \<squnion> (--x \<sqinter> z) = (--x \<squnion> y) \<sqinter> (-x \<squnion> z)"
proof -
have "(--x \<squnion> y) \<sqinter> (-x \<squnion> z) = (--x \<sqinter> z) \<squnion> (y \<sqinter> -x) \<squnion> (y \<sqinter> z)"
by (simp add: inf.commute inf_sup_distrib1 sup.assoc)
also have "... = (--x \<sqinter> z) \<squnion> (y \<sqinter> -x) \<squnion> (y \<sqinter> z \<sqinter> (-x \<squnion> --x))"
by simp
also have "... = (--x \<sqinter> z) \<squnion> ((y \<sqinter> -x) \<squnion> (y \<sqinter> -x \<sqinter> z)) \<squnion> (y \<sqinter> z \<sqinter> --x)"
using inf_sup_distrib1 sup_assoc inf_commute inf_assoc by presburger
also have "... = (--x \<sqinter> z) \<squnion> (y \<sqinter> -x) \<squnion> (y \<sqinter> z \<sqinter> --x)"
by simp
also have "... = ((--x \<sqinter> z) \<squnion> (--x \<sqinter> z \<sqinter> y)) \<squnion> (y \<sqinter> -x)"
by (simp add: inf_assoc inf_commute sup.left_commute sup_commute)
also have "... = (--x \<sqinter> z) \<squnion> (y \<sqinter> -x)"
by simp
finally show ?thesis
by (simp add: inf_commute sup_commute)
qed
lemma compl_inter_eq_pp:
"--x \<sqinter> y = --x \<sqinter> z \<Longrightarrow> -x \<sqinter> y = -x \<sqinter> z \<Longrightarrow> y = z"
by (metis inf_commute inf_p inf_sup_distrib1 inf_top.right_neutral p_bot p_dist_inf)
lemma maddux_3_21_pp [simp]:
"--x \<squnion> (-x \<sqinter> y) = --x \<squnion> y"
by (simp add: sup.commute sup_inf_distrib1)
lemma shunting_2_pp:
"x \<le> --y \<longleftrightarrow> -x \<squnion> --y = top"
by (metis inf_top_left p_bot p_dist_inf pseudo_complement)
lemma shunting_p:
"x \<sqinter> y \<le> -z \<longleftrightarrow> x \<le> -z \<squnion> -y"
by (metis inf.assoc p_dist_inf p_shunting_swap pseudo_complement)
text \<open>
The following weak shunting property is interesting as it does not require the element \<open>z\<close> on the right-hand side to be regular.
\<close>
lemma shunting_var_p:
"x \<sqinter> -y \<le> z \<longleftrightarrow> x \<le> z \<squnion> --y"
proof
assume "x \<sqinter> -y \<le> z"
hence "z \<squnion> --y = --y \<squnion> (z \<squnion> x \<sqinter> -y)"
by (simp add: sup.absorb1 sup.commute)
thus "x \<le> z \<squnion> --y"
by (metis inf_commute maddux_3_21_pp sup.commute sup.left_commute sup_left_divisibility)
next
assume "x \<le> z \<squnion> --y"
thus "x \<sqinter> -y \<le> z"
by (metis inf.mono maddux_3_12 sup_ge2)
qed
(* Whether conjugate_char_2_pp can be proved in pd_algebra or in p_algebra is unknown. *)
lemma conjugate_char_2_pp:
"conjugate f g \<longleftrightarrow> f bot = bot \<and> g bot = bot \<and> (\<forall>x y . f x \<sqinter> y \<le> --(f(x \<sqinter> --(g y))) \<and> g y \<sqinter> x \<le> --(g(y \<sqinter> --(f x))))"
proof
assume 1: "conjugate f g"
hence 2: "dual_additive (uminus \<circ> g)"
using conjugate_symmetric conjugate_dual_additive by auto
show "f bot = bot \<and> g bot = bot \<and> (\<forall>x y . f x \<sqinter> y \<le> --(f(x \<sqinter> --(g y))) \<and> g y \<sqinter> x \<le> --(g(y \<sqinter> --(f x))))"
proof (intro conjI)
show "f bot = bot"
using 1 by (metis conjugate_def inf_idem inf_bot_left)
next
show "g bot = bot"
using 1 by (metis conjugate_def inf_idem inf_bot_right)
next
show "\<forall>x y . f x \<sqinter> y \<le> --(f(x \<sqinter> --(g y))) \<and> g y \<sqinter> x \<le> --(g(y \<sqinter> --(f x)))"
proof (intro allI)
fix x y
have 3: "y \<le> -(f(x \<sqinter> -(g y)))"
using 1 by (simp add: conjugate_def pseudo_complement inf_commute)
have 4: "x \<le> -(g(y \<sqinter> -(f x)))"
using 1 conjugate_def inf.commute pseudo_complement by fastforce
have "y \<sqinter> -(f(x \<sqinter> --(g y))) = y \<sqinter> -(f(x \<sqinter> -(g y))) \<sqinter> -(f(x \<sqinter> --(g y)))"
using 3 by (simp add: inf.le_iff_sup inf_commute)
also have "... = y \<sqinter> -(f((x \<sqinter> -(g y)) \<squnion> (x \<sqinter> --(g y))))"
using 1 conjugate_dual_additive dual_additive_def inf_assoc by auto
also have "... = y \<sqinter> -(f x)"
by (simp add: maddux_3_11_pp)
also have "... \<le> -(f x)"
by simp
finally have 5: "f x \<sqinter> y \<le> --(f(x \<sqinter> --(g y)))"
by (simp add: inf_commute p_shunting_swap)
have "x \<sqinter> -(g(y \<sqinter> --(f x))) = x \<sqinter> -(g(y \<sqinter> -(f x))) \<sqinter> -(g(y \<sqinter> --(f x)))"
using 4 by (simp add: inf.le_iff_sup inf_commute)
also have "... = x \<sqinter> -(g((y \<sqinter> -(f x)) \<squnion> (y \<sqinter> --(f x))))"
using 2 by (simp add: dual_additive_def inf_assoc)
also have "... = x \<sqinter> -(g y)"
by (simp add: maddux_3_11_pp)
also have "... \<le> -(g y)"
by simp
finally have "g y \<sqinter> x \<le> --(g(y \<sqinter> --(f x)))"
by (simp add: inf_commute p_shunting_swap)
thus "f x \<sqinter> y \<le> --(f(x \<sqinter> --(g y))) \<and> g y \<sqinter> x \<le> --(g(y \<sqinter> --(f x)))"
using 5 by simp
qed
qed
next
assume "f bot = bot \<and> g bot = bot \<and> (\<forall>x y . f x \<sqinter> y \<le> --(f(x \<sqinter> --(g y))) \<and> g y \<sqinter> x \<le> --(g(y \<sqinter> --(f x))))"
thus "conjugate f g"
by (unfold conjugate_def, metis inf_commute le_bot pp_inf_bot_iff regular_closed_bot)
qed
lemma conjugate_char_2_pp_additive:
assumes "conjugate f g"
and "additive f"
and "additive g"
shows "f x \<sqinter> y \<le> f(x \<sqinter> --(g y)) \<and> g y \<sqinter> x \<le> g(y \<sqinter> --(f x))"
proof -
have "f x \<sqinter> y = f ((x \<sqinter> --g y) \<squnion> (x \<sqinter> -g y)) \<sqinter> y"
by (simp add: sup.commute sup_inf_distrib1)
also have "... = (f (x \<sqinter> --g y) \<sqinter> y) \<squnion> (f (x \<sqinter> -g y) \<sqinter> y)"
using assms(2) additive_def inf_sup_distrib2 by auto
also have "... = f (x \<sqinter> --g y) \<sqinter> y"
by (metis assms(1) conjugate_def inf_le2 pseudo_complement sup_bot.right_neutral)
finally have 2: "f x \<sqinter> y \<le> f (x \<sqinter> --g y)"
by simp
have "g y \<sqinter> x = g ((y \<sqinter> --f x) \<squnion> (y \<sqinter> -f x)) \<sqinter> x"
by (simp add: sup.commute sup_inf_distrib1)
also have "... = (g (y \<sqinter> --f x) \<sqinter> x) \<squnion> (g (y \<sqinter> -f x) \<sqinter> x)"
using assms(3) additive_def inf_sup_distrib2 by auto
also have "... = g (y \<sqinter> --f x) \<sqinter> x"
by (metis assms(1) conjugate_def inf.cobounded2 pseudo_complement sup_bot.right_neutral inf_commute)
finally have "g y \<sqinter> x \<le> g (y \<sqinter> --f x)"
by simp
thus ?thesis
using 2 by simp
qed
(*
lemma compl_le_swap2_iff: "-x \<le> y \<longleftrightarrow> -y \<le> x" nitpick [expect=genuine] oops
lemma huntington_3: "x = -(-x \<squnion> -y) \<squnion> -(-x \<squnion> y)" nitpick [expect=genuine] oops
lemma maddux_3_1: "x \<squnion> -x = y \<squnion> -y" nitpick [expect=genuine] oops
lemma maddux_3_4: "x \<squnion> (y \<squnion> -y) = z \<squnion> -z" nitpick [expect=genuine] oops
lemma maddux_3_11: "x = (x \<sqinter> y) \<squnion> (x \<sqinter> -y)" nitpick [expect=genuine] oops
lemma maddux_3_19: "(-x \<sqinter> y) \<squnion> (x \<sqinter> z) = (x \<squnion> y) \<sqinter> (-x \<squnion> z)" nitpick [expect=genuine] oops
lemma compl_inter_eq: "x \<sqinter> y = x \<sqinter> z \<and> -x \<sqinter> y = -x \<sqinter> z \<longrightarrow> y = z" nitpick [expect=genuine] oops
lemma maddux_3_21: "x \<squnion> y = x \<squnion> (-x \<sqinter> y)" nitpick [expect=genuine] oops
lemma shunting_1: "x \<le> y \<longleftrightarrow> x \<sqinter> -y = bot" nitpick [expect=genuine] oops
lemma shunting_2: "x \<le> y \<longleftrightarrow> -x \<squnion> y = top" nitpick [expect=genuine] oops
lemma conjugate_unique: "conjugate f g \<and> conjugate f h \<longrightarrow> g = h" nitpick [expect=genuine] oops
lemma conjugate_isotone_pp: "conjugate f g \<longrightarrow> isotone f" nitpick [expect=genuine] oops
lemma conjugate_char_1: "conjugate f g \<longleftrightarrow> (\<forall>x y . f(x \<sqinter> -(g y)) \<le> f x \<sqinter> -y \<and> g(y \<sqinter> -(f x)) \<le> g y \<sqinter> -x)" nitpick [expect=genuine] oops
lemma conjugate_char_2: "conjugate f g \<longleftrightarrow> f bot = bot \<and> g bot = bot \<and> (\<forall>x y . f x \<sqinter> y \<le> f(x \<sqinter> g y) \<and> g y \<sqinter> x \<le> g(y \<sqinter> f x))" nitpick [expect=genuine] oops
lemma shunting: "x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> z \<squnion> -y" nitpick [expect=genuine] oops
lemma shunting_var: "x \<sqinter> -y \<le> z \<longleftrightarrow> x \<le> z \<squnion> y" nitpick [expect=genuine] oops
lemma sup_compl_top: "x \<squnion> -x = top" nitpick [expect=genuine] oops
lemma selection_closed_p: "selection s x \<longrightarrow> selection (-s) x" nitpick [expect=genuine] oops
lemma selection_closed_pp: "selection s x \<longrightarrow> selection (--s) x" nitpick [expect=genuine] oops
*)
end
abbreviation stone_algebra_isomorphism :: "('a::stone_algebra \<Rightarrow> 'b::stone_algebra) \<Rightarrow> bool"
where "stone_algebra_isomorphism f \<equiv> sup_inf_top_bot_uminus_isomorphism f"
text \<open>
Every bounded linear order can be expanded to a Stone algebra.
The pseudocomplement takes \<open>bot\<close> to the \<open>top\<close> and every other element to \<open>bot\<close>.
\<close>
class linorder_stone_algebra_expansion = linorder_lattice_expansion + uminus +
assumes uminus_def [simp]: "-x = (if x = bot then top else bot)"
begin
subclass stone_algebra
apply unfold_locales
using bot_unique min_def top_le by auto
text \<open>
The regular elements are the least and greatest elements.
All elements except the least element are dense.
\<close>
lemma regular_bot_top:
"regular x \<longleftrightarrow> x = bot \<or> x = top"
by simp
lemma not_bot_dense:
"x \<noteq> bot \<Longrightarrow> --x = top"
by simp
end
subsection \<open>Heyting Algebras\<close>
text \<open>
In this section we add a relative pseudocomplement operation to semilattices and to lattices.
\<close>
subsubsection \<open>Heyting Semilattices\<close>
text \<open>
The pseudocomplement of an element \<open>y\<close> relative to an element \<open>z\<close> is the least element whose meet with \<open>y\<close> is below \<open>z\<close>.
This can be stated as a Galois connection.
Specialising \<open>z = bot\<close> gives (non-relative) pseudocomplements.
Many properties can already be shown if the underlying structure is just a semilattice.
\<close>
class implies =
fixes implies :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<leadsto>" 65)
class heyting_semilattice = semilattice_inf + implies +
assumes implies_galois: "x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> y \<leadsto> z"
begin
lemma implies_below_eq [simp]:
"y \<sqinter> (x \<leadsto> y) = y"
using implies_galois inf.absorb_iff1 inf.cobounded1 by blast
lemma implies_increasing:
"x \<le> y \<leadsto> x"
by (simp add: inf.orderI)
lemma implies_galois_swap:
"x \<le> y \<leadsto> z \<longleftrightarrow> y \<le> x \<leadsto> z"
by (metis implies_galois inf_commute)
lemma implies_galois_var:
"x \<sqinter> y \<le> z \<longleftrightarrow> y \<le> x \<leadsto> z"
by (simp add: implies_galois_swap implies_galois)
lemma implies_galois_increasing:
"x \<le> y \<leadsto> (x \<sqinter> y)"
using implies_galois by blast
lemma implies_galois_decreasing:
"(y \<leadsto> x) \<sqinter> y \<le> x"
using implies_galois by blast
lemma implies_mp_below:
"x \<sqinter> (x \<leadsto> y) \<le> y"
using implies_galois_decreasing inf_commute by auto
lemma implies_isotone:
"x \<le> y \<Longrightarrow> z \<leadsto> x \<le> z \<leadsto> y"
using implies_galois order_trans by blast
lemma implies_antitone:
"x \<le> y \<Longrightarrow> y \<leadsto> z \<le> x \<leadsto> z"
by (meson implies_galois_swap order_lesseq_imp)
lemma implies_isotone_inf:
"x \<leadsto> (y \<sqinter> z) \<le> x \<leadsto> y"
by (simp add: implies_isotone)
lemma implies_antitone_inf:
"x \<leadsto> z \<le> (x \<sqinter> y) \<leadsto> z"
by (simp add: implies_antitone)
lemma implies_curry:
"x \<leadsto> (y \<leadsto> z) = (x \<sqinter> y) \<leadsto> z"
- by (metis implies_galois_decreasing implies_galois inf_assoc antisym)
+ by (metis implies_galois_decreasing implies_galois inf_assoc order.antisym)
lemma implies_curry_flip:
"x \<leadsto> (y \<leadsto> z) = y \<leadsto> (x \<leadsto> z)"
by (simp add: implies_curry inf_commute)
lemma triple_implies [simp]:
"((x \<leadsto> y) \<leadsto> y) \<leadsto> y = x \<leadsto> y"
- using implies_antitone implies_galois_swap eq_iff by auto
+ using implies_antitone implies_galois_swap order.eq_iff by auto
lemma implies_mp_eq [simp]:
"x \<sqinter> (x \<leadsto> y) = x \<sqinter> y"
by (metis implies_below_eq implies_mp_below inf_left_commute inf.absorb2)
lemma implies_dist_implies:
"x \<leadsto> (y \<leadsto> z) \<le> (x \<leadsto> y) \<leadsto> (x \<leadsto> z)"
using implies_curry implies_curry_flip by auto
lemma implies_import_inf [simp]:
"x \<sqinter> ((x \<sqinter> y) \<leadsto> (x \<leadsto> z)) = x \<sqinter> (y \<leadsto> z)"
by (metis implies_curry implies_mp_eq inf_commute)
lemma implies_dist_inf:
"x \<leadsto> (y \<sqinter> z) = (x \<leadsto> y) \<sqinter> (x \<leadsto> z)"
proof -
have "(x \<leadsto> y) \<sqinter> (x \<leadsto> z) \<sqinter> x \<le> y \<sqinter> z"
by (simp add: implies_galois)
hence "(x \<leadsto> y) \<sqinter> (x \<leadsto> z) \<le> x \<leadsto> (y \<sqinter> z)"
using implies_galois by blast
thus ?thesis
- by (simp add: implies_isotone eq_iff)
+ by (simp add: implies_isotone order.eq_iff)
qed
lemma implies_itself_top:
"y \<le> x \<leadsto> x"
by (simp add: implies_galois_swap implies_increasing)
lemma inf_implies_top:
"z \<le> (x \<sqinter> y) \<leadsto> x"
using implies_galois_var le_infI1 by blast
lemma inf_inf_implies [simp]:
"z \<sqinter> ((x \<sqinter> y) \<leadsto> x) = z"
by (simp add: inf_implies_top inf_absorb1)
lemma le_implies_top:
"x \<le> y \<Longrightarrow> z \<le> x \<leadsto> y"
using implies_antitone implies_itself_top order.trans by blast
lemma le_iff_le_implies:
"x \<le> y \<longleftrightarrow> x \<le> x \<leadsto> y"
using implies_galois inf_idem by force
lemma implies_inf_isotone:
"x \<leadsto> y \<le> (x \<sqinter> z) \<leadsto> (y \<sqinter> z)"
by (metis implies_curry implies_galois_increasing implies_isotone)
lemma implies_transitive:
"(x \<leadsto> y) \<sqinter> (y \<leadsto> z) \<le> x \<leadsto> z"
using implies_dist_implies implies_galois_var implies_increasing order_lesseq_imp by blast
lemma implies_inf_absorb [simp]:
"x \<leadsto> (x \<sqinter> y) = x \<leadsto> y"
using implies_dist_inf implies_itself_top inf.absorb_iff2 by auto
lemma implies_implies_absorb [simp]:
"x \<leadsto> (x \<leadsto> y) = x \<leadsto> y"
by (simp add: implies_curry)
lemma implies_inf_identity:
"(x \<leadsto> y) \<sqinter> y = y"
by (simp add: inf_commute)
lemma implies_itself_same:
"x \<leadsto> x = y \<leadsto> y"
- by (simp add: le_implies_top eq_iff)
+ by (simp add: le_implies_top order.eq_iff)
end
text \<open>
The following class gives equational axioms for the relative pseudocomplement operation (inequalities can be written as equations).
\<close>
class heyting_semilattice_eq = semilattice_inf + implies +
assumes implies_mp_below: "x \<sqinter> (x \<leadsto> y) \<le> y"
and implies_galois_increasing: "x \<le> y \<leadsto> (x \<sqinter> y)"
and implies_isotone_inf: "x \<leadsto> (y \<sqinter> z) \<le> x \<leadsto> y"
begin
subclass heyting_semilattice
apply unfold_locales
apply (rule iffI)
apply (metis implies_galois_increasing implies_isotone_inf inf.absorb2 order_lesseq_imp)
by (metis implies_mp_below inf_commute order_trans inf_mono order_refl)
end
text \<open>
The following class allows us to explicitly give the pseudocomplement of an element relative to itself.
\<close>
class bounded_heyting_semilattice = bounded_semilattice_inf_top + heyting_semilattice
begin
lemma implies_itself [simp]:
"x \<leadsto> x = top"
using implies_galois inf_le2 top_le by blast
lemma implies_order:
"x \<le> y \<longleftrightarrow> x \<leadsto> y = top"
by (metis implies_galois inf_top.left_neutral top_unique)
lemma inf_implies [simp]:
"(x \<sqinter> y) \<leadsto> x = top"
using implies_order inf_le1 by blast
lemma top_implies [simp]:
"top \<leadsto> x = x"
by (metis implies_mp_eq inf_top.left_neutral)
end
subsubsection \<open>Heyting Lattices\<close>
text \<open>
We obtain further properties if the underlying structure is a lattice.
In particular, the lattice operations are automatically distributive in this case.
\<close>
class heyting_lattice = lattice + heyting_semilattice
begin
lemma sup_distrib_inf_le:
"(x \<squnion> y) \<sqinter> (x \<squnion> z) \<le> x \<squnion> (y \<sqinter> z)"
proof -
have "x \<squnion> z \<le> y \<leadsto> (x \<squnion> (y \<sqinter> z))"
using implies_galois_var implies_increasing sup.bounded_iff sup.cobounded2 by blast
hence "x \<squnion> y \<le> (x \<squnion> z) \<leadsto> (x \<squnion> (y \<sqinter> z))"
using implies_galois_swap implies_increasing le_sup_iff by blast
thus ?thesis
by (simp add: implies_galois)
qed
subclass distrib_lattice
apply unfold_locales
- using distrib_sup_le eq_iff sup_distrib_inf_le by auto
+ using distrib_sup_le order.eq_iff sup_distrib_inf_le by auto
lemma implies_isotone_sup:
"x \<leadsto> y \<le> x \<leadsto> (y \<squnion> z)"
by (simp add: implies_isotone)
lemma implies_antitone_sup:
"(x \<squnion> y) \<leadsto> z \<le> x \<leadsto> z"
by (simp add: implies_antitone)
lemma implies_sup:
"x \<leadsto> z \<le> (y \<leadsto> z) \<leadsto> ((x \<squnion> y) \<leadsto> z)"
proof -
have "(x \<leadsto> z) \<sqinter> (y \<leadsto> z) \<sqinter> y \<le> z"
by (simp add: implies_galois)
hence "(x \<leadsto> z) \<sqinter> (y \<leadsto> z) \<sqinter> (x \<squnion> y) \<le> z"
using implies_galois_swap implies_galois_var by fastforce
thus ?thesis
by (simp add: implies_galois)
qed
lemma implies_dist_sup:
"(x \<squnion> y) \<leadsto> z = (x \<leadsto> z) \<sqinter> (y \<leadsto> z)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: implies_antitone)
by (simp add: implies_sup implies_galois)
lemma implies_antitone_isotone:
"(x \<squnion> y) \<leadsto> (x \<sqinter> y) \<le> x \<leadsto> y"
by (simp add: implies_antitone_sup implies_dist_inf le_infI2)
lemma implies_antisymmetry:
"(x \<leadsto> y) \<sqinter> (y \<leadsto> x) = (x \<squnion> y) \<leadsto> (x \<sqinter> y)"
by (metis implies_dist_sup implies_inf_absorb inf.commute)
lemma sup_inf_implies [simp]:
"(x \<squnion> y) \<sqinter> (x \<leadsto> y) = y"
by (simp add: inf_sup_distrib2 sup.absorb2)
lemma implies_subdist_sup:
"(x \<leadsto> y) \<squnion> (x \<leadsto> z) \<le> x \<leadsto> (y \<squnion> z)"
by (simp add: implies_isotone)
lemma implies_subdist_inf:
"(x \<leadsto> z) \<squnion> (y \<leadsto> z) \<le> (x \<sqinter> y) \<leadsto> z"
by (simp add: implies_antitone)
lemma implies_sup_absorb:
"(x \<leadsto> y) \<squnion> z \<le> (x \<squnion> z) \<leadsto> (y \<squnion> z)"
by (metis implies_dist_sup implies_isotone_sup implies_increasing inf_inf_implies le_sup_iff sup_inf_implies)
lemma sup_below_implies_implies:
"x \<squnion> y \<le> (x \<leadsto> y) \<leadsto> y"
by (simp add: implies_dist_sup implies_galois_swap implies_increasing)
end
class bounded_heyting_lattice = bounded_lattice + heyting_lattice
begin
subclass bounded_heyting_semilattice ..
lemma implies_bot [simp]:
"bot \<leadsto> x = top"
using implies_galois top_unique by fastforce
end
subsubsection \<open>Heyting Algebras\<close>
text \<open>
The pseudocomplement operation can be defined in Heyting algebras, but it is typically not part of their signature.
We add the definition as an axiom so that we can use the class hierarchy, for example, to inherit results from the class \<open>pd_algebra\<close>.
\<close>
class heyting_algebra = bounded_heyting_lattice + uminus +
assumes uminus_eq: "-x = x \<leadsto> bot"
begin
subclass pd_algebra
apply unfold_locales
using bot_unique implies_galois uminus_eq by auto
lemma boolean_implies_below:
"-x \<squnion> y \<le> x \<leadsto> y"
by (simp add: implies_increasing implies_isotone uminus_eq)
lemma negation_implies:
"-(x \<leadsto> y) = --x \<sqinter> -y"
-proof (rule antisym)
+proof (rule order.antisym)
show "-(x \<leadsto> y) \<le> --x \<sqinter> -y"
using boolean_implies_below p_antitone by auto
next
have "x \<sqinter> -y \<sqinter> (x \<leadsto> y) = bot"
by (metis implies_mp_eq inf_p inf_bot_left inf_commute inf_left_commute)
hence "--x \<sqinter> -y \<sqinter> (x \<leadsto> y) = bot"
using pp_inf_bot_iff inf_assoc by auto
thus "--x \<sqinter> -y \<le> -(x \<leadsto> y)"
by (simp add: pseudo_complement)
qed
lemma double_negation_dist_implies:
"--(x \<leadsto> y) = --x \<leadsto> --y"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis pp_inf_below_iff implies_galois_decreasing implies_galois negation_implies ppp)
by (simp add: p_antitone_iff negation_implies)
(*
lemma stone: "-x \<squnion> --x = top" nitpick [expect=genuine] oops
*)
end
text \<open>
The following class gives equational axioms for Heyting algebras.
\<close>
class heyting_algebra_eq = bounded_lattice + implies + uminus +
assumes implies_mp_eq: "x \<sqinter> (x \<leadsto> y) = x \<sqinter> y"
and implies_import_inf: "x \<sqinter> ((x \<sqinter> y) \<leadsto> (x \<leadsto> z)) = x \<sqinter> (y \<leadsto> z)"
and inf_inf_implies: "z \<sqinter> ((x \<sqinter> y) \<leadsto> x) = z"
and uminus_eq_eq: "-x = x \<leadsto> bot"
begin
subclass heyting_algebra
apply unfold_locales
apply (rule iffI)
apply (metis implies_import_inf inf.sup_left_divisibility inf_inf_implies le_iff_inf)
apply (metis implies_mp_eq inf.commute inf.le_sup_iff inf.sup_right_isotone)
by (simp add: uminus_eq_eq)
end
text \<open>
A relative pseudocomplement is not enough to obtain the Stone equation, so we add it in the following class.
\<close>
class heyting_stone_algebra = heyting_algebra +
assumes heyting_stone: "-x \<squnion> --x = top"
begin
subclass stone_algebra
by unfold_locales (simp add: heyting_stone)
(*
lemma pre_linear: "(x \<leadsto> y) \<squnion> (y \<leadsto> x) = top" nitpick [expect=genuine] oops
*)
end
subsubsection \<open>Brouwer Algebras\<close>
text \<open>
Brouwer algebras are dual to Heyting algebras.
The dual pseudocomplement of an element \<open>y\<close> relative to an element \<open>x\<close> is the least element whose join with \<open>y\<close> is above \<open>x\<close>.
We can now use the binary operation provided by Boolean algebras in Isabelle/HOL because it is compatible with dual relative pseudocomplements (not relative pseudocomplements).
\<close>
class brouwer_algebra = bounded_lattice + minus + uminus +
assumes minus_galois: "x \<le> y \<squnion> z \<longleftrightarrow> x - y \<le> z"
and uminus_eq_minus: "-x = top - x"
begin
sublocale brouwer: heyting_algebra where inf = sup and less_eq = greater_eq and less = greater and sup = inf and bot = top and top = bot and implies = "\<lambda>x y . y - x"
apply unfold_locales
apply simp
apply simp
apply simp
apply simp
apply (metis minus_galois sup_commute)
by (simp add: uminus_eq_minus)
lemma curry_minus:
"x - (y \<squnion> z) = (x - y) - z"
by (simp add: brouwer.implies_curry sup_commute)
lemma minus_subdist_sup:
"(x - z) \<squnion> (y - z) \<le> (x \<squnion> y) - z"
by (simp add: brouwer.implies_dist_inf)
lemma inf_sup_minus:
"(x \<sqinter> y) \<squnion> (x - y) = x"
by (simp add: inf.absorb1 brouwer.inf_sup_distrib2)
end
subsection \<open>Boolean Algebras\<close>
text \<open>
This section integrates Boolean algebras in the above hierarchy.
In particular, we strengthen several results shown above.
\<close>
context boolean_algebra
begin
text \<open>
Every Boolean algebra is a Stone algebra, a Heyting algebra and a Brouwer algebra.
\<close>
subclass stone_algebra
apply unfold_locales
apply (rule iffI)
apply (metis compl_sup_top inf.orderI inf_bot_right inf_sup_distrib1 inf_top_right sup_inf_absorb)
using inf.commute inf.sup_right_divisibility apply fastforce
by simp
sublocale heyting: heyting_algebra where implies = "\<lambda>x y . -x \<squnion> y"
apply unfold_locales
apply (rule iffI)
using shunting_var_p sup_commute apply fastforce
using shunting_var_p sup_commute apply force
by simp
subclass brouwer_algebra
apply unfold_locales
apply (simp add: diff_eq shunting_var_p sup.commute)
by (simp add: diff_eq)
lemma huntington_3 [simp]:
"-(-x \<squnion> -y) \<squnion> -(-x \<squnion> y) = x"
using huntington_3_pp by auto
lemma maddux_3_1:
"x \<squnion> -x = y \<squnion> -y"
by simp
lemma maddux_3_4:
"x \<squnion> (y \<squnion> -y) = z \<squnion> -z"
by simp
lemma maddux_3_11 [simp]:
"(x \<sqinter> y) \<squnion> (x \<sqinter> -y) = x"
using brouwer.maddux_3_12 sup.commute by auto
lemma maddux_3_19:
"(-x \<sqinter> y) \<squnion> (x \<sqinter> z) = (x \<squnion> y) \<sqinter> (-x \<squnion> z)"
using maddux_3_19_pp by auto
lemma compl_inter_eq:
"x \<sqinter> y = x \<sqinter> z \<Longrightarrow> -x \<sqinter> y = -x \<sqinter> z \<Longrightarrow> y = z"
by (metis inf_commute maddux_3_11)
lemma maddux_3_21 [simp]:
"x \<squnion> (-x \<sqinter> y) = x \<squnion> y"
by (simp add: sup_inf_distrib1)
lemma shunting_1:
"x \<le> y \<longleftrightarrow> x \<sqinter> -y = bot"
by (simp add: pseudo_complement)
lemma uminus_involutive:
"uminus \<circ> uminus = id"
by auto
lemma uminus_injective:
"uminus \<circ> f = uminus \<circ> g \<Longrightarrow> f = g"
by (metis comp_assoc id_o minus_comp_minus)
lemma conjugate_unique:
"conjugate f g \<Longrightarrow> conjugate f h \<Longrightarrow> g = h"
using conjugate_unique_p uminus_injective by blast
lemma dual_additive_additive:
"dual_additive (uminus \<circ> f) \<Longrightarrow> additive f"
by (metis additive_def compl_eq_compl_iff dual_additive_def p_dist_sup o_def)
lemma conjugate_additive:
"conjugate f g \<Longrightarrow> additive f"
by (simp add: conjugate_dual_additive dual_additive_additive)
lemma conjugate_isotone:
"conjugate f g \<Longrightarrow> isotone f"
by (simp add: conjugate_additive additive_isotone)
lemma conjugate_char_1:
"conjugate f g \<longleftrightarrow> (\<forall>x y . f(x \<sqinter> -(g y)) \<le> f x \<sqinter> -y \<and> g(y \<sqinter> -(f x)) \<le> g y \<sqinter> -x)"
by (simp add: conjugate_char_1_pp)
lemma conjugate_char_2:
"conjugate f g \<longleftrightarrow> f bot = bot \<and> g bot = bot \<and> (\<forall>x y . f x \<sqinter> y \<le> f(x \<sqinter> g y) \<and> g y \<sqinter> x \<le> g(y \<sqinter> f x))"
by (simp add: conjugate_char_2_pp)
lemma shunting:
"x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> z \<squnion> -y"
by (simp add: heyting.implies_galois sup.commute)
lemma shunting_var:
"x \<sqinter> -y \<le> z \<longleftrightarrow> x \<le> z \<squnion> y"
by (simp add: shunting)
end
class non_trivial_stone_algebra = non_trivial_bounded_order + stone_algebra
class non_trivial_boolean_algebra = non_trivial_stone_algebra + boolean_algebra
end
diff --git a/thys/Stone_Algebras/Stone_Construction.thy b/thys/Stone_Algebras/Stone_Construction.thy
--- a/thys/Stone_Algebras/Stone_Construction.thy
+++ b/thys/Stone_Algebras/Stone_Construction.thy
@@ -1,2308 +1,2307 @@
(* Title: Construction of Stone Algebras
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Stone Construction\<close>
text \<open>
This theory proves the uniqueness theorem for the triple representation of Stone algebras and the construction theorem of Stone algebras \cite{ChenGraetzer1969,Katrinak1973}.
Every Stone algebra $S$ has an associated triple consisting of
\begin{itemize}
\item the set of regular elements $B(S)$ of $S$,
\item the set of dense elements $D(S)$ of $S$, and
\item the structure map $\varphi(S) : B(S) \to F(D(S))$ defined by $\varphi(x) = {\uparrow} x \cap D(S)$.
\end{itemize}
Here $F(X)$ is the set of filters of a partially ordered set $X$.
We first show that
\begin{itemize}
\item $B(S)$ is a Boolean algebra,
\item $D(S)$ is a distributive lattice with a greatest element, whence $F(D(S))$ is a bounded distributive lattice, and
\item $\varphi(S)$ is a bounded lattice homomorphism.
\end{itemize}
Next, from a triple $T = (B,D,\varphi)$ such that $B$ is a Boolean algebra, $D$ is a distributive lattice with a greatest element and $\varphi : B \to F(D)$ is a bounded lattice homomorphism, we construct a Stone algebra $S(T)$.
The elements of $S(T)$ are pairs taken from $B \times F(D)$ following the construction of \cite{Katrinak1973}.
We need to represent $S(T)$ as a type to be able to instantiate the Stone algebra class.
Because the pairs must satisfy a condition depending on $\varphi$, this would require dependent types.
Since Isabelle/HOL does not have dependent types, we use a function lifting instead.
The lifted pairs form a Stone algebra.
Next, we specialise the construction to start with the triple associated with a Stone algebra $S$, that is, we construct $S(B(S),D(S),\varphi(S))$.
In this case, we can instantiate the lifted pairs to obtain a type of pairs (that no longer implements a dependent type).
To achieve this, we construct an embedding of the type of pairs into the lifted pairs, so that we inherit the Stone algebra axioms (using a technique of universal algebra that works for universally quantified equations and equational implications).
Next, we show that the Stone algebras $S(B(S),D(S),\varphi(S))$ and $S$ are isomorphic.
We give explicit mappings in both directions.
This implies the uniqueness theorem for the triple representation of Stone algebras.
Finally, we show that the triples $(B(S(T)),D(S(T)),\varphi(S(T)))$ and $T$ are isomorphic.
This requires an isomorphism of the Boolean algebras $B$ and $B(S(T))$, an isomorphism of the distributive lattices $D$ and $D(S(T))$, and a proof that they preserve the structure maps.
We give explicit mappings of the Boolean algebra isomorphism and the distributive lattice isomorphism in both directions.
This implies the construction theorem of Stone algebras.
Because $S(T)$ is implemented by lifted pairs, so are $B(S(T))$ and $D(S(T))$; we therefore also lift $B$ and $D$ to establish the isomorphisms.
\<close>
theory Stone_Construction
imports P_Algebras Filters
begin
text \<open>
A triple consists of a Boolean algebra, a distributive lattice with a greatest element, and a structure map.
The Boolean algebra and the distributive lattice are represented as HOL types.
Because both occur in the type of the structure map, the triple is determined simply by the structure map and its HOL type.
The structure map needs to be a bounded lattice homomorphism.
\<close>
locale triple =
fixes phi :: "'a::boolean_algebra \<Rightarrow> 'b::distrib_lattice_top filter"
assumes hom: "bounded_lattice_homomorphism phi"
subsection \<open>The Triple of a Stone Algebra\<close>
text \<open>
In this section we construct the triple associated to a Stone algebra.
\<close>
subsubsection \<open>Regular Elements\<close>
text \<open>
The regular elements of a Stone algebra form a Boolean subalgebra.
\<close>
typedef (overloaded) 'a regular = "regular_elements::'a::stone_algebra set"
by auto
lemma simp_regular [simp]:
"\<exists>y . Rep_regular x = -y"
using Rep_regular by simp
setup_lifting type_definition_regular
instantiation regular :: (stone_algebra) boolean_algebra
begin
lift_definition sup_regular :: "'a regular \<Rightarrow> 'a regular \<Rightarrow> 'a regular" is sup
by (meson regular_in_p_image_iff regular_closed_sup)
lift_definition inf_regular :: "'a regular \<Rightarrow> 'a regular \<Rightarrow> 'a regular" is inf
by (meson regular_in_p_image_iff regular_closed_inf)
lift_definition minus_regular :: "'a regular \<Rightarrow> 'a regular \<Rightarrow> 'a regular" is "\<lambda>x y . x \<sqinter> -y"
by (meson regular_in_p_image_iff regular_closed_inf)
lift_definition uminus_regular :: "'a regular \<Rightarrow> 'a regular" is uminus
by auto
lift_definition bot_regular :: "'a regular" is bot
by (meson regular_in_p_image_iff regular_closed_bot)
lift_definition top_regular :: "'a regular" is top
by (meson regular_in_p_image_iff regular_closed_top)
lift_definition less_eq_regular :: "'a regular \<Rightarrow> 'a regular \<Rightarrow> bool" is less_eq .
lift_definition less_regular :: "'a regular \<Rightarrow> 'a regular \<Rightarrow> bool" is less .
instance
apply intro_classes
subgoal apply transfer by (simp add: less_le_not_le)
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by (simp add: sup_inf_distrib1)
subgoal apply transfer by simp
subgoal apply transfer by auto
subgoal apply transfer by simp
done
end
instantiation regular :: (non_trivial_stone_algebra) non_trivial_boolean_algebra
begin
instance
proof (intro_classes, rule ccontr)
assume "\<not>(\<exists>x y::'a regular . x \<noteq> y)"
hence "(bot::'a regular) = top"
by simp
hence "(bot::'a) = top"
by (metis bot_regular.rep_eq top_regular.rep_eq)
thus False
by (simp add: bot_not_top)
qed
end
subsubsection \<open>Dense Elements\<close>
text \<open>
The dense elements of a Stone algebra form a distributive lattice with a greatest element.
\<close>
typedef (overloaded) 'a dense = "dense_elements::'a::stone_algebra set"
using dense_closed_top by blast
lemma simp_dense [simp]:
"-Rep_dense x = bot"
using Rep_dense by simp
setup_lifting type_definition_dense
instantiation dense :: (stone_algebra) distrib_lattice_top
begin
lift_definition sup_dense :: "'a dense \<Rightarrow> 'a dense \<Rightarrow> 'a dense" is sup
by simp
lift_definition inf_dense :: "'a dense \<Rightarrow> 'a dense \<Rightarrow> 'a dense" is inf
by simp
lift_definition top_dense :: "'a dense" is top
by simp
lift_definition less_eq_dense :: "'a dense \<Rightarrow> 'a dense \<Rightarrow> bool" is less_eq .
lift_definition less_dense :: "'a dense \<Rightarrow> 'a dense \<Rightarrow> bool" is less .
instance
apply intro_classes
subgoal apply transfer by (simp add: inf.less_le_not_le)
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by (simp add: sup_inf_distrib1)
done
end
lemma up_filter_dense_antitone_dense:
"dense (x \<squnion> -x \<squnion> y) \<and> dense (x \<squnion> -x \<squnion> y \<squnion> z)"
by simp
lemma up_filter_dense_antitone:
"up_filter (Abs_dense (x \<squnion> -x \<squnion> y \<squnion> z)) \<le> up_filter (Abs_dense (x \<squnion> -x \<squnion> y))"
by (unfold up_filter_antitone[THEN sym]) (simp add: Abs_dense_inverse less_eq_dense.rep_eq)
text \<open>
The filters of dense elements of a Stone algebra form a bounded distributive lattice.
\<close>
type_synonym 'a dense_filter = "'a dense filter"
typedef (overloaded) 'a dense_filter_type = "{ x::'a dense_filter . True }"
using filter_top by blast
setup_lifting type_definition_dense_filter_type
instantiation dense_filter_type :: (stone_algebra) bounded_distrib_lattice
begin
lift_definition sup_dense_filter_type :: "'a dense_filter_type \<Rightarrow> 'a dense_filter_type \<Rightarrow> 'a dense_filter_type" is sup .
lift_definition inf_dense_filter_type :: "'a dense_filter_type \<Rightarrow> 'a dense_filter_type \<Rightarrow> 'a dense_filter_type" is inf .
lift_definition bot_dense_filter_type :: "'a dense_filter_type" is bot ..
lift_definition top_dense_filter_type :: "'a dense_filter_type" is top ..
lift_definition less_eq_dense_filter_type :: "'a dense_filter_type \<Rightarrow> 'a dense_filter_type \<Rightarrow> bool" is less_eq .
lift_definition less_dense_filter_type :: "'a dense_filter_type \<Rightarrow> 'a dense_filter_type \<Rightarrow> bool" is less .
instance
apply intro_classes
subgoal apply transfer by (simp add: inf.less_le_not_le)
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by simp
subgoal apply transfer by (simp add: sup_inf_distrib1)
done
end
subsubsection \<open>The Structure Map\<close>
text \<open>
The structure map of a Stone algebra is a bounded lattice homomorphism.
It maps a regular element \<open>x\<close> to the set of all dense elements above \<open>-x\<close>.
This set is a filter.
\<close>
abbreviation stone_phi_base :: "'a::stone_algebra regular \<Rightarrow> 'a dense set"
where "stone_phi_base x \<equiv> { y . -Rep_regular x \<le> Rep_dense y }"
lemma stone_phi_base_filter:
"filter (stone_phi_base x)"
apply (unfold filter_def, intro conjI)
apply (metis Collect_empty_eq top_dense.rep_eq top_greatest)
apply (metis inf_dense.rep_eq inf_le2 le_inf_iff mem_Collect_eq)
using order_trans less_eq_dense.rep_eq by blast
definition stone_phi :: "'a::stone_algebra regular \<Rightarrow> 'a dense_filter"
where "stone_phi x = Abs_filter (stone_phi_base x)"
text \<open>
To show that we obtain a triple, we only need to prove that \<open>stone_phi\<close> is a bounded lattice homomorphism.
The Boolean algebra and the distributive lattice requirements are taken care of by the type system.
\<close>
interpretation stone_phi: triple "stone_phi"
proof (unfold_locales, intro conjI)
have 1: "Rep_regular (Abs_regular bot) = bot"
by (metis bot_regular.rep_eq bot_regular_def)
show "stone_phi bot = bot"
apply (unfold stone_phi_def bot_regular_def 1 p_bot bot_filter_def)
by (metis (mono_tags, lifting) Collect_cong Rep_dense_inject order_refl singleton_conv top.extremum_uniqueI top_dense.rep_eq)
next
show "stone_phi top = top"
by (metis Collect_cong stone_phi_def UNIV_I bot.extremum dense_closed_top top_empty_eq top_filter.abs_eq top_regular.rep_eq top_set_def)
next
show "\<forall>x y::'a regular . stone_phi (x \<squnion> y) = stone_phi x \<squnion> stone_phi y"
proof (intro allI)
fix x y :: "'a regular"
have "stone_phi_base (x \<squnion> y) = filter_sup (stone_phi_base x) (stone_phi_base y)"
proof (rule set_eqI, rule iffI)
fix z
assume 2: "z \<in> stone_phi_base (x \<squnion> y)"
let ?t = "-Rep_regular x \<squnion> Rep_dense z"
let ?u = "-Rep_regular y \<squnion> Rep_dense z"
let ?v = "Abs_dense ?t"
let ?w = "Abs_dense ?u"
have 3: "?v \<in> stone_phi_base x \<and> ?w \<in> stone_phi_base y"
by (simp add: Abs_dense_inverse)
have "?v \<sqinter> ?w = Abs_dense (?t \<sqinter> ?u)"
by (simp add: eq_onp_def inf_dense.abs_eq)
also have "... = Abs_dense (-Rep_regular (x \<squnion> y) \<squnion> Rep_dense z)"
by (simp add: distrib(1) sup_commute sup_regular.rep_eq)
also have "... = Abs_dense (Rep_dense z)"
using 2 by (simp add: le_iff_sup)
also have "... = z"
by (simp add: Rep_dense_inverse)
finally show "z \<in> filter_sup (stone_phi_base x) (stone_phi_base y)"
using 3 mem_Collect_eq order_refl filter_sup_def by fastforce
next
fix z
assume "z \<in> filter_sup (stone_phi_base x) (stone_phi_base y)"
then obtain v w where 4: "v \<in> stone_phi_base x \<and> w \<in> stone_phi_base y \<and> v \<sqinter> w \<le> z"
unfolding filter_sup_def by auto
have "-Rep_regular (x \<squnion> y) = Rep_regular (-(x \<squnion> y))"
by (metis uminus_regular.rep_eq)
also have "... = -Rep_regular x \<sqinter> -Rep_regular y"
by (simp add: inf_regular.rep_eq uminus_regular.rep_eq)
also have "... \<le> Rep_dense v \<sqinter> Rep_dense w"
using 4 inf_mono mem_Collect_eq by blast
also have "... = Rep_dense (v \<sqinter> w)"
by (simp add: inf_dense.rep_eq)
also have "... \<le> Rep_dense z"
using 4 by (simp add: less_eq_dense.rep_eq)
finally show "z \<in> stone_phi_base (x \<squnion> y)"
by simp
qed
thus "stone_phi (x \<squnion> y) = stone_phi x \<squnion> stone_phi y"
by (simp add: stone_phi_def eq_onp_same_args stone_phi_base_filter sup_filter.abs_eq)
qed
next
show "\<forall>x y::'a regular . stone_phi (x \<sqinter> y) = stone_phi x \<sqinter> stone_phi y"
proof (intro allI)
fix x y :: "'a regular"
have "\<forall>z . -Rep_regular (x \<sqinter> y) \<le> Rep_dense z \<longleftrightarrow> -Rep_regular x \<le> Rep_dense z \<and> -Rep_regular y \<le> Rep_dense z"
by (simp add: inf_regular.rep_eq)
hence "stone_phi_base (x \<sqinter> y) = (stone_phi_base x) \<inter> (stone_phi_base y)"
by auto
thus "stone_phi (x \<sqinter> y) = stone_phi x \<sqinter> stone_phi y"
by (simp add: stone_phi_def eq_onp_same_args stone_phi_base_filter inf_filter.abs_eq)
qed
qed
subsection \<open>Properties of Triples\<close>
text \<open>
In this section we construct a certain set of pairs from a triple, introduce operations on these pairs and develop their properties.
The given set and operations will form a Stone algebra.
\<close>
context triple
begin
lemma phi_bot:
"phi bot = Abs_filter {top}"
by (metis hom bot_filter_def)
lemma phi_top:
"phi top = Abs_filter UNIV"
by (metis hom top_filter_def)
text \<open>
The occurrence of \<open>phi\<close> in the following definition of the pairs creates a need for dependent types.
\<close>
definition pairs :: "('a \<times> 'b filter) set"
where "pairs = { (x,y) . \<exists>z . y = phi (-x) \<squnion> up_filter z }"
text \<open>
Operations on pairs are defined in the following.
They will be used to establish that the pairs form a Stone algebra.
\<close>
fun pairs_less_eq :: "('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter) \<Rightarrow> bool"
where "pairs_less_eq (x,y) (z,w) = (x \<le> z \<and> w \<le> y)"
fun pairs_less :: "('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter) \<Rightarrow> bool"
where "pairs_less (x,y) (z,w) = (pairs_less_eq (x,y) (z,w) \<and> \<not> pairs_less_eq (z,w) (x,y))"
fun pairs_sup :: "('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter)"
where "pairs_sup (x,y) (z,w) = (x \<squnion> z,y \<sqinter> w)"
fun pairs_inf :: "('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter)"
where "pairs_inf (x,y) (z,w) = (x \<sqinter> z,y \<squnion> w)"
fun pairs_minus :: "('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter)"
where "pairs_minus (x,y) (z,w) = (x \<sqinter> -z,y \<squnion> phi z)"
fun pairs_uminus :: "('a \<times> 'b filter) \<Rightarrow> ('a \<times> 'b filter)"
where "pairs_uminus (x,y) = (-x,phi x)"
abbreviation pairs_bot :: "('a \<times> 'b filter)"
where "pairs_bot \<equiv> (bot,Abs_filter UNIV)"
abbreviation pairs_top :: "('a \<times> 'b filter)"
where "pairs_top \<equiv> (top,Abs_filter {top})"
lemma pairs_top_in_set:
"(x,y) \<in> pairs \<Longrightarrow> top \<in> Rep_filter y"
by simp
lemma phi_complemented:
"complement (phi x) (phi (-x))"
by (metis hom inf_compl_bot sup_compl_top)
lemma phi_inf_principal:
"\<exists>z . up_filter z = phi x \<sqinter> up_filter y"
proof -
let ?F = "Rep_filter (phi x)"
let ?G = "Rep_filter (phi (-x))"
have 1: "eq_onp filter ?F ?F \<and> eq_onp filter (\<up>y) (\<up>y)"
by (simp add: eq_onp_def)
have "filter_complements ?F ?G"
apply (intro conjI)
apply simp
apply simp
apply (metis (no_types) phi_complemented sup_filter.rep_eq top_filter.rep_eq)
by (metis (no_types) phi_complemented inf_filter.rep_eq bot_filter.rep_eq)
hence "is_principal_up (?F \<inter> \<up>y)"
using complemented_filter_inf_principal by blast
then obtain z where "\<up>z = ?F \<inter> \<up>y"
by auto
hence "up_filter z = Abs_filter (?F \<inter> \<up>y)"
by simp
also have "... = Abs_filter ?F \<sqinter> up_filter y"
using 1 inf_filter.abs_eq by force
also have "... = phi x \<sqinter> up_filter y"
by (simp add: Rep_filter_inverse)
finally show ?thesis
by auto
qed
text \<open>
Quite a bit of filter theory is involved in showing that the intersection of \<open>phi x\<close> with a principal filter is a principal filter, so the following function can extract its least element.
\<close>
fun rho :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
where "rho x y = (SOME z . up_filter z = phi x \<sqinter> up_filter y)"
lemma rho_char:
"up_filter (rho x y) = phi x \<sqinter> up_filter y"
by (metis (mono_tags) someI_ex rho.simps phi_inf_principal)
text \<open>
The following results show that the pairs are closed under the given operations.
\<close>
lemma pairs_sup_closed:
assumes "(x,y) \<in> pairs"
and "(z,w) \<in> pairs"
shows "pairs_sup (x,y) (z,w) \<in> pairs"
proof -
from assms obtain u v where "y = phi (-x) \<squnion> up_filter u \<and> w = phi (-z) \<squnion> up_filter v"
using pairs_def by auto
hence "pairs_sup (x,y) (z,w) = (x \<squnion> z,(phi (-x) \<squnion> up_filter u) \<sqinter> (phi (-z) \<squnion> up_filter v))"
by simp
also have "... = (x \<squnion> z,(phi (-x) \<sqinter> phi (-z)) \<squnion> (phi (-x) \<sqinter> up_filter v) \<squnion> (up_filter u \<sqinter> phi (-z)) \<squnion> (up_filter u \<sqinter> up_filter v))"
by (simp add: inf.sup_commute inf_sup_distrib1 sup_commute sup_left_commute)
also have "... = (x \<squnion> z,phi (-(x \<squnion> z)) \<squnion> (phi (-x) \<sqinter> up_filter v) \<squnion> (up_filter u \<sqinter> phi (-z)) \<squnion> (up_filter u \<sqinter> up_filter v))"
using hom by simp
also have "... = (x \<squnion> z,phi (-(x \<squnion> z)) \<squnion> up_filter (rho (-x) v) \<squnion> up_filter (rho (-z) u) \<squnion> (up_filter u \<sqinter> up_filter v))"
by (metis inf.sup_commute rho_char)
also have "... = (x \<squnion> z,phi (-(x \<squnion> z)) \<squnion> up_filter (rho (-x) v) \<squnion> up_filter (rho (-z) u) \<squnion> up_filter (u \<squnion> v))"
by (metis up_filter_dist_sup)
also have "... = (x \<squnion> z,phi (-(x \<squnion> z)) \<squnion> up_filter (rho (-x) v \<sqinter> rho (-z) u \<sqinter> (u \<squnion> v)))"
by (simp add: sup_commute sup_left_commute up_filter_dist_inf)
finally show ?thesis
using pairs_def by auto
qed
lemma pairs_inf_closed:
assumes "(x,y) \<in> pairs"
and "(z,w) \<in> pairs"
shows "pairs_inf (x,y) (z,w) \<in> pairs"
proof -
from assms obtain u v where "y = phi (-x) \<squnion> up_filter u \<and> w = phi (-z) \<squnion> up_filter v"
using pairs_def by auto
hence "pairs_inf (x,y) (z,w) = (x \<sqinter> z,(phi (-x) \<squnion> up_filter u) \<squnion> (phi (-z) \<squnion> up_filter v))"
by simp
also have "... = (x \<sqinter> z,(phi (-x) \<squnion> phi (-z)) \<squnion> (up_filter u \<squnion> up_filter v))"
by (simp add: sup_commute sup_left_commute)
also have "... = (x \<sqinter> z,phi (-(x \<sqinter> z)) \<squnion> (up_filter u \<squnion> up_filter v))"
using hom by simp
also have "... = (x \<sqinter> z,phi (-(x \<sqinter> z)) \<squnion> up_filter (u \<sqinter> v))"
by (simp add: up_filter_dist_inf)
finally show ?thesis
using pairs_def by auto
qed
lemma pairs_uminus_closed:
"pairs_uminus (x,y) \<in> pairs"
proof -
have "pairs_uminus (x,y) = (-x,phi (--x) \<squnion> bot)"
by simp
also have "... = (-x,phi (--x) \<squnion> up_filter top)"
by (simp add: bot_filter.abs_eq)
finally show ?thesis
by (metis (mono_tags, lifting) mem_Collect_eq old.prod.case pairs_def)
qed
lemma pairs_bot_closed:
"pairs_bot \<in> pairs"
using pairs_def phi_top triple.hom triple_axioms by fastforce
lemma pairs_top_closed:
"pairs_top \<in> pairs"
by (metis p_bot pairs_uminus.simps pairs_uminus_closed phi_bot)
text \<open>
We prove enough properties of the pair operations so that we can later show they form a Stone algebra.
\<close>
lemma pairs_sup_dist_inf:
"(x,y) \<in> pairs \<Longrightarrow> (z,w) \<in> pairs \<Longrightarrow> (u,v) \<in> pairs \<Longrightarrow> pairs_sup (x,y) (pairs_inf (z,w) (u,v)) = pairs_inf (pairs_sup (x,y) (z,w)) (pairs_sup (x,y) (u,v))"
using sup_inf_distrib1 inf_sup_distrib1 by auto
lemma pairs_phi_less_eq:
"(x,y) \<in> pairs \<Longrightarrow> phi (-x) \<le> y"
using pairs_def by auto
lemma pairs_uminus_galois:
assumes "(x,y) \<in> pairs"
and "(z,w) \<in> pairs"
shows "pairs_inf (x,y) (z,w) = pairs_bot \<longleftrightarrow> pairs_less_eq (x,y) (pairs_uminus (z,w))"
proof -
have 1: "x \<sqinter> z = bot \<and> y \<squnion> w = Abs_filter UNIV \<longrightarrow> phi z \<le> y"
by (metis (no_types, lifting) assms(1) heyting.implies_inf_absorb hom le_supE pairs_phi_less_eq sup_bot_right)
have 2: "x \<le> -z \<and> phi z \<le> y \<longrightarrow> y \<squnion> w = Abs_filter UNIV"
proof
assume 3: "x \<le> -z \<and> phi z \<le> y"
have "Abs_filter UNIV = phi z \<squnion> phi (-z)"
using hom phi_complemented phi_top by auto
also have "... \<le> y \<squnion> w"
using 3 assms(2) sup_mono pairs_phi_less_eq by auto
finally show "y \<squnion> w = Abs_filter UNIV"
using hom phi_top top.extremum_uniqueI by auto
qed
have "x \<sqinter> z = bot \<longleftrightarrow> x \<le> -z"
by (simp add: shunting_1)
thus ?thesis
using 1 2 Pair_inject pairs_inf.simps pairs_less_eq.simps pairs_uminus.simps by auto
qed
lemma pairs_stone:
"(x,y) \<in> pairs \<Longrightarrow> pairs_sup (pairs_uminus (x,y)) (pairs_uminus (pairs_uminus (x,y))) = pairs_top"
by (metis hom pairs_sup.simps pairs_uminus.simps phi_bot phi_complemented stone)
text \<open>
The following results show how the regular elements and the dense elements among the pairs look like.
\<close>
abbreviation "dense_pairs \<equiv> { (x,y) . (x,y) \<in> pairs \<and> pairs_uminus (x,y) = pairs_bot }"
abbreviation "regular_pairs \<equiv> { (x,y) . (x,y) \<in> pairs \<and> pairs_uminus (pairs_uminus (x,y)) = (x,y) }"
abbreviation "is_principal_up_filter x \<equiv> \<exists>y . x = up_filter y"
lemma dense_pairs:
"dense_pairs = { (x,y) . x = top \<and> is_principal_up_filter y }"
proof -
have "dense_pairs = { (x,y) . (x,y) \<in> pairs \<and> x = top }"
by (metis Pair_inject compl_bot_eq double_compl pairs_uminus.simps phi_top)
also have "... = { (x,y) . (\<exists>z . y = up_filter z) \<and> x = top }"
using hom pairs_def by auto
finally show ?thesis
by auto
qed
lemma regular_pairs:
"regular_pairs = { (x,y) . y = phi (-x) }"
using pairs_def pairs_uminus_closed by fastforce
text \<open>
The following extraction function will be used in defining one direction of the Stone algebra isomorphism.
\<close>
fun rho_pair :: "'a \<times> 'b filter \<Rightarrow> 'b"
where "rho_pair (x,y) = (SOME z . up_filter z = phi x \<sqinter> y)"
lemma get_rho_pair_char:
assumes "(x,y) \<in> pairs"
shows "up_filter (rho_pair (x,y)) = phi x \<sqinter> y"
proof -
from assms obtain w where "y = phi (-x) \<squnion> up_filter w"
using pairs_def by auto
hence "phi x \<sqinter> y = phi x \<sqinter> up_filter w"
by (simp add: inf_sup_distrib1 phi_complemented)
thus ?thesis
using rho_char by auto
qed
lemma sa_iso_pair:
"(--x,phi (-x) \<squnion> up_filter y) \<in> pairs"
using pairs_def by auto
end
subsection \<open>The Stone Algebra of a Triple\<close>
text \<open>
In this section we prove that the set of pairs constructed in a triple forms a Stone Algebra.
The following type captures the parameter \<open>phi\<close> on which the type of triples depends.
This parameter is the structure map that occurs in the definition of the set of pairs.
The set of all structure maps is the set of all bounded lattice homomorphisms (of appropriate type).
In order to make it a HOL type, we need to show that at least one such structure map exists.
To this end we use the ultrafilter lemma: the required bounded lattice homomorphism is essentially the characteristic map of an ultrafilter, but the latter must exist.
In particular, the underlying Boolean algebra must contain at least two elements.
\<close>
typedef (overloaded) ('a,'b) phi = "{ f::'a::non_trivial_boolean_algebra \<Rightarrow> 'b::distrib_lattice_top filter . bounded_lattice_homomorphism f }"
proof -
from ultra_filter_exists obtain F :: "'a set" where 1: "ultra_filter F"
by auto
hence 2: "prime_filter F"
using ultra_filter_prime by auto
let ?f = "\<lambda>x . if x\<in>F then top else bot::'b filter"
have "bounded_lattice_homomorphism ?f"
proof (intro conjI)
show "?f bot = bot"
using 1 by (meson bot.extremum filter_def subset_eq top.extremum_unique)
next
show "?f top = top"
using 1 by simp
next
show "\<forall>x y . ?f (x \<squnion> y) = ?f x \<squnion> ?f y"
proof (intro allI)
fix x y
show "?f (x \<squnion> y) = ?f x \<squnion> ?f y"
apply (cases "x \<in> F"; cases "y \<in> F")
using 1 filter_def apply fastforce
using 1 filter_def apply fastforce
using 1 filter_def apply fastforce
using 2 sup_bot_left by auto
qed
next
show "\<forall>x y . ?f (x \<sqinter> y) = ?f x \<sqinter> ?f y"
proof (intro allI)
fix x y
show "?f (x \<sqinter> y) = ?f x \<sqinter> ?f y"
apply (cases "x \<in> F"; cases "y \<in> F")
using 1 apply (simp add: filter_inf_closed)
using 1 apply (metis (mono_tags, lifting) brouwer.inf_sup_ord(4) inf_top_left filter_def)
using 1 apply (metis (mono_tags, lifting) brouwer.inf_sup_ord(3) inf_top_right filter_def)
using 1 filter_def by force
qed
qed
hence "?f \<in> {f . bounded_lattice_homomorphism f}"
by simp
thus ?thesis
by meson
qed
lemma simp_phi [simp]:
"bounded_lattice_homomorphism (Rep_phi x)"
using Rep_phi by simp
setup_lifting type_definition_phi
text \<open>
The following implements the dependent type of pairs depending on structure maps.
It uses functions from structure maps to pairs with the requirement that, for each structure map, the corresponding pair is contained in the set of pairs constructed for a triple with that structure map.
If this type could be defined in the locale \<open>triple\<close> and instantiated to Stone algebras there, there would be no need for the lifting and we could work with triples directly.
\<close>
typedef (overloaded) ('a,'b) lifted_pair = "{ pf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) phi \<Rightarrow> 'a \<times> 'b filter . \<forall>f . pf f \<in> triple.pairs (Rep_phi f) }"
proof -
have "\<forall>f::('a,'b) phi . triple.pairs_bot \<in> triple.pairs (Rep_phi f)"
proof
fix f :: "('a,'b) phi"
have "triple (Rep_phi f)"
by (simp add: triple_def)
thus "triple.pairs_bot \<in> triple.pairs (Rep_phi f)"
using triple.regular_pairs triple.phi_top by fastforce
qed
thus ?thesis
by auto
qed
lemma simp_lifted_pair [simp]:
"\<forall>f . Rep_lifted_pair pf f \<in> triple.pairs (Rep_phi f)"
using Rep_lifted_pair by simp
setup_lifting type_definition_lifted_pair
text \<open>
The lifted pairs form a Stone algebra.
\<close>
instantiation lifted_pair :: (non_trivial_boolean_algebra,distrib_lattice_top) stone_algebra
begin
text \<open>
All operations are lifted point-wise.
\<close>
lift_definition sup_lifted_pair :: "('a,'b) lifted_pair \<Rightarrow> ('a,'b) lifted_pair \<Rightarrow> ('a,'b) lifted_pair" is "\<lambda>xf yf f . triple.pairs_sup (xf f) (yf f)"
by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_sup_closed prod.collapse)
lift_definition inf_lifted_pair :: "('a,'b) lifted_pair \<Rightarrow> ('a,'b) lifted_pair \<Rightarrow> ('a,'b) lifted_pair" is "\<lambda>xf yf f . triple.pairs_inf (xf f) (yf f)"
by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_inf_closed prod.collapse)
lift_definition uminus_lifted_pair :: "('a,'b) lifted_pair \<Rightarrow> ('a,'b) lifted_pair" is "\<lambda>xf f . triple.pairs_uminus (Rep_phi f) (xf f)"
by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_uminus_closed prod.collapse)
lift_definition bot_lifted_pair :: "('a,'b) lifted_pair" is "\<lambda>f . triple.pairs_bot"
by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_bot_closed)
lift_definition top_lifted_pair :: "('a,'b) lifted_pair" is "\<lambda>f . triple.pairs_top"
by (metis (no_types, hide_lams) simp_phi triple_def triple.pairs_top_closed)
lift_definition less_eq_lifted_pair :: "('a,'b) lifted_pair \<Rightarrow> ('a,'b) lifted_pair \<Rightarrow> bool" is "\<lambda>xf yf . \<forall>f . triple.pairs_less_eq (xf f) (yf f)" .
lift_definition less_lifted_pair :: "('a,'b) lifted_pair \<Rightarrow> ('a,'b) lifted_pair \<Rightarrow> bool" is "\<lambda>xf yf . (\<forall>f . triple.pairs_less_eq (xf f) (yf f)) \<and> \<not> (\<forall>f . triple.pairs_less_eq (yf f) (xf f))" .
instance
proof intro_classes
fix xf yf :: "('a,'b) lifted_pair"
show "xf < yf \<longleftrightarrow> xf \<le> yf \<and> \<not> yf \<le> xf"
by (simp add: less_lifted_pair.rep_eq less_eq_lifted_pair.rep_eq)
next
fix xf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
obtain x1 x2 where "(x1,x2) = ?x"
using prod.collapse by blast
hence "triple.pairs_less_eq ?x ?x"
using 1 by (metis triple.pairs_less_eq.simps order_refl)
}
thus "xf \<le> xf"
by (simp add: less_eq_lifted_pair.rep_eq)
next
fix xf yf zf :: "('a,'b) lifted_pair"
assume 1: "xf \<le> yf" and 2: "yf \<le> zf"
{
fix f :: "('a,'b) phi"
have 3: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
let ?z = "Rep_lifted_pair zf f"
obtain x1 x2 y1 y2 z1 z2 where 4: "(x1,x2) = ?x \<and> (y1,y2) = ?y \<and> (z1,z2) = ?z"
using prod.collapse by blast
have "triple.pairs_less_eq ?x ?y \<and> triple.pairs_less_eq ?y ?z"
using 1 2 3 less_eq_lifted_pair.rep_eq by simp
hence "triple.pairs_less_eq ?x ?z"
using 3 4 by (metis (mono_tags, lifting) triple.pairs_less_eq.simps order_trans)
}
thus "xf \<le> zf"
by (simp add: less_eq_lifted_pair.rep_eq)
next
fix xf yf :: "('a,'b) lifted_pair"
assume 1: "xf \<le> yf" and 2: "yf \<le> xf"
{
fix f :: "('a,'b) phi"
have 3: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
obtain x1 x2 y1 y2 where 4: "(x1,x2) = ?x \<and> (y1,y2) = ?y"
using prod.collapse by blast
have "triple.pairs_less_eq ?x ?y \<and> triple.pairs_less_eq ?y ?x"
using 1 2 3 less_eq_lifted_pair.rep_eq by simp
hence "?x = ?y"
using 3 4 by (metis (mono_tags, lifting) triple.pairs_less_eq.simps antisym)
}
thus "xf = yf"
by (metis Rep_lifted_pair_inverse ext)
next
fix xf yf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
obtain x1 x2 y1 y2 where "(x1,x2) = ?x \<and> (y1,y2) = ?y"
using prod.collapse by blast
hence "triple.pairs_less_eq (triple.pairs_inf ?x ?y) ?y"
using 1 by (metis (mono_tags, lifting) inf_sup_ord(2) sup.cobounded2 triple.pairs_inf.simps triple.pairs_less_eq.simps inf_lifted_pair.rep_eq)
}
thus "xf \<sqinter> yf \<le> yf"
by (simp add: less_eq_lifted_pair.rep_eq inf_lifted_pair.rep_eq)
next
fix xf yf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
obtain x1 x2 y1 y2 where "(x1,x2) = ?x \<and> (y1,y2) = ?y"
using prod.collapse by blast
hence "triple.pairs_less_eq (triple.pairs_inf ?x ?y) ?x"
using 1 by (metis (mono_tags, lifting) inf_sup_ord(1) sup.cobounded1 triple.pairs_inf.simps triple.pairs_less_eq.simps inf_lifted_pair.rep_eq)
}
thus "xf \<sqinter> yf \<le> xf"
by (simp add: less_eq_lifted_pair.rep_eq inf_lifted_pair.rep_eq)
next
fix xf yf zf :: "('a,'b) lifted_pair"
assume 1: "xf \<le> yf" and 2: "xf \<le> zf"
{
fix f :: "('a,'b) phi"
have 3: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
let ?z = "Rep_lifted_pair zf f"
obtain x1 x2 y1 y2 z1 z2 where 4: "(x1,x2) = ?x \<and> (y1,y2) = ?y \<and> (z1,z2) = ?z"
using prod.collapse by blast
have "triple.pairs_less_eq ?x ?y \<and> triple.pairs_less_eq ?x ?z"
using 1 2 3 less_eq_lifted_pair.rep_eq by simp
hence "triple.pairs_less_eq ?x (triple.pairs_inf ?y ?z)"
using 3 4 by (metis (mono_tags, lifting) le_inf_iff sup.bounded_iff triple.pairs_inf.simps triple.pairs_less_eq.simps)
}
thus "xf \<le> yf \<sqinter> zf"
by (simp add: less_eq_lifted_pair.rep_eq inf_lifted_pair.rep_eq)
next
fix xf yf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
obtain x1 x2 y1 y2 where "(x1,x2) = ?x \<and> (y1,y2) = ?y"
using prod.collapse by blast
hence "triple.pairs_less_eq ?x (triple.pairs_sup ?x ?y)"
using 1 by (metis (no_types, lifting) inf_commute sup.cobounded1 inf.cobounded2 triple.pairs_sup.simps triple.pairs_less_eq.simps sup_lifted_pair.rep_eq)
}
thus "xf \<le> xf \<squnion> yf"
by (simp add: less_eq_lifted_pair.rep_eq sup_lifted_pair.rep_eq)
next
fix xf yf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
obtain x1 x2 y1 y2 where "(x1,x2) = ?x \<and> (y1,y2) = ?y"
using prod.collapse by blast
hence "triple.pairs_less_eq ?y (triple.pairs_sup ?x ?y)"
using 1 by (metis (no_types, lifting) sup.cobounded2 inf.cobounded2 triple.pairs_sup.simps triple.pairs_less_eq.simps sup_lifted_pair.rep_eq)
}
thus "yf \<le> xf \<squnion> yf"
by (simp add: less_eq_lifted_pair.rep_eq sup_lifted_pair.rep_eq)
next
fix xf yf zf :: "('a,'b) lifted_pair"
assume 1: "yf \<le> xf" and 2: "zf \<le> xf"
{
fix f :: "('a,'b) phi"
have 3: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
let ?z = "Rep_lifted_pair zf f"
obtain x1 x2 y1 y2 z1 z2 where 4: "(x1,x2) = ?x \<and> (y1,y2) = ?y \<and> (z1,z2) = ?z"
using prod.collapse by blast
have "triple.pairs_less_eq ?y ?x \<and> triple.pairs_less_eq ?z ?x"
using 1 2 3 less_eq_lifted_pair.rep_eq by simp
hence "triple.pairs_less_eq (triple.pairs_sup ?y ?z) ?x"
using 3 4 by (metis (mono_tags, lifting) le_inf_iff sup.bounded_iff triple.pairs_sup.simps triple.pairs_less_eq.simps)
}
thus "yf \<squnion> zf \<le> xf"
by (simp add: less_eq_lifted_pair.rep_eq sup_lifted_pair.rep_eq)
next
fix xf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
obtain x1 x2 where "(x1,x2) = ?x"
using prod.collapse by blast
hence "triple.pairs_less_eq triple.pairs_bot ?x"
using 1 by (metis bot.extremum top_greatest top_filter.abs_eq triple.pairs_less_eq.simps)
}
thus "bot \<le> xf"
by (simp add: less_eq_lifted_pair.rep_eq bot_lifted_pair.rep_eq)
next
fix xf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
obtain x1 x2 where "(x1,x2) = ?x"
using prod.collapse by blast
hence "triple.pairs_less_eq ?x triple.pairs_top"
using 1 by (metis top.extremum bot_least bot_filter.abs_eq triple.pairs_less_eq.simps)
}
thus "xf \<le> top"
by (simp add: less_eq_lifted_pair.rep_eq top_lifted_pair.rep_eq)
next
fix xf yf zf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
let ?z = "Rep_lifted_pair zf f"
obtain x1 x2 y1 y2 z1 z2 where "(x1,x2) = ?x \<and> (y1,y2) = ?y \<and> (z1,z2) = ?z"
using prod.collapse by blast
hence "triple.pairs_sup ?x (triple.pairs_inf ?y ?z) = triple.pairs_inf (triple.pairs_sup ?x ?y) (triple.pairs_sup ?x ?z)"
using 1 by (metis (no_types) sup_inf_distrib1 inf_sup_distrib1 triple.pairs_sup.simps triple.pairs_inf.simps)
}
thus "xf \<squnion> (yf \<sqinter> zf) = (xf \<squnion> yf) \<sqinter> (xf \<squnion> zf)"
by (metis Rep_lifted_pair_inverse ext sup_lifted_pair.rep_eq inf_lifted_pair.rep_eq)
next
fix xf yf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
let ?y = "Rep_lifted_pair yf f"
obtain x1 x2 y1 y2 where 2: "(x1,x2) = ?x \<and> (y1,y2) = ?y"
using prod.collapse by blast
have "?x \<in> triple.pairs (Rep_phi f) \<and> ?y \<in> triple.pairs (Rep_phi f)"
by simp
hence "(triple.pairs_inf ?x ?y = triple.pairs_bot) \<longleftrightarrow> triple.pairs_less_eq ?x (triple.pairs_uminus (Rep_phi f) ?y)"
using 1 2 by (metis triple.pairs_uminus_galois)
}
hence "\<forall>f . (Rep_lifted_pair (xf \<sqinter> yf) f = Rep_lifted_pair bot f) \<longleftrightarrow> triple.pairs_less_eq (Rep_lifted_pair xf f) (Rep_lifted_pair (-yf) f)"
using bot_lifted_pair.rep_eq inf_lifted_pair.rep_eq uminus_lifted_pair.rep_eq by simp
hence "(Rep_lifted_pair (xf \<sqinter> yf) = Rep_lifted_pair bot) \<longleftrightarrow> xf \<le> -yf"
using less_eq_lifted_pair.rep_eq by auto
thus "(xf \<sqinter> yf = bot) \<longleftrightarrow> (xf \<le> -yf)"
by (simp add: Rep_lifted_pair_inject)
next
fix xf :: "('a,'b) lifted_pair"
{
fix f :: "('a,'b) phi"
have 1: "triple (Rep_phi f)"
by (simp add: triple_def)
let ?x = "Rep_lifted_pair xf f"
obtain x1 x2 where "(x1,x2) = ?x"
using prod.collapse by blast
hence "triple.pairs_sup (triple.pairs_uminus (Rep_phi f) ?x) (triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) ?x)) = triple.pairs_top"
using 1 by (metis simp_lifted_pair triple.pairs_stone)
}
hence "Rep_lifted_pair (-xf \<squnion> --xf) = Rep_lifted_pair top"
using sup_lifted_pair.rep_eq uminus_lifted_pair.rep_eq top_lifted_pair.rep_eq by simp
thus "-xf \<squnion> --xf = top"
by (simp add: Rep_lifted_pair_inject)
qed
end
subsection \<open>The Stone Algebra of the Triple of a Stone Algebra\<close>
text \<open>
In this section we specialise the above construction to a particular structure map, namely the one obtained in the triple of a Stone algebra.
For this particular structure map (as well as for any other particular structure map) the resulting type is no longer a dependent type.
It is just the set of pairs obtained for the given structure map.
\<close>
typedef (overloaded) 'a stone_phi_pair = "triple.pairs (stone_phi::'a::stone_algebra regular \<Rightarrow> 'a dense_filter)"
using stone_phi.pairs_bot_closed by auto
setup_lifting type_definition_stone_phi_pair
instantiation stone_phi_pair :: (stone_algebra) sup_inf_top_bot_uminus_ord
begin
lift_definition sup_stone_phi_pair :: "'a stone_phi_pair \<Rightarrow> 'a stone_phi_pair \<Rightarrow> 'a stone_phi_pair" is triple.pairs_sup
using stone_phi.pairs_sup_closed by auto
lift_definition inf_stone_phi_pair :: "'a stone_phi_pair \<Rightarrow> 'a stone_phi_pair \<Rightarrow> 'a stone_phi_pair" is triple.pairs_inf
using stone_phi.pairs_inf_closed by auto
lift_definition uminus_stone_phi_pair :: "'a stone_phi_pair \<Rightarrow> 'a stone_phi_pair" is "triple.pairs_uminus stone_phi"
using stone_phi.pairs_uminus_closed by auto
lift_definition bot_stone_phi_pair :: "'a stone_phi_pair" is "triple.pairs_bot"
by (rule stone_phi.pairs_bot_closed)
lift_definition top_stone_phi_pair :: "'a stone_phi_pair" is "triple.pairs_top"
by (rule stone_phi.pairs_top_closed)
lift_definition less_eq_stone_phi_pair :: "'a stone_phi_pair \<Rightarrow> 'a stone_phi_pair \<Rightarrow> bool" is triple.pairs_less_eq .
lift_definition less_stone_phi_pair :: "'a stone_phi_pair \<Rightarrow> 'a stone_phi_pair \<Rightarrow> bool" is triple.pairs_less .
instance ..
end
(*
instantiation stone_phi_pair :: (stone_algebra) stone_algebra
begin
instance
apply intro_classes
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by (metis (no_types, lifting) Pair_inject compl_bot_eq heyting.implies_order stone_phi.pairs_less_eq.elims(3) stone_phi.phi_top stone_phi.triple_axioms sup_top_left top_greatest triple_def)
subgoal apply transfer by (metis (no_types, lifting) Pair_inject stone_phi.pairs_less_eq.elims(3) top.extremum bot_least bot_filter.abs_eq)
subgoal apply transfer using stone_phi.triple_axioms triple.pairs_sup_dist_inf by fastforce
subgoal apply transfer using stone_phi.pairs_uminus_galois by fastforce
subgoal apply transfer using stone_phi.pairs_stone by fastforce
done
end
*)
text \<open>
The result is a Stone algebra and could be proved so by repeating and specialising the above proof for lifted pairs.
We choose a different approach, namely by embedding the type of pairs into the lifted type.
The embedding injects a pair \<open>x\<close> into a function as the value at the given structure map; this makes the embedding injective.
The value of the function at any other structure map needs to be carefully chosen so that the resulting function is a Stone algebra homomorphism.
We use \<open>--x\<close>, which is essentially a projection to the regular element component of \<open>x\<close>, whence the image has the structure of a Boolean algebra.
\<close>
fun stone_phi_embed :: "'a::non_trivial_stone_algebra stone_phi_pair \<Rightarrow> ('a regular,'a dense) lifted_pair"
where "stone_phi_embed x = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x)))"
text \<open>
The following lemma shows that in both cases the value of the function is a valid pair for the given structure map.
\<close>
lemma stone_phi_embed_triple_pair:
"(if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x))) \<in> triple.pairs (Rep_phi f)"
by (metis (no_types, hide_lams) Rep_stone_phi_pair simp_phi surj_pair triple.pairs_uminus_closed triple_def)
text \<open>
The following result shows that the embedding preserves the operations of Stone algebras.
Of course, it is not (yet) a Stone algebra homomorphism as we do not know (yet) that the domain of the embedding is a Stone algebra.
To establish the latter is the purpose of the embedding.
\<close>
lemma stone_phi_embed_homomorphism:
"sup_inf_top_bot_uminus_ord_homomorphism stone_phi_embed"
proof (intro conjI)
let ?p = "\<lambda>f . triple.pairs_uminus (Rep_phi f)"
let ?pp = "\<lambda>f x . ?p f (?p f x)"
let ?q = "\<lambda>f x . ?pp f (Rep_stone_phi_pair x)"
show "\<forall>x y::'a stone_phi_pair . stone_phi_embed (x \<squnion> y) = stone_phi_embed x \<squnion> stone_phi_embed y"
proof (intro allI)
fix x y :: "'a stone_phi_pair"
have 1: "\<forall>f . triple.pairs_sup (?q f x) (?q f y) = ?q f (x \<squnion> y)"
proof
fix f :: "('a regular,'a dense) phi"
let ?r = "Rep_phi f"
obtain x1 x2 y1 y2 where 2: "(x1,x2) = Rep_stone_phi_pair x \<and> (y1,y2) = Rep_stone_phi_pair y"
using prod.collapse by blast
hence "triple.pairs_sup (?q f x) (?q f y) = triple.pairs_sup (?pp f (x1,x2)) (?pp f (y1,y2))"
by simp
also have "... = triple.pairs_sup (--x1,?r (-x1)) (--y1,?r (-y1))"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = (--x1 \<squnion> --y1,?r (-x1) \<sqinter> ?r (-y1))"
by simp
also have "... = (--(x1 \<squnion> y1),?r (-(x1 \<squnion> y1)))"
by simp
also have "... = ?pp f (x1 \<squnion> y1,x2 \<sqinter> y2)"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = ?pp f (triple.pairs_sup (x1,x2) (y1,y2))"
by simp
also have "... = ?q f (x \<squnion> y)"
using 2 by (simp add: sup_stone_phi_pair.rep_eq)
finally show "triple.pairs_sup (?q f x) (?q f y) = ?q f (x \<squnion> y)"
.
qed
have "stone_phi_embed x \<squnion> stone_phi_embed y = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) \<squnion> Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)"
by simp
also have "... = Abs_lifted_pair (\<lambda>f . triple.pairs_sup (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y))"
by (rule sup_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then triple.pairs_sup (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else triple.pairs_sup (?q f x) (?q f y))"
by (simp add: if_distrib_2)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then triple.pairs_sup (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else ?q f (x \<squnion> y))"
using 1 by meson
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair (x \<squnion> y) else ?q f (x \<squnion> y))"
by (metis sup_stone_phi_pair.rep_eq)
also have "... = stone_phi_embed (x \<squnion> y)"
by simp
finally show "stone_phi_embed (x \<squnion> y) = stone_phi_embed x \<squnion> stone_phi_embed y"
by simp
qed
next
let ?p = "\<lambda>f . triple.pairs_uminus (Rep_phi f)"
let ?pp = "\<lambda>f x . ?p f (?p f x)"
let ?q = "\<lambda>f x . ?pp f (Rep_stone_phi_pair x)"
show "\<forall>x y::'a stone_phi_pair . stone_phi_embed (x \<sqinter> y) = stone_phi_embed x \<sqinter> stone_phi_embed y"
proof (intro allI)
fix x y :: "'a stone_phi_pair"
have 1: "\<forall>f . triple.pairs_inf (?q f x) (?q f y) = ?q f (x \<sqinter> y)"
proof
fix f :: "('a regular,'a dense) phi"
let ?r = "Rep_phi f"
obtain x1 x2 y1 y2 where 2: "(x1,x2) = Rep_stone_phi_pair x \<and> (y1,y2) = Rep_stone_phi_pair y"
using prod.collapse by blast
hence "triple.pairs_inf (?q f x) (?q f y) = triple.pairs_inf (?pp f (x1,x2)) (?pp f (y1,y2))"
by simp
also have "... = triple.pairs_inf (--x1,?r (-x1)) (--y1,?r (-y1))"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = (--x1 \<sqinter> --y1,?r (-x1) \<squnion> ?r (-y1))"
by simp
also have "... = (--(x1 \<sqinter> y1),?r (-(x1 \<sqinter> y1)))"
by simp
also have "... = ?pp f (x1 \<sqinter> y1,x2 \<squnion> y2)"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = ?pp f (triple.pairs_inf (x1,x2) (y1,y2))"
by simp
also have "... = ?q f (x \<sqinter> y)"
using 2 by (simp add: inf_stone_phi_pair.rep_eq)
finally show "triple.pairs_inf (?q f x) (?q f y) = ?q f (x \<sqinter> y)"
.
qed
have "stone_phi_embed x \<sqinter> stone_phi_embed y = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) \<sqinter> Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)"
by simp
also have "... = Abs_lifted_pair (\<lambda>f . triple.pairs_inf (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y))"
by (rule inf_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then triple.pairs_inf (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else triple.pairs_inf (?q f x) (?q f y))"
by (simp add: if_distrib_2)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then triple.pairs_inf (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else ?q f (x \<sqinter> y))"
using 1 by meson
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair (x \<sqinter> y) else ?q f (x \<sqinter> y))"
by (metis inf_stone_phi_pair.rep_eq)
also have "... = stone_phi_embed (x \<sqinter> y)"
by simp
finally show "stone_phi_embed (x \<sqinter> y) = stone_phi_embed x \<sqinter> stone_phi_embed y"
by simp
qed
next
have "stone_phi_embed (top::'a stone_phi_pair) = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair top else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair top)))"
by simp
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then (top,bot) else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (top,bot)))"
by (metis (no_types, hide_lams) bot_filter.abs_eq top_stone_phi_pair.rep_eq)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then (top,bot) else triple.pairs_uminus (Rep_phi f) (bot,top))"
by (metis (no_types, hide_lams) dense_closed_top simp_phi triple.pairs_uminus.simps triple_def)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then (top,bot) else (top,bot))"
by (metis (no_types, hide_lams) p_bot simp_phi triple.pairs_uminus.simps triple_def)
also have "... = Abs_lifted_pair (\<lambda>f . (top,Abs_filter {top}))"
by (simp add: bot_filter.abs_eq)
also have "... = top"
by (rule top_lifted_pair.abs_eq[THEN sym])
finally show "stone_phi_embed (top::'a stone_phi_pair) = top"
.
next
have "stone_phi_embed (bot::'a stone_phi_pair) = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair bot else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair bot)))"
by simp
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then (bot,top) else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (bot,top)))"
by (metis (no_types, hide_lams) top_filter.abs_eq bot_stone_phi_pair.rep_eq)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then (bot,top) else triple.pairs_uminus (Rep_phi f) (top,bot))"
by (metis (no_types, hide_lams) p_bot simp_phi triple.pairs_uminus.simps triple_def)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then (bot,top) else (bot,top))"
by (metis (no_types, hide_lams) p_top simp_phi triple.pairs_uminus.simps triple_def)
also have "... = Abs_lifted_pair (\<lambda>f . (bot,Abs_filter UNIV))"
by (simp add: top_filter.abs_eq)
also have "... = bot"
by (rule bot_lifted_pair.abs_eq[THEN sym])
finally show "stone_phi_embed (bot::'a stone_phi_pair) = bot"
.
next
let ?p = "\<lambda>f . triple.pairs_uminus (Rep_phi f)"
let ?pp = "\<lambda>f x . ?p f (?p f x)"
let ?q = "\<lambda>f x . ?pp f (Rep_stone_phi_pair x)"
show "\<forall>x::'a stone_phi_pair . stone_phi_embed (-x) = -stone_phi_embed x"
proof (intro allI)
fix x :: "'a stone_phi_pair"
have 1: "\<forall>f . triple.pairs_uminus (Rep_phi f) (?q f x) = ?q f (-x)"
proof
fix f :: "('a regular,'a dense) phi"
let ?r = "Rep_phi f"
obtain x1 x2 where 2: "(x1,x2) = Rep_stone_phi_pair x"
using prod.collapse by blast
hence "triple.pairs_uminus (Rep_phi f) (?q f x) = triple.pairs_uminus (Rep_phi f) (?pp f (x1,x2))"
by simp
also have "... = triple.pairs_uminus (Rep_phi f) (--x1,?r (-x1))"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = (---x1,?r (--x1))"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = ?pp f (-x1,stone_phi x1)"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = ?pp f (triple.pairs_uminus stone_phi (x1,x2))"
by simp
also have "... = ?q f (-x)"
using 2 by (simp add: uminus_stone_phi_pair.rep_eq)
finally show "triple.pairs_uminus (Rep_phi f) (?q f x) = ?q f (-x)"
.
qed
have "-stone_phi_embed x = -Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x)"
by simp
also have "... = Abs_lifted_pair (\<lambda>f . triple.pairs_uminus (Rep_phi f) (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x))"
by (rule uminus_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x) else triple.pairs_uminus (Rep_phi f) (?q f x))"
by (simp add: if_distrib)
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x) else ?q f (-x))"
using 1 by meson
also have "... = Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair (-x) else ?q f (-x))"
by (metis uminus_stone_phi_pair.rep_eq)
also have "... = stone_phi_embed (-x)"
by simp
finally show "stone_phi_embed (-x) = -stone_phi_embed x"
by simp
qed
next
let ?p = "\<lambda>f . triple.pairs_uminus (Rep_phi f)"
let ?pp = "\<lambda>f x . ?p f (?p f x)"
let ?q = "\<lambda>f x . ?pp f (Rep_stone_phi_pair x)"
show "\<forall>x y::'a stone_phi_pair . x \<le> y \<longrightarrow> stone_phi_embed x \<le> stone_phi_embed y"
proof (intro allI, rule impI)
fix x y :: "'a stone_phi_pair"
assume 1: "x \<le> y"
have "\<forall>f . triple.pairs_less_eq (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)"
proof
fix f :: "('a regular,'a dense) phi"
let ?r = "Rep_phi f"
obtain x1 x2 y1 y2 where 2: "(x1,x2) = Rep_stone_phi_pair x \<and> (y1,y2) = Rep_stone_phi_pair y"
using prod.collapse by blast
have "x1 \<le> y1"
using 1 2 by (metis less_eq_stone_phi_pair.rep_eq stone_phi.pairs_less_eq.simps)
hence "--x1 \<le> --y1 \<and> ?r (-y1) \<le> ?r (-x1)"
by (metis compl_le_compl_iff le_iff_sup simp_phi)
hence "triple.pairs_less_eq (--x1,?r (-x1)) (--y1,?r (-y1))"
by simp
hence "triple.pairs_less_eq (?pp f (x1,x2)) (?pp f (y1,y2))"
by (simp add: triple.pairs_uminus.simps triple_def)
hence "triple.pairs_less_eq (?q f x) (?q f y)"
using 2 by simp
hence "if ?r = stone_phi then triple.pairs_less_eq (Rep_stone_phi_pair x) (Rep_stone_phi_pair y) else triple.pairs_less_eq (?q f x) (?q f y)"
using 1 by (simp add: less_eq_stone_phi_pair.rep_eq)
thus "triple.pairs_less_eq (if ?r = stone_phi then Rep_stone_phi_pair x else ?q f x) (if ?r = stone_phi then Rep_stone_phi_pair y else ?q f y)"
by (simp add: if_distrib_2)
qed
hence "Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else ?q f x) \<le> Abs_lifted_pair (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else ?q f y)"
by (subst less_eq_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair)
thus "stone_phi_embed x \<le> stone_phi_embed y"
by simp
qed
qed
text \<open>
The following lemmas show that the embedding is injective and reflects the order.
The latter allows us to easily inherit properties involving inequalities from the target of the embedding, without transforming them to equations.
\<close>
lemma stone_phi_embed_injective:
"inj stone_phi_embed"
proof (rule injI)
fix x y :: "'a stone_phi_pair"
have 1: "Rep_phi (Abs_phi stone_phi) = stone_phi"
by (simp add: Abs_phi_inverse stone_phi.hom)
assume 2: "stone_phi_embed x = stone_phi_embed y"
have "\<forall>x::'a stone_phi_pair . Rep_lifted_pair (stone_phi_embed x) = (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x)))"
by (simp add: Abs_lifted_pair_inverse stone_phi_embed_triple_pair)
hence "(\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x))) = (\<lambda>f . if Rep_phi f = stone_phi then Rep_stone_phi_pair y else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair y)))"
using 2 by metis
hence "Rep_stone_phi_pair x = Rep_stone_phi_pair y"
using 1 by metis
thus "x = y"
by (simp add: Rep_stone_phi_pair_inject)
qed
lemma stone_phi_embed_order_injective:
assumes "stone_phi_embed x \<le> stone_phi_embed y"
shows "x \<le> y"
proof -
let ?f = "Abs_phi stone_phi"
have "\<forall>f . triple.pairs_less_eq (if Rep_phi f = stone_phi then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair x))) (if Rep_phi f = stone_phi then Rep_stone_phi_pair y else triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_stone_phi_pair y)))"
using assms by (subst less_eq_lifted_pair.abs_eq[THEN sym]) (simp_all add: eq_onp_same_args stone_phi_embed_triple_pair)
hence "triple.pairs_less_eq (if Rep_phi ?f = (stone_phi::'a regular \<Rightarrow> 'a dense_filter) then Rep_stone_phi_pair x else triple.pairs_uminus (Rep_phi ?f) (triple.pairs_uminus (Rep_phi ?f) (Rep_stone_phi_pair x))) (if Rep_phi ?f = (stone_phi::'a regular \<Rightarrow> 'a dense_filter) then Rep_stone_phi_pair y else triple.pairs_uminus (Rep_phi ?f) (triple.pairs_uminus (Rep_phi ?f) (Rep_stone_phi_pair y)))"
by simp
hence "triple.pairs_less_eq (Rep_stone_phi_pair x) (Rep_stone_phi_pair y)"
by (simp add: Abs_phi_inverse stone_phi.hom)
thus "x \<le> y"
by (simp add: less_eq_stone_phi_pair.rep_eq)
qed
lemma stone_phi_embed_strict_order_isomorphism:
"x < y \<longleftrightarrow> stone_phi_embed x < stone_phi_embed y"
by (smt less_eq_stone_phi_pair.rep_eq less_le_not_le less_stone_phi_pair.rep_eq stone_phi.pairs_less.elims(2,3) stone_phi_embed_homomorphism stone_phi_embed_order_injective)
text \<open>
Now all Stone algebra axioms can be inherited using the embedding.
This is due to the fact that the axioms are universally quantified equations or conditional equations (or inequalities); this is called a quasivariety in universal algebra.
It would be useful to have this construction available for arbitrary quasivarieties.
\<close>
instantiation stone_phi_pair :: (non_trivial_stone_algebra) stone_algebra
begin
instance
apply intro_classes
apply (metis (mono_tags, lifting) stone_phi_embed_homomorphism stone_phi_embed_strict_order_isomorphism stone_phi_embed_order_injective less_le_not_le)
apply (simp add: stone_phi_embed_order_injective)
apply (meson order.trans stone_phi_embed_homomorphism stone_phi_embed_order_injective)
apply (meson stone_phi_embed_homomorphism antisym stone_phi_embed_injective injD)
apply (metis inf.sup_ge1 stone_phi_embed_homomorphism stone_phi_embed_order_injective)
apply (metis inf.sup_ge2 stone_phi_embed_homomorphism stone_phi_embed_order_injective)
apply (metis inf_greatest stone_phi_embed_homomorphism stone_phi_embed_order_injective)
apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective sup_ge1)
apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective sup.cobounded2)
apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective sup_least)
apply (metis bot.extremum stone_phi_embed_homomorphism stone_phi_embed_order_injective)
apply (metis stone_phi_embed_homomorphism stone_phi_embed_order_injective top_greatest)
apply (metis (mono_tags, lifting) stone_phi_embed_homomorphism sup_inf_distrib1 stone_phi_embed_injective injD)
apply (metis stone_phi_embed_homomorphism stone_phi_embed_injective injD stone_phi_embed_order_injective pseudo_complement)
by (metis injD stone_phi_embed_homomorphism stone_phi_embed_injective stone)
end
subsection \<open>Stone Algebra Isomorphism\<close>
text \<open>
In this section we prove that the Stone algebra of the triple of a Stone algebra is isomorphic to the original Stone algebra.
The following two definitions give the isomorphism.
\<close>
abbreviation sa_iso_inv :: "'a::non_trivial_stone_algebra stone_phi_pair \<Rightarrow> 'a"
where "sa_iso_inv \<equiv> \<lambda>p . Rep_regular (fst (Rep_stone_phi_pair p)) \<sqinter> Rep_dense (triple.rho_pair stone_phi (Rep_stone_phi_pair p))"
abbreviation sa_iso :: "'a::non_trivial_stone_algebra \<Rightarrow> 'a stone_phi_pair"
where "sa_iso \<equiv> \<lambda>x . Abs_stone_phi_pair (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x)))"
lemma sa_iso_triple_pair:
"(Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<in> triple.pairs stone_phi"
by (metis (mono_tags, lifting) double_compl eq_onp_same_args stone_phi.sa_iso_pair uminus_regular.abs_eq)
lemma stone_phi_inf_dense:
"stone_phi (Abs_regular (-x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y)) \<le> up_filter (Abs_dense (y \<squnion> -y \<squnion> x))"
proof -
have "Rep_filter (stone_phi (Abs_regular (-x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y))) \<le> \<up>(Abs_dense (y \<squnion> -y \<squnion> x))"
proof
fix z :: "'a dense"
let ?r = "Rep_dense z"
assume "z \<in> Rep_filter (stone_phi (Abs_regular (-x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y)))"
also have "... = Rep_filter (stone_phi (Abs_regular (-x))) \<inter> Rep_filter (up_filter (Abs_dense (y \<squnion> -y)))"
by (simp add: inf_filter.rep_eq)
also have "... = stone_phi_base (Abs_regular (-x)) \<inter> \<up>(Abs_dense (y \<squnion> -y))"
by (metis Abs_filter_inverse mem_Collect_eq up_filter stone_phi_base_filter stone_phi_def)
finally have "--x \<le> ?r \<and> Abs_dense (y \<squnion> -y) \<le> z"
by (metis (mono_tags, lifting) Abs_regular_inverse Int_Collect mem_Collect_eq)
hence "--x \<le> ?r \<and> y \<squnion> -y \<le> ?r"
by (simp add: Abs_dense_inverse less_eq_dense.rep_eq)
hence "y \<squnion> -y \<squnion> x \<le> ?r"
using order_trans pp_increasing by auto
hence "Abs_dense (y \<squnion> -y \<squnion> x) \<le> Abs_dense ?r"
by (subst less_eq_dense.abs_eq) (simp_all add: eq_onp_same_args)
thus "z \<in> \<up>(Abs_dense (y \<squnion> -y \<squnion> x))"
by (simp add: Rep_dense_inverse)
qed
hence "Abs_filter (Rep_filter (stone_phi (Abs_regular (-x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y)))) \<le> up_filter (Abs_dense (y \<squnion> -y \<squnion> x))"
by (simp add: eq_onp_same_args less_eq_filter.abs_eq)
thus ?thesis
by (simp add: Rep_filter_inverse)
qed
lemma stone_phi_complement:
"complement (stone_phi (Abs_regular (-x))) (stone_phi (Abs_regular (--x)))"
by (metis (mono_tags, lifting) eq_onp_same_args stone_phi.phi_complemented uminus_regular.abs_eq)
lemma up_dense_stone_phi:
"up_filter (Abs_dense (x \<squnion> -x)) \<le> stone_phi (Abs_regular (--x))"
proof -
have "\<up>(Abs_dense (x \<squnion> -x)) \<le> stone_phi_base (Abs_regular (--x))"
proof
fix z :: "'a dense"
let ?r = "Rep_dense z"
assume "z \<in> \<up>(Abs_dense (x \<squnion> -x))"
hence "---x \<le> ?r"
by (simp add: Abs_dense_inverse less_eq_dense.rep_eq)
hence "-Rep_regular (Abs_regular (--x)) \<le> ?r"
by (metis (mono_tags, lifting) Abs_regular_inverse mem_Collect_eq)
thus "z \<in> stone_phi_base (Abs_regular (--x))"
by simp
qed
thus ?thesis
by (unfold stone_phi_def, subst less_eq_filter.abs_eq, simp_all add: eq_onp_same_args stone_phi_base_filter)
qed
text \<open>
The following two results prove that the isomorphisms are mutually inverse.
\<close>
lemma sa_iso_left_invertible:
"sa_iso_inv (sa_iso x) = x"
proof -
have "up_filter (triple.rho_pair stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x)))) = stone_phi (Abs_regular (--x)) \<sqinter> (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x)))"
using sa_iso_triple_pair stone_phi.get_rho_pair_char by blast
also have "... = stone_phi (Abs_regular (--x)) \<sqinter> up_filter (Abs_dense (x \<squnion> -x))"
by (simp add: inf.sup_commute inf_sup_distrib1 stone_phi_complement)
also have "... = up_filter (Abs_dense (x \<squnion> -x))"
using up_dense_stone_phi inf.absorb2 by auto
finally have 1: "triple.rho_pair stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) = Abs_dense (x \<squnion> -x)"
using up_filter_injective by auto
have "sa_iso_inv (sa_iso x) = (\<lambda>p . Rep_regular (fst p) \<sqinter> Rep_dense (triple.rho_pair stone_phi p)) (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x)))"
by (simp add: Abs_stone_phi_pair_inverse sa_iso_triple_pair)
also have "... = Rep_regular (Abs_regular (--x)) \<sqinter> Rep_dense (triple.rho_pair stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))))"
by simp
also have "... = --x \<sqinter> Rep_dense (Abs_dense (x \<squnion> -x))"
using 1 by (subst Abs_regular_inverse) auto
also have "... = --x \<sqinter> (x \<squnion> -x)"
by (subst Abs_dense_inverse) simp_all
also have "... = x"
by simp
finally show ?thesis
by auto
qed
lemma sa_iso_right_invertible:
"sa_iso (sa_iso_inv p) = p"
proof -
obtain x y where 1: "(x,y) = Rep_stone_phi_pair p"
using prod.collapse by blast
hence 2: "(x,y) \<in> triple.pairs stone_phi"
by (simp add: Rep_stone_phi_pair)
hence 3: "stone_phi (-x) \<le> y"
by (simp add: stone_phi.pairs_phi_less_eq)
have 4: "\<forall>z . z \<in> Rep_filter (stone_phi x \<sqinter> y) \<longrightarrow> -Rep_regular x \<le> Rep_dense z"
proof (rule allI, rule impI)
fix z :: "'a dense"
let ?r = "Rep_dense z"
assume "z \<in> Rep_filter (stone_phi x \<sqinter> y)"
hence "z \<in> Rep_filter (stone_phi x)"
by (simp add: inf_filter.rep_eq)
also have "... = stone_phi_base x"
by (simp add: stone_phi_def Abs_filter_inverse stone_phi_base_filter)
finally show "-Rep_regular x \<le> ?r"
by simp
qed
have "triple.rho_pair stone_phi (x,y) \<in> \<up>(triple.rho_pair stone_phi (x,y))"
by simp
also have "... = Rep_filter (Abs_filter (\<up>(triple.rho_pair stone_phi (x,y))))"
by (simp add: Abs_filter_inverse)
also have "... = Rep_filter (stone_phi x \<sqinter> y)"
using 2 stone_phi.get_rho_pair_char by fastforce
finally have "triple.rho_pair stone_phi (x,y) \<in> Rep_filter (stone_phi x \<sqinter> y)"
by simp
hence 5: "-Rep_regular x \<le> Rep_dense (triple.rho_pair stone_phi (x,y))"
using 4 by simp
have 6: "sa_iso_inv p = Rep_regular x \<sqinter> Rep_dense (triple.rho_pair stone_phi (x,y))"
using 1 by (metis fstI)
hence "-sa_iso_inv p = -Rep_regular x"
by simp
hence "sa_iso (sa_iso_inv p) = Abs_stone_phi_pair (Abs_regular (--Rep_regular x),stone_phi (Abs_regular (-Rep_regular x)) \<squnion> up_filter (Abs_dense ((Rep_regular x \<sqinter> Rep_dense (triple.rho_pair stone_phi (x,y))) \<squnion> -Rep_regular x)))"
using 6 by simp
also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \<squnion> up_filter (Abs_dense ((Rep_regular x \<sqinter> Rep_dense (triple.rho_pair stone_phi (x,y))) \<squnion> -Rep_regular x)))"
by (metis (mono_tags, lifting) Rep_regular_inverse double_compl uminus_regular.rep_eq)
also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \<squnion> up_filter (Abs_dense (Rep_dense (triple.rho_pair stone_phi (x,y)) \<squnion> -Rep_regular x)))"
by (metis inf_sup_aci(5) maddux_3_21_pp simp_regular)
also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \<squnion> up_filter (Abs_dense (Rep_dense (triple.rho_pair stone_phi (x,y)))))"
using 5 by (simp add: sup.absorb1)
also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \<squnion> up_filter (triple.rho_pair stone_phi (x,y)))"
by (simp add: Rep_dense_inverse)
also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \<squnion> (stone_phi x \<sqinter> y))"
using 2 stone_phi.get_rho_pair_char by fastforce
also have "... = Abs_stone_phi_pair (x,stone_phi (-x) \<squnion> y)"
by (simp add: stone_phi.phi_complemented sup.commute sup_inf_distrib1)
also have "... = Abs_stone_phi_pair (x,y)"
using 3 by (simp add: le_iff_sup)
also have "... = p"
using 1 by (simp add: Rep_stone_phi_pair_inverse)
finally show ?thesis
.
qed
text \<open>
It remains to show the homomorphism properties, which is done in the following result.
\<close>
lemma sa_iso:
"stone_algebra_isomorphism sa_iso"
proof (intro conjI)
have "Abs_stone_phi_pair (Abs_regular (--bot),stone_phi (Abs_regular (-bot)) \<squnion> up_filter (Abs_dense (bot \<squnion> -bot))) = Abs_stone_phi_pair (bot,stone_phi top \<squnion> up_filter top)"
by (simp add: bot_regular.abs_eq top_regular.abs_eq top_dense.abs_eq)
also have "... = Abs_stone_phi_pair (bot,stone_phi top)"
by (simp add: stone_phi.hom)
also have "... = bot"
by (simp add: bot_stone_phi_pair_def stone_phi.phi_top)
finally show "sa_iso bot = bot"
.
next
have "Abs_stone_phi_pair (Abs_regular (--top),stone_phi (Abs_regular (-top)) \<squnion> up_filter (Abs_dense (top \<squnion> -top))) = Abs_stone_phi_pair (top,stone_phi bot \<squnion> up_filter top)"
by (simp add: bot_regular.abs_eq top_regular.abs_eq top_dense.abs_eq)
also have "... = top"
by (simp add: stone_phi.phi_bot top_stone_phi_pair_def)
finally show "sa_iso top = top"
.
next
have 1: "\<forall>x y::'a . dense (x \<squnion> -x \<squnion> y)"
by simp
have 2: "\<forall>x y::'a . up_filter (Abs_dense (x \<squnion> -x \<squnion> y)) \<le> (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<sqinter> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
proof (intro allI)
fix x y :: 'a
let ?u = "Abs_dense (x \<squnion> -x \<squnion> --y)"
let ?v = "Abs_dense (y \<squnion> -y)"
have "\<up>(Abs_dense (x \<squnion> -x \<squnion> y)) \<le> Rep_filter (stone_phi (Abs_regular (-y)) \<squnion> up_filter ?v)"
proof
fix z
assume "z \<in> \<up>(Abs_dense (x \<squnion> -x \<squnion> y))"
hence "Abs_dense (x \<squnion> -x \<squnion> y) \<le> z"
by simp
hence 3: "x \<squnion> -x \<squnion> y \<le> Rep_dense z"
by (simp add: Abs_dense_inverse less_eq_dense.rep_eq)
have "y \<le> x \<squnion> -x \<squnion> --y"
by (simp add: le_supI2 pp_increasing)
hence "(x \<squnion> -x \<squnion> --y) \<sqinter> (y \<squnion> -y) = y \<squnion> ((x \<squnion> -x \<squnion> --y) \<sqinter> -y)"
by (simp add: le_iff_sup sup_inf_distrib1)
also have "... = y \<squnion> ((x \<squnion> -x) \<sqinter> -y)"
by (simp add: inf_commute inf_sup_distrib1)
also have "... \<le> Rep_dense z"
using 3 by (meson le_infI1 sup.bounded_iff)
finally have "Abs_dense ((x \<squnion> -x \<squnion> --y) \<sqinter> (y \<squnion> -y)) \<le> z"
by (simp add: Abs_dense_inverse less_eq_dense.rep_eq)
hence 4: "?u \<sqinter> ?v \<le> z"
by (simp add: eq_onp_same_args inf_dense.abs_eq)
have "-Rep_regular (Abs_regular (-y)) = --y"
by (metis (mono_tags, lifting) mem_Collect_eq Abs_regular_inverse)
also have "... \<le> Rep_dense ?u"
by (simp add: Abs_dense_inverse)
finally have "?u \<in> stone_phi_base (Abs_regular (-y))"
by simp
hence 5: "?u \<in> Rep_filter (stone_phi (Abs_regular (-y)))"
by (metis mem_Collect_eq stone_phi_def stone_phi_base_filter Abs_filter_inverse)
have "?v \<in> \<up>?v"
by simp
hence "?v \<in> Rep_filter (up_filter ?v)"
by (metis Abs_filter_inverse mem_Collect_eq up_filter)
thus "z \<in> Rep_filter (stone_phi (Abs_regular (-y)) \<squnion> up_filter ?v)"
using 4 5 sup_filter.rep_eq filter_sup_def by blast
qed
hence "up_filter (Abs_dense (x \<squnion> -x \<squnion> y)) \<le> Abs_filter (Rep_filter (stone_phi (Abs_regular (-y)) \<squnion> up_filter ?v))"
by (simp add: eq_onp_same_args less_eq_filter.abs_eq)
also have "... = stone_phi (Abs_regular (-y)) \<squnion> up_filter ?v"
by (simp add: Rep_filter_inverse)
finally show "up_filter (Abs_dense (x \<squnion> -x \<squnion> y)) \<le> (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<sqinter> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
by (metis le_infI le_supI2 sup_bot.right_neutral up_filter_dense_antitone)
qed
have 6: "\<forall>x::'a . in_p_image (-x)"
by auto
show "\<forall>x y::'a . sa_iso (x \<squnion> y) = sa_iso x \<squnion> sa_iso y"
proof (intro allI)
fix x y :: 'a
have 7: "up_filter (Abs_dense (x \<squnion> -x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y)) \<le> up_filter (Abs_dense (y \<squnion> -y \<squnion> x))"
proof -
have "up_filter (Abs_dense (x \<squnion> -x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y)) = up_filter (Abs_dense (x \<squnion> -x) \<squnion> Abs_dense (y \<squnion> -y))"
by (metis up_filter_dist_sup)
also have "... = up_filter (Abs_dense (x \<squnion> -x \<squnion> (y \<squnion> -y)))"
by (subst sup_dense.abs_eq) (simp_all add: eq_onp_same_args)
also have "... = up_filter (Abs_dense (y \<squnion> -y \<squnion> x \<squnion> -x))"
by (simp add: sup_commute sup_left_commute)
also have "... \<le> up_filter (Abs_dense (y \<squnion> -y \<squnion> x))"
using up_filter_dense_antitone by auto
finally show ?thesis
.
qed
have "Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y)) = Abs_dense ((x \<squnion> -x \<squnion> y) \<sqinter> (y \<squnion> -y \<squnion> x))"
by (simp add: sup_commute sup_inf_distrib1 sup_left_commute)
also have "... = Abs_dense (x \<squnion> -x \<squnion> y) \<sqinter> Abs_dense (y \<squnion> -y \<squnion> x)"
using 1 by (metis (mono_tags, lifting) Abs_dense_inverse Rep_dense_inverse inf_dense.rep_eq mem_Collect_eq)
finally have 8: "up_filter (Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y))) = up_filter (Abs_dense (x \<squnion> -x \<squnion> y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y \<squnion> x))"
by (simp add: up_filter_dist_inf)
also have "... \<le> (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<sqinter> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
using 2 by (simp add: inf.sup_commute le_sup_iff)
finally have 9: "(stone_phi (Abs_regular (-x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> up_filter (Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y))) \<le> ..."
by (simp add: le_supI1)
have "... = (stone_phi (Abs_regular (-x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> (stone_phi (Abs_regular (-x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y))) \<squnion> ((up_filter (Abs_dense (x \<squnion> -x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> (up_filter (Abs_dense (x \<squnion> -x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y))))"
by (metis (no_types) inf_sup_distrib1 inf_sup_distrib2)
also have "... \<le> (stone_phi (Abs_regular (-x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> up_filter (Abs_dense (y \<squnion> -y \<squnion> x)) \<squnion> ((up_filter (Abs_dense (x \<squnion> -x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> (up_filter (Abs_dense (x \<squnion> -x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y))))"
by (meson sup_left_isotone sup_right_isotone stone_phi_inf_dense)
also have "... \<le> (stone_phi (Abs_regular (-x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> up_filter (Abs_dense (y \<squnion> -y \<squnion> x)) \<squnion> (up_filter (Abs_dense (x \<squnion> -x \<squnion> y)) \<squnion> (up_filter (Abs_dense (x \<squnion> -x)) \<sqinter> up_filter (Abs_dense (y \<squnion> -y))))"
by (metis inf.commute sup_left_isotone sup_right_isotone stone_phi_inf_dense)
also have "... \<le> (stone_phi (Abs_regular (-x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> up_filter (Abs_dense (y \<squnion> -y \<squnion> x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x \<squnion> y))"
using 7 by (simp add: sup.absorb1 sup_commute sup_left_commute)
also have "... = (stone_phi (Abs_regular (-x)) \<sqinter> stone_phi (Abs_regular (-y))) \<squnion> up_filter (Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y)))"
using 8 by (simp add: sup.commute sup.left_commute)
finally have "(stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<sqinter> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y))) = ..."
using 9 using antisym by blast
also have "... = stone_phi (Abs_regular (-x) \<sqinter> Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y)))"
by (simp add: stone_phi.hom)
also have "... = stone_phi (Abs_regular (-(x \<squnion> y))) \<squnion> up_filter (Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y)))"
using 6 by (subst inf_regular.abs_eq) (simp_all add: eq_onp_same_args)
finally have 10: "stone_phi (Abs_regular (-(x \<squnion> y))) \<squnion> up_filter (Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y))) = (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<sqinter> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
by simp
have "Abs_regular (--(x \<squnion> y)) = Abs_regular (--x) \<squnion> Abs_regular (--y)"
using 6 by (subst sup_regular.abs_eq) (simp_all add: eq_onp_same_args)
hence "Abs_stone_phi_pair (Abs_regular (--(x \<squnion> y)),stone_phi (Abs_regular (-(x \<squnion> y))) \<squnion> up_filter (Abs_dense (x \<squnion> y \<squnion> -(x \<squnion> y)))) = Abs_stone_phi_pair (triple.pairs_sup (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) (Abs_regular (--y),stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y))))"
using 10 by auto
also have "... = Abs_stone_phi_pair (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<squnion> Abs_stone_phi_pair (Abs_regular (--y),stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
by (rule sup_stone_phi_pair.abs_eq[THEN sym]) (simp_all add: eq_onp_same_args sa_iso_triple_pair)
finally show "sa_iso (x \<squnion> y) = sa_iso x \<squnion> sa_iso y"
.
qed
next
have 1: "\<forall>x y::'a . dense (x \<squnion> -x \<squnion> y)"
by simp
have 2: "\<forall>x::'a . in_p_image (-x)"
by auto
have 3: "\<forall>x y::'a . stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> -x)) = stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> -x \<squnion> -y))"
proof (intro allI)
fix x y :: 'a
have 4: "up_filter (Abs_dense (x \<squnion> -x)) \<le> stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> -x \<squnion> -y))"
by (metis (no_types, lifting) complement_shunting stone_phi_inf_dense stone_phi_complement complement_symmetric)
have "up_filter (Abs_dense (x \<squnion> -x \<squnion> -y)) \<le> up_filter (Abs_dense (x \<squnion> -x))"
by (metis sup_idem up_filter_dense_antitone)
thus "stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> -x)) = stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> -x \<squnion> -y))"
using 4 by (simp add: le_iff_sup sup_commute sup_left_commute)
qed
show "\<forall>x y::'a . sa_iso (x \<sqinter> y) = sa_iso x \<sqinter> sa_iso y"
proof (intro allI)
fix x y :: 'a
have "Abs_dense ((x \<sqinter> y) \<squnion> -(x \<sqinter> y)) = Abs_dense ((x \<squnion> -x \<squnion> -y) \<sqinter> (y \<squnion> -y \<squnion> -x))"
by (simp add: sup_commute sup_inf_distrib1 sup_left_commute)
also have "... = Abs_dense (x \<squnion> -x \<squnion> -y) \<sqinter> Abs_dense (y \<squnion> -y \<squnion> -x)"
using 1 by (metis (mono_tags, lifting) Abs_dense_inverse Rep_dense_inverse inf_dense.rep_eq mem_Collect_eq)
finally have 5: "up_filter (Abs_dense ((x \<sqinter> y) \<squnion> -(x \<sqinter> y))) = up_filter (Abs_dense (x \<squnion> -x \<squnion> -y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y \<squnion> -x))"
by (simp add: up_filter_dist_inf)
have "(stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<squnion> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y))) = (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<squnion> (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
by (simp add: inf_sup_aci(6) sup_left_commute)
also have "... = (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (x \<squnion> -x \<squnion> -y))) \<squnion> (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (y \<squnion> -y \<squnion> -x)))"
using 3 by simp
also have "... = (stone_phi (Abs_regular (-x)) \<squnion> stone_phi (Abs_regular (-y))) \<squnion> (up_filter (Abs_dense (x \<squnion> -x \<squnion> -y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y \<squnion> -x)))"
by (simp add: inf_sup_aci(6) sup_left_commute)
also have "... = (stone_phi (Abs_regular (-x)) \<squnion> stone_phi (Abs_regular (-y))) \<squnion> up_filter (Abs_dense ((x \<sqinter> y) \<squnion> -(x \<sqinter> y)))"
using 5 by (simp add: sup.commute sup.left_commute)
finally have "(stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<squnion> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y))) = ..."
by simp
also have "... = stone_phi (Abs_regular (-x) \<squnion> Abs_regular (-y)) \<squnion> up_filter (Abs_dense ((x \<sqinter> y) \<squnion> -(x \<sqinter> y)))"
by (simp add: stone_phi.hom)
also have "... = stone_phi (Abs_regular (-(x \<sqinter> y))) \<squnion> up_filter (Abs_dense ((x \<sqinter> y) \<squnion> -(x \<sqinter> y)))"
using 2 by (subst sup_regular.abs_eq) (simp_all add: eq_onp_same_args)
finally have 6: "stone_phi (Abs_regular (-(x \<sqinter> y))) \<squnion> up_filter (Abs_dense ((x \<sqinter> y) \<squnion> -(x \<sqinter> y))) = (stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<squnion> (stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
by simp
have "Abs_regular (--(x \<sqinter> y)) = Abs_regular (--x) \<sqinter> Abs_regular (--y)"
using 2 by (subst inf_regular.abs_eq) (simp_all add: eq_onp_same_args)
hence "Abs_stone_phi_pair (Abs_regular (--(x \<sqinter> y)),stone_phi (Abs_regular (-(x \<sqinter> y))) \<squnion> up_filter (Abs_dense ((x \<sqinter> y) \<squnion> -(x \<sqinter> y)))) = Abs_stone_phi_pair (triple.pairs_inf (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) (Abs_regular (--y),stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y))))"
using 6 by auto
also have "... = Abs_stone_phi_pair (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))) \<sqinter> Abs_stone_phi_pair (Abs_regular (--y),stone_phi (Abs_regular (-y)) \<squnion> up_filter (Abs_dense (y \<squnion> -y)))"
by (rule inf_stone_phi_pair.abs_eq[THEN sym]) (simp_all add: eq_onp_same_args sa_iso_triple_pair)
finally show "sa_iso (x \<sqinter> y) = sa_iso x \<sqinter> sa_iso y"
.
qed
next
show "\<forall>x::'a . sa_iso (-x) = -sa_iso x"
proof
fix x :: 'a
have "sa_iso (-x) = Abs_stone_phi_pair (Abs_regular (---x),stone_phi (Abs_regular (--x)) \<squnion> up_filter top)"
by (simp add: top_dense_def)
also have "... = Abs_stone_phi_pair (Abs_regular (---x),stone_phi (Abs_regular (--x)))"
by (metis bot_filter.abs_eq sup_bot.right_neutral up_top)
also have "... = Abs_stone_phi_pair (triple.pairs_uminus stone_phi (Abs_regular (--x),stone_phi (Abs_regular (-x)) \<squnion> up_filter (Abs_dense (x \<squnion> -x))))"
by (subst uminus_regular.abs_eq[THEN sym], unfold eq_onp_same_args) auto
also have "... = -sa_iso x"
by (simp add: eq_onp_def sa_iso_triple_pair uminus_stone_phi_pair.abs_eq)
finally show "sa_iso (-x) = -sa_iso x"
by simp
qed
next
show "bij sa_iso"
by (metis (mono_tags, lifting) sa_iso_left_invertible sa_iso_right_invertible invertible_bij[where g=sa_iso_inv])
qed
subsection \<open>Triple Isomorphism\<close>
text \<open>
In this section we prove that the triple of the Stone algebra of a triple is isomorphic to the original triple.
The notion of isomorphism for triples is described in \cite{ChenGraetzer1969}.
It amounts to an isomorphism of Boolean algebras, an isomorphism of distributive lattices with a greatest element, and a commuting diagram involving the structure maps.
\<close>
subsubsection \<open>Boolean Algebra Isomorphism\<close>
text \<open>
We first define and prove the isomorphism of Boolean algebras.
Because the Stone algebra of a triple is implemented as a lifted pair, we also lift the Boolean algebra.
\<close>
typedef (overloaded) ('a,'b) lifted_boolean_algebra = "{ xf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) phi \<Rightarrow> 'a . True }"
by simp
setup_lifting type_definition_lifted_boolean_algebra
instantiation lifted_boolean_algebra :: (non_trivial_boolean_algebra,distrib_lattice_top) boolean_algebra
begin
lift_definition sup_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra" is "\<lambda>xf yf f . sup (xf f) (yf f)" .
lift_definition inf_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra" is "\<lambda>xf yf f . inf (xf f) (yf f)" .
lift_definition minus_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra" is "\<lambda>xf yf f . minus (xf f) (yf f)" .
lift_definition uminus_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra" is "\<lambda>xf f . uminus (xf f)" .
lift_definition bot_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra" is "\<lambda>f . bot" ..
lift_definition top_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra" is "\<lambda>f . top" ..
lift_definition less_eq_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra \<Rightarrow> bool" is "\<lambda>xf yf . \<forall>f . less_eq (xf f) (yf f)" .
lift_definition less_lifted_boolean_algebra :: "('a,'b) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_boolean_algebra \<Rightarrow> bool" is "\<lambda>xf yf . (\<forall>f . less_eq (xf f) (yf f)) \<and> \<not> (\<forall>f . less_eq (yf f) (xf f))" .
instance
apply intro_classes
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer using order_trans by blast
subgoal apply transfer using antisym ext by blast
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by (simp add: sup_inf_distrib1)
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by (simp add: diff_eq)
done
end
text \<open>
The following two definitions give the Boolean algebra isomorphism.
\<close>
abbreviation ba_iso_inv :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_boolean_algebra \<Rightarrow> ('a,'b) lifted_pair regular"
where "ba_iso_inv \<equiv> \<lambda>xf . Abs_regular (Abs_lifted_pair (\<lambda>f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f))))"
abbreviation ba_iso :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair regular \<Rightarrow> ('a,'b) lifted_boolean_algebra"
where "ba_iso \<equiv> \<lambda>pf . Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f))"
lemma ba_iso_inv_lifted_pair:
"(Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)) \<in> triple.pairs (Rep_phi f)"
by (metis (no_types, hide_lams) double_compl simp_phi triple.pairs_uminus.simps triple_def triple.pairs_uminus_closed)
lemma ba_iso_inv_regular:
"regular (Abs_lifted_pair (\<lambda>f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f))))"
proof -
have "\<forall>f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)) = triple.pairs_uminus (Rep_phi f) (triple.pairs_uminus (Rep_phi f) (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)))"
by (simp add: triple.pairs_uminus.simps triple_def)
hence "Abs_lifted_pair (\<lambda>f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f))) = --Abs_lifted_pair (\<lambda>f . (Rep_lifted_boolean_algebra xf f,Rep_phi f (-Rep_lifted_boolean_algebra xf f)))"
by (simp add: triple.pairs_uminus_closed triple_def eq_onp_def uminus_lifted_pair.abs_eq ba_iso_inv_lifted_pair)
thus ?thesis
by simp
qed
text \<open>
The following two results prove that the isomorphisms are mutually inverse.
\<close>
lemma ba_iso_left_invertible:
"ba_iso_inv (ba_iso pf) = pf"
proof -
have 1: "\<forall>f . snd (Rep_lifted_pair (Rep_regular pf) f) = Rep_phi f (-fst (Rep_lifted_pair (Rep_regular pf) f))"
proof
fix f :: "('a,'b) phi"
let ?r = "Rep_phi f"
have "triple ?r"
by (simp add: triple_def)
hence 2: "\<forall>p . triple.pairs_uminus ?r p = (-fst p,?r (fst p))"
by (metis prod.collapse triple.pairs_uminus.simps)
have 3: "Rep_regular pf = --Rep_regular pf"
by (simp add: regular_in_p_image_iff)
show "snd (Rep_lifted_pair (Rep_regular pf) f) = ?r (-fst (Rep_lifted_pair (Rep_regular pf) f))"
using 2 3 by (metis fstI sndI uminus_lifted_pair.rep_eq)
qed
have "ba_iso_inv (ba_iso pf) = Abs_regular (Abs_lifted_pair (\<lambda>f . (fst (Rep_lifted_pair (Rep_regular pf) f),Rep_phi f (-fst (Rep_lifted_pair (Rep_regular pf) f)))))"
by (simp add: Abs_lifted_boolean_algebra_inverse)
also have "... = Abs_regular (Abs_lifted_pair (Rep_lifted_pair (Rep_regular pf)))"
using 1 by (metis prod.collapse)
also have "... = pf"
by (simp add: Rep_regular_inverse Rep_lifted_pair_inverse)
finally show ?thesis
.
qed
lemma ba_iso_right_invertible:
"ba_iso (ba_iso_inv xf) = xf"
proof -
let ?rf = "Rep_lifted_boolean_algebra xf"
have 1: "\<forall>f . (-?rf f,Rep_phi f (?rf f)) \<in> triple.pairs (Rep_phi f) \<and> (?rf f,Rep_phi f (-?rf f)) \<in> triple.pairs (Rep_phi f)"
proof
fix f
have "up_filter top = bot"
by (simp add: bot_filter.abs_eq)
hence "(\<exists>z . Rep_phi f (?rf f) = Rep_phi f (?rf f) \<squnion> up_filter z) \<and> (\<exists>z . Rep_phi f (-?rf f) = Rep_phi f (-?rf f) \<squnion> up_filter z)"
by (metis sup_bot_right)
thus "(-?rf f,Rep_phi f (?rf f)) \<in> triple.pairs (Rep_phi f) \<and> (?rf f,Rep_phi f (-?rf f)) \<in> triple.pairs (Rep_phi f)"
by (simp add: triple_def triple.pairs_def)
qed
have "regular (Abs_lifted_pair (\<lambda>f . (?rf f,Rep_phi f (-?rf f))))"
proof -
have "--Abs_lifted_pair (\<lambda>f . (?rf f,Rep_phi f (-?rf f))) = -Abs_lifted_pair (\<lambda>f . triple.pairs_uminus (Rep_phi f) (?rf f,Rep_phi f (-?rf f)))"
using 1 by (simp add: eq_onp_same_args uminus_lifted_pair.abs_eq)
also have "... = -Abs_lifted_pair (\<lambda>f . (-?rf f,Rep_phi f (?rf f)))"
by (metis (no_types, lifting) simp_phi triple_def triple.pairs_uminus.simps)
also have "... = Abs_lifted_pair (\<lambda>f . triple.pairs_uminus (Rep_phi f) (-?rf f,Rep_phi f (?rf f)))"
using 1 by (simp add: eq_onp_same_args uminus_lifted_pair.abs_eq)
also have "... = Abs_lifted_pair (\<lambda>f . (?rf f,Rep_phi f (-?rf f)))"
by (metis (no_types, lifting) simp_phi triple_def triple.pairs_uminus.simps double_compl)
finally show ?thesis
by simp
qed
hence "in_p_image (Abs_lifted_pair (\<lambda>f . (?rf f,Rep_phi f (-?rf f))))"
by blast
thus ?thesis
using 1 by (simp add: Rep_lifted_boolean_algebra_inverse Abs_lifted_pair_inverse Abs_regular_inverse)
qed
text \<open>
The isomorphism is established by proving the remaining Boolean algebra homomorphism properties.
\<close>
lemma ba_iso:
"boolean_algebra_isomorphism ba_iso"
proof (intro conjI)
show "Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular bot) f)) = bot"
by (simp add: bot_lifted_boolean_algebra_def bot_regular.rep_eq bot_lifted_pair.rep_eq)
show "Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular top) f)) = top"
by (simp add: top_lifted_boolean_algebra_def top_regular.rep_eq top_lifted_pair.rep_eq)
show "\<forall>pf qf . Abs_lifted_boolean_algebra (\<lambda>f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (pf \<squnion> qf)) f)) = Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f)) \<squnion> Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular qf) f))"
proof (intro allI)
fix pf qf :: "('a,'b) lifted_pair regular"
{
fix f
obtain x y z w where 1: "(x,y) = Rep_lifted_pair (Rep_regular pf) f \<and> (z,w) = Rep_lifted_pair (Rep_regular qf) f"
using prod.collapse by blast
have "triple (Rep_phi f)"
by (simp add: triple_def)
hence "fst (triple.pairs_sup (x,y) (z,w)) = fst (x,y) \<squnion> fst (z,w)"
using triple.pairs_sup.simps by force
hence "fst (triple.pairs_sup (Rep_lifted_pair (Rep_regular pf) f) (Rep_lifted_pair (Rep_regular qf) f)) = fst (Rep_lifted_pair (Rep_regular pf) f) \<squnion> fst (Rep_lifted_pair (Rep_regular qf) f)"
using 1 by simp
hence "fst (Rep_lifted_pair (Rep_regular (pf \<squnion> qf)) f) = fst (Rep_lifted_pair (Rep_regular pf) f) \<squnion> fst (Rep_lifted_pair (Rep_regular qf) f)"
by (unfold sup_regular.rep_eq sup_lifted_pair.rep_eq) simp
}
thus "Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular (pf \<squnion> qf)) f)) = Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f)) \<squnion> Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular qf) f))"
by (simp add: eq_onp_same_args sup_lifted_boolean_algebra.abs_eq sup_regular.rep_eq sup_lifted_boolean_algebra.rep_eq)
qed
show 1: "\<forall>pf qf . Abs_lifted_boolean_algebra (\<lambda>f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (pf \<sqinter> qf)) f)) = Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f)) \<sqinter> Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular qf) f))"
proof (intro allI)
fix pf qf :: "('a,'b) lifted_pair regular"
{
fix f
obtain x y z w where 1: "(x,y) = Rep_lifted_pair (Rep_regular pf) f \<and> (z,w) = Rep_lifted_pair (Rep_regular qf) f"
using prod.collapse by blast
have "triple (Rep_phi f)"
by (simp add: triple_def)
hence "fst (triple.pairs_inf (x,y) (z,w)) = fst (x,y) \<sqinter> fst (z,w)"
using triple.pairs_inf.simps by force
hence "fst (triple.pairs_inf (Rep_lifted_pair (Rep_regular pf) f) (Rep_lifted_pair (Rep_regular qf) f)) = fst (Rep_lifted_pair (Rep_regular pf) f) \<sqinter> fst (Rep_lifted_pair (Rep_regular qf) f)"
using 1 by simp
hence "fst (Rep_lifted_pair (Rep_regular (pf \<sqinter> qf)) f) = fst (Rep_lifted_pair (Rep_regular pf) f) \<sqinter> fst (Rep_lifted_pair (Rep_regular qf) f)"
by (unfold inf_regular.rep_eq inf_lifted_pair.rep_eq) simp
}
thus "Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular (pf \<sqinter> qf)) f)) = Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f)) \<sqinter> Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular qf) f))"
by (simp add: eq_onp_same_args inf_lifted_boolean_algebra.abs_eq inf_regular.rep_eq inf_lifted_boolean_algebra.rep_eq)
qed
show "\<forall>pf . Abs_lifted_boolean_algebra (\<lambda>f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (-pf)) f)) = -Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f))"
proof
fix pf :: "('a,'b) lifted_pair regular"
{
fix f
obtain x y where 1: "(x,y) = Rep_lifted_pair (Rep_regular pf) f"
using prod.collapse by blast
have "triple (Rep_phi f)"
by (simp add: triple_def)
hence "fst (triple.pairs_uminus (Rep_phi f) (x,y)) = -fst (x,y)"
using triple.pairs_uminus.simps by force
hence "fst (triple.pairs_uminus (Rep_phi f) (Rep_lifted_pair (Rep_regular pf) f)) = -fst (Rep_lifted_pair (Rep_regular pf) f)"
using 1 by simp
hence "fst (Rep_lifted_pair (Rep_regular (-pf)) f) = -fst (Rep_lifted_pair (Rep_regular pf) f)"
by (unfold uminus_regular.rep_eq uminus_lifted_pair.rep_eq) simp
}
thus "Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular (-pf)) f)) = -Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f))"
by (simp add: eq_onp_same_args uminus_lifted_boolean_algebra.abs_eq uminus_regular.rep_eq uminus_lifted_boolean_algebra.rep_eq)
qed
thus "\<forall>pf qf . Abs_lifted_boolean_algebra (\<lambda>f::('a,'b) phi . fst (Rep_lifted_pair (Rep_regular (pf - qf)) f)) = Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular pf) f)) - Abs_lifted_boolean_algebra (\<lambda>f . fst (Rep_lifted_pair (Rep_regular qf) f))"
using 1 by (simp add: diff_eq)
show "bij ba_iso"
by (rule invertible_bij[where g=ba_iso_inv]) (simp_all add: ba_iso_left_invertible ba_iso_right_invertible)
qed
subsubsection \<open>Distributive Lattice Isomorphism\<close>
text \<open>
We carry out a similar development for the isomorphism of distributive lattices.
Again, the original distributive lattice with a greatest element needs to be lifted to match the lifted pairs.
\<close>
typedef (overloaded) ('a,'b) lifted_distrib_lattice_top = "{ xf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) phi \<Rightarrow> 'b . True }"
by simp
setup_lifting type_definition_lifted_distrib_lattice_top
instantiation lifted_distrib_lattice_top :: (non_trivial_boolean_algebra,distrib_lattice_top) distrib_lattice_top
begin
lift_definition sup_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \<Rightarrow> ('a,'b) lifted_distrib_lattice_top \<Rightarrow> ('a,'b) lifted_distrib_lattice_top" is "\<lambda>xf yf f . sup (xf f) (yf f)" .
lift_definition inf_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \<Rightarrow> ('a,'b) lifted_distrib_lattice_top \<Rightarrow> ('a,'b) lifted_distrib_lattice_top" is "\<lambda>xf yf f . inf (xf f) (yf f)" .
lift_definition top_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top" is "\<lambda>f . top" ..
lift_definition less_eq_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \<Rightarrow> ('a,'b) lifted_distrib_lattice_top \<Rightarrow> bool" is "\<lambda>xf yf . \<forall>f . less_eq (xf f) (yf f)" .
lift_definition less_lifted_distrib_lattice_top :: "('a,'b) lifted_distrib_lattice_top \<Rightarrow> ('a,'b) lifted_distrib_lattice_top \<Rightarrow> bool" is "\<lambda>xf yf . (\<forall>f . less_eq (xf f) (yf f)) \<and> \<not> (\<forall>f . less_eq (yf f) (xf f))" .
instance
apply intro_classes
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer using order_trans by blast
subgoal apply transfer using antisym ext by blast
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by auto
subgoal apply transfer by (simp add: sup_inf_distrib1)
done
end
text \<open>
The following function extracts the least element of the filter of a dense pair, which turns out to be a principal filter.
It is used to define one of the isomorphisms below.
\<close>
fun get_dense :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense \<Rightarrow> ('a,'b) phi \<Rightarrow> 'b"
where "get_dense pf f = (SOME z . Rep_lifted_pair (Rep_dense pf) f = (top,up_filter z))"
lemma get_dense_char:
"Rep_lifted_pair (Rep_dense pf) f = (top,up_filter (get_dense pf f))"
proof -
obtain x y where 1: "(x,y) = Rep_lifted_pair (Rep_dense pf) f \<and> (x,y) \<in> triple.pairs (Rep_phi f) \<and> triple.pairs_uminus (Rep_phi f) (x,y) = triple.pairs_bot"
by (metis bot_lifted_pair.rep_eq prod.collapse simp_dense simp_lifted_pair uminus_lifted_pair.rep_eq)
hence 2: "x = top"
by (simp add: triple.intro triple.pairs_uminus.simps dense_pp)
have "triple (Rep_phi f)"
by (simp add: triple_def)
hence "\<exists>z. y = Rep_phi f (-x) \<squnion> up_filter z"
using 1 triple.pairs_def by blast
then obtain z where "y = up_filter z"
using 2 by auto
hence "Rep_lifted_pair (Rep_dense pf) f = (top,up_filter z)"
using 1 2 by simp
thus ?thesis
by (metis (mono_tags, lifting) tfl_some get_dense.simps)
qed
text \<open>
The following two definitions give the distributive lattice isomorphism.
\<close>
abbreviation dl_iso_inv :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_distrib_lattice_top \<Rightarrow> ('a,'b) lifted_pair dense"
where "dl_iso_inv \<equiv> \<lambda>xf . Abs_dense (Abs_lifted_pair (\<lambda>f . (top,up_filter (Rep_lifted_distrib_lattice_top xf f))))"
abbreviation dl_iso :: "('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense \<Rightarrow> ('a,'b) lifted_distrib_lattice_top"
where "dl_iso \<equiv> \<lambda>pf . Abs_lifted_distrib_lattice_top (get_dense pf)"
lemma dl_iso_inv_lifted_pair:
"(top,up_filter (Rep_lifted_distrib_lattice_top xf f)) \<in> triple.pairs (Rep_phi f)"
by (metis (no_types, hide_lams) compl_bot_eq double_compl simp_phi sup_bot.left_neutral triple.sa_iso_pair triple_def)
lemma dl_iso_inv_dense:
"dense (Abs_lifted_pair (\<lambda>f . (top,up_filter (Rep_lifted_distrib_lattice_top xf f))))"
proof -
have "\<forall>f . triple.pairs_uminus (Rep_phi f) (top,up_filter (Rep_lifted_distrib_lattice_top xf f)) = triple.pairs_bot"
by (simp add: top_filter.abs_eq triple.pairs_uminus.simps triple_def)
hence "bot = -Abs_lifted_pair (\<lambda>f . (top,up_filter (Rep_lifted_distrib_lattice_top xf f)))"
by (simp add: eq_onp_def uminus_lifted_pair.abs_eq dl_iso_inv_lifted_pair bot_lifted_pair_def)
thus ?thesis
by simp
qed
text \<open>
The following two results prove that the isomorphisms are mutually inverse.
\<close>
lemma dl_iso_left_invertible:
"dl_iso_inv (dl_iso pf) = pf"
proof -
have "dl_iso_inv (dl_iso pf) = Abs_dense (Abs_lifted_pair (\<lambda>f . (top,up_filter (get_dense pf f))))"
by (metis Abs_lifted_distrib_lattice_top_inverse UNIV_I UNIV_def)
also have "... = Abs_dense (Abs_lifted_pair (Rep_lifted_pair (Rep_dense pf)))"
by (metis get_dense_char)
also have "... = pf"
by (simp add: Rep_dense_inverse Rep_lifted_pair_inverse)
finally show ?thesis
.
qed
lemma dl_iso_right_invertible:
"dl_iso (dl_iso_inv xf) = xf"
proof -
let ?rf = "Rep_lifted_distrib_lattice_top xf"
let ?pf = "Abs_dense (Abs_lifted_pair (\<lambda>f . (top,up_filter (?rf f))))"
have 1: "\<forall>f . (top,up_filter (?rf f)) \<in> triple.pairs (Rep_phi f)"
proof
fix f :: "('a,'b) phi"
have "triple (Rep_phi f)"
by (simp add: triple_def)
thus "(top,up_filter (?rf f)) \<in> triple.pairs (Rep_phi f)"
using triple.pairs_def by force
qed
have 2: "dense (Abs_lifted_pair (\<lambda>f . (top,up_filter (?rf f))))"
proof -
have "-Abs_lifted_pair (\<lambda>f . (top,up_filter (?rf f))) = Abs_lifted_pair (\<lambda>f . triple.pairs_uminus (Rep_phi f) (top,up_filter (?rf f)))"
using 1 by (simp add: eq_onp_same_args uminus_lifted_pair.abs_eq)
also have "... = Abs_lifted_pair (\<lambda>f . (bot,Rep_phi f top))"
by (simp add: triple.pairs_uminus.simps triple_def)
also have "... = Abs_lifted_pair (\<lambda>f . triple.pairs_bot)"
by (metis (no_types, hide_lams) simp_phi triple.phi_top triple_def)
also have "... = bot"
by (simp add: bot_lifted_pair_def)
finally show ?thesis
by simp
qed
have "get_dense ?pf = ?rf"
proof
fix f
have "(top,up_filter (get_dense ?pf f)) = Rep_lifted_pair (Rep_dense ?pf) f"
by (metis get_dense_char)
also have "... = Rep_lifted_pair (Abs_lifted_pair (\<lambda>f . (top,up_filter (?rf f)))) f"
using Abs_dense_inverse 2 by force
also have "... = (top,up_filter (?rf f))"
using 1 by (simp add: Abs_lifted_pair_inverse)
finally show "get_dense ?pf f = ?rf f"
using up_filter_injective by auto
qed
thus ?thesis
by (simp add: Rep_lifted_distrib_lattice_top_inverse)
qed
text \<open>
To obtain the isomorphism, it remains to show the homomorphism properties of lattices with a greatest element.
\<close>
lemma dl_iso:
"bounded_lattice_top_isomorphism dl_iso"
proof (intro conjI)
have "get_dense top = (\<lambda>f::('a,'b) phi . top)"
proof
fix f :: "('a,'b) phi"
have "Rep_lifted_pair (Rep_dense top) f = (top,Abs_filter {top})"
by (simp add: top_dense.rep_eq top_lifted_pair.rep_eq)
hence "up_filter (get_dense top f) = Abs_filter {top}"
by (metis prod.inject get_dense_char)
hence "Rep_filter (up_filter (get_dense top f)) = {top}"
by (metis bot_filter.abs_eq bot_filter.rep_eq)
thus "get_dense top f = top"
by (metis self_in_upset singletonD Abs_filter_inverse mem_Collect_eq up_filter)
qed
thus "Abs_lifted_distrib_lattice_top (get_dense top::('a,'b) phi \<Rightarrow> 'b) = top"
by (metis top_lifted_distrib_lattice_top_def)
next
show "\<forall>pf qf :: ('a,'b) lifted_pair dense . Abs_lifted_distrib_lattice_top (get_dense (pf \<squnion> qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \<squnion> Abs_lifted_distrib_lattice_top (get_dense qf)"
proof (intro allI)
fix pf qf :: "('a,'b) lifted_pair dense"
have 1: "Abs_lifted_distrib_lattice_top (get_dense pf) \<squnion> Abs_lifted_distrib_lattice_top (get_dense qf) = Abs_lifted_distrib_lattice_top (\<lambda>f . get_dense pf f \<squnion> get_dense qf f)"
by (simp add: eq_onp_same_args sup_lifted_distrib_lattice_top.abs_eq)
have "(\<lambda>f . get_dense (pf \<squnion> qf) f) = (\<lambda>f . get_dense pf f \<squnion> get_dense qf f)"
proof
fix f
have "(top,up_filter (get_dense (pf \<squnion> qf) f)) = Rep_lifted_pair (Rep_dense (pf \<squnion> qf)) f"
by (metis get_dense_char)
also have "... = triple.pairs_sup (Rep_lifted_pair (Rep_dense pf) f) (Rep_lifted_pair (Rep_dense qf) f)"
by (simp add: sup_lifted_pair.rep_eq sup_dense.rep_eq)
also have "... = triple.pairs_sup (top,up_filter (get_dense pf f)) (top,up_filter (get_dense qf f))"
by (metis get_dense_char)
also have "... = (top,up_filter (get_dense pf f) \<sqinter> up_filter (get_dense qf f))"
by (metis (no_types, lifting) calculation prod.simps(1) simp_phi triple.pairs_sup.simps triple_def)
also have "... = (top,up_filter (get_dense pf f \<squnion> get_dense qf f))"
by (metis up_filter_dist_sup)
finally show "get_dense (pf \<squnion> qf) f = get_dense pf f \<squnion> get_dense qf f"
using up_filter_injective by blast
qed
thus "Abs_lifted_distrib_lattice_top (get_dense (pf \<squnion> qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \<squnion> Abs_lifted_distrib_lattice_top (get_dense qf)"
using 1 by metis
qed
next
show "\<forall>pf qf :: ('a,'b) lifted_pair dense . Abs_lifted_distrib_lattice_top (get_dense (pf \<sqinter> qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \<sqinter> Abs_lifted_distrib_lattice_top (get_dense qf)"
proof (intro allI)
fix pf qf :: "('a,'b) lifted_pair dense"
have 1: "Abs_lifted_distrib_lattice_top (get_dense pf) \<sqinter> Abs_lifted_distrib_lattice_top (get_dense qf) = Abs_lifted_distrib_lattice_top (\<lambda>f . get_dense pf f \<sqinter> get_dense qf f)"
by (simp add: eq_onp_same_args inf_lifted_distrib_lattice_top.abs_eq)
have "(\<lambda>f . get_dense (pf \<sqinter> qf) f) = (\<lambda>f . get_dense pf f \<sqinter> get_dense qf f)"
proof
fix f
have "(top,up_filter (get_dense (pf \<sqinter> qf) f)) = Rep_lifted_pair (Rep_dense (pf \<sqinter> qf)) f"
by (metis get_dense_char)
also have "... = triple.pairs_inf (Rep_lifted_pair (Rep_dense pf) f) (Rep_lifted_pair (Rep_dense qf) f)"
by (simp add: inf_lifted_pair.rep_eq inf_dense.rep_eq)
also have "... = triple.pairs_inf (top,up_filter (get_dense pf f)) (top,up_filter (get_dense qf f))"
by (metis get_dense_char)
also have "... = (top,up_filter (get_dense pf f) \<squnion> up_filter (get_dense qf f))"
by (metis (no_types, lifting) calculation prod.simps(1) simp_phi triple.pairs_inf.simps triple_def)
also have "... = (top,up_filter (get_dense pf f \<sqinter> get_dense qf f))"
by (metis up_filter_dist_inf)
finally show "get_dense (pf \<sqinter> qf) f = get_dense pf f \<sqinter> get_dense qf f"
using up_filter_injective by blast
qed
thus "Abs_lifted_distrib_lattice_top (get_dense (pf \<sqinter> qf)) = Abs_lifted_distrib_lattice_top (get_dense pf) \<sqinter> Abs_lifted_distrib_lattice_top (get_dense qf)"
using 1 by metis
qed
next
show "bij dl_iso"
by (rule invertible_bij[where g=dl_iso_inv]) (simp_all add: dl_iso_left_invertible dl_iso_right_invertible)
qed
subsubsection \<open>Structure Map Preservation\<close>
text \<open>
We finally show that the isomorphisms are compatible with the structure maps.
This involves lifting the distributive lattice isomorphism to filters of distributive lattices (as these are the targets of the structure maps).
To this end, we first show that the lifted isomorphism preserves filters.
\<close>
lemma phi_iso_filter:
"filter ((\<lambda>qf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense . Rep_lifted_distrib_lattice_top (dl_iso qf) f) ` Rep_filter (stone_phi pf))"
proof (rule filter_map_filter)
show "mono (\<lambda>qf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense . Rep_lifted_distrib_lattice_top (dl_iso qf) f)"
by (metis (no_types, lifting) mono_def dl_iso le_iff_sup sup_lifted_distrib_lattice_top.rep_eq)
next
show "\<forall>qf y . Rep_lifted_distrib_lattice_top (dl_iso qf) f \<le> y \<longrightarrow> (\<exists>rf . qf \<le> rf \<and> y = Rep_lifted_distrib_lattice_top (dl_iso rf) f)"
proof (intro allI, rule impI)
fix qf :: "('a,'b) lifted_pair dense"
fix y :: 'b
assume 1: "Rep_lifted_distrib_lattice_top (dl_iso qf) f \<le> y"
let ?rf = "Abs_dense (Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))"
have 2: "\<forall>g . (if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g) \<in> triple.pairs (Rep_phi g)"
by (metis Abs_lifted_distrib_lattice_top_inverse dl_iso_inv_lifted_pair mem_Collect_eq simp_lifted_pair)
hence "-Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g) = Abs_lifted_pair (\<lambda>g . triple.pairs_uminus (Rep_phi g) (if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))"
by (simp add: eq_onp_def uminus_lifted_pair.abs_eq)
also have "... = Abs_lifted_pair (\<lambda>g . if g = f then triple.pairs_uminus (Rep_phi g) (top,up_filter y) else triple.pairs_uminus (Rep_phi g) (Rep_lifted_pair (Rep_dense qf) g))"
by (simp add: if_distrib)
also have "... = Abs_lifted_pair (\<lambda>g . if g = f then (bot,top) else triple.pairs_uminus (Rep_phi g) (Rep_lifted_pair (Rep_dense qf) g))"
by (subst triple.pairs_uminus.simps, simp add: triple_def, metis compl_top_eq simp_phi)
also have "... = Abs_lifted_pair (\<lambda>g . if g = f then (bot,top) else (bot,top))"
by (metis bot_lifted_pair.rep_eq simp_dense top_filter.abs_eq uminus_lifted_pair.rep_eq)
also have "... = bot"
by (simp add: bot_lifted_pair.abs_eq top_filter.abs_eq)
finally have 3: "Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g) \<in> dense_elements"
by blast
hence "(top,up_filter (get_dense (Abs_dense (Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))) f)) = Rep_lifted_pair (Rep_dense (Abs_dense (Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)))) f"
by (metis (mono_tags, lifting) get_dense_char)
also have "... = Rep_lifted_pair (Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)) f"
using 3 by (simp add: Abs_dense_inverse)
also have "... = (top,up_filter y)"
using 2 by (simp add: Abs_lifted_pair_inverse)
finally have "get_dense (Abs_dense (Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g))) f = y"
using up_filter_injective by blast
hence 4: "Rep_lifted_distrib_lattice_top (dl_iso ?rf) f = y"
by (simp add: Abs_lifted_distrib_lattice_top_inverse)
{
fix g
have "Rep_lifted_distrib_lattice_top (dl_iso qf) g \<le> Rep_lifted_distrib_lattice_top (dl_iso ?rf) g"
proof (cases "g = f")
assume "g = f"
thus ?thesis
using 1 4 by simp
next
assume 5: "g \<noteq> f"
have "(top,up_filter (get_dense ?rf g)) = Rep_lifted_pair (Rep_dense (Abs_dense (Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)))) g"
by (metis (mono_tags, lifting) get_dense_char)
also have "... = Rep_lifted_pair (Abs_lifted_pair (\<lambda>g . if g = f then (top,up_filter y) else Rep_lifted_pair (Rep_dense qf) g)) g"
using 3 by (simp add: Abs_dense_inverse)
also have "... = Rep_lifted_pair (Rep_dense qf) g"
using 2 5 by (simp add: Abs_lifted_pair_inverse)
also have "... = (top,up_filter (get_dense qf g))"
using get_dense_char by auto
finally have "get_dense ?rf g = get_dense qf g"
using up_filter_injective by blast
thus "Rep_lifted_distrib_lattice_top (dl_iso qf) g \<le> Rep_lifted_distrib_lattice_top (dl_iso ?rf) g"
by (simp add: Abs_lifted_distrib_lattice_top_inverse)
qed
}
hence "Rep_lifted_distrib_lattice_top (dl_iso qf) \<le> Rep_lifted_distrib_lattice_top (dl_iso ?rf)"
by (simp add: le_funI)
hence 6: "dl_iso qf \<le> dl_iso ?rf"
by (simp add: le_funD less_eq_lifted_distrib_lattice_top.rep_eq)
hence "qf \<le> ?rf"
by (metis (no_types, lifting) dl_iso sup_isomorphism_ord_isomorphism)
thus "\<exists>rf . qf \<le> rf \<and> y = Rep_lifted_distrib_lattice_top (dl_iso rf) f"
using 4 by auto
qed
qed
text \<open>
The commutativity property states that the same result is obtained in two ways by starting with a regular lifted pair \<open>pf\<close>:
\begin{itemize}
\item apply the Boolean algebra isomorphism to the pair; then apply a structure map \<open>f\<close> to obtain a filter of dense elements; or,
\item apply the structure map \<open>stone_phi\<close> to the pair; then apply the distributive lattice isomorphism lifted to the resulting filter.
\end{itemize}
\<close>
lemma phi_iso:
"Rep_phi f (Rep_lifted_boolean_algebra (ba_iso pf) f) = filter_map (\<lambda>qf::('a::non_trivial_boolean_algebra,'b::distrib_lattice_top) lifted_pair dense . Rep_lifted_distrib_lattice_top (dl_iso qf) f) (stone_phi pf)"
proof -
let ?r = "Rep_phi f"
let ?ppf = "\<lambda>g . triple.pairs_uminus (Rep_phi g) (Rep_lifted_pair (Rep_regular pf) g)"
have 1: "triple ?r"
by (simp add: triple_def)
have 2: "Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f))) \<subseteq> { z . \<exists>qf . -Rep_regular pf \<le> Rep_dense qf \<and> z = get_dense qf f }"
proof
fix z
obtain x where 3: "x = fst (Rep_lifted_pair (Rep_regular pf) f)"
by simp
assume "z \<in> Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))"
hence "\<up>z \<subseteq> Rep_filter (?r x)"
using 3 filter_def by fastforce
hence 4: "up_filter z \<le> ?r x"
by (metis Rep_filter_cases Rep_filter_inverse less_eq_filter.rep_eq mem_Collect_eq up_filter)
have 5: "\<forall>g . ?ppf g \<in> triple.pairs (Rep_phi g)"
by (metis (no_types) simp_lifted_pair uminus_lifted_pair.rep_eq)
let ?zf = "\<lambda>g . if g = f then (top,up_filter z) else triple.pairs_top"
have 6: "\<forall>g . ?zf g \<in> triple.pairs (Rep_phi g)"
proof
fix g :: "('a,'b) phi"
have "triple (Rep_phi g)"
by (simp add: triple_def)
hence "(top,up_filter z) \<in> triple.pairs (Rep_phi g)"
using triple.pairs_def by force
thus "?zf g \<in> triple.pairs (Rep_phi g)"
by (metis simp_lifted_pair top_lifted_pair.rep_eq)
qed
hence "-Abs_lifted_pair ?zf = Abs_lifted_pair (\<lambda>g . triple.pairs_uminus (Rep_phi g) (?zf g))"
by (subst uminus_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args)
also have "... = Abs_lifted_pair (\<lambda>g . if g = f then triple.pairs_uminus (Rep_phi g) (top,up_filter z) else triple.pairs_uminus (Rep_phi g) triple.pairs_top)"
by (rule arg_cong[where f=Abs_lifted_pair]) auto
also have "... = Abs_lifted_pair (\<lambda>g . triple.pairs_bot)"
using 1 by (metis bot_lifted_pair.rep_eq dense_closed_top top_lifted_pair.rep_eq triple.pairs_uminus.simps uminus_lifted_pair.rep_eq)
finally have 7: "Abs_lifted_pair ?zf \<in> dense_elements"
by (simp add: bot_lifted_pair.abs_eq)
let ?qf = "Abs_dense (Abs_lifted_pair ?zf)"
have "\<forall>g . triple.pairs_less_eq (?ppf g) (?zf g)"
proof
fix g
show "triple.pairs_less_eq (?ppf g) (?zf g)"
proof (cases "g = f")
assume 8: "g = f"
hence 9: "?ppf g = (-x,?r x)"
using 1 3 by (metis prod.collapse triple.pairs_uminus.simps)
have "triple.pairs_less_eq (-x,?r x) (top,up_filter z)"
using 1 4 by (meson inf.bot_least triple.pairs_less_eq.simps)
thus ?thesis
using 8 9 by simp
next
assume 10: "g \<noteq> f"
have "triple.pairs_less_eq (?ppf g) triple.pairs_top"
using 1 by (metis (no_types, hide_lams) bot.extremum top_greatest prod.collapse triple_def triple.pairs_less_eq.simps triple.phi_bot)
thus ?thesis
using 10 by simp
qed
qed
hence "Abs_lifted_pair ?ppf \<le> Abs_lifted_pair ?zf"
using 5 6 by (subst less_eq_lifted_pair.abs_eq) (simp_all add: eq_onp_same_args)
hence 11: "-Rep_regular pf \<le> Rep_dense ?qf"
using 7 by (simp add: uminus_lifted_pair_def Abs_dense_inverse)
have "(top,up_filter (get_dense ?qf f)) = Rep_lifted_pair (Rep_dense ?qf) f"
by (metis get_dense_char)
also have "... = (top,up_filter z)"
using 6 7 Abs_dense_inverse Abs_lifted_pair_inverse by force
finally have "z = get_dense ?qf f"
using up_filter_injective by force
thus "z \<in> { z . \<exists>qf . -Rep_regular pf \<le> Rep_dense qf \<and> z = get_dense qf f }"
using 11 by auto
qed
have 12: "Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f))) \<supseteq> { z . \<exists>qf . -Rep_regular pf \<le> Rep_dense qf \<and> z = get_dense qf f }"
proof
fix z
assume "z \<in> { z . \<exists>qf . -Rep_regular pf \<le> Rep_dense qf \<and> z = get_dense qf f }"
hence "\<exists>qf . -Rep_regular pf \<le> Rep_dense qf \<and> z = get_dense qf f"
by auto
hence "triple.pairs_less_eq (Rep_lifted_pair (-Rep_regular pf) f) (top,up_filter z)"
by (metis less_eq_lifted_pair.rep_eq get_dense_char)
hence "up_filter z \<le> snd (Rep_lifted_pair (-Rep_regular pf) f)"
using 1 by (metis (no_types, hide_lams) prod.collapse triple.pairs_less_eq.simps)
also have "... = snd (?ppf f)"
by (metis uminus_lifted_pair.rep_eq)
also have "... = ?r (fst (Rep_lifted_pair (Rep_regular pf) f))"
using 1 by (metis (no_types) prod.collapse prod.inject triple.pairs_uminus.simps)
finally have "Rep_filter (up_filter z) \<subseteq> Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))"
by (simp add: less_eq_filter.rep_eq)
hence "\<up>z \<subseteq> Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))"
by (metis Abs_filter_inverse mem_Collect_eq up_filter)
thus "z \<in> Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f)))"
by blast
qed
have 13: "\<forall>qf\<in>Rep_filter (stone_phi pf) . Rep_lifted_distrib_lattice_top (Abs_lifted_distrib_lattice_top (get_dense qf)) f = get_dense qf f"
by (metis Abs_lifted_distrib_lattice_top_inverse UNIV_I UNIV_def)
have "Rep_filter (?r (fst (Rep_lifted_pair (Rep_regular pf) f))) = { z . \<exists>qf\<in>stone_phi_base pf . z = get_dense qf f }"
using 2 12 by simp
hence "?r (fst (Rep_lifted_pair (Rep_regular pf) f)) = Abs_filter { z . \<exists>qf\<in>stone_phi_base pf . z = get_dense qf f }"
by (metis Rep_filter_inverse)
hence "?r (Rep_lifted_boolean_algebra (ba_iso pf) f) = Abs_filter { z . \<exists>qf\<in>Rep_filter (stone_phi pf) . z = Rep_lifted_distrib_lattice_top (dl_iso qf) f }"
using 13 by (simp add: Abs_filter_inverse stone_phi_base_filter stone_phi_def Abs_lifted_boolean_algebra_inverse)
thus ?thesis
by (simp add: image_def)
qed
end
-
diff --git a/thys/Stone_Kleene_Relation_Algebras/Iterings.thy b/thys/Stone_Kleene_Relation_Algebras/Iterings.thy
--- a/thys/Stone_Kleene_Relation_Algebras/Iterings.thy
+++ b/thys/Stone_Kleene_Relation_Algebras/Iterings.thy
@@ -1,858 +1,858 @@
(* Title: Iterings
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Iterings\<close>
text \<open>
This theory introduces algebraic structures with an operation that describes iteration in various relational computation models.
An iteration describes the repeated sequential execution of a computation.
This is typically modelled by fixpoints, but different computation models use different fixpoints in the refinement order.
We therefore look at equational and simulation axioms rather than induction axioms.
Our development is based on \cite{Guttmann2012c} and the proposed algebras generalise Kleene algebras.
We first consider a variant of Conway semirings \cite{BloomEsik1993a} based on idempotent left semirings.
Conway semirings expand semirings by an iteration operation satisfying Conway's sumstar and productstar axioms \cite{Conway1971}.
Many properties of iteration follow already from these equational axioms.
Next we introduce iterings, which use generalised versions of simulation axioms in addition to sumstar and productstar.
Unlike the induction axioms of the Kleene star, which hold only in partial-correctness models, the simulation axioms are also valid in total and general correctness models.
They are still powerful enough to prove the correctness of complex results such as separation theorems of \cite{Cohen2000} and Back's atomicity refinement theorem \cite{BackWright1999,Wright2004}.
\<close>
theory Iterings
imports Stone_Relation_Algebras.Semirings
begin
subsection \<open>Conway Semirings\<close>
text \<open>
In this section, we consider equational axioms for iteration.
The algebraic structures are based on idempotent left semirings, which are expanded by a unary iteration operation.
We start with an unfold property, one inequality of the sliding rule and distributivity over joins, which is similar to Conway's sumstar.
\<close>
class circ =
fixes circ :: "'a \<Rightarrow> 'a" ("_\<^sup>\<circ>" [100] 100)
class left_conway_semiring = idempotent_left_semiring + circ +
assumes circ_left_unfold: "1 \<squnion> x * x\<^sup>\<circ> = x\<^sup>\<circ>"
assumes circ_left_slide: "(x * y)\<^sup>\<circ> * x \<le> x * (y * x)\<^sup>\<circ>"
assumes circ_sup_1: "(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * (y * x\<^sup>\<circ>)\<^sup>\<circ>"
begin
text \<open>
We obtain one inequality of Conway's productstar, as well as of the other unfold rule.
\<close>
lemma circ_mult_sub:
"1 \<squnion> x * (y * x)\<^sup>\<circ> * y \<le> (x * y)\<^sup>\<circ>"
by (metis sup_right_isotone circ_left_slide circ_left_unfold mult_assoc mult_right_isotone)
lemma circ_right_unfold_sub:
"1 \<squnion> x\<^sup>\<circ> * x \<le> x\<^sup>\<circ>"
by (metis circ_mult_sub mult_1_left mult_1_right)
lemma circ_zero:
"bot\<^sup>\<circ> = 1"
by (metis sup_monoid.add_0_right circ_left_unfold mult_left_zero)
lemma circ_increasing:
"x \<le> x\<^sup>\<circ>"
by (metis le_supI2 circ_left_unfold circ_right_unfold_sub mult_1_left mult_right_sub_dist_sup_left order_trans)
lemma circ_reflexive:
"1 \<le> x\<^sup>\<circ>"
by (metis sup_left_divisibility circ_left_unfold)
lemma circ_mult_increasing:
"x \<le> x * x\<^sup>\<circ>"
by (metis circ_reflexive mult_right_isotone mult_1_right)
lemma circ_mult_increasing_2:
"x \<le> x\<^sup>\<circ> * x"
by (metis circ_reflexive mult_left_isotone mult_1_left)
lemma circ_transitive_equal:
"x\<^sup>\<circ> * x\<^sup>\<circ> = x\<^sup>\<circ>"
by (metis sup_idem circ_sup_1 circ_left_unfold mult_assoc)
text \<open>
While iteration is not idempotent, a fixpoint is reached after applying this operation twice.
Iteration is idempotent for the unit.
\<close>
lemma circ_circ_circ:
"x\<^sup>\<circ>\<^sup>\<circ>\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>"
by (metis sup_idem circ_sup_1 circ_increasing circ_transitive_equal le_iff_sup)
lemma circ_one:
"1\<^sup>\<circ> = 1\<^sup>\<circ>\<^sup>\<circ>"
by (metis circ_circ_circ circ_zero)
lemma circ_sup_sub:
"(x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ> \<le> (x \<squnion> y)\<^sup>\<circ>"
by (metis circ_sup_1 circ_left_slide)
lemma circ_plus_one:
"x\<^sup>\<circ> = 1 \<squnion> x\<^sup>\<circ>"
by (metis le_iff_sup circ_reflexive)
text \<open>
Iteration satisfies a characteristic property of reflexive transitive closures.
\<close>
lemma circ_rtc_2:
"1 \<squnion> x \<squnion> x\<^sup>\<circ> * x\<^sup>\<circ> = x\<^sup>\<circ>"
by (metis sup_assoc circ_increasing circ_plus_one circ_transitive_equal le_iff_sup)
lemma mult_zero_circ:
"(x * bot)\<^sup>\<circ> = 1 \<squnion> x * bot"
by (metis circ_left_unfold mult_assoc mult_left_zero)
lemma mult_zero_sup_circ:
"(x \<squnion> y * bot)\<^sup>\<circ> = x\<^sup>\<circ> * (y * bot)\<^sup>\<circ>"
by (metis circ_sup_1 mult_assoc mult_left_zero)
lemma circ_plus_sub:
"x\<^sup>\<circ> * x \<le> x * x\<^sup>\<circ>"
by (metis circ_left_slide mult_1_left mult_1_right)
lemma circ_loop_fixpoint:
"y * (y\<^sup>\<circ> * z) \<squnion> z = y\<^sup>\<circ> * z"
by (metis sup_commute circ_left_unfold mult_assoc mult_1_left mult_right_dist_sup)
lemma left_plus_below_circ:
"x * x\<^sup>\<circ> \<le> x\<^sup>\<circ>"
by (metis sup.cobounded2 circ_left_unfold)
lemma right_plus_below_circ:
"x\<^sup>\<circ> * x \<le> x\<^sup>\<circ>"
using circ_right_unfold_sub by auto
lemma circ_sup_upper_bound:
"x \<le> z\<^sup>\<circ> \<Longrightarrow> y \<le> z\<^sup>\<circ> \<Longrightarrow> x \<squnion> y \<le> z\<^sup>\<circ>"
by simp
lemma circ_mult_upper_bound:
"x \<le> z\<^sup>\<circ> \<Longrightarrow> y \<le> z\<^sup>\<circ> \<Longrightarrow> x * y \<le> z\<^sup>\<circ>"
by (metis mult_isotone circ_transitive_equal)
lemma circ_sub_dist:
"x\<^sup>\<circ> \<le> (x \<squnion> y)\<^sup>\<circ>"
by (metis circ_sup_sub circ_plus_one mult_1_left mult_right_sub_dist_sup_left order_trans)
lemma circ_sub_dist_1:
"x \<le> (x \<squnion> y)\<^sup>\<circ>"
using circ_increasing le_supE by blast
lemma circ_sub_dist_2:
"x * y \<le> (x \<squnion> y)\<^sup>\<circ>"
by (metis sup_commute circ_mult_upper_bound circ_sub_dist_1)
lemma circ_sub_dist_3:
"x\<^sup>\<circ> * y\<^sup>\<circ> \<le> (x \<squnion> y)\<^sup>\<circ>"
by (metis sup_commute circ_mult_upper_bound circ_sub_dist)
lemma circ_isotone:
"x \<le> y \<Longrightarrow> x\<^sup>\<circ> \<le> y\<^sup>\<circ>"
by (metis circ_sub_dist le_iff_sup)
lemma circ_sup_2:
"(x \<squnion> y)\<^sup>\<circ> \<le> (x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ>"
by (metis sup.bounded_iff circ_increasing circ_isotone circ_reflexive mult_isotone mult_1_left mult_1_right)
lemma circ_sup_one_left_unfold:
"1 \<le> x \<Longrightarrow> x * x\<^sup>\<circ> = x\<^sup>\<circ>"
- by (metis antisym le_iff_sup mult_1_left mult_right_sub_dist_sup_left left_plus_below_circ)
+ by (metis order.antisym le_iff_sup mult_1_left mult_right_sub_dist_sup_left left_plus_below_circ)
lemma circ_sup_one_right_unfold:
"1 \<le> x \<Longrightarrow> x\<^sup>\<circ> * x = x\<^sup>\<circ>"
- by (metis antisym le_iff_sup mult_left_sub_dist_sup_left mult_1_right right_plus_below_circ)
+ by (metis order.antisym le_iff_sup mult_left_sub_dist_sup_left mult_1_right right_plus_below_circ)
lemma circ_decompose_4:
"(x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ> = x\<^sup>\<circ> * (y\<^sup>\<circ> * x\<^sup>\<circ>)\<^sup>\<circ>"
by (metis sup_assoc sup_commute circ_sup_1 circ_loop_fixpoint circ_plus_one circ_rtc_2 circ_transitive_equal mult_assoc)
lemma circ_decompose_5:
"(x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ> = (y\<^sup>\<circ> * x\<^sup>\<circ>)\<^sup>\<circ>"
- by (metis circ_decompose_4 circ_loop_fixpoint antisym mult_right_sub_dist_sup_right mult_assoc)
+ by (metis circ_decompose_4 circ_loop_fixpoint order.antisym mult_right_sub_dist_sup_right mult_assoc)
lemma circ_decompose_6:
"x\<^sup>\<circ> * (y * x\<^sup>\<circ>)\<^sup>\<circ> = y\<^sup>\<circ> * (x * y\<^sup>\<circ>)\<^sup>\<circ>"
by (metis sup_commute circ_sup_1)
lemma circ_decompose_7:
"(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ> * (x \<squnion> y)\<^sup>\<circ>"
by (metis circ_sup_1 circ_decompose_6 circ_transitive_equal mult_assoc)
lemma circ_decompose_8:
"(x \<squnion> y)\<^sup>\<circ> = (x \<squnion> y)\<^sup>\<circ> * x\<^sup>\<circ> * y\<^sup>\<circ>"
- by (metis antisym eq_refl mult_assoc mult_isotone mult_1_right circ_mult_upper_bound circ_reflexive circ_sub_dist_3)
+ by (metis order.antisym eq_refl mult_assoc mult_isotone mult_1_right circ_mult_upper_bound circ_reflexive circ_sub_dist_3)
lemma circ_decompose_9:
"(x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ> * (x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ>"
by (metis circ_decompose_4 mult_assoc)
lemma circ_decompose_10:
"(x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ> = (x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ> * x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis sup_ge2 circ_loop_fixpoint circ_reflexive circ_sup_one_right_unfold mult_assoc order_trans)
lemma circ_back_loop_prefixpoint:
"(z * y\<^sup>\<circ>) * y \<squnion> z \<le> z * y\<^sup>\<circ>"
by (metis sup.bounded_iff circ_left_unfold mult_assoc mult_left_sub_dist_sup_left mult_right_isotone mult_1_right right_plus_below_circ)
text \<open>
We obtain the fixpoint and prefixpoint properties of iteration, but not least or greatest fixpoint properties.
\<close>
lemma circ_loop_is_fixpoint:
"is_fixpoint (\<lambda>x . y * x \<squnion> z) (y\<^sup>\<circ> * z)"
by (metis circ_loop_fixpoint is_fixpoint_def)
lemma circ_back_loop_is_prefixpoint:
"is_prefixpoint (\<lambda>x . x * y \<squnion> z) (z * y\<^sup>\<circ>)"
by (metis circ_back_loop_prefixpoint is_prefixpoint_def)
lemma circ_circ_sup:
"(1 \<squnion> x)\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>"
by (metis sup_commute circ_sup_1 circ_decompose_4 circ_zero mult_1_right)
lemma circ_circ_mult_sub:
"x\<^sup>\<circ> * 1\<^sup>\<circ> \<le> x\<^sup>\<circ>\<^sup>\<circ>"
by (metis circ_increasing circ_isotone circ_mult_upper_bound circ_reflexive)
lemma left_plus_circ:
"(x * x\<^sup>\<circ>)\<^sup>\<circ> = x\<^sup>\<circ>"
by (metis circ_left_unfold circ_sup_1 mult_1_right mult_sub_right_one sup.absorb1 mult_assoc)
lemma right_plus_circ:
"(x\<^sup>\<circ> * x)\<^sup>\<circ> = x\<^sup>\<circ>"
- by (metis sup_commute circ_isotone circ_loop_fixpoint circ_plus_sub circ_sub_dist eq_iff left_plus_circ)
+ by (metis sup_commute circ_isotone circ_loop_fixpoint circ_plus_sub circ_sub_dist order.eq_iff left_plus_circ)
lemma circ_square:
"(x * x)\<^sup>\<circ> \<le> x\<^sup>\<circ>"
by (metis circ_increasing circ_isotone left_plus_circ mult_right_isotone)
lemma circ_mult_sub_sup:
"(x * y)\<^sup>\<circ> \<le> (x \<squnion> y)\<^sup>\<circ>"
by (metis sup_ge1 sup_ge2 circ_isotone circ_square mult_isotone order_trans)
lemma circ_sup_mult_zero:
"x\<^sup>\<circ> * y = (x \<squnion> y * bot)\<^sup>\<circ> * y"
proof -
have "(x \<squnion> y * bot)\<^sup>\<circ> * y = x\<^sup>\<circ> * (1 \<squnion> y * bot) * y"
by (metis mult_zero_sup_circ mult_zero_circ)
also have "... = x\<^sup>\<circ> * (y \<squnion> y * bot)"
by (metis mult_assoc mult_1_left mult_left_zero mult_right_dist_sup)
also have "... = x\<^sup>\<circ> * y"
by (metis sup_commute le_iff_sup zero_right_mult_decreasing)
finally show ?thesis
by simp
qed
lemma troeger_1:
"(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * (1 \<squnion> y * (x \<squnion> y)\<^sup>\<circ>)"
by (metis circ_sup_1 circ_left_unfold mult_assoc)
lemma troeger_2:
"(x \<squnion> y)\<^sup>\<circ> * z = x\<^sup>\<circ> * (y * (x \<squnion> y)\<^sup>\<circ> * z \<squnion> z)"
by (metis circ_sup_1 circ_loop_fixpoint mult_assoc)
lemma troeger_3:
"(x \<squnion> y * bot)\<^sup>\<circ> = x\<^sup>\<circ> * (1 \<squnion> y * bot)"
by (metis mult_zero_sup_circ mult_zero_circ)
lemma circ_sup_sub_sup_one_1:
"x \<squnion> y \<le> x\<^sup>\<circ> * (1 \<squnion> y)"
by (metis circ_increasing circ_left_unfold mult_1_left mult_1_right mult_left_sub_dist_sup mult_right_sub_dist_sup_left order_trans sup_mono)
lemma circ_sup_sub_sup_one_2:
"x\<^sup>\<circ> * (x \<squnion> y) \<le> x\<^sup>\<circ> * (1 \<squnion> y)"
by (metis circ_sup_sub_sup_one_1 circ_transitive_equal mult_assoc mult_right_isotone)
lemma circ_sup_sub_sup_one:
"x * x\<^sup>\<circ> * (x \<squnion> y) \<le> x * x\<^sup>\<circ> * (1 \<squnion> y)"
by (metis circ_sup_sub_sup_one_2 mult_assoc mult_right_isotone)
lemma circ_square_2:
"(x * x)\<^sup>\<circ> * (x \<squnion> 1) \<le> x\<^sup>\<circ>"
by (metis sup.bounded_iff circ_increasing circ_mult_upper_bound circ_reflexive circ_square)
lemma circ_extra_circ:
"(y * x\<^sup>\<circ>)\<^sup>\<circ> = (y * y\<^sup>\<circ> * x\<^sup>\<circ>)\<^sup>\<circ>"
by (metis circ_decompose_6 circ_transitive_equal left_plus_circ mult_assoc)
lemma circ_circ_sub_mult:
"1\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ>\<^sup>\<circ>"
by (metis circ_increasing circ_isotone circ_mult_upper_bound circ_reflexive)
lemma circ_decompose_11:
"(x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ> = (x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ> * x\<^sup>\<circ>"
by (metis circ_decompose_10 circ_decompose_4 circ_decompose_5 circ_decompose_9 left_plus_circ)
lemma circ_mult_below_circ_circ:
"(x * y)\<^sup>\<circ> \<le> (x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ>"
by (metis circ_increasing circ_isotone circ_reflexive dual_order.trans mult_left_isotone mult_right_isotone mult_1_right)
(*
lemma circ_right_unfold: "1 \<squnion> x\<^sup>\<circ> * x = x\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma circ_mult: "1 \<squnion> x * (y * x)\<^sup>\<circ> * y = (x * y)\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma circ_slide: "(x * y)\<^sup>\<circ> * x = x * (y * x)\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma circ_plus_same: "x\<^sup>\<circ> * x = x * x\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "1\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * 1\<^sup>\<circ>" nitpick [expect=genuine,card=7] oops
lemma circ_circ_mult_1: "x\<^sup>\<circ> * 1\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>" nitpick [expect=genuine,card=7] oops
lemma "x\<^sup>\<circ> * 1\<^sup>\<circ> \<le> 1\<^sup>\<circ> * x\<^sup>\<circ>" nitpick [expect=genuine,card=7] oops
lemma circ_circ_mult: "1\<^sup>\<circ> * x\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>" nitpick [expect=genuine,card=7] oops
lemma circ_sup: "(x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ> = (x \<squnion> y)\<^sup>\<circ>" nitpick [expect=genuine,card=8] oops
lemma circ_unfold_sum: "(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> \<squnion> x\<^sup>\<circ> * y * (x \<squnion> y)\<^sup>\<circ>" nitpick [expect=genuine,card=7] oops
lemma mult_zero_sup_circ_2: "(x \<squnion> y * bot)\<^sup>\<circ> = x\<^sup>\<circ> \<squnion> x\<^sup>\<circ> * y * bot" nitpick [expect=genuine,card=7] oops
lemma sub_mult_one_circ: "x * 1\<^sup>\<circ> \<le> 1\<^sup>\<circ> * x" nitpick [expect=genuine] oops
lemma circ_back_loop_fixpoint: "(z * y\<^sup>\<circ>) * y \<squnion> z = z * y\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma circ_back_loop_is_fixpoint: "is_fixpoint (\<lambda>x . x * y \<squnion> z) (z * y\<^sup>\<circ>)" nitpick [expect=genuine] oops
lemma "x\<^sup>\<circ> * y\<^sup>\<circ> \<le> (x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ>" nitpick [expect=genuine,card=7] oops
*)
end
text \<open>
The next class considers the interaction of iteration with a greatest element.
\<close>
class bounded_left_conway_semiring = bounded_idempotent_left_semiring + left_conway_semiring
begin
lemma circ_top:
"top\<^sup>\<circ> = top"
- by (simp add: antisym circ_increasing)
+ by (simp add: order.antisym circ_increasing)
lemma circ_right_top:
"x\<^sup>\<circ> * top = top"
by (metis sup_right_top circ_loop_fixpoint)
lemma circ_left_top:
"top * x\<^sup>\<circ> = top"
by (metis circ_right_top circ_top circ_decompose_11)
lemma mult_top_circ:
"(x * top)\<^sup>\<circ> = 1 \<squnion> x * top"
by (metis circ_left_top circ_left_unfold mult_assoc)
end
class left_zero_conway_semiring = idempotent_left_zero_semiring + left_conway_semiring
begin
lemma mult_zero_sup_circ_2:
"(x \<squnion> y * bot)\<^sup>\<circ> = x\<^sup>\<circ> \<squnion> x\<^sup>\<circ> * y * bot"
by (metis mult_assoc mult_left_dist_sup mult_1_right troeger_3)
lemma circ_unfold_sum:
"(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> \<squnion> x\<^sup>\<circ> * y * (x \<squnion> y)\<^sup>\<circ>"
by (metis mult_assoc mult_left_dist_sup mult_1_right troeger_1)
end
text \<open>
The next class assumes the full sliding equation.
\<close>
class left_conway_semiring_1 = left_conway_semiring +
assumes circ_right_slide: "x * (y * x)\<^sup>\<circ> \<le> (x * y)\<^sup>\<circ> * x"
begin
lemma circ_slide_1:
"x * (y * x)\<^sup>\<circ> = (x * y)\<^sup>\<circ> * x"
- by (metis antisym circ_left_slide circ_right_slide)
+ by (metis order.antisym circ_left_slide circ_right_slide)
text \<open>
This implies the full unfold rules and Conway's productstar.
\<close>
lemma circ_right_unfold_1:
"1 \<squnion> x\<^sup>\<circ> * x = x\<^sup>\<circ>"
by (metis circ_left_unfold circ_slide_1 mult_1_left mult_1_right)
lemma circ_mult_1:
"(x * y)\<^sup>\<circ> = 1 \<squnion> x * (y * x)\<^sup>\<circ> * y"
by (metis circ_left_unfold circ_slide_1 mult_assoc)
lemma circ_sup_9:
"(x \<squnion> y)\<^sup>\<circ> = (x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ>"
by (metis circ_sup_1 circ_slide_1)
lemma circ_plus_same:
"x\<^sup>\<circ> * x = x * x\<^sup>\<circ>"
by (metis circ_slide_1 mult_1_left mult_1_right)
lemma circ_decompose_12:
"x\<^sup>\<circ> * y\<^sup>\<circ> \<le> (x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ>"
by (metis circ_sup_9 circ_sub_dist_3)
end
class left_zero_conway_semiring_1 = left_zero_conway_semiring + left_conway_semiring_1
begin
lemma circ_back_loop_fixpoint:
"(z * y\<^sup>\<circ>) * y \<squnion> z = z * y\<^sup>\<circ>"
by (metis sup_commute circ_left_unfold circ_plus_same mult_assoc mult_left_dist_sup mult_1_right)
lemma circ_back_loop_is_fixpoint:
"is_fixpoint (\<lambda>x . x * y \<squnion> z) (z * y\<^sup>\<circ>)"
by (metis circ_back_loop_fixpoint is_fixpoint_def)
lemma circ_elimination:
"x * y = bot \<Longrightarrow> x * y\<^sup>\<circ> \<le> x"
by (metis sup_monoid.add_0_left circ_back_loop_fixpoint circ_plus_same mult_assoc mult_left_zero order_refl)
end
subsection \<open>Iterings\<close>
text \<open>
This section adds simulation axioms to Conway semirings.
We consider several classes with increasingly general simulation axioms.
\<close>
class itering_1 = left_conway_semiring_1 +
assumes circ_simulate: "z * x \<le> y * z \<longrightarrow> z * x\<^sup>\<circ> \<le> y\<^sup>\<circ> * z"
begin
lemma circ_circ_mult:
"1\<^sup>\<circ> * x\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>"
- by (metis antisym circ_circ_sup circ_reflexive circ_simulate circ_sub_dist_3 circ_sup_one_left_unfold circ_transitive_equal mult_1_left order_refl)
+ by (metis order.antisym circ_circ_sup circ_reflexive circ_simulate circ_sub_dist_3 circ_sup_one_left_unfold circ_transitive_equal mult_1_left order_refl)
lemma sub_mult_one_circ:
"x * 1\<^sup>\<circ> \<le> 1\<^sup>\<circ> * x"
by (metis circ_simulate mult_1_left mult_1_right order_refl)
text \<open>
The left simulation axioms is enough to prove a basic import property of tests.
\<close>
lemma circ_import:
assumes "p \<le> p * p"
and "p \<le> 1"
and "p * x \<le> x * p"
shows "p * x\<^sup>\<circ> = p * (p * x)\<^sup>\<circ>"
proof -
have "p * x \<le> p * (p * x * p) * p"
- by (metis assms coreflexive_transitive eq_iff test_preserves_equation mult_assoc)
+ by (metis assms coreflexive_transitive order.eq_iff test_preserves_equation mult_assoc)
hence "p * x\<^sup>\<circ> \<le> p * (p * x)\<^sup>\<circ>"
by (metis (no_types) assms circ_simulate circ_slide_1 test_preserves_equation)
thus ?thesis
- by (metis assms(2) circ_isotone mult_left_isotone mult_1_left mult_right_isotone antisym)
+ by (metis assms(2) circ_isotone mult_left_isotone mult_1_left mult_right_isotone order.antisym)
qed
end
text \<open>
Including generalisations of both simulation axioms allows us to prove separation rules.
\<close>
class itering_2 = left_conway_semiring_1 +
assumes circ_simulate_right: "z * x \<le> y * z \<squnion> w \<longrightarrow> z * x\<^sup>\<circ> \<le> y\<^sup>\<circ> * (z \<squnion> w * x\<^sup>\<circ>)"
assumes circ_simulate_left: "x * z \<le> z * y \<squnion> w \<longrightarrow> x\<^sup>\<circ> * z \<le> (z \<squnion> x\<^sup>\<circ> * w) * y\<^sup>\<circ>"
begin
subclass itering_1
apply unfold_locales
by (metis sup_monoid.add_0_right circ_simulate_right mult_left_zero)
lemma circ_simulate_left_1:
"x * z \<le> z * y \<Longrightarrow> x\<^sup>\<circ> * z \<le> z * y\<^sup>\<circ> \<squnion> x\<^sup>\<circ> * bot"
by (metis sup_monoid.add_0_right circ_simulate_left mult_assoc mult_left_zero mult_right_dist_sup)
lemma circ_separate_1:
assumes "y * x \<le> x * y"
shows "(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ>"
proof -
have "y\<^sup>\<circ> * x \<le> x * y\<^sup>\<circ> \<squnion> y\<^sup>\<circ> * bot"
by (metis assms circ_simulate_left_1)
hence "y\<^sup>\<circ> * x * y\<^sup>\<circ> \<le> x * y\<^sup>\<circ> * y\<^sup>\<circ> \<squnion> y\<^sup>\<circ> * bot * y\<^sup>\<circ>"
by (metis mult_assoc mult_left_isotone mult_right_dist_sup)
also have "... = x * y\<^sup>\<circ> \<squnion> y\<^sup>\<circ> * bot"
by (metis circ_transitive_equal mult_assoc mult_left_zero)
finally have "y\<^sup>\<circ> * (x * y\<^sup>\<circ>)\<^sup>\<circ> \<le> x\<^sup>\<circ> * (y\<^sup>\<circ> \<squnion> y\<^sup>\<circ> * bot)"
using circ_simulate_right mult_assoc by fastforce
also have "... = x\<^sup>\<circ> * y\<^sup>\<circ>"
by (simp add: sup_absorb1 zero_right_mult_decreasing)
finally have "(x \<squnion> y)\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
by (simp add: circ_decompose_6 circ_sup_1)
thus ?thesis
- by (simp add: antisym circ_sub_dist_3)
+ by (simp add: order.antisym circ_sub_dist_3)
qed
lemma circ_circ_mult_1:
"x\<^sup>\<circ> * 1\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>"
by (metis sup_commute circ_circ_sup circ_separate_1 mult_1_left mult_1_right order_refl)
end
text \<open>
With distributivity, we also get Back's atomicity refinement theorem.
\<close>
class itering_3 = itering_2 + left_zero_conway_semiring_1
begin
lemma circ_simulate_1:
assumes "y * x \<le> x * y"
shows "y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
proof -
have "y * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y"
by (metis assms circ_simulate)
hence "y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ> \<squnion> y\<^sup>\<circ> * bot"
by (metis circ_simulate_left_1)
thus ?thesis
by (metis sup_assoc sup_monoid.add_0_right circ_loop_fixpoint mult_assoc mult_left_zero mult_zero_sup_circ_2)
qed
lemma atomicity_refinement:
assumes "s = s * q"
and "x = q * x"
and "q * b = bot"
and "r * b \<le> b * r"
and "r * l \<le> l * r"
and "x * l \<le> l * x"
and "b * l \<le> l * b"
and "q * l \<le> l * q"
and "r\<^sup>\<circ> * q \<le> q * r\<^sup>\<circ>"
and "q \<le> 1"
shows "s * (x \<squnion> b \<squnion> r \<squnion> l)\<^sup>\<circ> * q \<le> s * (x * b\<^sup>\<circ> * q \<squnion> r \<squnion> l)\<^sup>\<circ>"
proof -
have "(x \<squnion> b \<squnion> r) * l \<le> l * (x \<squnion> b \<squnion> r)"
using assms(5-7) mult_left_dist_sup mult_right_dist_sup semiring.add_mono by presburger
hence "s * (x \<squnion> b \<squnion> r \<squnion> l)\<^sup>\<circ> * q = s * l\<^sup>\<circ> * (x \<squnion> b \<squnion> r)\<^sup>\<circ> * q"
by (metis sup_commute circ_separate_1 mult_assoc)
also have "... = s * l\<^sup>\<circ> * b\<^sup>\<circ> * r\<^sup>\<circ> * q * (x * b\<^sup>\<circ> * r\<^sup>\<circ> * q)\<^sup>\<circ>"
proof -
have "(b \<squnion> r)\<^sup>\<circ> = b\<^sup>\<circ> * r\<^sup>\<circ>"
by (simp add: assms(4) circ_separate_1)
hence "b\<^sup>\<circ> * r\<^sup>\<circ> * (q * (x * b\<^sup>\<circ> * r\<^sup>\<circ>))\<^sup>\<circ> = (x \<squnion> b \<squnion> r)\<^sup>\<circ>"
by (metis (full_types) assms(2) circ_sup_1 sup_assoc sup_commute mult_assoc)
thus ?thesis
by (metis circ_slide_1 mult_assoc)
qed
also have "... \<le> s * l\<^sup>\<circ> * b\<^sup>\<circ> * r\<^sup>\<circ> * q * (x * b\<^sup>\<circ> * q * r\<^sup>\<circ>)\<^sup>\<circ>"
by (metis assms(9) circ_isotone mult_assoc mult_right_isotone)
also have "... \<le> s * q * l\<^sup>\<circ> * b\<^sup>\<circ> * r\<^sup>\<circ> * (x * b\<^sup>\<circ> * q * r\<^sup>\<circ>)\<^sup>\<circ>"
by (metis assms(1,10) mult_left_isotone mult_right_isotone mult_1_right)
also have "... \<le> s * l\<^sup>\<circ> * q * b\<^sup>\<circ> * r\<^sup>\<circ> * (x * b\<^sup>\<circ> * q * r\<^sup>\<circ>)\<^sup>\<circ>"
by (metis assms(1,8) circ_simulate mult_assoc mult_left_isotone mult_right_isotone)
also have "... \<le> s * l\<^sup>\<circ> * r\<^sup>\<circ> * (x * b\<^sup>\<circ> * q * r\<^sup>\<circ>)\<^sup>\<circ>"
by (metis assms(3,10) sup_monoid.add_0_left circ_back_loop_fixpoint circ_plus_same mult_assoc mult_left_zero mult_left_isotone mult_right_isotone mult_1_right)
also have "... \<le> s * (x * b\<^sup>\<circ> * q \<squnion> r \<squnion> l)\<^sup>\<circ>"
by (metis sup_commute circ_sup_1 circ_sub_dist_3 mult_assoc mult_right_isotone)
finally show ?thesis
.
qed
end
text \<open>
The following class contains the most general simulation axioms we consider.
They allow us to prove further separation properties.
\<close>
class itering = idempotent_left_zero_semiring + circ +
assumes circ_sup: "(x \<squnion> y)\<^sup>\<circ> = (x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ>"
assumes circ_mult: "(x * y)\<^sup>\<circ> = 1 \<squnion> x * (y * x)\<^sup>\<circ> * y"
assumes circ_simulate_right_plus: "z * x \<le> y * y\<^sup>\<circ> * z \<squnion> w \<longrightarrow> z * x\<^sup>\<circ> \<le> y\<^sup>\<circ> * (z \<squnion> w * x\<^sup>\<circ>)"
assumes circ_simulate_left_plus: "x * z \<le> z * y\<^sup>\<circ> \<squnion> w \<longrightarrow> x\<^sup>\<circ> * z \<le> (z \<squnion> x\<^sup>\<circ> * w) * y\<^sup>\<circ>"
begin
lemma circ_right_unfold:
"1 \<squnion> x\<^sup>\<circ> * x = x\<^sup>\<circ>"
by (metis circ_mult mult_1_left mult_1_right)
lemma circ_slide:
"x * (y * x)\<^sup>\<circ> = (x * y)\<^sup>\<circ> * x"
proof -
have "x * (y * x)\<^sup>\<circ> = Rf x (y * 1 \<squnion> y * (x * (y * x)\<^sup>\<circ> * y)) * x"
by (metis (no_types) circ_mult mult_1_left mult_1_right mult_left_dist_sup mult_right_dist_sup mult_assoc)
thus ?thesis
by (metis (no_types) circ_mult mult_1_right mult_left_dist_sup mult_assoc)
qed
subclass itering_3
apply unfold_locales
apply (metis circ_mult mult_1_left mult_1_right)
apply (metis circ_slide order_refl)
apply (metis circ_sup circ_slide)
apply (metis circ_slide order_refl)
apply (metis sup_left_isotone circ_right_unfold mult_left_isotone mult_left_sub_dist_sup_left mult_1_right order_trans circ_simulate_right_plus)
by (metis sup_commute sup_ge1 sup_right_isotone circ_mult mult_right_isotone mult_1_right order_trans circ_simulate_left_plus)
lemma circ_simulate_right_plus_1:
"z * x \<le> y * y\<^sup>\<circ> * z \<Longrightarrow> z * x\<^sup>\<circ> \<le> y\<^sup>\<circ> * z"
by (metis sup_monoid.add_0_right circ_simulate_right_plus mult_left_zero)
lemma circ_simulate_left_plus_1:
"x * z \<le> z * y\<^sup>\<circ> \<Longrightarrow> x\<^sup>\<circ> * z \<le> z * y\<^sup>\<circ> \<squnion> x\<^sup>\<circ> * bot"
by (metis sup_monoid.add_0_right circ_simulate_left_plus mult_assoc mult_left_zero mult_right_dist_sup)
lemma circ_simulate_2:
"y * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ> \<longleftrightarrow> y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
apply (rule iffI)
apply (metis sup_assoc sup_monoid.add_0_right circ_loop_fixpoint circ_simulate_left_plus_1 mult_assoc mult_left_zero mult_zero_sup_circ_2)
by (metis circ_increasing mult_left_isotone order_trans)
lemma circ_simulate_absorb:
"y * x \<le> x \<Longrightarrow> y\<^sup>\<circ> * x \<le> x \<squnion> y\<^sup>\<circ> * bot"
by (metis circ_simulate_left_plus_1 circ_zero mult_1_right)
lemma circ_simulate_3:
"y * x\<^sup>\<circ> \<le> x\<^sup>\<circ> \<Longrightarrow> y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis sup.bounded_iff circ_reflexive circ_simulate_2 le_iff_sup mult_right_isotone mult_1_right)
lemma circ_separate_mult_1:
"y * x \<le> x * y \<Longrightarrow> (x * y)\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis circ_mult_sub_sup circ_separate_1)
lemma circ_separate_unfold:
"(y * x\<^sup>\<circ>)\<^sup>\<circ> = y\<^sup>\<circ> \<squnion> y\<^sup>\<circ> * y * x * x\<^sup>\<circ> * (y * x\<^sup>\<circ>)\<^sup>\<circ>"
by (metis circ_back_loop_fixpoint circ_plus_same circ_unfold_sum sup_commute mult_assoc)
lemma separation:
assumes "y * x \<le> x * y\<^sup>\<circ>"
shows "(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ>"
proof -
have "y\<^sup>\<circ> * x * y\<^sup>\<circ> \<le> x * y\<^sup>\<circ> \<squnion> y\<^sup>\<circ> * bot"
by (metis assms circ_simulate_left_plus_1 circ_transitive_equal mult_assoc mult_left_isotone)
thus ?thesis
by (metis sup_commute circ_sup_1 circ_simulate_right circ_sub_dist_3 le_iff_sup mult_assoc mult_left_zero zero_right_mult_decreasing)
qed
lemma simulation:
"y * x \<le> x * y\<^sup>\<circ> \<Longrightarrow> y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis sup_ge2 circ_isotone circ_mult_upper_bound circ_sub_dist separation)
lemma circ_simulate_4:
assumes "y * x \<le> x * x\<^sup>\<circ> * (1 \<squnion> y)"
shows "y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
proof -
have "x \<squnion> (x * x\<^sup>\<circ> * x * x \<squnion> x * x) = x * x\<^sup>\<circ>"
by (metis (no_types) circ_back_loop_fixpoint mult_right_dist_sup sup_commute)
hence "x \<le> x * x\<^sup>\<circ> * 1 \<squnion> x * x\<^sup>\<circ> * y"
by (metis mult_1_right sup_assoc sup_ge1)
hence "(1 \<squnion> y) * x \<le> x * x\<^sup>\<circ> * (1 \<squnion> y)"
using assms mult_left_dist_sup mult_right_dist_sup by force
hence "y * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis circ_sup_upper_bound circ_increasing circ_reflexive circ_simulate_right_plus_1 mult_right_isotone mult_right_sub_dist_sup_right order_trans)
thus ?thesis
by (metis circ_simulate_2)
qed
lemma circ_simulate_5:
"y * x \<le> x * x\<^sup>\<circ> * (x \<squnion> y) \<Longrightarrow> y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis circ_sup_sub_sup_one circ_simulate_4 order_trans)
lemma circ_simulate_6:
"y * x \<le> x * (x \<squnion> y) \<Longrightarrow> y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis sup_commute circ_back_loop_fixpoint circ_simulate_5 mult_right_sub_dist_sup_left order_trans)
lemma circ_separate_4:
assumes "y * x \<le> x * x\<^sup>\<circ> * (1 \<squnion> y)"
shows "(x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ>"
proof -
have "y * x * x\<^sup>\<circ> \<le> x * x\<^sup>\<circ> * (1 \<squnion> y) * x\<^sup>\<circ>"
by (simp add: assms mult_left_isotone)
also have "... = x * x\<^sup>\<circ> \<squnion> x * x\<^sup>\<circ> * y * x\<^sup>\<circ>"
by (simp add: circ_transitive_equal mult_left_dist_sup mult_right_dist_sup mult_assoc)
also have "... \<le> x * x\<^sup>\<circ> \<squnion> x * x\<^sup>\<circ> * x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis assms sup_right_isotone circ_simulate_2 circ_simulate_4 mult_assoc mult_right_isotone)
finally have "y * x * x\<^sup>\<circ> \<le> x * x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis circ_reflexive circ_transitive_equal le_iff_sup mult_assoc mult_right_isotone mult_1_right)
thus ?thesis
by (metis circ_sup_1 left_plus_circ mult_assoc separation)
qed
lemma circ_separate_5:
"y * x \<le> x * x\<^sup>\<circ> * (x \<squnion> y) \<Longrightarrow> (x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis circ_sup_sub_sup_one circ_separate_4 order_trans)
lemma circ_separate_6:
"y * x \<le> x * (x \<squnion> y) \<Longrightarrow> (x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ>"
by (metis sup_commute circ_back_loop_fixpoint circ_separate_5 mult_right_sub_dist_sup_left order_trans)
end
class bounded_itering = bounded_idempotent_left_zero_semiring + itering
begin
subclass bounded_left_conway_semiring ..
(*
lemma "1 = x\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x = x\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x = x * x\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x * x\<^sup>\<circ> = x\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "(x * y)\<^sup>\<circ> = (x \<squnion> y)\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x\<^sup>\<circ> * y\<^sup>\<circ> = (x \<squnion> y)\<^sup>\<circ>" nitpick [expect=genuine,card=6] oops
lemma "(x \<squnion> y)\<^sup>\<circ> = (x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "1 = 1\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "1 = (x * bot)\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "1 \<squnion> x * bot = x\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x\<^sup>\<circ> = x\<^sup>\<circ> * 1\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "z \<squnion> y * x = x \<longrightarrow> y\<^sup>\<circ> * z \<le> x" nitpick [expect=genuine] oops
lemma "y * x = x \<longrightarrow> y\<^sup>\<circ> * x \<le> x" nitpick [expect=genuine] oops
lemma "z \<squnion> x * y = x \<longrightarrow> z * y\<^sup>\<circ> \<le> x" nitpick [expect=genuine] oops
lemma "x * y = x \<longrightarrow> x * y\<^sup>\<circ> \<le> x" nitpick [expect=genuine] oops
lemma "x = z \<squnion> y * x \<longrightarrow> x \<le> y\<^sup>\<circ> * z" nitpick [expect=genuine] oops
lemma "x = y * x \<longrightarrow> x \<le> y\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x * z = z * y \<longrightarrow> x\<^sup>\<circ> * z \<le> z * y\<^sup>\<circ>" nitpick [expect=genuine] oops
lemma "x\<^sup>\<circ> = (x * x)\<^sup>\<circ> * (x \<squnion> 1)" oops
lemma "y\<^sup>\<circ> * x\<^sup>\<circ> \<le> x\<^sup>\<circ> * y\<^sup>\<circ> \<longrightarrow> (x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ>" oops
lemma "y * x \<le> (1 \<squnion> x) * y\<^sup>\<circ> \<longrightarrow> (x \<squnion> y)\<^sup>\<circ> = x\<^sup>\<circ> * y\<^sup>\<circ>" oops
lemma "y * x \<le> x \<longrightarrow> y\<^sup>\<circ> * x \<le> 1\<^sup>\<circ> * x" oops
*)
end
text \<open>
We finally expand Conway semirings and iterings by an element that corresponds to the endless loop.
\<close>
class L =
fixes L :: "'a"
class left_conway_semiring_L = left_conway_semiring + L +
assumes one_circ_mult_split: "1\<^sup>\<circ> * x = L \<squnion> x"
assumes L_split_sup: "x * (y \<squnion> L) \<le> x * y \<squnion> L"
begin
lemma L_def:
"L = 1\<^sup>\<circ> * bot"
by (metis sup_monoid.add_0_right one_circ_mult_split)
lemma one_circ_split:
"1\<^sup>\<circ> = L \<squnion> 1"
by (metis mult_1_right one_circ_mult_split)
lemma one_circ_circ_split:
"1\<^sup>\<circ>\<^sup>\<circ> = L \<squnion> 1"
by (metis circ_one one_circ_split)
lemma sub_mult_one_circ:
"x * 1\<^sup>\<circ> \<le> 1\<^sup>\<circ> * x"
by (metis L_split_sup sup_commute mult_1_right one_circ_mult_split)
lemma one_circ_mult_split_2:
"1\<^sup>\<circ> * x = x * 1\<^sup>\<circ> \<squnion> L"
proof -
have 1: "x * 1\<^sup>\<circ> \<le> L \<squnion> x"
using one_circ_mult_split sub_mult_one_circ by presburger
have "x \<squnion> x * 1\<^sup>\<circ> = x * 1\<^sup>\<circ>"
by (meson circ_back_loop_prefixpoint le_iff_sup sup.boundedE)
thus ?thesis
using 1 by (simp add: le_iff_sup one_circ_mult_split sup_assoc sup_commute)
qed
lemma sub_mult_one_circ_split:
"x * 1\<^sup>\<circ> \<le> x \<squnion> L"
by (metis sup_commute one_circ_mult_split sub_mult_one_circ)
lemma sub_mult_one_circ_split_2:
"x * 1\<^sup>\<circ> \<le> x \<squnion> 1\<^sup>\<circ>"
by (metis L_def sup_right_isotone order_trans sub_mult_one_circ_split zero_right_mult_decreasing)
lemma L_split:
"x * L \<le> x * bot \<squnion> L"
by (metis L_split_sup sup_monoid.add_0_left)
lemma L_left_zero:
"L * x = L"
by (metis L_def mult_assoc mult_left_zero)
lemma one_circ_L:
"1\<^sup>\<circ> * L = L"
by (metis L_def circ_transitive_equal mult_assoc)
lemma mult_L_circ:
"(x * L)\<^sup>\<circ> = 1 \<squnion> x * L"
by (metis L_left_zero circ_left_unfold mult_assoc)
lemma mult_L_circ_mult:
"(x * L)\<^sup>\<circ> * y = y \<squnion> x * L"
by (metis L_left_zero mult_L_circ mult_assoc mult_1_left mult_right_dist_sup)
lemma circ_L:
"L\<^sup>\<circ> = L \<squnion> 1"
by (metis L_left_zero sup_commute circ_left_unfold)
lemma L_below_one_circ:
"L \<le> 1\<^sup>\<circ>"
by (metis L_def zero_right_mult_decreasing)
lemma circ_circ_mult_1:
"x\<^sup>\<circ> * 1\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>"
by (metis L_left_zero sup_commute circ_sup_1 circ_circ_sup mult_zero_circ one_circ_split)
lemma circ_circ_mult:
"1\<^sup>\<circ> * x\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>"
- by (metis antisym circ_circ_mult_1 circ_circ_sub_mult sub_mult_one_circ)
+ by (metis order.antisym circ_circ_mult_1 circ_circ_sub_mult sub_mult_one_circ)
lemma circ_circ_split:
"x\<^sup>\<circ>\<^sup>\<circ> = L \<squnion> x\<^sup>\<circ>"
by (metis circ_circ_mult one_circ_mult_split)
lemma circ_sup_6:
"L \<squnion> (x \<squnion> y)\<^sup>\<circ> = (x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<circ>"
by (metis sup_assoc sup_commute circ_sup_1 circ_circ_sup circ_circ_split circ_decompose_4)
end
class itering_L = itering + L +
assumes L_def: "L = 1\<^sup>\<circ> * bot"
begin
lemma one_circ_split:
"1\<^sup>\<circ> = L \<squnion> 1"
- by (metis L_def sup_commute antisym circ_sup_upper_bound circ_reflexive circ_simulate_absorb mult_1_right order_refl zero_right_mult_decreasing)
+ by (metis L_def sup_commute order.antisym circ_sup_upper_bound circ_reflexive circ_simulate_absorb mult_1_right order_refl zero_right_mult_decreasing)
lemma one_circ_mult_split:
"1\<^sup>\<circ> * x = L \<squnion> x"
by (metis L_def sup_commute circ_loop_fixpoint mult_assoc mult_left_zero mult_zero_circ one_circ_split)
lemma sub_mult_one_circ_split:
"x * 1\<^sup>\<circ> \<le> x \<squnion> L"
by (metis sup_commute one_circ_mult_split sub_mult_one_circ)
lemma sub_mult_one_circ_split_2:
"x * 1\<^sup>\<circ> \<le> x \<squnion> 1\<^sup>\<circ>"
by (metis L_def sup_right_isotone order_trans sub_mult_one_circ_split zero_right_mult_decreasing)
lemma L_split:
"x * L \<le> x * bot \<squnion> L"
by (metis L_def mult_assoc mult_left_isotone mult_right_dist_sup sub_mult_one_circ_split_2)
subclass left_conway_semiring_L
apply unfold_locales
apply (metis L_def sup_commute circ_loop_fixpoint mult_assoc mult_left_zero mult_zero_circ one_circ_split)
by (metis sup_commute mult_assoc mult_left_isotone one_circ_mult_split sub_mult_one_circ)
lemma circ_left_induct_mult_L:
"L \<le> x \<Longrightarrow> x * y \<le> x \<Longrightarrow> x * y\<^sup>\<circ> \<le> x"
by (metis circ_one circ_simulate le_iff_sup one_circ_mult_split)
lemma circ_left_induct_mult_iff_L:
"L \<le> x \<Longrightarrow> x * y \<le> x \<longleftrightarrow> x * y\<^sup>\<circ> \<le> x"
by (metis sup.bounded_iff circ_back_loop_fixpoint circ_left_induct_mult_L le_iff_sup)
lemma circ_left_induct_L:
"L \<le> x \<Longrightarrow> x * y \<squnion> z \<le> x \<Longrightarrow> z * y\<^sup>\<circ> \<le> x"
by (metis sup.bounded_iff circ_left_induct_mult_L le_iff_sup mult_right_dist_sup)
end
end
diff --git a/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy
--- a/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy
+++ b/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy
@@ -1,750 +1,750 @@
(* Title: Kleene Algebras
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Kleene Algebras\<close>
text \<open>
Kleene algebras have been axiomatised by Kozen to describe the equational theory of regular languages \cite{Kozen1994}.
Binary relations are another important model.
This theory implements variants of Kleene algebras based on idempotent left semirings \cite{Moeller2007}.
The weakening of some semiring axioms allows the treatment of further computation models.
The presented algebras are special cases of iterings, so many results can be inherited.
\<close>
theory Kleene_Algebras
imports Iterings
begin
text \<open>
We start with left Kleene algebras, which use the left unfold and left induction axioms of Kleene algebras.
\<close>
class star =
fixes star :: "'a \<Rightarrow> 'a" ("_\<^sup>\<star>" [100] 100)
class left_kleene_algebra = idempotent_left_semiring + star +
assumes star_left_unfold : "1 \<squnion> y * y\<^sup>\<star> \<le> y\<^sup>\<star>"
assumes star_left_induct : "z \<squnion> y * x \<le> x \<longrightarrow> y\<^sup>\<star> * z \<le> x"
begin
no_notation
trancl ("(_\<^sup>+)" [1000] 999)
abbreviation tc ("_\<^sup>+" [100] 100) where "tc x \<equiv> x * x\<^sup>\<star>"
lemma star_left_unfold_equal:
"1 \<squnion> x * x\<^sup>\<star> = x\<^sup>\<star>"
- by (metis sup_right_isotone antisym mult_right_isotone mult_1_right star_left_induct star_left_unfold)
+ by (metis sup_right_isotone order.antisym mult_right_isotone mult_1_right star_left_induct star_left_unfold)
text \<open>
This means that for some properties of Kleene algebras, only one inequality can be derived, as exemplified by the following sliding rule.
\<close>
lemma star_left_slide:
"(x * y)\<^sup>\<star> * x \<le> x * (y * x)\<^sup>\<star>"
by (metis mult_assoc mult_left_sub_dist_sup mult_1_right star_left_induct star_left_unfold_equal)
lemma star_isotone:
"x \<le> y \<Longrightarrow> x\<^sup>\<star> \<le> y\<^sup>\<star>"
by (metis sup_right_isotone mult_left_isotone order_trans star_left_unfold mult_1_right star_left_induct)
lemma star_sup_1:
"(x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "y * x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star> \<le> (y * x\<^sup>\<star>)\<^sup>\<star>"
using sup_right_divisibility star_left_unfold_equal by auto
also have "... \<le> x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>"
using mult_left_isotone sup_left_divisibility star_left_unfold_equal by fastforce
finally have "(x \<squnion> y) * (x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>) \<le> x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>"
by (metis le_supI mult_right_dist_sup mult_right_sub_dist_sup_right mult_assoc star_left_unfold_equal)
hence "1 \<squnion> (x \<squnion> y) * (x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>) \<le> x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>"
using reflexive_mult_closed star_left_unfold by auto
thus "(x \<squnion> y)\<^sup>\<star> \<le> x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>"
using star_left_induct by force
next
have "x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star> \<le> x\<^sup>\<star> * (y * (x \<squnion> y)\<^sup>\<star>)\<^sup>\<star>"
by (simp add: mult_right_isotone star_isotone)
also have "... \<le> x\<^sup>\<star> * ((x \<squnion> y) * (x \<squnion> y)\<^sup>\<star>)\<^sup>\<star>"
by (simp add: mult_right_isotone mult_right_sub_dist_sup_right star_isotone)
also have "... \<le> x\<^sup>\<star> * (x \<squnion> y)\<^sup>\<star>\<^sup>\<star>"
using mult_right_isotone star_left_unfold star_isotone by auto
also have "... \<le> (x \<squnion> y)\<^sup>\<star> * (x \<squnion> y)\<^sup>\<star>\<^sup>\<star>"
by (simp add: mult_left_isotone star_isotone)
also have "... \<le> (x \<squnion> y)\<^sup>\<star>"
by (metis sup.bounded_iff mult_1_right star_left_induct star_left_unfold)
finally show "x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star> \<le> (x \<squnion> y)\<^sup>\<star>"
by simp
qed
end
text \<open>
We now show that left Kleene algebras form iterings.
A sublocale is used instead of a subclass, because iterings use a different iteration operation.
\<close>
sublocale left_kleene_algebra < star: left_conway_semiring where circ = star
apply unfold_locales
apply (rule star_left_unfold_equal)
apply (rule star_left_slide)
by (rule star_sup_1)
context left_kleene_algebra
begin
text \<open>
A number of lemmas in this class are taken from Georg Struth's Kleene algebra theory \cite{ArmstrongGomesStruthWeber2016}.
\<close>
lemma star_sub_one:
"x \<le> 1 \<Longrightarrow> x\<^sup>\<star> = 1"
- by (metis sup_right_isotone eq_iff le_iff_sup mult_1_right star.circ_plus_one star_left_induct)
+ by (metis sup_right_isotone order.eq_iff le_iff_sup mult_1_right star.circ_plus_one star_left_induct)
lemma star_one:
"1\<^sup>\<star> = 1"
by (simp add: star_sub_one)
lemma star_left_induct_mult:
"x * y \<le> y \<Longrightarrow> x\<^sup>\<star> * y \<le> y"
by (simp add: star_left_induct)
lemma star_left_induct_mult_iff:
"x * y \<le> y \<longleftrightarrow> x\<^sup>\<star> * y \<le> y"
using mult_left_isotone order_trans star.circ_increasing star_left_induct_mult by blast
lemma star_involutive:
"x\<^sup>\<star> = x\<^sup>\<star>\<^sup>\<star>"
using star.circ_circ_sup star_sup_1 star_one by auto
lemma star_sup_one:
"(1 \<squnion> x)\<^sup>\<star> = x\<^sup>\<star>"
using star.circ_circ_sup star_involutive by auto
lemma star_plus_loops:
"x\<^sup>\<star> \<squnion> 1 = x\<^sup>+ \<squnion> 1"
using star.circ_plus_one star_left_unfold_equal sup_commute by auto
lemma star_left_induct_equal:
"z \<squnion> x * y = y \<Longrightarrow> x\<^sup>\<star> * z \<le> y"
by (simp add: star_left_induct)
lemma star_left_induct_mult_equal:
"x * y = y \<Longrightarrow> x\<^sup>\<star> * y \<le> y"
by (simp add: star_left_induct_mult)
lemma star_star_upper_bound:
"x\<^sup>\<star> \<le> z\<^sup>\<star> \<Longrightarrow> x\<^sup>\<star>\<^sup>\<star> \<le> z\<^sup>\<star>"
using star_involutive by auto
lemma star_simulation_left:
assumes "x * z \<le> z * y"
shows "x\<^sup>\<star> * z \<le> z * y\<^sup>\<star>"
proof -
have "x * z * y\<^sup>\<star> \<le> z * y * y\<^sup>\<star>"
by (simp add: assms mult_left_isotone)
also have "... \<le> z * y\<^sup>\<star>"
by (simp add: mult_right_isotone star.left_plus_below_circ mult_assoc)
finally have "z \<squnion> x * z * y\<^sup>\<star> \<le> z * y\<^sup>\<star>"
using star.circ_back_loop_prefixpoint by auto
thus ?thesis
by (simp add: star_left_induct mult_assoc)
qed
lemma quasicomm_1:
"y * x \<le> x * (x \<squnion> y)\<^sup>\<star> \<longleftrightarrow> y\<^sup>\<star> * x \<le> x * (x \<squnion> y)\<^sup>\<star>"
by (metis mult_isotone order_refl order_trans star.circ_increasing star_involutive star_simulation_left)
lemma star_rtc_3:
"1 \<squnion> x \<squnion> y * y = y \<Longrightarrow> x\<^sup>\<star> \<le> y"
by (metis sup.bounded_iff le_iff_sup mult_left_sub_dist_sup_left mult_1_right star_left_induct_mult_iff star.circ_sub_dist)
lemma star_decompose_1:
"(x \<squnion> y)\<^sup>\<star> = (x\<^sup>\<star> * y\<^sup>\<star>)\<^sup>\<star>"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (simp add: star.circ_sup_2)
using star.circ_sub_dist_3 star_isotone star_involutive by fastforce
lemma star_sum:
"(x \<squnion> y)\<^sup>\<star> = (x\<^sup>\<star> \<squnion> y\<^sup>\<star>)\<^sup>\<star>"
using star_decompose_1 star_involutive by auto
lemma star_decompose_3:
"(x\<^sup>\<star> * y\<^sup>\<star>)\<^sup>\<star> = x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>"
using star_sup_1 star_decompose_1 by auto
text \<open>
In contrast to iterings, we now obtain that the iteration operation results in least fixpoints.
\<close>
lemma star_loop_least_fixpoint:
"y * x \<squnion> z = x \<Longrightarrow> y\<^sup>\<star> * z \<le> x"
by (simp add: sup_commute star_left_induct_equal)
lemma star_loop_is_least_fixpoint:
"is_least_fixpoint (\<lambda>x . y * x \<squnion> z) (y\<^sup>\<star> * z)"
by (simp add: is_least_fixpoint_def star.circ_loop_fixpoint star_loop_least_fixpoint)
lemma star_loop_mu:
"\<mu> (\<lambda>x . y * x \<squnion> z) = y\<^sup>\<star> * z"
by (metis least_fixpoint_same star_loop_is_least_fixpoint)
lemma affine_has_least_fixpoint:
"has_least_fixpoint (\<lambda>x . y * x \<squnion> z)"
by (metis has_least_fixpoint_def star_loop_is_least_fixpoint)
lemma star_outer_increasing:
"x \<le> y\<^sup>\<star> * x * y\<^sup>\<star>"
by (metis star.circ_back_loop_prefixpoint star.circ_loop_fixpoint sup.boundedE)
(*
lemma circ_sup: "(x\<^sup>\<star> * y)\<^sup>\<star> * x\<^sup>\<star> = (x \<squnion> y)\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_mult: "1 \<squnion> x * (y * x)\<^sup>\<star> * y = (x * y)\<^sup>\<star>" nitpick [expect=genuine] oops
lemma circ_plus_same: "x\<^sup>\<star> * x = x * x\<^sup>\<star>" nitpick [expect=genuine] oops
lemma circ_unfold_sum: "(x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> \<squnion> x\<^sup>\<star> * y * (x \<squnion> y)\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma mult_zero_sup_circ_2: "(x \<squnion> y * bot)\<^sup>\<star> = x\<^sup>\<star> \<squnion> x\<^sup>\<star> * y * bot" nitpick [expect=genuine,card=7] oops
lemma circ_simulate_left: "x * z \<le> z * y \<squnion> w \<longrightarrow> x\<^sup>\<star> * z \<le> (z \<squnion> x\<^sup>\<star> * w) * y\<^sup>\<star>" nitpick [expect=genuine] oops
lemma circ_simulate_1: "y * x \<le> x * y \<longrightarrow> y\<^sup>\<star> * x\<^sup>\<star> \<le> x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_separate_1: "y * x \<le> x * y \<longrightarrow> (x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma atomicity_refinement: "s = s * q \<and> x = q * x \<and> q * b = bot \<and> r * b \<le> b * r \<and> r * l \<le> l * r \<and> x * l \<le> l * x \<and> b * l \<le> l * b \<and> q * l \<le> l * q \<and> r\<^sup>\<star> * q \<le> q * r\<^sup>\<star> \<and> q \<le> 1 \<longrightarrow> s * (x \<squnion> b \<squnion> r \<squnion> l)\<^sup>\<star> * q \<le> s * (x * b\<^sup>\<star> * q \<squnion> r \<squnion> l)\<^sup>\<star>" nitpick [expect=genuine] oops
lemma circ_simulate_left_plus: "x * z \<le> z * y\<^sup>\<star> \<squnion> w \<longrightarrow> x\<^sup>\<star> * z \<le> (z \<squnion> x\<^sup>\<star> * w) * y\<^sup>\<star>" nitpick [expect=genuine] oops
lemma circ_separate_unfold: "(y * x\<^sup>\<star>)\<^sup>\<star> = y\<^sup>\<star> \<squnion> y\<^sup>\<star> * y * x * x\<^sup>\<star> * (y * x\<^sup>\<star>)\<^sup>\<star>" nitpick [expect=genuine] oops
lemma separation: "y * x \<le> x * y\<^sup>\<star> \<longrightarrow> (x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_simulate_4: "y * x \<le> x * x\<^sup>\<star> * (1 \<squnion> y) \<longrightarrow> y\<^sup>\<star> * x\<^sup>\<star> \<le> x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_simulate_5: "y * x \<le> x * x\<^sup>\<star> * (x \<squnion> y) \<longrightarrow> y\<^sup>\<star> * x\<^sup>\<star> \<le> x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_simulate_6: "y * x \<le> x * (x \<squnion> y) \<longrightarrow> y\<^sup>\<star> * x\<^sup>\<star> \<le> x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_separate_4: "y * x \<le> x * x\<^sup>\<star> * (1 \<squnion> y) \<longrightarrow> (x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_separate_5: "y * x \<le> x * x\<^sup>\<star> * (x \<squnion> y) \<longrightarrow> (x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
lemma circ_separate_6: "y * x \<le> x * (x \<squnion> y) \<longrightarrow> (x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
*)
end
text \<open>
We next add the right induction rule, which allows us to strengthen many inequalities of left Kleene algebras to equalities.
\<close>
class strong_left_kleene_algebra = left_kleene_algebra +
assumes star_right_induct: "z \<squnion> x * y \<le> x \<longrightarrow> z * y\<^sup>\<star> \<le> x"
begin
lemma star_plus:
"y\<^sup>\<star> * y = y * y\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "y\<^sup>\<star> * y \<le> y * y\<^sup>\<star>"
by (simp add: star.circ_plus_sub)
next
have "y\<^sup>\<star> * y * y \<le> y\<^sup>\<star> * y"
by (simp add: mult_left_isotone star.right_plus_below_circ)
hence "y \<squnion> y\<^sup>\<star> * y * y \<le> y\<^sup>\<star> * y"
by (simp add: star.circ_mult_increasing_2)
thus "y * y\<^sup>\<star> \<le> y\<^sup>\<star> * y"
using star_right_induct by blast
qed
lemma star_slide:
"(x * y)\<^sup>\<star> * x = x * (y * x)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "(x * y)\<^sup>\<star> * x \<le> x * (y * x)\<^sup>\<star>"
by (rule star_left_slide)
next
have "x \<squnion> (x * y)\<^sup>\<star> * x * y * x \<le> (x * y)\<^sup>\<star> * x"
by (metis (full_types) sup.commute eq_refl star.circ_loop_fixpoint mult.assoc star_plus)
thus "x * (y * x)\<^sup>\<star> \<le> (x * y)\<^sup>\<star> * x"
by (simp add: mult_assoc star_right_induct)
qed
lemma star_simulation_right:
assumes "z * x \<le> y * z"
shows "z * x\<^sup>\<star> \<le> y\<^sup>\<star> * z"
proof -
have "y\<^sup>\<star> * z * x \<le> y\<^sup>\<star> * z"
by (metis assms dual_order.trans mult_isotone mult_left_sub_dist_sup_right star.circ_loop_fixpoint star.circ_transitive_equal sup.cobounded1 mult_assoc)
thus ?thesis
by (metis le_supI star.circ_loop_fixpoint star_right_induct sup.cobounded2)
qed
end
text \<open>
Again we inherit results from the itering hierarchy.
\<close>
sublocale strong_left_kleene_algebra < star: itering_1 where circ = star
apply unfold_locales
apply (simp add: star_slide)
by (simp add: star_simulation_right)
context strong_left_kleene_algebra
begin
lemma star_right_induct_mult:
"y * x \<le> y \<Longrightarrow> y * x\<^sup>\<star> \<le> y"
by (simp add: star_right_induct)
lemma star_right_induct_mult_iff:
"y * x \<le> y \<longleftrightarrow> y * x\<^sup>\<star> \<le> y"
using mult_right_isotone order_trans star.circ_increasing star_right_induct_mult by blast
lemma star_simulation_right_equal:
"z * x = y * z \<Longrightarrow> z * x\<^sup>\<star> = y\<^sup>\<star> * z"
- by (metis eq_iff star_simulation_left star_simulation_right)
+ by (metis order.eq_iff star_simulation_left star_simulation_right)
lemma star_simulation_star:
"x * y \<le> y * x \<Longrightarrow> x\<^sup>\<star> * y\<^sup>\<star> \<le> y\<^sup>\<star> * x\<^sup>\<star>"
by (simp add: star_simulation_left star_simulation_right)
lemma star_right_induct_equal:
"z \<squnion> y * x = y \<Longrightarrow> z * x\<^sup>\<star> \<le> y"
by (simp add: star_right_induct)
lemma star_right_induct_mult_equal:
"y * x = y \<Longrightarrow> y * x\<^sup>\<star> \<le> y"
by (simp add: star_right_induct_mult)
lemma star_back_loop_least_fixpoint:
"x * y \<squnion> z = x \<Longrightarrow> z * y\<^sup>\<star> \<le> x"
by (simp add: sup_commute star_right_induct_equal)
lemma star_back_loop_is_least_fixpoint:
"is_least_fixpoint (\<lambda>x . x * y \<squnion> z) (z * y\<^sup>\<star>)"
proof (unfold is_least_fixpoint_def, rule conjI)
have "(z * y\<^sup>\<star> * y \<squnion> z) * y \<le> z * y\<^sup>\<star> * y \<squnion> z"
using le_supI1 mult_left_isotone star.circ_back_loop_prefixpoint by auto
hence "z * y\<^sup>\<star> \<le> z * y\<^sup>\<star> * y \<squnion> z"
by (simp add: star_right_induct)
thus "z * y\<^sup>\<star> * y \<squnion> z = z * y\<^sup>\<star>"
- using antisym star.circ_back_loop_prefixpoint by auto
+ using order.antisym star.circ_back_loop_prefixpoint by auto
next
show "\<forall>x. x * y \<squnion> z = x \<longrightarrow> z * y\<^sup>\<star> \<le> x"
by (simp add: star_back_loop_least_fixpoint)
qed
lemma star_back_loop_mu:
"\<mu> (\<lambda>x . x * y \<squnion> z) = z * y\<^sup>\<star>"
by (metis least_fixpoint_same star_back_loop_is_least_fixpoint)
lemma star_square:
"x\<^sup>\<star> = (1 \<squnion> x) * (x * x)\<^sup>\<star>"
proof -
let ?f = "\<lambda>y . y * x \<squnion> 1"
have 1: "isotone ?f"
by (metis sup_left_isotone isotone_def mult_left_isotone)
have "?f \<circ> ?f = (\<lambda>y . y * (x * x) \<squnion> (1 \<squnion> x))"
by (simp add: sup_assoc sup_commute mult_assoc mult_right_dist_sup o_def)
thus ?thesis
using 1 by (metis mu_square mult_left_one star_back_loop_mu has_least_fixpoint_def star_back_loop_is_least_fixpoint)
qed
lemma star_square_2:
"x\<^sup>\<star> = (x * x)\<^sup>\<star> * (x \<squnion> 1)"
proof -
have "(1 \<squnion> x) * (x * x)\<^sup>\<star> = (x * x)\<^sup>\<star> * 1 \<squnion> x * (x * x)\<^sup>\<star>"
using mult_right_dist_sup by force
thus ?thesis
- by (metis (no_types) antisym mult_left_sub_dist_sup star.circ_square_2 star_slide sup_commute star_square)
+ by (metis (no_types) order.antisym mult_left_sub_dist_sup star.circ_square_2 star_slide sup_commute star_square)
qed
lemma star_circ_simulate_right_plus:
assumes "z * x \<le> y * y\<^sup>\<star> * z \<squnion> w"
shows "z * x\<^sup>\<star> \<le> y\<^sup>\<star> * (z \<squnion> w * x\<^sup>\<star>)"
proof -
have "(z \<squnion> w * x\<^sup>\<star>) * x \<le> z * x \<squnion> w * x\<^sup>\<star>"
using mult_right_dist_sup star.circ_back_loop_prefixpoint sup_right_isotone by auto
also have "... \<le> y * y\<^sup>\<star> * z \<squnion> w \<squnion> w * x\<^sup>\<star>"
using assms sup_left_isotone by blast
also have "... \<le> y * y\<^sup>\<star> * z \<squnion> w * x\<^sup>\<star>"
using le_supI1 star.circ_back_loop_prefixpoint sup_commute by auto
also have "... \<le> y\<^sup>\<star> * (z \<squnion> w * x\<^sup>\<star>)"
by (metis sup.bounded_iff mult_isotone mult_left_isotone mult_left_one mult_left_sub_dist_sup_left star.circ_reflexive star.left_plus_below_circ)
finally have "y\<^sup>\<star> * (z \<squnion> w * x\<^sup>\<star>) * x \<le> y\<^sup>\<star> * (z \<squnion> w * x\<^sup>\<star>)"
by (metis mult_assoc mult_right_isotone star.circ_transitive_equal)
thus ?thesis
by (metis sup.bounded_iff star_right_induct mult_left_sub_dist_sup_left star.circ_loop_fixpoint)
qed
lemma transitive_star:
"x * x \<le> x \<Longrightarrow> x\<^sup>\<star> = 1 \<squnion> x"
by (metis order.antisym star.circ_mult_increasing_2 star.circ_plus_same star_left_induct_mult star_left_unfold_equal)
lemma star_sup_2:
assumes "x * x \<le> x"
and "x * y \<le> x"
shows "(x \<squnion> y)\<^sup>\<star> = y\<^sup>\<star> * (x \<squnion> 1)"
proof -
have "(x \<squnion> y)\<^sup>\<star> = y\<^sup>\<star> * (x * y\<^sup>\<star>)\<^sup>\<star>"
by (simp add: star.circ_decompose_6 star_sup_1)
also have "... = y\<^sup>\<star> * x\<^sup>\<star>"
using assms(2) dual_order.antisym star.circ_back_loop_prefixpoint star_right_induct_mult by fastforce
also have "... = y\<^sup>\<star> * (x \<squnion> 1)"
by (simp add: assms(1) sup_commute transitive_star)
finally show ?thesis
.
qed
(*
lemma star_circ_simulate_left_plus: "x * z \<le> z * y\<^sup>\<star> \<squnion> w \<longrightarrow> x\<^sup>\<star> * z \<le> (z \<squnion> x\<^sup>\<star> * w) * y\<^sup>\<star>" nitpick [expect=genuine,card=7] oops
*)
end
text \<open>
The following class contains a generalisation of Kleene algebras, which lacks the right zero axiom.
\<close>
class left_zero_kleene_algebra = idempotent_left_zero_semiring + strong_left_kleene_algebra
begin
lemma star_star_absorb:
"y\<^sup>\<star> * (y\<^sup>\<star> * x)\<^sup>\<star> * y\<^sup>\<star> = (y\<^sup>\<star> * x)\<^sup>\<star> * y\<^sup>\<star>"
by (metis star.circ_transitive_equal star_slide mult_assoc)
lemma star_circ_simulate_left_plus:
assumes "x * z \<le> z * y\<^sup>\<star> \<squnion> w"
shows "x\<^sup>\<star> * z \<le> (z \<squnion> x\<^sup>\<star> * w) * y\<^sup>\<star>"
proof -
have "x * (x\<^sup>\<star> * (w * y\<^sup>\<star>)) \<le> x\<^sup>\<star> * (w * y\<^sup>\<star>)"
by (metis (no_types) mult_right_sub_dist_sup_left star.circ_loop_fixpoint mult_assoc)
hence "x * ((z \<squnion> x\<^sup>\<star> * w) * y\<^sup>\<star>) \<le> x * z * y\<^sup>\<star> \<squnion> x\<^sup>\<star> * w * y\<^sup>\<star>"
using mult_left_dist_sup mult_right_dist_sup sup_right_isotone mult_assoc by presburger
also have "... \<le> (z * y\<^sup>\<star> \<squnion> w) * y\<^sup>\<star> \<squnion> x\<^sup>\<star> * w * y\<^sup>\<star>"
using assms mult_isotone semiring.add_right_mono by blast
also have "... = z * y\<^sup>\<star> \<squnion> w * y\<^sup>\<star> \<squnion> x\<^sup>\<star> * w * y\<^sup>\<star>"
by (simp add: mult_right_dist_sup star.circ_transitive_equal mult_assoc)
also have "... = (z \<squnion> w \<squnion> x\<^sup>\<star> * w) * y\<^sup>\<star>"
by (simp add: mult_right_dist_sup)
also have "... = (z \<squnion> x\<^sup>\<star> * w) * y\<^sup>\<star>"
by (metis sup_assoc sup_ge2 le_iff_sup star.circ_loop_fixpoint)
finally show ?thesis
by (metis sup.bounded_iff mult_left_sub_dist_sup_left mult_1_right star.circ_right_unfold_1 star_left_induct)
qed
lemma star_one_sup_below:
"x * y\<^sup>\<star> * (1 \<squnion> z) \<le> x * (y \<squnion> z)\<^sup>\<star>"
proof -
have "y\<^sup>\<star> * z \<le> (y \<squnion> z)\<^sup>\<star>"
using sup_ge2 order_trans star.circ_increasing star.circ_mult_upper_bound star.circ_sub_dist by blast
hence "y\<^sup>\<star> \<squnion> y\<^sup>\<star> * z \<le> (y \<squnion> z)\<^sup>\<star>"
by (simp add: star.circ_sup_upper_bound star.circ_sub_dist)
hence "y\<^sup>\<star> * (1 \<squnion> z) \<le> (y \<squnion> z)\<^sup>\<star>"
by (simp add: mult_left_dist_sup)
thus ?thesis
by (metis mult_right_isotone mult_assoc)
qed
text \<open>
The following theorem is similar to the puzzle where persons insert themselves always in the middle between two groups of people in a line.
Here, however, items in the middle annihilate each other, leaving just one group of items behind.
\<close>
lemma cancel_separate:
assumes "x * y \<le> 1"
shows "x\<^sup>\<star> * y\<^sup>\<star> \<le> x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
proof -
have "x * y\<^sup>\<star> = x \<squnion> x * y * y\<^sup>\<star>"
by (metis mult_assoc mult_left_dist_sup mult_1_right star_left_unfold_equal)
also have "... \<le> x \<squnion> y\<^sup>\<star>"
by (meson assms dual_order.trans order.refl star.circ_mult_upper_bound star.circ_reflexive sup_right_isotone)
also have "... \<le> x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
using star.circ_increasing sup_left_isotone by auto
finally have 1: "x * y\<^sup>\<star> \<le> x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
.
have "x * (x\<^sup>\<star> \<squnion> y\<^sup>\<star>) = x * x\<^sup>\<star> \<squnion> x * y\<^sup>\<star>"
by (simp add: mult_left_dist_sup)
also have "... \<le> x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
using 1 by (metis sup.bounded_iff sup_ge1 order_trans star.left_plus_below_circ)
finally have 2: "x * (x\<^sup>\<star> \<squnion> y\<^sup>\<star>) \<le> x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
.
have "y\<^sup>\<star> \<le> x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
by simp
hence "y\<^sup>\<star> \<squnion> x * (x\<^sup>\<star> \<squnion> y\<^sup>\<star>) \<le> x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
using 2 sup.bounded_iff by blast
thus ?thesis
by (metis star_left_induct)
qed
lemma star_separate:
assumes "x * y = bot"
and "y * y = bot"
shows "(x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>"
proof -
have 1: "y\<^sup>\<star> = 1 \<squnion> y"
using assms(2) by (simp add: transitive_star)
have "(x \<squnion> y)\<^sup>\<star> = y\<^sup>\<star> * (x * y\<^sup>\<star>)\<^sup>\<star>"
by (simp add: star.circ_decompose_6 star_sup_1)
also have "... = y\<^sup>\<star> * (x * (1 \<squnion> y * y\<^sup>\<star>))\<^sup>\<star>"
by (simp add: star_left_unfold_equal)
also have "... = (1 \<squnion> y) * x\<^sup>\<star>"
using 1 by (simp add: assms mult_left_dist_sup)
also have "... = x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>"
by (simp add: mult_right_dist_sup)
finally show ?thesis
.
qed
end
text \<open>
We can now inherit from the strongest variant of iterings.
\<close>
sublocale left_zero_kleene_algebra < star: itering where circ = star
apply unfold_locales
apply (metis star.circ_sup_9)
apply (metis star.circ_mult_1)
apply (simp add: star_circ_simulate_right_plus)
by (simp add: star_circ_simulate_left_plus)
context left_zero_kleene_algebra
begin
lemma star_absorb:
"x * y = bot \<Longrightarrow> x * y\<^sup>\<star> = x"
by (metis sup.bounded_iff antisym_conv star.circ_back_loop_prefixpoint star.circ_elimination)
lemma star_separate_2:
assumes "x * z\<^sup>+ * y = bot"
and "y * z\<^sup>+ * y = bot"
and "z * x = bot"
shows "(x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star> = z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star>"
proof -
have 1: "x\<^sup>\<star> * z\<^sup>+ * y = z\<^sup>+ * y"
by (metis assms mult_assoc mult_1_left mult_left_zero star.circ_zero star_simulation_right_equal)
have 2: "z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>+ \<le> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star>"
by (simp add: mult_right_isotone star.left_plus_below_circ)
have "z\<^sup>\<star> * z\<^sup>+ * y * x\<^sup>\<star> \<le> z\<^sup>\<star> * y * x\<^sup>\<star>"
by (metis mult_left_isotone star.left_plus_below_circ star.right_plus_circ star_plus)
also have "... \<le> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>)"
by (simp add: mult_assoc mult_left_sub_dist_sup_right)
also have "... \<le> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star>"
using sup_right_divisibility star.circ_back_loop_fixpoint by blast
finally have 3: "z\<^sup>\<star> * z\<^sup>+ * y * x\<^sup>\<star> \<le> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star>"
.
have "z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star> * (z * (1 \<squnion> y * x\<^sup>\<star>)) = z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>+ \<squnion> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>+ * y * x\<^sup>\<star>"
by (metis mult_1_right semiring.distrib_left star.circ_plus_same mult_assoc)
also have "... = z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>+ \<squnion> z\<^sup>\<star> * (1 \<squnion> y) * x\<^sup>\<star> * z\<^sup>+ * y * x\<^sup>\<star>"
by (simp add: semiring.distrib_right mult_assoc)
also have "... = z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>+ \<squnion> z\<^sup>\<star> * (1 \<squnion> y) * z\<^sup>+ * y * x\<^sup>\<star>"
using 1 by (simp add: mult_assoc)
also have "... = z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>+ \<squnion> z\<^sup>\<star> * z\<^sup>+ * y * x\<^sup>\<star> \<squnion> z\<^sup>\<star> * y * z\<^sup>+ * y * x\<^sup>\<star>"
using mult_left_dist_sup mult_right_dist_sup sup_assoc by auto
also have "... = z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>+ \<squnion> z\<^sup>\<star> * z\<^sup>+ * y * x\<^sup>\<star>"
by (metis assms(2) mult_left_dist_sup mult_left_zero sup_commute sup_monoid.add_0_left mult_assoc)
also have "... \<le> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star>"
using 2 3 by simp
finally have "(x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) \<squnion> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star> * (z * (1 \<squnion> y * x\<^sup>\<star>)) \<le> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star>"
by (simp add: star_outer_increasing)
hence 4: "(x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star> \<le> z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star>"
by (simp add: star_right_induct)
have 5: "(x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star> \<le> (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star>"
by (metis sup_ge1 mult_right_isotone mult_1_right star_isotone)
have "z * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) = z * x\<^sup>\<star> \<squnion> z * y * x\<^sup>\<star>"
by (simp add: mult_assoc mult_left_dist_sup)
also have "... = z \<squnion> z * y * x\<^sup>\<star>"
by (simp add: assms star_absorb)
also have "... = z * (1 \<squnion> y * x\<^sup>\<star>)"
by (simp add: mult_assoc mult_left_dist_sup)
also have "... \<le> (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star>"
by (simp add: star.circ_increasing)
also have "... \<le> (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star>"
by (metis le_supE mult_right_sub_dist_sup_left star.circ_loop_fixpoint)
finally have "z * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) \<le> (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star>"
.
hence "z * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star> \<le> (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star>"
by (metis mult_assoc mult_left_isotone star.circ_transitive_equal)
hence "z\<^sup>\<star> * (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * z\<^sup>\<star> \<le> (x\<^sup>\<star> \<squnion> y * x\<^sup>\<star>) * (z * (1 \<squnion> y * x\<^sup>\<star>))\<^sup>\<star>"
using 5 by (metis star_left_induct sup.bounded_iff mult_assoc)
thus ?thesis
using 4 by (simp add: antisym)
qed
lemma cancel_separate_eq:
"x * y \<le> 1 \<Longrightarrow> x\<^sup>\<star> * y\<^sup>\<star> = x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
- by (metis antisym cancel_separate star.circ_plus_one star.circ_sup_sub_sup_one_1 star_involutive)
+ by (metis order.antisym cancel_separate star.circ_plus_one star.circ_sup_sub_sup_one_1 star_involutive)
lemma cancel_separate_1:
assumes "x * y \<le> 1"
shows "(x \<squnion> y)\<^sup>\<star> = y\<^sup>\<star> * x\<^sup>\<star>"
proof -
have "y\<^sup>\<star> * x\<^sup>\<star> * y = y\<^sup>\<star> * x\<^sup>\<star> * x * y \<squnion> y\<^sup>\<star> * y"
by (metis mult_right_dist_sup star.circ_back_loop_fixpoint)
also have "... \<le> y\<^sup>\<star> * x\<^sup>\<star> \<squnion> y\<^sup>\<star> * y"
by (metis assms semiring.add_right_mono mult_right_isotone mult_1_right mult_assoc)
also have "... \<le> y\<^sup>\<star> * x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
using semiring.add_left_mono star.right_plus_below_circ by simp
also have "... = y\<^sup>\<star> * x\<^sup>\<star>"
by (metis star.circ_back_loop_fixpoint sup.left_idem sup_commute)
finally have "y\<^sup>\<star> * x\<^sup>\<star> * y \<le> y\<^sup>\<star> * x\<^sup>\<star>"
by simp
hence "y\<^sup>\<star> * x\<^sup>\<star> * (x \<squnion> y) \<le> y\<^sup>\<star> * x\<^sup>\<star> * x \<squnion> y\<^sup>\<star> * x\<^sup>\<star>"
using mult_left_dist_sup order_lesseq_imp by fastforce
also have "... = y\<^sup>\<star> * x\<^sup>\<star>"
by (metis star.circ_back_loop_fixpoint sup.left_idem)
finally have "(x \<squnion> y)\<^sup>\<star> \<le> y\<^sup>\<star> * x\<^sup>\<star>"
by (metis star.circ_decompose_7 star_right_induct_mult sup_commute)
thus ?thesis
- using antisym star.circ_sub_dist_3 sup_commute by fastforce
+ using order.antisym star.circ_sub_dist_3 sup_commute by fastforce
qed
lemma plus_sup:
"(x \<squnion> y)\<^sup>+ = (x\<^sup>\<star> * y)\<^sup>\<star> * x\<^sup>+ \<squnion> (x\<^sup>\<star> * y)\<^sup>+"
by (metis semiring.distrib_left star.circ_sup_9 star_plus mult_assoc)
lemma plus_half_bot:
"x * y * x = bot \<Longrightarrow> (x * y)\<^sup>+ = x * y"
by (metis star_absorb star_slide mult_assoc)
lemma cancel_separate_1_sup:
assumes "x * y \<le> 1"
and "y * x \<le> 1"
shows "(x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> \<squnion> y\<^sup>\<star>"
by (simp add: assms cancel_separate_1 cancel_separate_eq sup_commute)
text \<open>Lemma \<open>star_separate_3\<close> was contributed by Nicolas Robinson-O'Brien.\<close>
lemma star_separate_3:
assumes "y * x\<^sup>\<star> * y \<le> y"
shows "(x \<squnion> y)\<^sup>\<star> = x\<^sup>\<star> \<squnion> x\<^sup>\<star> * y * x\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x\<^sup>\<star> * y * (x\<^sup>\<star> * y)\<^sup>\<star> * x\<^sup>\<star> \<le> x\<^sup>\<star> * y * x\<^sup>\<star>"
by (metis assms mult_left_isotone mult_right_isotone star_right_induct_mult mult_assoc)
thus "(x \<squnion> y)\<^sup>\<star> \<le> x\<^sup>\<star> \<squnion> x\<^sup>\<star> * y * x\<^sup>\<star>"
- by (metis antisym semiring.add_left_mono star.circ_sup_2 star.circ_sup_sub star.circ_unfold_sum star_decompose_3 star_slide mult_assoc)
+ by (metis order.antisym semiring.add_left_mono star.circ_sup_2 star.circ_sup_sub star.circ_unfold_sum star_decompose_3 star_slide mult_assoc)
next
show "x\<^sup>\<star> \<squnion> x\<^sup>\<star> * y * x\<^sup>\<star> \<le> (x \<squnion> y)\<^sup>\<star>"
using mult_isotone star.circ_increasing star.circ_sub_dist star.circ_sup_9 by auto
qed
end
text \<open>
A Kleene algebra is obtained by requiring an idempotent semiring.
\<close>
class kleene_algebra = left_zero_kleene_algebra + idempotent_semiring
text \<open>
The following classes are variants of Kleene algebras expanded by an additional iteration operation.
This is useful to study the Kleene star in computation models that do not use least fixpoints in the refinement order as the semantics of recursion.
\<close>
class left_kleene_conway_semiring = left_kleene_algebra + left_conway_semiring
begin
lemma star_below_circ:
"x\<^sup>\<star> \<le> x\<^sup>\<circ>"
by (metis circ_left_unfold mult_1_right order_refl star_left_induct)
lemma star_zero_below_circ_mult:
"x\<^sup>\<star> * bot \<le> x\<^sup>\<circ> * y"
by (simp add: mult_isotone star_below_circ)
lemma star_mult_circ:
"x\<^sup>\<star> * x\<^sup>\<circ> = x\<^sup>\<circ>"
- by (metis sup_right_divisibility antisym circ_left_unfold star_left_induct_mult star.circ_loop_fixpoint)
+ by (metis sup_right_divisibility order.antisym circ_left_unfold star_left_induct_mult star.circ_loop_fixpoint)
lemma circ_mult_star:
"x\<^sup>\<circ> * x\<^sup>\<star> = x\<^sup>\<circ>"
- by (metis sup_assoc sup.bounded_iff circ_left_unfold circ_rtc_2 eq_iff left_plus_circ star.circ_sup_sub star.circ_back_loop_prefixpoint star.circ_increasing star_below_circ star_mult_circ star_sup_one)
+ by (metis sup_assoc sup.bounded_iff circ_left_unfold circ_rtc_2 order.eq_iff left_plus_circ star.circ_sup_sub star.circ_back_loop_prefixpoint star.circ_increasing star_below_circ star_mult_circ star_sup_one)
lemma circ_star:
"x\<^sup>\<circ>\<^sup>\<star> = x\<^sup>\<circ>"
- by (metis antisym circ_reflexive circ_transitive_equal star.circ_increasing star.circ_sup_one_right_unfold star_left_induct_mult_equal)
+ by (metis order.antisym circ_reflexive circ_transitive_equal star.circ_increasing star.circ_sup_one_right_unfold star_left_induct_mult_equal)
lemma star_circ:
"x\<^sup>\<star>\<^sup>\<circ> = x\<^sup>\<circ>\<^sup>\<circ>"
- by (metis antisym circ_circ_sup circ_sub_dist le_iff_sup star.circ_rtc_2 star_below_circ)
+ by (metis order.antisym circ_circ_sup circ_sub_dist le_iff_sup star.circ_rtc_2 star_below_circ)
lemma circ_sup_3:
"(x\<^sup>\<circ> * y\<^sup>\<circ>)\<^sup>\<star> \<le> (x \<squnion> y)\<^sup>\<circ>"
using circ_star circ_sub_dist_3 star_isotone by fastforce
end
class left_zero_kleene_conway_semiring = left_zero_kleene_algebra + itering
begin
subclass left_kleene_conway_semiring ..
lemma circ_isolate:
"x\<^sup>\<circ> = x\<^sup>\<circ> * bot \<squnion> x\<^sup>\<star>"
- by (metis sup_commute antisym circ_sup_upper_bound circ_mult_star circ_simulate_absorb star.left_plus_below_circ star_below_circ zero_right_mult_decreasing)
+ by (metis sup_commute order.antisym circ_sup_upper_bound circ_mult_star circ_simulate_absorb star.left_plus_below_circ star_below_circ zero_right_mult_decreasing)
lemma circ_isolate_mult:
"x\<^sup>\<circ> * y = x\<^sup>\<circ> * bot \<squnion> x\<^sup>\<star> * y"
by (metis circ_isolate mult_assoc mult_left_zero mult_right_dist_sup)
lemma circ_isolate_mult_sub:
"x\<^sup>\<circ> * y \<le> x\<^sup>\<circ> \<squnion> x\<^sup>\<star> * y"
by (metis sup_left_isotone circ_isolate_mult zero_right_mult_decreasing)
lemma circ_sub_decompose:
"(x\<^sup>\<circ> * y)\<^sup>\<circ> \<le> (x\<^sup>\<star> * y)\<^sup>\<circ> * x\<^sup>\<circ>"
proof -
have "x\<^sup>\<star> * y \<squnion> x\<^sup>\<circ> * bot = x\<^sup>\<circ> * y"
by (metis sup.commute circ_isolate_mult)
hence "(x\<^sup>\<star> * y)\<^sup>\<circ> * x\<^sup>\<circ> = ((x\<^sup>\<circ> * y)\<^sup>\<circ> \<squnion> x\<^sup>\<circ>)\<^sup>\<star>"
by (metis circ_star circ_sup_9 circ_sup_mult_zero star_decompose_1)
thus ?thesis
by (metis circ_star le_iff_sup star.circ_decompose_7 star.circ_unfold_sum)
qed
lemma circ_sup_4:
"(x \<squnion> y)\<^sup>\<circ> = (x\<^sup>\<star> * y)\<^sup>\<circ> * x\<^sup>\<circ>"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis circ_sup circ_sub_decompose circ_transitive_equal mult_assoc mult_left_isotone)
by (metis circ_sup circ_isotone mult_left_isotone star_below_circ)
lemma circ_sup_5:
"(x\<^sup>\<circ> * y)\<^sup>\<circ> * x\<^sup>\<circ> = (x\<^sup>\<star> * y)\<^sup>\<circ> * x\<^sup>\<circ>"
using circ_sup_4 circ_sup_9 by auto
lemma plus_circ:
"(x\<^sup>\<star> * x)\<^sup>\<circ> = x\<^sup>\<circ>"
by (metis sup_idem circ_sup_4 circ_decompose_7 circ_star star.circ_decompose_5 star.right_plus_circ)
(*
lemma "(x\<^sup>\<star> * y * x\<^sup>\<star>)\<^sup>\<circ> = (x\<^sup>\<star> * y)\<^sup>\<circ>" nitpick [expect=genuine] oops
*)
end
text \<open>
The following classes add a greatest element.
\<close>
class bounded_left_kleene_algebra = bounded_idempotent_left_semiring + left_kleene_algebra
sublocale bounded_left_kleene_algebra < star: bounded_left_conway_semiring where circ = star ..
class bounded_left_zero_kleene_algebra = bounded_idempotent_left_semiring + left_zero_kleene_algebra
sublocale bounded_left_zero_kleene_algebra < star: bounded_itering where circ = star ..
class bounded_kleene_algebra = bounded_idempotent_semiring + kleene_algebra
sublocale bounded_kleene_algebra < star: bounded_itering where circ = star ..
text \<open>
We conclude with an alternative axiomatisation of Kleene algebras.
\<close>
class kleene_algebra_var = idempotent_semiring + star +
assumes star_left_unfold_var : "1 \<squnion> y * y\<^sup>\<star> \<le> y\<^sup>\<star>"
assumes star_left_induct_var : "y * x \<le> x \<longrightarrow> y\<^sup>\<star> * x \<le> x"
assumes star_right_induct_var : "x * y \<le> x \<longrightarrow> x * y\<^sup>\<star> \<le> x"
begin
subclass kleene_algebra
apply unfold_locales
apply (rule star_left_unfold_var)
apply (meson sup.bounded_iff mult_right_isotone order_trans star_left_induct_var)
by (meson sup.bounded_iff mult_left_isotone order_trans star_right_induct_var)
end
end
diff --git a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy
--- a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy
+++ b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy
@@ -1,3269 +1,3269 @@
(* Title: Kleene Relation Algebras
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Kleene Relation Algebras\<close>
text \<open>
This theory combines Kleene algebras with Stone relation algebras.
Relation algebras with transitive closure have been studied by \cite{Ng1984}.
The weakening to Stone relation algebras allows us to talk about reachability in weighted graphs, for example.
Many results in this theory are used in the correctness proof of Prim's minimum spanning tree algorithm.
In particular, they are concerned with the exchange property, preservation of parts of the invariant and with establishing parts of the postcondition.
\<close>
theory Kleene_Relation_Algebras
imports Stone_Relation_Algebras.Relation_Algebras Kleene_Algebras
begin
text \<open>
We first note that bounded distributive lattices can be expanded to Kleene algebras by reusing some of the operations.
\<close>
sublocale bounded_distrib_lattice < comp_inf: bounded_kleene_algebra where star = "\<lambda>x . top" and one = top and times = inf
apply unfold_locales
apply (simp add: inf.assoc)
apply simp
apply simp
apply (simp add: le_infI2)
apply (simp add: inf_sup_distrib2)
apply simp
apply simp
apply simp
apply simp
apply simp
apply (simp add: inf_sup_distrib1)
apply simp
apply simp
by (simp add: inf_assoc)
text \<open>
We add the Kleene star operation to each of bounded distributive allegories, pseudocomplemented distributive allegories and Stone relation algebras.
We start with single-object bounded distributive allegories.
\<close>
class bounded_distrib_kleene_allegory = bounded_distrib_allegory + kleene_algebra
begin
subclass bounded_kleene_algebra ..
lemma conv_star_conv:
"x\<^sup>\<star> \<le> x\<^sup>T\<^sup>\<star>\<^sup>T"
proof -
have "x\<^sup>T\<^sup>\<star> * x\<^sup>T \<le> x\<^sup>T\<^sup>\<star>"
by (simp add: star.right_plus_below_circ)
hence 1: "x * x\<^sup>T\<^sup>\<star>\<^sup>T \<le> x\<^sup>T\<^sup>\<star>\<^sup>T"
using conv_dist_comp conv_isotone by fastforce
have "1 \<le> x\<^sup>T\<^sup>\<star>\<^sup>T"
by (simp add: reflexive_conv_closed star.circ_reflexive)
hence "1 \<squnion> x * x\<^sup>T\<^sup>\<star>\<^sup>T \<le> x\<^sup>T\<^sup>\<star>\<^sup>T"
using 1 by simp
thus ?thesis
using star_left_induct by fastforce
qed
text \<open>
It follows that star and converse commute.
\<close>
lemma conv_star_commute:
"x\<^sup>\<star>\<^sup>T = x\<^sup>T\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
show "x\<^sup>\<star>\<^sup>T \<le> x\<^sup>T\<^sup>\<star>"
using conv_star_conv conv_isotone by fastforce
next
show "x\<^sup>T\<^sup>\<star> \<le> x\<^sup>\<star>\<^sup>T"
by (metis conv_star_conv conv_involutive)
qed
lemma conv_plus_commute:
"x\<^sup>+\<^sup>T = x\<^sup>T\<^sup>+"
by (simp add: conv_dist_comp conv_star_commute star_plus)
text \<open>Lemma \<open>reflexive_inf_star\<close> was contributed by Nicolas Robinson-O'Brien.\<close>
lemma reflexive_inf_star:
assumes "reflexive y"
shows "y \<sqinter> x\<^sup>\<star> = 1 \<squnion> (y \<sqinter> x\<^sup>+)"
by (simp add: assms star_left_unfold_equal sup.absorb2 sup_inf_distrib1)
text \<open>
The following results are variants of a separation lemma of Kleene algebras.
\<close>
lemma cancel_separate_2:
assumes "x * y \<le> 1"
shows "((w \<sqinter> x) \<squnion> (z \<sqinter> y))\<^sup>\<star> = (z \<sqinter> y)\<^sup>\<star> * (w \<sqinter> x)\<^sup>\<star>"
proof -
have "(w \<sqinter> x) * (z \<sqinter> y) \<le> 1"
by (meson assms comp_isotone order.trans inf.cobounded2)
thus ?thesis
using cancel_separate_1 sup_commute by simp
qed
lemma cancel_separate_3:
assumes "x * y \<le> 1"
shows "(w \<sqinter> x)\<^sup>\<star> * (z \<sqinter> y)\<^sup>\<star> = (w \<sqinter> x)\<^sup>\<star> \<squnion> (z \<sqinter> y)\<^sup>\<star>"
proof -
have "(w \<sqinter> x) * (z \<sqinter> y) \<le> 1"
by (meson assms comp_isotone order.trans inf.cobounded2)
thus ?thesis
by (simp add: cancel_separate_eq)
qed
lemma cancel_separate_4:
assumes "z * y \<le> 1"
and "w \<le> y \<squnion> z"
and "x \<le> y \<squnion> z"
shows "w\<^sup>\<star> * x\<^sup>\<star> = (w \<sqinter> y)\<^sup>\<star> * ((w \<sqinter> z)\<^sup>\<star> \<squnion> (x \<sqinter> y)\<^sup>\<star>) * (x \<sqinter> z)\<^sup>\<star>"
proof -
have "w\<^sup>\<star> * x\<^sup>\<star> = ((w \<sqinter> y) \<squnion> (w \<sqinter> z))\<^sup>\<star> * ((x \<sqinter> y) \<squnion> (x \<sqinter> z))\<^sup>\<star>"
by (metis assms(2,3) inf.orderE inf_sup_distrib1)
also have "... = (w \<sqinter> y)\<^sup>\<star> * ((w \<sqinter> z)\<^sup>\<star> * (x \<sqinter> y)\<^sup>\<star>) * (x \<sqinter> z)\<^sup>\<star>"
by (metis assms(1) cancel_separate_2 sup_commute mult_assoc)
finally show ?thesis
by (simp add: assms(1) cancel_separate_3)
qed
lemma cancel_separate_5:
assumes "w * z\<^sup>T \<le> 1"
shows "w \<sqinter> x * (y \<sqinter> z) \<le> y"
proof -
have "w \<sqinter> x * (y \<sqinter> z) \<le> (x \<sqinter> w * (y \<sqinter> z)\<^sup>T) * (y \<sqinter> z)"
by (metis dedekind_2 inf_commute)
also have "... \<le> w * z\<^sup>T * (y \<sqinter> z)"
by (simp add: conv_dist_inf inf.coboundedI2 mult_left_isotone mult_right_isotone)
also have "... \<le> y \<sqinter> z"
by (metis assms mult_1_left mult_left_isotone)
finally show ?thesis
by simp
qed
lemma cancel_separate_6:
assumes "z * y \<le> 1"
and "w \<le> y \<squnion> z"
and "x \<le> y \<squnion> z"
and "v * z\<^sup>T \<le> 1"
and "v \<sqinter> y\<^sup>\<star> = bot"
shows "v \<sqinter> w\<^sup>\<star> * x\<^sup>\<star> \<le> x \<squnion> w"
proof -
have "v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * (x \<sqinter> y)\<^sup>\<star> \<le> v \<sqinter> y\<^sup>\<star> * (x \<sqinter> y)\<^sup>\<star>"
using comp_inf.mult_right_isotone mult_left_isotone star_isotone by simp
also have "... \<le> v \<sqinter> y\<^sup>\<star>"
by (simp add: inf.coboundedI2 star.circ_increasing star.circ_mult_upper_bound star_right_induct_mult)
finally have 1: "v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * (x \<sqinter> y)\<^sup>\<star> = bot"
using assms(5) le_bot by simp
have "v \<sqinter> w\<^sup>\<star> * x\<^sup>\<star> = v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * ((w \<sqinter> z)\<^sup>\<star> \<squnion> (x \<sqinter> y)\<^sup>\<star>) * (x \<sqinter> z)\<^sup>\<star>"
using assms(1-3) cancel_separate_4 by simp
also have "... = (v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * ((w \<sqinter> z)\<^sup>\<star> \<squnion> (x \<sqinter> y)\<^sup>\<star>) * (x \<sqinter> z)\<^sup>\<star> * (x \<sqinter> z)) \<squnion> (v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * ((w \<sqinter> z)\<^sup>\<star> \<squnion> (x \<sqinter> y)\<^sup>\<star>))"
by (metis inf_sup_distrib1 star.circ_back_loop_fixpoint)
also have "... \<le> x \<squnion> (v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * ((w \<sqinter> z)\<^sup>\<star> \<squnion> (x \<sqinter> y)\<^sup>\<star>))"
using assms(4) cancel_separate_5 semiring.add_right_mono by simp
also have "... = x \<squnion> (v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * (w \<sqinter> z)\<^sup>\<star>)"
using 1 by (simp add: inf_sup_distrib1 mult_left_dist_sup sup_monoid.add_assoc)
also have "... = x \<squnion> (v \<sqinter> (w \<sqinter> y)\<^sup>\<star> * (w \<sqinter> z)\<^sup>\<star> * (w \<sqinter> z)) \<squnion> (v \<sqinter> (w \<sqinter> y)\<^sup>\<star>)"
by (metis comp_inf.semiring.distrib_left star.circ_back_loop_fixpoint sup_assoc)
also have "... \<le> x \<squnion> w \<squnion> (v \<sqinter> (w \<sqinter> y)\<^sup>\<star>)"
using assms(4) cancel_separate_5 sup_left_isotone sup_right_isotone by simp
also have "... \<le> x \<squnion> w \<squnion> (v \<sqinter> y\<^sup>\<star>)"
using comp_inf.mult_right_isotone star_isotone sup_right_isotone by simp
finally show ?thesis
using assms(5) le_bot by simp
qed
text \<open>
We show several results about the interaction of vectors and the Kleene star.
\<close>
lemma vector_star_1:
assumes "vector x"
shows "x\<^sup>T * (x * x\<^sup>T)\<^sup>\<star> \<le> x\<^sup>T"
proof -
have "x\<^sup>T * (x * x\<^sup>T)\<^sup>\<star> = (x\<^sup>T * x)\<^sup>\<star> * x\<^sup>T"
by (simp add: star_slide)
also have "... \<le> top * x\<^sup>T"
by (simp add: mult_left_isotone)
also have "... = x\<^sup>T"
using assms vector_conv_covector by auto
finally show ?thesis
.
qed
lemma vector_star_2:
"vector x \<Longrightarrow> x\<^sup>T * (x * x\<^sup>T)\<^sup>\<star> \<le> x\<^sup>T * bot\<^sup>\<star>"
by (simp add: star_absorb vector_star_1)
lemma vector_vector_star:
"vector v \<Longrightarrow> (v * v\<^sup>T)\<^sup>\<star> = 1 \<squnion> v * v\<^sup>T"
by (simp add: transitive_star vv_transitive)
lemma equivalence_star_closed:
"equivalence x \<Longrightarrow> equivalence (x\<^sup>\<star>)"
by (simp add: conv_star_commute star.circ_reflexive star.circ_transitive_equal)
lemma equivalence_plus_closed:
"equivalence x \<Longrightarrow> equivalence (x\<^sup>+)"
by (simp add: conv_star_commute star.circ_reflexive star.circ_sup_one_left_unfold star.circ_transitive_equal)
text \<open>
The following equivalence relation characterises the component trees of a forest.
This is a special case of undirected reachability in a directed graph.
\<close>
abbreviation "forest_components f \<equiv> f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>"
lemma forest_components_equivalence:
"injective x \<Longrightarrow> equivalence (forest_components x)"
apply (intro conjI)
apply (simp add: reflexive_mult_closed star.circ_reflexive)
- apply (metis cancel_separate_1 eq_iff star.circ_transitive_equal)
+ apply (metis cancel_separate_1 order.eq_iff star.circ_transitive_equal)
by (simp add: conv_dist_comp conv_star_commute)
lemma forest_components_increasing:
"x \<le> forest_components x"
by (metis order.trans mult_left_isotone mult_left_one star.circ_increasing star.circ_reflexive)
lemma forest_components_isotone:
"x \<le> y \<Longrightarrow> forest_components x \<le> forest_components y"
by (simp add: comp_isotone conv_isotone star_isotone)
lemma forest_components_idempotent:
"injective x \<Longrightarrow> forest_components (forest_components x) = forest_components x"
by (metis forest_components_equivalence cancel_separate_1 star.circ_transitive_equal star_involutive)
lemma forest_components_star:
"injective x \<Longrightarrow> (forest_components x)\<^sup>\<star> = forest_components x"
using forest_components_equivalence forest_components_idempotent star.circ_transitive_equal by simp
text \<open>
The following lemma shows that the nodes reachable in the graph can be reached by only using edges between reachable nodes.
\<close>
lemma reachable_restrict:
assumes "vector r"
shows "r\<^sup>T * g\<^sup>\<star> = r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star>"
proof -
have 1: "r\<^sup>T \<le> r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star>"
using mult_right_isotone mult_1_right star.circ_reflexive by fastforce
have 2: "covector (r\<^sup>T * g\<^sup>\<star>)"
using assms covector_mult_closed vector_conv_covector by auto
have "r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> * g \<le> r\<^sup>T * g\<^sup>\<star> * g"
by (simp add: mult_left_isotone mult_right_isotone star_isotone)
also have "... \<le> r\<^sup>T * g\<^sup>\<star>"
by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus)
finally have "r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> * g = r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> * g \<sqinter> r\<^sup>T * g\<^sup>\<star>"
by (simp add: le_iff_inf)
also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> * (g \<sqinter> r\<^sup>T * g\<^sup>\<star>)"
using assms covector_comp_inf covector_mult_closed vector_conv_covector by auto
also have "... = (r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> \<sqinter> r\<^sup>T * g\<^sup>\<star>) * (g \<sqinter> r\<^sup>T * g\<^sup>\<star>)"
by (simp add: inf.absorb2 inf_commute mult_right_isotone star_isotone)
also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> * (g \<sqinter> r\<^sup>T * g\<^sup>\<star> \<sqinter> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T)"
using 2 by (metis comp_inf_vector_1)
also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T \<sqinter> r\<^sup>T * g\<^sup>\<star> \<sqinter> g)"
using inf_commute inf_assoc by simp
also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star> * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)"
using 2 by (metis covector_conv_vector inf_top.right_neutral vector_inf_comp)
also have "... \<le> r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star>"
by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus)
finally have "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g)\<^sup>\<star>"
using 1 star_right_induct by auto
thus ?thesis
- by (simp add: inf.eq_iff mult_right_isotone star_isotone)
+ by (simp add: order.eq_iff mult_right_isotone star_isotone)
qed
lemma kruskal_acyclic_inv_1:
assumes "injective f"
and "e * forest_components f * e = bot"
shows "(f \<sqinter> top * e * f\<^sup>T\<^sup>\<star>)\<^sup>T * f\<^sup>\<star> * e = bot"
proof -
let ?q = "top * e * f\<^sup>T\<^sup>\<star>"
let ?F = "forest_components f"
have "(f \<sqinter> ?q)\<^sup>T * f\<^sup>\<star> * e = ?q\<^sup>T \<sqinter> f\<^sup>T * f\<^sup>\<star> * e"
by (metis (mono_tags) comp_associative conv_dist_inf covector_conv_vector inf_vector_comp vector_top_closed)
also have "... \<le> ?q\<^sup>T \<sqinter> ?F * e"
using comp_inf.mult_right_isotone mult_left_isotone star.circ_increasing by simp
also have "... = f\<^sup>\<star> * e\<^sup>T * top \<sqinter> ?F * e"
by (simp add: conv_dist_comp conv_star_commute mult_assoc)
also have "... \<le> ?F * e\<^sup>T * top \<sqinter> ?F * e"
by (metis conv_dist_comp conv_star_commute conv_top inf.sup_left_isotone star.circ_right_top star_outer_increasing mult_assoc)
also have "... = ?F * (e\<^sup>T * top \<sqinter> ?F * e)"
by (metis assms(1) forest_components_equivalence equivalence_comp_dist_inf mult_assoc)
also have "... = (?F \<sqinter> top * e) * ?F * e"
by (simp add: comp_associative comp_inf_vector_1 conv_dist_comp inf_vector_comp)
also have "... \<le> top * e * ?F * e"
by (simp add: mult_left_isotone)
also have "... = bot"
using assms(2) mult_assoc by simp
finally show ?thesis
by (simp add: bot_unique)
qed
lemma kruskal_forest_components_inf_1:
assumes "f \<le> w \<squnion> w\<^sup>T"
and "injective w"
and "f \<le> forest_components g"
shows "f * forest_components (forest_components g \<sqinter> w) \<le> forest_components (forest_components g \<sqinter> w)"
proof -
let ?f = "forest_components g"
let ?w = "forest_components (?f \<sqinter> w)"
have "f * ?w = (f \<sqinter> (w \<squnion> w\<^sup>T)) * ?w"
by (simp add: assms(1) inf.absorb1)
also have "... = (f \<sqinter> w) * ?w \<squnion> (f \<sqinter> w\<^sup>T) * ?w"
by (simp add: inf_sup_distrib1 semiring.distrib_right)
also have "... \<le> (?f \<sqinter> w) * ?w \<squnion> (f \<sqinter> w\<^sup>T) * ?w"
using assms(3) inf.sup_left_isotone mult_left_isotone sup_left_isotone by simp
also have "... \<le> (?f \<sqinter> w) * ?w \<squnion> (?f \<sqinter> w\<^sup>T) * ?w"
using assms(3) inf.sup_left_isotone mult_left_isotone sup_right_isotone by simp
also have "... = (?f \<sqinter> w) * ?w \<squnion> (?f \<sqinter> w)\<^sup>T * ?w"
by (simp add: conv_dist_comp conv_dist_inf conv_star_commute)
also have "... \<le> (?f \<sqinter> w) * ?w \<squnion> ?w"
by (metis star.circ_loop_fixpoint sup_ge1 sup_right_isotone)
also have "... = ?w \<squnion> (?f \<sqinter> w) * (?f \<sqinter> w)\<^sup>\<star> \<squnion> (?f \<sqinter> w) * (?f \<sqinter> w)\<^sup>T\<^sup>+ * (?f \<sqinter> w)\<^sup>\<star>"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute sup_assoc)
also have "... \<le> ?w \<squnion> (?f \<sqinter> w)\<^sup>\<star> \<squnion> (?f \<sqinter> w) * (?f \<sqinter> w)\<^sup>T\<^sup>+ * (?f \<sqinter> w)\<^sup>\<star>"
using star.left_plus_below_circ sup_left_isotone sup_right_isotone by auto
also have "... = ?w \<squnion> (?f \<sqinter> w) * (?f \<sqinter> w)\<^sup>T\<^sup>+ * (?f \<sqinter> w)\<^sup>\<star>"
by (metis star.circ_loop_fixpoint sup.right_idem)
also have "... \<le> ?w \<squnion> w * w\<^sup>T * ?w"
using comp_associative conv_dist_inf mult_isotone sup_right_isotone by simp
also have "... = ?w"
by (metis assms(2) coreflexive_comp_top_inf inf.cobounded2 sup.orderE)
finally show ?thesis
by simp
qed
lemma kruskal_forest_components_inf:
assumes "f \<le> w \<squnion> w\<^sup>T"
and "injective w"
shows "forest_components f \<le> forest_components (forest_components f \<sqinter> w)"
proof -
let ?f = "forest_components f"
let ?w = "forest_components (?f \<sqinter> w)"
have 1: "1 \<le> ?w"
by (simp add: reflexive_mult_closed star.circ_reflexive)
have "f * ?w \<le> ?w"
using assms forest_components_increasing kruskal_forest_components_inf_1 by simp
hence 2: "f\<^sup>\<star> \<le> ?w"
using 1 star_left_induct by fastforce
have "f\<^sup>T * ?w \<le> ?w"
apply (rule kruskal_forest_components_inf_1)
apply (metis assms(1) conv_dist_sup conv_involutive conv_isotone sup_commute)
apply (simp add: assms(2))
by (metis le_supI2 star.circ_back_loop_fixpoint star.circ_increasing)
thus "?f \<le> ?w"
using 2 star_left_induct by simp
qed
end
text \<open>
We next add the Kleene star to single-object pseudocomplemented distributive allegories.
\<close>
class pd_kleene_allegory = pd_allegory + bounded_distrib_kleene_allegory
begin
text \<open>
The following definitions and results concern acyclic graphs and forests.
\<close>
abbreviation acyclic :: "'a \<Rightarrow> bool" where "acyclic x \<equiv> x\<^sup>+ \<le> -1"
abbreviation forest :: "'a \<Rightarrow> bool" where "forest x \<equiv> injective x \<and> acyclic x"
lemma forest_bot:
"forest bot"
by simp
lemma acyclic_down_closed:
"x \<le> y \<Longrightarrow> acyclic y \<Longrightarrow> acyclic x"
using comp_isotone star_isotone by fastforce
lemma forest_down_closed:
"x \<le> y \<Longrightarrow> forest y \<Longrightarrow> forest x"
using conv_isotone mult_isotone star_isotone by fastforce
lemma acyclic_star_below_complement:
"acyclic w \<longleftrightarrow> w\<^sup>T\<^sup>\<star> \<le> -w"
by (simp add: conv_star_commute schroeder_4_p)
lemma acyclic_star_below_complement_1:
"acyclic w \<longleftrightarrow> w\<^sup>\<star> \<sqinter> w\<^sup>T = bot"
using pseudo_complement schroeder_5_p by force
lemma acyclic_star_inf_conv:
assumes "acyclic w"
shows "w\<^sup>\<star> \<sqinter> w\<^sup>T\<^sup>\<star> = 1"
proof -
have "w\<^sup>+ \<sqinter> w\<^sup>T\<^sup>\<star> \<le> (w \<sqinter> w\<^sup>T\<^sup>\<star>) * w\<^sup>\<star>"
by (metis conv_star_commute dedekind_2 star.circ_transitive_equal)
also have "... = bot"
by (metis assms conv_star_commute p_antitone_iff pseudo_complement schroeder_4_p semiring.mult_not_zero star.circ_circ_mult star_involutive star_one)
finally have "w\<^sup>\<star> \<sqinter> w\<^sup>T\<^sup>\<star> \<le> 1"
- by (metis eq_iff le_bot mult_left_zero star.circ_plus_one star.circ_zero star_left_unfold_equal sup_inf_distrib1)
+ by (metis order.eq_iff le_bot mult_left_zero star.circ_plus_one star.circ_zero star_left_unfold_equal sup_inf_distrib1)
thus ?thesis
- by (simp add: inf.antisym star.circ_reflexive)
+ by (simp add: order.antisym star.circ_reflexive)
qed
lemma acyclic_asymmetric:
"acyclic w \<Longrightarrow> asymmetric w"
by (simp add: dual_order.trans pseudo_complement schroeder_5_p star.circ_increasing)
lemma forest_separate:
assumes "forest x"
shows "x\<^sup>\<star> * x\<^sup>T\<^sup>\<star> \<sqinter> x\<^sup>T * x \<le> 1"
proof -
have "x\<^sup>\<star> * 1 \<le> -x\<^sup>T"
using assms schroeder_5_p by force
hence 1: "x\<^sup>\<star> \<sqinter> x\<^sup>T = bot"
by (simp add: pseudo_complement)
have "x\<^sup>\<star> \<sqinter> x\<^sup>T * x = (1 \<squnion> x\<^sup>\<star> * x) \<sqinter> x\<^sup>T * x"
using star.circ_right_unfold_1 by simp
also have "... = (1 \<sqinter> x\<^sup>T * x) \<squnion> (x\<^sup>\<star> * x \<sqinter> x\<^sup>T * x)"
by (simp add: inf_sup_distrib2)
also have "... \<le> 1 \<squnion> (x\<^sup>\<star> * x \<sqinter> x\<^sup>T * x)"
using sup_left_isotone by simp
also have "... = 1 \<squnion> (x\<^sup>\<star> \<sqinter> x\<^sup>T) * x"
by (simp add: assms injective_comp_right_dist_inf)
also have "... = 1"
using 1 by simp
finally have 2: "x\<^sup>\<star> \<sqinter> x\<^sup>T * x \<le> 1"
.
hence 3: "x\<^sup>T\<^sup>\<star> \<sqinter> x\<^sup>T * x \<le> 1"
by (metis (mono_tags, lifting) conv_star_commute conv_dist_comp conv_dist_inf conv_involutive coreflexive_symmetric)
have "x\<^sup>\<star> * x\<^sup>T\<^sup>\<star> \<sqinter> x\<^sup>T * x \<le> (x\<^sup>\<star> \<squnion> x\<^sup>T\<^sup>\<star>) \<sqinter> x\<^sup>T * x"
using assms cancel_separate inf.sup_left_isotone by simp
also have "... \<le> 1"
using 2 3 by (simp add: inf_sup_distrib2)
finally show ?thesis
.
qed
text \<open>
The following definition captures the components of undirected weighted graphs.
\<close>
abbreviation "components g \<equiv> (--g)\<^sup>\<star>"
lemma components_equivalence:
"symmetric x \<Longrightarrow> equivalence (components x)"
by (simp add: conv_star_commute conv_complement star.circ_reflexive star.circ_transitive_equal)
lemma components_increasing:
"x \<le> components x"
using order_trans pp_increasing star.circ_increasing by blast
lemma components_isotone:
"x \<le> y \<Longrightarrow> components x \<le> components y"
by (simp add: pp_isotone star_isotone)
lemma cut_reachable:
assumes "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "t \<le> g"
shows "v * -v\<^sup>T \<sqinter> g \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>)"
proof -
have "v * -v\<^sup>T \<sqinter> g \<le> v * top \<sqinter> g"
using inf.sup_left_isotone mult_right_isotone top_greatest by blast
also have "... = (r\<^sup>T * t\<^sup>\<star>)\<^sup>T * top \<sqinter> g"
by (metis assms(1) conv_involutive)
also have "... \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * top \<sqinter> g"
using assms(2) conv_isotone inf.sup_left_isotone mult_left_isotone mult_right_isotone star_isotone by auto
also have "... \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * ((r\<^sup>T * g\<^sup>\<star>) * g)"
by (metis conv_involutive dedekind_1 inf_top.left_neutral)
also have "... \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>)"
by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus)
finally show ?thesis
.
qed
text \<open>
The following lemma shows that the predecessors of visited nodes in the minimum spanning tree extending the current tree have all been visited.
\<close>
lemma predecessors_reachable:
assumes "vector r"
and "injective r"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "forest w"
and "t \<le> w"
and "w \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g"
and "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
shows "w * v \<le> v"
proof -
have "w * r \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) * r"
using assms(6) mult_left_isotone by auto
also have "... \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * top"
by (simp add: mult_assoc mult_right_isotone)
also have "... = (r\<^sup>T * g\<^sup>\<star>)\<^sup>T"
by (simp add: assms(1) comp_associative conv_dist_comp)
also have "... \<le> (r\<^sup>T * w\<^sup>\<star>)\<^sup>T"
by (simp add: assms(7) conv_isotone)
also have "... = w\<^sup>T\<^sup>\<star> * r"
by (simp add: conv_dist_comp conv_star_commute)
also have "... \<le> -w * r"
using assms(4) by (simp add: mult_left_isotone acyclic_star_below_complement)
also have "... \<le> -(w * r)"
by (simp add: assms(2) comp_injective_below_complement)
finally have 1: "w * r = bot"
by (simp add: le_iff_inf)
have "v = t\<^sup>T\<^sup>\<star> * r"
by (metis assms(3) conv_dist_comp conv_involutive conv_star_commute)
also have "... = t\<^sup>T * v \<squnion> r"
by (simp add: calculation star.circ_loop_fixpoint)
also have "... \<le> w\<^sup>T * v \<squnion> r"
using assms(5) comp_isotone conv_isotone semiring.add_right_mono by auto
finally have "w * v \<le> w * w\<^sup>T * v \<squnion> w * r"
by (simp add: comp_left_dist_sup mult_assoc mult_right_isotone)
also have "... = w * w\<^sup>T * v"
using 1 by simp
also have "... \<le> v"
using assms(4) by (simp add: star_left_induct_mult_iff star_sub_one)
finally show ?thesis
.
qed
subsection \<open>Prim's Algorithm\<close>
text \<open>
The following results are used for proving the correctness of Prim's minimum spanning tree algorithm.
\<close>
subsubsection \<open>Preservation of Invariant\<close>
text \<open>
We first treat the preservation of the invariant.
The following lemma shows that the while-loop preserves that \<open>v\<close> represents the nodes of the constructed tree.
The remaining lemmas in this section show that \<open>t\<close> is a spanning tree.
The exchange property is treated in the following two sections.
\<close>
lemma reachable_inv:
assumes "vector v"
and "e \<le> v * -v\<^sup>T"
and "e * t = bot"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
shows "(v \<squnion> e\<^sup>T * top)\<^sup>T = r\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
proof -
have 1: "v\<^sup>T \<le> r\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
by (simp add: assms(4) mult_right_isotone star.circ_sub_dist)
have 2: "(e\<^sup>T * top)\<^sup>T = top * e"
by (simp add: conv_dist_comp)
also have "... = top * (v * -v\<^sup>T \<sqinter> e)"
by (simp add: assms(2) inf_absorb2)
also have "... \<le> top * (v * top \<sqinter> e)"
using inf.sup_left_isotone mult_right_isotone top_greatest by blast
also have "... = top * v\<^sup>T * e"
by (simp add: comp_inf_vector inf.sup_monoid.add_commute)
also have "... = v\<^sup>T * e"
using assms(1) vector_conv_covector by auto
also have "... \<le> r\<^sup>T * (t \<squnion> e)\<^sup>\<star> * e"
using 1 by (simp add: mult_left_isotone)
also have "... \<le> r\<^sup>T * (t \<squnion> e)\<^sup>\<star> * (t \<squnion> e)"
by (simp add: mult_right_isotone)
also have "... \<le> r\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
by (simp add: comp_associative mult_right_isotone star.right_plus_below_circ)
finally have 3: "(v \<squnion> e\<^sup>T * top)\<^sup>T \<le> r\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
using 1 by (simp add: conv_dist_sup)
have "r\<^sup>T \<le> r\<^sup>T * t\<^sup>\<star>"
using sup.bounded_iff star.circ_back_loop_prefixpoint by blast
also have "... \<le> (v \<squnion> e\<^sup>T * top)\<^sup>T"
by (metis assms(4) conv_isotone sup_ge1)
finally have 4: "r\<^sup>T \<le> (v \<squnion> e\<^sup>T * top)\<^sup>T"
.
have "(v \<squnion> e\<^sup>T * top)\<^sup>T * (t \<squnion> e) = (v \<squnion> e\<^sup>T * top)\<^sup>T * t \<squnion> (v \<squnion> e\<^sup>T * top)\<^sup>T * e"
by (simp add: mult_left_dist_sup)
also have "... \<le> (v \<squnion> e\<^sup>T * top)\<^sup>T * t \<squnion> top * e"
using comp_isotone semiring.add_left_mono by auto
also have "... = v\<^sup>T * t \<squnion> top * e * t \<squnion> top * e"
using 2 by (simp add: conv_dist_sup mult_right_dist_sup)
also have "... = v\<^sup>T * t \<squnion> top * e"
by (simp add: assms(3) comp_associative)
also have "... \<le> r\<^sup>T * t\<^sup>\<star> \<squnion> top * e"
by (metis assms(4) star.circ_back_loop_fixpoint sup_ge1 sup_left_isotone)
also have "... = v\<^sup>T \<squnion> top * e"
by (simp add: assms(4))
finally have 5: "(v \<squnion> e\<^sup>T * top)\<^sup>T * (t \<squnion> e) \<le> (v \<squnion> e\<^sup>T * top)\<^sup>T"
using 2 by (simp add: conv_dist_sup)
have "r\<^sup>T * (t \<squnion> e)\<^sup>\<star> \<le> (v \<squnion> e\<^sup>T * top)\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
using 4 by (simp add: mult_left_isotone)
also have "... \<le> (v \<squnion> e\<^sup>T * top)\<^sup>T"
using 5 by (simp add: star_right_induct_mult)
finally show ?thesis
- using 3 by (simp add: inf.eq_iff)
+ using 3 by (simp add: order.eq_iff)
qed
text \<open>
The next result is used to show that the while-loop preserves acyclicity of the constructed tree.
\<close>
lemma acyclic_inv:
assumes "acyclic t"
and "vector v"
and "e \<le> v * -v\<^sup>T"
and "t \<le> v * v\<^sup>T"
shows "acyclic (t \<squnion> e)"
proof -
have "t\<^sup>+ * e \<le> t\<^sup>+ * v * -v\<^sup>T"
by (simp add: assms(3) comp_associative mult_right_isotone)
also have "... \<le> v * v\<^sup>T * t\<^sup>\<star> * v * -v\<^sup>T"
by (simp add: assms(4) mult_left_isotone)
also have "... \<le> v * top * -v\<^sup>T"
by (metis mult_assoc mult_left_isotone mult_right_isotone top_greatest)
also have "... = v * -v\<^sup>T"
by (simp add: assms(2))
also have "... \<le> -1"
by (simp add: pp_increasing schroeder_3_p)
finally have 1: "t\<^sup>+ * e \<le> -1"
.
have 2: "e * t\<^sup>\<star> = e"
using assms(2-4) et(1) star_absorb by blast
have "e\<^sup>\<star> = 1 \<squnion> e \<squnion> e * e * e\<^sup>\<star>"
by (metis star.circ_loop_fixpoint star_square_2 sup_commute)
also have "... = 1 \<squnion> e"
using assms(2,3) ee comp_left_zero bot_least sup_absorb1 by simp
finally have 3: "e\<^sup>\<star> = 1 \<squnion> e"
.
have "e \<le> v * -v\<^sup>T"
by (simp add: assms(3))
also have "... \<le> -1"
by (simp add: pp_increasing schroeder_3_p)
finally have 4: "t\<^sup>+ * e \<squnion> e \<le> -1"
using 1 by simp
have "(t \<squnion> e)\<^sup>+ = (t \<squnion> e) * t\<^sup>\<star> * (e * t\<^sup>\<star>)\<^sup>\<star>"
using star_sup_1 mult_assoc by simp
also have "... = (t \<squnion> e) * t\<^sup>\<star> * (1 \<squnion> e)"
using 2 3 by simp
also have "... = t\<^sup>+ * (1 \<squnion> e) \<squnion> e * t\<^sup>\<star> * (1 \<squnion> e)"
by (simp add: comp_right_dist_sup)
also have "... = t\<^sup>+ * (1 \<squnion> e) \<squnion> e * (1 \<squnion> e)"
using 2 by simp
also have "... = t\<^sup>+ * (1 \<squnion> e) \<squnion> e"
using 3 by (metis star_absorb assms(2,3) ee)
also have "... = t\<^sup>+ \<squnion> t\<^sup>+ * e \<squnion> e"
by (simp add: mult_left_dist_sup)
also have "... \<le> -1"
using 4 by (metis assms(1) sup.absorb1 sup.orderI sup_assoc)
finally show ?thesis
.
qed
text \<open>
The following lemma shows that the extended tree is in the component reachable from the root.
\<close>
lemma mst_subgraph_inv_2:
assumes "regular (v * v\<^sup>T)"
and "t \<le> v * v\<^sup>T \<sqinter> --g"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "e \<le> v * -v\<^sup>T \<sqinter> --g"
and "vector v"
and "regular ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T)"
shows "t \<squnion> e \<le> (r\<^sup>T * (--((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g))\<^sup>\<star>)\<^sup>T * (r\<^sup>T * (--((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g))\<^sup>\<star>)"
proof -
let ?v = "v \<squnion> e\<^sup>T * top"
let ?G = "?v * ?v\<^sup>T \<sqinter> g"
let ?c = "r\<^sup>T * (--?G)\<^sup>\<star>"
have "v\<^sup>T \<le> r\<^sup>T * (--(v * v\<^sup>T \<sqinter> g))\<^sup>\<star>"
using assms(1-3) inf_pp_commute mult_right_isotone star_isotone by auto
also have "... \<le> ?c"
using comp_inf.mult_right_isotone comp_isotone conv_isotone inf.commute mult_right_isotone pp_isotone star_isotone sup.cobounded1 by presburger
finally have 2: "v\<^sup>T \<le> ?c \<and> v \<le> ?c\<^sup>T"
by (metis conv_isotone conv_involutive)
have "t \<le> v * v\<^sup>T"
using assms(2) by auto
hence 3: "t \<le> ?c\<^sup>T * ?c"
using 2 order_trans mult_isotone by blast
have "e \<le> v * top \<sqinter> --g"
by (metis assms(4,5) inf.bounded_iff inf.sup_left_divisibility mult_right_isotone top.extremum)
hence "e \<le> v * top \<sqinter> top * e \<sqinter> --g"
by (simp add: top_left_mult_increasing inf.boundedI)
hence "e \<le> v * top * e \<sqinter> --g"
by (metis comp_inf_covector inf.absorb2 mult_assoc top.extremum)
hence "t \<squnion> e \<le> (v * v\<^sup>T \<sqinter> --g) \<squnion> (v * top * e \<sqinter> --g)"
using assms(2) sup_mono by blast
also have "... = v * ?v\<^sup>T \<sqinter> --g"
by (simp add: inf_sup_distrib2 mult_assoc mult_left_dist_sup conv_dist_comp conv_dist_sup)
also have "... \<le> --?G"
using assms(6) comp_left_increasing_sup inf.sup_left_isotone pp_dist_inf by auto
finally have 4: "t \<squnion> e \<le> --?G"
.
have "e \<le> e * e\<^sup>T * e"
by (simp add: ex231c)
also have "... \<le> v * -v\<^sup>T * -v * v\<^sup>T * e"
by (metis assms(4) mult_left_isotone conv_isotone conv_dist_comp mult_assoc mult_isotone conv_involutive conv_complement inf.boundedE)
also have "... \<le> v * top * v\<^sup>T * e"
by (metis mult_assoc mult_left_isotone mult_right_isotone top.extremum)
also have "... = v * r\<^sup>T * t\<^sup>\<star> * e"
using assms(3,5) by (simp add: mult_assoc)
also have "... \<le> v * r\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
by (simp add: comp_associative mult_right_isotone star.circ_mult_upper_bound star.circ_sub_dist_1 star_isotone sup_commute)
also have "... \<le> v * ?c"
using 4 by (simp add: mult_assoc mult_right_isotone star_isotone)
also have "... \<le> ?c\<^sup>T * ?c"
using 2 by (simp add: mult_left_isotone)
finally show ?thesis
using 3 by simp
qed
lemma span_inv:
assumes "e \<le> v * -v\<^sup>T"
and "vector v"
and "arc e"
and "t \<le> (v * v\<^sup>T) \<sqinter> g"
and "g\<^sup>T = g"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "injective r"
and "r\<^sup>T \<le> v\<^sup>T"
and "r\<^sup>T * ((v * v\<^sup>T) \<sqinter> g)\<^sup>\<star> \<le> r\<^sup>T * t\<^sup>\<star>"
shows "r\<^sup>T * (((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T) \<sqinter> g)\<^sup>\<star> \<le> r\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
proof -
let ?d = "(v * v\<^sup>T) \<sqinter> g"
have 1: "(v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T = v * v\<^sup>T \<squnion> v * v\<^sup>T * e \<squnion> e\<^sup>T * v * v\<^sup>T \<squnion> e\<^sup>T * e"
using assms(1-3) ve_dist by simp
have "t\<^sup>T \<le> ?d\<^sup>T"
using assms(4) conv_isotone by simp
also have "... = (v * v\<^sup>T) \<sqinter> g\<^sup>T"
by (simp add: conv_dist_comp conv_dist_inf)
also have "... = ?d"
by (simp add: assms(5))
finally have 2: "t\<^sup>T \<le> ?d"
.
have "v * v\<^sup>T = (r\<^sup>T * t\<^sup>\<star>)\<^sup>T * (r\<^sup>T * t\<^sup>\<star>)"
by (metis assms(6) conv_involutive)
also have "... = t\<^sup>T\<^sup>\<star> * (r * r\<^sup>T) * t\<^sup>\<star>"
by (simp add: comp_associative conv_dist_comp conv_star_commute)
also have "... \<le> t\<^sup>T\<^sup>\<star> * 1 * t\<^sup>\<star>"
by (simp add: assms(7) mult_left_isotone star_right_induct_mult_iff star_sub_one)
also have "... = t\<^sup>T\<^sup>\<star> * t\<^sup>\<star>"
by simp
also have "... \<le> ?d\<^sup>\<star> * t\<^sup>\<star>"
using 2 by (simp add: comp_left_isotone star.circ_isotone)
also have "... \<le> ?d\<^sup>\<star> * ?d\<^sup>\<star>"
using assms(4) mult_right_isotone star_isotone by simp
also have 3: "... = ?d\<^sup>\<star>"
by (simp add: star.circ_transitive_equal)
finally have 4: "v * v\<^sup>T \<le> ?d\<^sup>\<star>"
.
have 5: "r\<^sup>T * ?d\<^sup>\<star> * (v * v\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star>"
by (simp add: comp_associative mult_right_isotone star.circ_plus_same star.left_plus_below_circ)
have "r\<^sup>T * ?d\<^sup>\<star> * (v * v\<^sup>T * e \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * v * v\<^sup>T * e"
by (simp add: comp_associative comp_right_isotone)
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> * e"
using 3 4 by (metis comp_associative comp_isotone eq_refl)
finally have 6: "r\<^sup>T * ?d\<^sup>\<star> * (v * v\<^sup>T * e \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e"
.
have 7: "\<forall>x . r\<^sup>T * (1 \<squnion> v * v\<^sup>T) * e\<^sup>T * x = bot"
proof
fix x
have "r\<^sup>T * (1 \<squnion> v * v\<^sup>T) * e\<^sup>T * x \<le> r\<^sup>T * (1 \<squnion> v * v\<^sup>T) * e\<^sup>T * top"
by (simp add: mult_right_isotone)
also have "... = r\<^sup>T * e\<^sup>T * top \<squnion> r\<^sup>T * v * v\<^sup>T * e\<^sup>T * top"
by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup)
also have "... = r\<^sup>T * e\<^sup>T * top"
by (metis assms(1,2) mult_assoc mult_right_dist_sup mult_right_zero sup_bot_right vTeT)
also have "... \<le> v\<^sup>T * e\<^sup>T * top"
by (simp add: assms(8) comp_isotone)
also have "... = bot"
using vTeT assms(1,2) by simp
finally show "r\<^sup>T * (1 \<squnion> v * v\<^sup>T) * e\<^sup>T * x = bot"
by (simp add: le_bot)
qed
have "r\<^sup>T * ?d\<^sup>\<star> * (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e\<^sup>T * v * v\<^sup>T"
by (simp add: comp_associative comp_right_isotone)
also have "... \<le> r\<^sup>T * (1 \<squnion> v * v\<^sup>T) * e\<^sup>T * v * v\<^sup>T"
by (metis assms(2) star.circ_isotone vector_vector_star inf_le1 comp_associative comp_right_isotone comp_left_isotone)
also have "... = bot"
using 7 by simp
finally have 8: "r\<^sup>T * ?d\<^sup>\<star> * (e\<^sup>T * v * v\<^sup>T \<sqinter> g) = bot"
by (simp add: le_bot)
have "r\<^sup>T * ?d\<^sup>\<star> * (e\<^sup>T * e \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e\<^sup>T * e"
by (simp add: comp_associative comp_right_isotone)
also have "... \<le> r\<^sup>T * (1 \<squnion> v * v\<^sup>T) * e\<^sup>T * e"
by (metis assms(2) star.circ_isotone vector_vector_star inf_le1 comp_associative comp_right_isotone comp_left_isotone)
also have "... = bot"
using 7 by simp
finally have 9: "r\<^sup>T * ?d\<^sup>\<star> * (e\<^sup>T * e \<sqinter> g) = bot"
by (simp add: le_bot)
have "r\<^sup>T * ?d\<^sup>\<star> * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) = r\<^sup>T * ?d\<^sup>\<star> * ((v * v\<^sup>T \<squnion> v * v\<^sup>T * e \<squnion> e\<^sup>T * v * v\<^sup>T \<squnion> e\<^sup>T * e) \<sqinter> g)"
using 1 by simp
also have "... = r\<^sup>T * ?d\<^sup>\<star> * ((v * v\<^sup>T \<sqinter> g) \<squnion> (v * v\<^sup>T * e \<sqinter> g) \<squnion> (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<squnion> (e\<^sup>T * e \<sqinter> g))"
by (simp add: inf_sup_distrib2)
also have "... = r\<^sup>T * ?d\<^sup>\<star> * (v * v\<^sup>T \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * (v * v\<^sup>T * e \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * (e\<^sup>T * e \<sqinter> g)"
by (simp add: comp_left_dist_sup)
also have "... = r\<^sup>T * ?d\<^sup>\<star> * (v * v\<^sup>T \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * (v * v\<^sup>T * e \<sqinter> g)"
using 8 9 by simp
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> \<squnion> r\<^sup>T * ?d\<^sup>\<star> * e"
using 5 6 sup.mono by simp
also have "... = r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
by (simp add: mult_left_dist_sup)
finally have 10: "r\<^sup>T * ?d\<^sup>\<star> * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
by simp
have "r\<^sup>T * ?d\<^sup>\<star> * e * (v * v\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e * v * v\<^sup>T"
by (simp add: comp_associative comp_right_isotone)
also have "... = bot"
by (metis assms(1,2) comp_associative comp_right_zero ev comp_left_zero)
finally have 11: "r\<^sup>T * ?d\<^sup>\<star> * e * (v * v\<^sup>T \<sqinter> g) = bot"
by (simp add: le_bot)
have "r\<^sup>T * ?d\<^sup>\<star> * e * (v * v\<^sup>T * e \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e * v * v\<^sup>T * e"
by (simp add: comp_associative comp_right_isotone)
also have "... = bot"
by (metis assms(1,2) comp_associative comp_right_zero ev comp_left_zero)
finally have 12: "r\<^sup>T * ?d\<^sup>\<star> * e * (v * v\<^sup>T * e \<sqinter> g) = bot"
by (simp add: le_bot)
have "r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e * e\<^sup>T * v * v\<^sup>T"
by (simp add: comp_associative comp_right_isotone)
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> * 1 * v * v\<^sup>T"
by (metis assms(3) arc_injective comp_associative comp_left_isotone comp_right_isotone)
also have "... = r\<^sup>T * ?d\<^sup>\<star> * v * v\<^sup>T"
by simp
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> * ?d\<^sup>\<star>"
using 4 by (simp add: mult_right_isotone mult_assoc)
also have "... = r\<^sup>T * ?d\<^sup>\<star>"
by (simp add: star.circ_transitive_equal comp_associative)
finally have 13: "r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star>"
.
have "r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * e \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e * e\<^sup>T * e"
by (simp add: comp_associative comp_right_isotone)
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> * 1 * e"
by (metis assms(3) arc_injective comp_associative comp_left_isotone comp_right_isotone)
also have "... = r\<^sup>T * ?d\<^sup>\<star> * e"
by simp
finally have 14: "r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * e \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * e"
.
have "r\<^sup>T * ?d\<^sup>\<star> * e * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) = r\<^sup>T * ?d\<^sup>\<star> * e * ((v * v\<^sup>T \<squnion> v * v\<^sup>T * e \<squnion> e\<^sup>T * v * v\<^sup>T \<squnion> e\<^sup>T * e) \<sqinter> g)"
using 1 by simp
also have "... = r\<^sup>T * ?d\<^sup>\<star> * e * ((v * v\<^sup>T \<sqinter> g) \<squnion> (v * v\<^sup>T * e \<sqinter> g) \<squnion> (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<squnion> (e\<^sup>T * e \<sqinter> g))"
by (simp add: inf_sup_distrib2)
also have "... = r\<^sup>T * ?d\<^sup>\<star> * e * (v * v\<^sup>T \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * e * (v * v\<^sup>T * e \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * e \<sqinter> g)"
by (simp add: comp_left_dist_sup)
also have "... = r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * v * v\<^sup>T \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * e * (e\<^sup>T * e \<sqinter> g)"
using 11 12 by simp
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> \<squnion> r\<^sup>T * ?d\<^sup>\<star> * e"
using 13 14 sup_mono by simp
also have "... = r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
by (simp add: mult_left_dist_sup)
finally have 15: "r\<^sup>T * ?d\<^sup>\<star> * e * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
by simp
have "r\<^sup>T \<le> r\<^sup>T * ?d\<^sup>\<star>"
using mult_right_isotone star.circ_reflexive by fastforce
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
by (simp add: semiring.distrib_left)
finally have 16: "r\<^sup>T \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
.
have "r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e) * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) = r\<^sup>T * ?d\<^sup>\<star> * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) \<squnion> r\<^sup>T * ?d\<^sup>\<star> * e * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g)"
by (simp add: semiring.distrib_left semiring.distrib_right)
also have "... \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
using 10 15 le_supI by simp
finally have "r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e) * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
.
hence "r\<^sup>T \<squnion> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e) * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g) \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
using 16 sup_least by simp
hence "r\<^sup>T * ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g)\<^sup>\<star> \<le> r\<^sup>T * ?d\<^sup>\<star> * (1 \<squnion> e)"
by (simp add: star_right_induct)
also have "... \<le> r\<^sup>T * t\<^sup>\<star> * (1 \<squnion> e)"
by (simp add: assms(9) mult_left_isotone)
also have "... \<le> r\<^sup>T * (t \<squnion> e)\<^sup>\<star>"
by (simp add: star_one_sup_below)
finally show ?thesis
.
qed
subsubsection \<open>Exchange gives Spanning Trees\<close>
text \<open>
The following abbreviations are used in the spanning tree application using Prim's algorithm to construct the new tree for the exchange property.
It is obtained by replacing an edge with one that has minimal weight and reversing the path connecting these edges.
Here, w represents a weighted graph, v represents a set of nodes and e represents an edge.
\<close>
abbreviation prim_E :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where "prim_E w v e \<equiv> w \<sqinter> --v * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
abbreviation prim_P :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where "prim_P w v e \<equiv> w \<sqinter> -v * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
abbreviation prim_EP :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where "prim_EP w v e \<equiv> w \<sqinter> -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
abbreviation prim_W :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where "prim_W w v e \<equiv> (w \<sqinter> -(prim_EP w v e)) \<squnion> (prim_P w v e)\<^sup>T \<squnion> e"
text \<open>
The lemmas in this section are used to show that the relation after exchange represents a spanning tree.
The results in the next section are used to show that it is a minimum spanning tree.
\<close>
lemma exchange_injective_3:
assumes "e \<le> v * -v\<^sup>T"
and "vector v"
shows "(w \<sqinter> -(prim_EP w v e)) * e\<^sup>T = bot"
proof -
have 1: "top * e \<le> -v\<^sup>T"
by (simp add: assms schroeder_4_p vTeT)
have "top * e \<le> top * e * w\<^sup>T\<^sup>\<star>"
using sup_right_divisibility star.circ_back_loop_fixpoint by blast
hence "top * e \<le> -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
using 1 by simp
hence "top * e \<le> -(w \<sqinter> -prim_EP w v e)"
by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff)
hence "(w \<sqinter> -(prim_EP w v e)) * e\<^sup>T \<le> bot"
using p_top schroeder_4_p by blast
thus ?thesis
using le_bot by simp
qed
lemma exchange_injective_6:
assumes "arc e"
and "forest w"
shows "(prim_P w v e)\<^sup>T * e\<^sup>T = bot"
proof -
have "e\<^sup>T * top * e \<le> --1"
by (simp add: assms(1) p_antitone p_antitone_iff point_injective)
hence 1: "e * -1 * e\<^sup>T \<le> bot"
by (metis conv_involutive p_top triple_schroeder_p)
have "(prim_P w v e)\<^sup>T * e\<^sup>T \<le> (w \<sqinter> top * e * w\<^sup>T\<^sup>\<star>)\<^sup>T * e\<^sup>T"
using comp_inf.mult_left_isotone conv_dist_inf mult_left_isotone by simp
also have "... = (w\<^sup>T \<sqinter> w\<^sup>T\<^sup>\<star>\<^sup>T * e\<^sup>T * top) * e\<^sup>T"
by (simp add: comp_associative conv_dist_comp conv_dist_inf)
also have "... = w\<^sup>\<star> * e\<^sup>T * top \<sqinter> w\<^sup>T * e\<^sup>T"
by (simp add: conv_star_commute inf_vector_comp)
also have "... \<le> (w\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top * e) * (e\<^sup>T \<sqinter> w\<^sup>+ * e\<^sup>T * top)"
by (metis dedekind mult_assoc conv_involutive inf_commute)
also have "... \<le> (w\<^sup>\<star> * e\<^sup>T * top * e) * (w\<^sup>+ * e\<^sup>T * top)"
by (simp add: mult_isotone)
also have "... \<le> (top * e) * (w\<^sup>+ * e\<^sup>T * top)"
by (simp add: mult_left_isotone)
also have "... = top * e * w\<^sup>+ * e\<^sup>T * top"
using mult_assoc by simp
also have "... \<le> top * e * -1 * e\<^sup>T * top"
using assms(2) mult_left_isotone mult_right_isotone by simp
also have "... \<le> bot"
using 1 by (metis le_bot semiring.mult_not_zero mult_assoc)
finally show ?thesis
using le_bot by simp
qed
text \<open>
The graph after exchanging is injective.
\<close>
lemma exchange_injective:
assumes "arc e"
and "e \<le> v * -v\<^sup>T"
and "forest w"
and "vector v"
shows "injective (prim_W w v e)"
proof -
have 1: "(w \<sqinter> -(prim_EP w v e)) * (w \<sqinter> -(prim_EP w v e))\<^sup>T \<le> 1"
proof -
have "(w \<sqinter> -(prim_EP w v e)) * (w \<sqinter> -(prim_EP w v e))\<^sup>T \<le> w * w\<^sup>T"
by (simp add: comp_isotone conv_isotone)
also have "... \<le> 1"
by (simp add: assms(3))
finally show ?thesis
.
qed
have 2: "(w \<sqinter> -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \<le> 1"
proof -
have "top * (prim_P w v e)\<^sup>T = top * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>T\<^sup>\<star>\<^sup>T * e\<^sup>T * top)"
by (simp add: comp_associative conv_complement conv_dist_comp conv_dist_inf)
also have "... = top * e * w\<^sup>T\<^sup>\<star> * (w\<^sup>T \<sqinter> -v * -v\<^sup>T)"
by (metis comp_inf_vector conv_dist_comp conv_involutive inf_top_left mult_assoc)
also have "... \<le> top * e * w\<^sup>T\<^sup>\<star> * (w\<^sup>T \<sqinter> top * -v\<^sup>T)"
using comp_inf.mult_right_isotone mult_left_isotone mult_right_isotone by simp
also have "... = top * e * w\<^sup>T\<^sup>\<star> * w\<^sup>T \<sqinter> -v\<^sup>T"
by (metis assms(4) comp_inf_covector vector_conv_compl)
also have "... \<le> -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
by (simp add: comp_associative comp_isotone inf.coboundedI1 star.circ_plus_same star.left_plus_below_circ)
finally have "top * (prim_P w v e)\<^sup>T \<le> -(w \<sqinter> -prim_EP w v e)"
by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff)
hence "(w \<sqinter> -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \<le> bot"
using p_top schroeder_4_p by blast
thus ?thesis
by (simp add: bot_unique)
qed
have 3: "(w \<sqinter> -(prim_EP w v e)) * e\<^sup>T \<le> 1"
by (metis assms(2,4) exchange_injective_3 bot_least)
have 4: "(prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>T \<le> 1"
using 2 conv_dist_comp coreflexive_symmetric by fastforce
have 5: "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \<le> 1"
proof -
have "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \<le> (top * e * w\<^sup>T\<^sup>\<star>)\<^sup>T * (top * e * w\<^sup>T\<^sup>\<star>)"
by (simp add: conv_dist_inf mult_isotone)
also have "... = w\<^sup>\<star> * e\<^sup>T * top * top * e * w\<^sup>T\<^sup>\<star>"
using conv_star_commute conv_dist_comp conv_involutive conv_top mult_assoc by presburger
also have "... = w\<^sup>\<star> * e\<^sup>T * top * e * w\<^sup>T\<^sup>\<star>"
by (simp add: comp_associative)
also have "... \<le> w\<^sup>\<star> * 1 * w\<^sup>T\<^sup>\<star>"
by (metis comp_left_isotone comp_right_isotone mult_assoc assms(1) point_injective)
finally have "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \<le> w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> w\<^sup>T * w"
by (simp add: conv_isotone inf.left_commute inf.sup_monoid.add_commute mult_isotone)
also have "... \<le> 1"
by (simp add: assms(3) forest_separate)
finally show ?thesis
.
qed
have 6: "(prim_P w v e)\<^sup>T * e\<^sup>T \<le> 1"
using assms exchange_injective_6 bot_least by simp
have 7: "e * (w \<sqinter> -(prim_EP w v e))\<^sup>T \<le> 1"
using 3 by (metis conv_dist_comp conv_involutive coreflexive_symmetric)
have 8: "e * (prim_P w v e)\<^sup>T\<^sup>T \<le> 1"
using 6 conv_dist_comp coreflexive_symmetric by fastforce
have 9: "e * e\<^sup>T \<le> 1"
by (simp add: assms(1) arc_injective)
have "(prim_W w v e) * (prim_W w v e)\<^sup>T = (w \<sqinter> -(prim_EP w v e)) * (w \<sqinter> -(prim_EP w v e))\<^sup>T \<squnion> (w \<sqinter> -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \<squnion> (w \<sqinter> -(prim_EP w v e)) * e\<^sup>T \<squnion> (prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>T \<squnion> (prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \<squnion> (prim_P w v e)\<^sup>T * e\<^sup>T \<squnion> e * (w \<sqinter> -(prim_EP w v e))\<^sup>T \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>T \<squnion> e * e\<^sup>T"
using comp_left_dist_sup comp_right_dist_sup conv_dist_sup sup.assoc by simp
also have "... \<le> 1"
using 1 2 3 4 5 6 7 8 9 by simp
finally show ?thesis
.
qed
lemma pv:
assumes "vector v"
shows "(prim_P w v e)\<^sup>T * v = bot"
proof -
have "(prim_P w v e)\<^sup>T * v \<le> (-v * -v\<^sup>T)\<^sup>T * v"
by (meson conv_isotone inf_le1 inf_le2 mult_left_isotone order_trans)
also have "... = -v * -v\<^sup>T * v"
by (simp add: conv_complement conv_dist_comp)
also have "... = bot"
by (simp add: assms covector_vector_comp mult_assoc)
finally show ?thesis
- by (simp add: antisym)
+ by (simp add: order.antisym)
qed
lemma vector_pred_inv:
assumes "arc e"
and "e \<le> v * -v\<^sup>T"
and "forest w"
and "vector v"
and "w * v \<le> v"
shows "(prim_W w v e) * (v \<squnion> e\<^sup>T * top) \<le> v \<squnion> e\<^sup>T * top"
proof -
have "(prim_W w v e) * e\<^sup>T * top = (w \<sqinter> -(prim_EP w v e)) * e\<^sup>T * top \<squnion> (prim_P w v e)\<^sup>T * e\<^sup>T * top \<squnion> e * e\<^sup>T * top"
by (simp add: mult_right_dist_sup)
also have "... = e * e\<^sup>T * top"
using assms exchange_injective_3 exchange_injective_6 comp_left_zero by simp
also have "... \<le> v * -v\<^sup>T * e\<^sup>T * top"
by (simp add: assms(2) comp_isotone)
also have "... \<le> v * top"
by (simp add: comp_associative mult_right_isotone)
also have "... = v"
by (simp add: assms(4))
finally have 1: "(prim_W w v e) * e\<^sup>T * top \<le> v"
.
have "(prim_W w v e) * v = (w \<sqinter> -(prim_EP w v e)) * v \<squnion> (prim_P w v e)\<^sup>T * v \<squnion> e * v"
by (simp add: mult_right_dist_sup)
also have "... = (w \<sqinter> -(prim_EP w v e)) * v"
by (metis assms(2,4) pv ev sup_bot_right)
also have "... \<le> w * v"
by (simp add: mult_left_isotone)
finally have 2: "(prim_W w v e) * v \<le> v"
using assms(5) order_trans by blast
have "(prim_W w v e) * (v \<squnion> e\<^sup>T * top) = (prim_W w v e) * v \<squnion> (prim_W w v e) * e\<^sup>T * top"
by (simp add: semiring.distrib_left mult_assoc)
also have "... \<le> v"
using 1 2 by simp
also have "... \<le> v \<squnion> e\<^sup>T * top"
by simp
finally show ?thesis
.
qed
text \<open>
The graph after exchanging is acyclic.
\<close>
lemma exchange_acyclic:
assumes "vector v"
and "e \<le> v * -v\<^sup>T"
and "w * v \<le> v"
and "acyclic w"
shows "acyclic (prim_W w v e)"
proof -
have 1: "(prim_P w v e)\<^sup>T * e = bot"
proof -
have "(prim_P w v e)\<^sup>T * e \<le> (-v * -v\<^sup>T)\<^sup>T * e"
by (meson conv_order dual_order.trans inf.cobounded1 inf.cobounded2 mult_left_isotone)
also have "... = -v * -v\<^sup>T * e"
by (simp add: conv_complement conv_dist_comp)
also have "... \<le> -v * -v\<^sup>T * v * -v\<^sup>T"
by (simp add: assms(2) comp_associative mult_right_isotone)
also have "... = bot"
by (simp add: assms(1) covector_vector_comp mult_assoc)
finally show ?thesis
by (simp add: bot_unique)
qed
have 2: "e * e = bot"
using assms(1,2) ee by auto
have 3: "(w \<sqinter> -(prim_EP w v e)) * (prim_P w v e)\<^sup>T = bot"
proof -
have "top * (prim_P w v e) \<le> top * (-v * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>)"
using comp_inf.mult_semi_associative mult_right_isotone by auto
also have "... \<le> top * -v * -v\<^sup>T \<sqinter> top * top * e * w\<^sup>T\<^sup>\<star>"
by (simp add: comp_inf_covector mult_assoc)
also have "... \<le> top * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
using mult_left_isotone top.extremum inf_mono by presburger
also have "... = -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
by (simp add: assms(1) vector_conv_compl)
finally have "top * (prim_P w v e) \<le> -(w \<sqinter> -prim_EP w v e)"
by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff)
hence "(w \<sqinter> -(prim_EP w v e)) * (prim_P w v e)\<^sup>T \<le> bot"
using p_top schroeder_4_p by blast
thus ?thesis
using bot_unique by blast
qed
hence 4: "(w \<sqinter> -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>\<star> = w \<sqinter> -(prim_EP w v e)"
using star_absorb by blast
hence 5: "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\<star> = (w \<sqinter> -(prim_EP w v e))\<^sup>+"
by (metis star_plus mult_assoc)
hence 6: "(w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> = (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>\<star>"
by (metis star.circ_loop_fixpoint mult_assoc)
have 7: "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * e \<le> v * top"
proof -
have "e \<le> v * top"
using assms(2) dual_order.trans mult_right_isotone top_greatest by blast
hence 8: "e \<squnion> w * v * top \<le> v * top"
by (simp add: assms(1,3) comp_associative)
have "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * e \<le> w\<^sup>+ * e"
by (simp add: comp_isotone star_isotone)
also have "... \<le> w\<^sup>\<star> * e"
by (simp add: mult_left_isotone star.left_plus_below_circ)
also have "... \<le> v * top"
using 8 by (simp add: comp_associative star_left_induct)
finally show ?thesis
.
qed
have 9: "(prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * e = bot"
proof -
have "(prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * e \<le> (prim_P w v e)\<^sup>T * v * top"
using 7 by (simp add: mult_assoc mult_right_isotone)
also have "... = bot"
by (simp add: assms(1) pv)
finally show ?thesis
using bot_unique by blast
qed
have 10: "e * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * e = bot"
proof -
have "e * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * e \<le> e * v * top"
using 7 by (simp add: mult_assoc mult_right_isotone)
also have "... \<le> v * -v\<^sup>T * v * top"
by (simp add: assms(2) mult_left_isotone)
also have "... = bot"
by (simp add: assms(1) covector_vector_comp mult_assoc)
finally show ?thesis
using bot_unique by blast
qed
have 11: "e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> v * -v\<^sup>T"
proof -
have 12: "-v\<^sup>T * w \<le> -v\<^sup>T"
by (metis assms(3) conv_complement order_lesseq_imp pp_increasing schroeder_6_p)
have "v * -v\<^sup>T * (w \<sqinter> -(prim_EP w v e)) \<le> v * -v\<^sup>T * w"
by (simp add: comp_isotone star_isotone)
also have "... \<le> v * -v\<^sup>T"
using 12 by (simp add: comp_isotone comp_associative)
finally have 13: "v * -v\<^sup>T * (w \<sqinter> -(prim_EP w v e)) \<le> v * -v\<^sup>T"
.
have 14: "(prim_P w v e)\<^sup>T \<le> -v * -v\<^sup>T"
by (metis conv_complement conv_dist_comp conv_involutive conv_order inf_le1 inf_le2 order_trans)
have "e * (prim_P w v e)\<^sup>T\<^sup>\<star> \<le> v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\<star>"
by (simp add: assms(2) mult_left_isotone)
also have "... = v * -v\<^sup>T \<squnion> v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>+"
by (metis mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute)
also have "... = v * -v\<^sup>T \<squnion> v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\<star> * (prim_P w v e)\<^sup>T"
by (simp add: mult_assoc star_plus)
also have "... \<le> v * -v\<^sup>T \<squnion> v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\<star> * -v * -v\<^sup>T"
using 14 mult_assoc mult_right_isotone sup_right_isotone by simp
also have "... \<le> v * -v\<^sup>T \<squnion> v * top * -v\<^sup>T"
by (metis top_greatest mult_right_isotone mult_left_isotone mult_assoc sup_right_isotone)
also have "... = v * -v\<^sup>T"
by (simp add: assms(1))
finally have "e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> v * -v\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (simp add: mult_left_isotone)
also have "... \<le> v * -v\<^sup>T"
using 13 by (simp add: star_right_induct_mult)
finally show ?thesis
.
qed
have 15: "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
proof -
have "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> = (w \<sqinter> -(prim_EP w v e))\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
using 5 by simp
also have "... = (w \<sqinter> -(prim_EP w v e))\<^sup>+"
by (simp add: mult_assoc star.circ_transitive_equal)
also have "... \<le> w\<^sup>+"
by (simp add: comp_isotone star_isotone)
finally show ?thesis
using assms(4) by simp
qed
have 16: "(prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
proof -
have "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>+ \<le> (w \<sqinter> -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\<star>"
by (simp add: mult_right_isotone star.left_plus_below_circ)
also have "... = (w \<sqinter> -(prim_EP w v e))\<^sup>+"
using 5 by simp
also have "... \<le> w\<^sup>+"
by (simp add: comp_isotone star_isotone)
finally have "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>+ \<le> -1"
using assms(4) by simp
hence 17: "(prim_P w v e)\<^sup>T\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<le> -1"
by (simp add: comp_commute_below_diversity)
have "(prim_P w v e)\<^sup>T\<^sup>+ \<le> w\<^sup>T\<^sup>+"
by (simp add: comp_isotone conv_dist_inf inf.left_commute inf.sup_monoid.add_commute star_isotone)
also have "... = w\<^sup>+\<^sup>T"
by (simp add: conv_dist_comp conv_star_commute star_plus)
also have "... \<le> -1"
using assms(4) conv_complement conv_isotone by force
finally have 18: "(prim_P w v e)\<^sup>T\<^sup>+ \<le> -1"
.
have "(prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> = (prim_P w v e)\<^sup>T * ((w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>\<star>) * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
using 6 by (simp add: comp_associative)
also have "... = (prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<squnion> (prim_P w v e)\<^sup>T\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (simp add: mult_left_dist_sup mult_right_dist_sup)
also have "... = (prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (simp add: mult_assoc star.circ_transitive_equal)
also have "... = (prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>+ * (1 \<squnion> (w \<sqinter> -(prim_EP w v e))\<^sup>+)"
using star_left_unfold_equal by simp
also have "... = (prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>+"
by (simp add: mult_left_dist_sup sup.left_commute sup_commute)
also have "... = ((prim_P w v e)\<^sup>T \<squnion> (prim_P w v e)\<^sup>T\<^sup>+) * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>+"
by (simp add: mult_right_dist_sup)
also have "... = (prim_P w v e)\<^sup>T\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>+"
using star.circ_mult_increasing by (simp add: le_iff_sup)
also have "... \<le> -1"
using 17 18 by simp
finally show ?thesis
.
qed
have 19: "e * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
proof -
have "e * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> = e * ((w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> (prim_P w v e)\<^sup>T\<^sup>\<star>) * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
using 6 by (simp add: mult_assoc)
also have "... = e * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (simp add: mult_left_dist_sup mult_right_dist_sup)
also have "... = e * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (simp add: mult_assoc star.circ_transitive_equal)
also have "... \<le> e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>+ \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (metis mult_right_sub_dist_sup_right semiring.add_right_mono star.circ_back_loop_fixpoint)
also have "... \<le> e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
using mult_right_isotone star.left_plus_below_circ by auto
also have "... \<le> v * -v\<^sup>T"
using 11 by simp
also have "... \<le> -1"
by (simp add: pp_increasing schroeder_3_p)
finally show ?thesis
.
qed
have 20: "(prim_W w v e) * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
using 15 16 19 by (simp add: comp_right_dist_sup)
have 21: "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
proof -
have "(w \<sqinter> -(prim_EP w v e)) * v * -v\<^sup>T \<le> w * v * -v\<^sup>T"
by (simp add: comp_isotone star_isotone)
also have "... \<le> v * -v\<^sup>T"
by (simp add: assms(3) mult_left_isotone)
finally have 22: "(w \<sqinter> -(prim_EP w v e)) * v * -v\<^sup>T \<le> v * -v\<^sup>T"
.
have "(w \<sqinter> -(prim_EP w v e))\<^sup>+ * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> (w \<sqinter> -(prim_EP w v e))\<^sup>+ * v * -v\<^sup>T"
using 11 by (simp add: mult_right_isotone mult_assoc)
also have "... \<le> (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * v * -v\<^sup>T"
using mult_left_isotone star.left_plus_below_circ by blast
also have "... \<le> v * -v\<^sup>T"
using 22 by (simp add: star_left_induct_mult mult_assoc)
also have "... \<le> -1"
by (simp add: pp_increasing schroeder_3_p)
finally show ?thesis
.
qed
have 23: "(prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
proof -
have "(prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e = (prim_P w v e)\<^sup>T * e \<squnion> (prim_P w v e)\<^sup>T * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * e"
using comp_left_dist_sup mult_assoc star.circ_loop_fixpoint sup_commute by auto
also have "... = bot"
using 1 9 by simp
finally show ?thesis
by simp
qed
have 24: "e * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
proof -
have "e * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e = e * e \<squnion> e * (w \<sqinter> -(prim_EP w v e))\<^sup>+ * e"
using comp_left_dist_sup mult_assoc star.circ_loop_fixpoint sup_commute by auto
also have "... = bot"
using 2 10 by simp
finally show ?thesis
by simp
qed
have 25: "(prim_W w v e) * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<le> -1"
using 21 23 24 by (simp add: comp_right_dist_sup)
have "(prim_W w v e)\<^sup>\<star> = ((prim_P w v e)\<^sup>T \<squnion> e)\<^sup>\<star> * ((w \<sqinter> -(prim_EP w v e)) * ((prim_P w v e)\<^sup>T \<squnion> e)\<^sup>\<star>)\<^sup>\<star>"
by (metis star_sup_1 sup.left_commute sup_commute)
also have "... = ((prim_P w v e)\<^sup>T\<^sup>\<star> \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star>) * ((w \<sqinter> -(prim_EP w v e)) * ((prim_P w v e)\<^sup>T\<^sup>\<star> \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star>))\<^sup>\<star>"
using 1 2 star_separate by auto
also have "... = ((prim_P w v e)\<^sup>T\<^sup>\<star> \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star>) * ((w \<sqinter> -(prim_EP w v e)) * (1 \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star>))\<^sup>\<star>"
using 4 mult_left_dist_sup by auto
also have "... = (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * ((prim_P w v e)\<^sup>T\<^sup>\<star> \<squnion> e * (prim_P w v e)\<^sup>T\<^sup>\<star>) * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
using 3 9 10 star_separate_2 by blast
also have "... = (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<squnion> (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (simp add: semiring.distrib_left semiring.distrib_right mult_assoc)
finally have "(prim_W w v e)\<^sup>+ = (prim_W w v e) * ((w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<squnion> (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>)"
by simp
also have "... = (prim_W w v e) * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> \<squnion> (prim_W w v e) * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e))\<^sup>\<star>"
by (simp add: comp_left_dist_sup comp_associative)
also have "... \<le> -1"
using 20 25 by simp
finally show ?thesis
.
qed
text \<open>
The following lemma shows that an edge across the cut between visited nodes and unvisited nodes does not leave the component of visited nodes.
\<close>
lemma mst_subgraph_inv:
assumes "e \<le> v * -v\<^sup>T \<sqinter> g"
and "t \<le> g"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
shows "e \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g"
proof -
have "e \<le> v * -v\<^sup>T \<sqinter> g"
by (rule assms(1))
also have "... \<le> v * (-v\<^sup>T \<sqinter> v\<^sup>T * g) \<sqinter> g"
by (simp add: dedekind_1)
also have "... \<le> v * v\<^sup>T * g \<sqinter> g"
by (simp add: comp_associative comp_right_isotone inf_commute le_infI2)
also have "... = v * (r\<^sup>T * t\<^sup>\<star>) * g \<sqinter> g"
by (simp add: assms(3))
also have "... = (r\<^sup>T * t\<^sup>\<star>)\<^sup>T * (r\<^sup>T * t\<^sup>\<star>) * g \<sqinter> g"
by (metis assms(3) conv_involutive)
also have "... \<le> (r\<^sup>T * t\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) * g \<sqinter> g"
using assms(2) comp_inf.mult_left_isotone comp_isotone star_isotone by auto
also have "... \<le> (r\<^sup>T * t\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g"
using inf.sup_right_isotone inf_commute mult_assoc mult_right_isotone star.left_plus_below_circ star_plus by presburger
also have "... \<le> (r\<^sup>T * g\<^sup>\<star>)\<^sup>T * (r\<^sup>T * g\<^sup>\<star>) \<sqinter> g"
using assms(2) comp_inf.mult_left_isotone conv_dist_comp conv_isotone mult_left_isotone star_isotone by auto
finally show ?thesis
.
qed
text \<open>
The following lemmas show that the tree after exchanging contains the currently constructed and tree and its extension by the chosen edge.
\<close>
lemma mst_extends_old_tree:
assumes "t \<le> w"
and "t \<le> v * v\<^sup>T"
and "vector v"
shows "t \<le> prim_W w v e"
proof -
have "t \<sqinter> prim_EP w v e \<le> t \<sqinter> -v\<^sup>T"
by (simp add: inf.coboundedI2 inf.sup_monoid.add_assoc)
also have "... \<le> v * v\<^sup>T \<sqinter> -v\<^sup>T"
by (simp add: assms(2) inf.coboundedI1)
also have "... \<le> bot"
by (simp add: assms(3) covector_vector_comp eq_refl schroeder_2)
finally have "t \<le> -(prim_EP w v e)"
using le_bot pseudo_complement by blast
hence "t \<le> w \<sqinter> -(prim_EP w v e)"
using assms(1) by simp
thus ?thesis
using le_supI1 by blast
qed
lemma mst_extends_new_tree:
"t \<le> w \<Longrightarrow> t \<le> v * v\<^sup>T \<Longrightarrow> vector v \<Longrightarrow> t \<squnion> e \<le> prim_W w v e"
using mst_extends_old_tree by auto
text \<open>Lemmas \<open>forests_bot_1\<close>, \<open>forests_bot_2\<close>, \<open>forests_bot_3\<close> and \<open>fc_comp_eq_fc\<close> were contributed by Nicolas Robinson-O'Brien.\<close>
lemma forests_bot_1:
assumes "equivalence e"
and "forest f"
shows "(-e \<sqinter> f) * (e \<sqinter> f)\<^sup>T = bot"
proof -
have "f * f\<^sup>T \<le> e"
using assms dual_order.trans by blast
hence "f * (e \<sqinter> f)\<^sup>T \<le> e"
by (metis conv_dist_inf inf.boundedE inf.cobounded2 inf.orderE mult_right_isotone)
hence "-e \<sqinter> f * (e \<sqinter> f)\<^sup>T = bot"
by (simp add: p_antitone pseudo_complement)
thus ?thesis
by (metis assms(1) comp_isotone conv_dist_inf equivalence_comp_right_complement inf.boundedI inf.cobounded1 inf.cobounded2 le_bot)
qed
lemma forests_bot_2:
assumes "equivalence e"
and "forest f"
shows "(-e \<sqinter> f\<^sup>T) * x \<sqinter> (e \<sqinter> f\<^sup>T) * y = bot"
proof -
have "(-e \<sqinter> f) * (e \<sqinter> f\<^sup>T) = bot"
using assms forests_bot_1 conv_dist_inf by simp
thus ?thesis
by (smt assms(1) comp_associative comp_inf.semiring.mult_not_zero conv_complement conv_dist_comp conv_dist_inf conv_involutive dedekind_1 inf.cobounded2 inf.sup_monoid.add_commute le_bot mult_right_zero p_antitone_iff pseudo_complement semiring.mult_not_zero symmetric_top_closed top.extremum)
qed
lemma forests_bot_3:
assumes "equivalence e"
and "forest f"
shows "x * (-e \<sqinter> f) \<sqinter> y * (e \<sqinter> f) = bot"
proof -
have "(e \<sqinter> f) * (-e \<sqinter> f\<^sup>T) = bot"
using assms forests_bot_1 conv_dist_inf conv_complement by (smt conv_dist_comp conv_involutive conv_order coreflexive_bot_closed coreflexive_symmetric)
hence "y * (e \<sqinter> f) * (-e \<sqinter> f\<^sup>T) = bot"
by (simp add: comp_associative)
hence 1: "x \<sqinter> y * (e \<sqinter> f) * (-e \<sqinter> f\<^sup>T) = bot"
using comp_inf.semiring.mult_not_zero by blast
hence "(x \<sqinter> y * (e \<sqinter> f) * (-e \<sqinter> f\<^sup>T)) * (-e \<sqinter> f) = bot"
using semiring.mult_not_zero by blast
hence "x * (-e \<sqinter> f\<^sup>T)\<^sup>T \<sqinter> y * (e \<sqinter> f) = bot"
using 1 dedekind_2 inf_commute schroeder_2 by auto
thus ?thesis
by (simp add: assms(1) conv_complement conv_dist_inf)
qed
end
text \<open>
We finally add the Kleene star to Stone relation algebras.
Kleene star and the relational operations are reasonably independent.
The only additional axiom we need in the generalisation to Stone-Kleene relation algebras is that star distributes over double complement.
\<close>
class stone_kleene_relation_algebra = stone_relation_algebra + pd_kleene_allegory +
assumes pp_dist_star: "--(x\<^sup>\<star>) = (--x)\<^sup>\<star>"
begin
lemma reachable_without_loops:
"x\<^sup>\<star> = (x \<sqinter> -1)\<^sup>\<star>"
-proof (rule antisym)
+proof (rule order.antisym)
have "x * (x \<sqinter> -1)\<^sup>\<star> = (x \<sqinter> 1) * (x \<sqinter> -1)\<^sup>\<star> \<squnion> (x \<sqinter> -1) * (x \<sqinter> -1)\<^sup>\<star>"
by (metis maddux_3_11_pp mult_right_dist_sup regular_one_closed)
also have "... \<le> (x \<sqinter> -1)\<^sup>\<star>"
by (metis inf.cobounded2 le_supI mult_left_isotone star.circ_circ_mult star.left_plus_below_circ star_involutive star_one)
finally show "x\<^sup>\<star> \<le> (x \<sqinter> -1)\<^sup>\<star>"
by (metis inf.cobounded2 maddux_3_11_pp regular_one_closed star.circ_circ_mult star.circ_sup_2 star_involutive star_sub_one)
next
show "(x \<sqinter> -1)\<^sup>\<star> \<le> x\<^sup>\<star>"
by (simp add: star_isotone)
qed
lemma plus_reachable_without_loops:
"x\<^sup>+ = (x \<sqinter> -1)\<^sup>+ \<squnion> (x \<sqinter> 1)"
by (metis comp_associative maddux_3_11_pp regular_one_closed star.circ_back_loop_fixpoint star.circ_loop_fixpoint sup_assoc reachable_without_loops)
lemma star_plus_without_loops:
"x\<^sup>\<star> \<sqinter> -1 = x\<^sup>+ \<sqinter> -1"
by (metis maddux_3_13 star_left_unfold_equal)
lemma regular_closed_star:
"regular x \<Longrightarrow> regular (x\<^sup>\<star>)"
by (simp add: pp_dist_star)
lemma components_idempotent:
"components (components x) = components x"
using pp_dist_star star_involutive by auto
lemma fc_comp_eq_fc:
"-forest_components (--f) = -forest_components f"
by (metis conv_complement p_comp_pp p_pp_comp pp_dist_star)
text \<open>
The following lemma shows that the nodes reachable in the tree after exchange contain the nodes reachable in the tree before exchange.
\<close>
lemma mst_reachable_inv:
assumes "regular (prim_EP w v e)"
and "vector r"
and "e \<le> v * -v\<^sup>T"
and "vector v"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "t \<le> w"
and "t \<le> v * v\<^sup>T"
and "w * v \<le> v"
shows "r\<^sup>T * w\<^sup>\<star> \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star>"
proof -
have 1: "r\<^sup>T \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star>"
using sup.bounded_iff star.circ_back_loop_prefixpoint by blast
have "top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> * w\<^sup>T \<sqinter> -v\<^sup>T = top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> * (w\<^sup>T \<sqinter> -v\<^sup>T)"
by (simp add: assms(4) covector_comp_inf vector_conv_compl)
also have "... \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star>"
by (simp add: comp_isotone mult_assoc star.circ_plus_same star.left_plus_below_circ)
finally have 2: "top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> * w\<^sup>T \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T"
by (simp add: shunting_var_p)
have 3: "--v\<^sup>T * w\<^sup>T \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T"
by (metis assms(8) conv_dist_comp conv_order mult_assoc order.trans pp_comp_semi_commute pp_isotone sup.coboundedI1 sup_commute)
have 4: "top * e \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T"
using sup_right_divisibility star.circ_back_loop_fixpoint le_supI1 by blast
have "(top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T) * w\<^sup>T = top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> * w\<^sup>T \<squnion> --v\<^sup>T * w\<^sup>T"
by (simp add: comp_right_dist_sup)
also have "... \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T"
using 2 3 by simp
finally have "top * e \<squnion> (top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T) * w\<^sup>T \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T"
using 4 by simp
hence 5: "top * e * w\<^sup>T\<^sup>\<star> \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star> \<squnion> --v\<^sup>T"
by (simp add: star_right_induct)
have 6: "top * e \<le> top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>"
using sup_right_divisibility star.circ_back_loop_fixpoint by blast
have "(top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>)\<^sup>T \<le> (top * e * w\<^sup>T\<^sup>\<star>)\<^sup>T"
by (simp add: star_isotone mult_right_isotone conv_isotone inf_assoc)
also have "... = w\<^sup>\<star> * e\<^sup>T * top"
by (simp add: conv_dist_comp conv_star_commute mult_assoc)
finally have 7: "(top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>)\<^sup>T \<le> w\<^sup>\<star> * e\<^sup>T * top"
.
have "(top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>)\<^sup>T \<le> (top * e * (-v * -v\<^sup>T)\<^sup>\<star>)\<^sup>T"
by (simp add: conv_isotone inf_commute mult_right_isotone star_isotone le_infI2)
also have "... \<le> (top * v * -v\<^sup>T * (-v * -v\<^sup>T)\<^sup>\<star>)\<^sup>T"
by (metis assms(3) conv_isotone mult_left_isotone mult_right_isotone mult_assoc)
also have "... = (top * v * (-v\<^sup>T * -v)\<^sup>\<star> * -v\<^sup>T)\<^sup>T"
by (simp add: mult_assoc star_slide)
also have "... \<le> (top * -v\<^sup>T)\<^sup>T"
using conv_order mult_left_isotone by auto
also have "... = -v"
by (simp add: assms(4) conv_complement vector_conv_compl)
finally have 8: "(top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>)\<^sup>T \<le> w\<^sup>\<star> * e\<^sup>T * top \<sqinter> -v"
using 7 by simp
have "covector (top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>)"
by (simp add: covector_mult_closed)
hence "top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star> * (w\<^sup>T \<sqinter> -v\<^sup>T) = top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star> * (w\<^sup>T \<sqinter> -v\<^sup>T \<sqinter> (top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>)\<^sup>T)"
by (metis comp_inf_vector_1 inf.idem)
also have "... \<le> top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star> * (w\<^sup>T \<sqinter> -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top \<sqinter> -v)"
using 8 mult_right_isotone inf.sup_right_isotone inf_assoc by simp
also have "... = top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star> * (w\<^sup>T \<sqinter> (-v \<sqinter> -v\<^sup>T) \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)"
using inf_assoc inf_commute by (simp add: inf_assoc)
also have "... = top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star> * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)"
using assms(4) conv_complement vector_complement_closed vector_covector by fastforce
also have "... \<le> top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>"
by (simp add: comp_associative comp_isotone star.circ_plus_same star.left_plus_below_circ)
finally have 9: "top * e \<squnion> top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star> * (w\<^sup>T \<sqinter> -v\<^sup>T) \<le> top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>"
using 6 by simp
have "prim_EP w v e \<le> -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
using inf.sup_left_isotone by auto
also have "... \<le> top * e * (w\<^sup>T \<sqinter> -v\<^sup>T)\<^sup>\<star>"
using 5 by (metis inf_commute shunting_var_p)
also have "... \<le> top * e * (w\<^sup>T \<sqinter> -v * -v\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>\<star>"
using 9 by (simp add: star_right_induct)
finally have 10: "prim_EP w v e \<le> top * e * (prim_P w v e)\<^sup>T\<^sup>\<star>"
by (simp add: conv_complement conv_dist_comp conv_dist_inf conv_star_commute mult_assoc)
have "top * e = top * (v * -v\<^sup>T \<sqinter> e)"
by (simp add: assms(3) inf.absorb2)
also have "... \<le> top * (v * top \<sqinter> e)"
using inf.sup_right_isotone inf_commute mult_right_isotone top_greatest by presburger
also have "... = (top \<sqinter> (v * top)\<^sup>T) * e"
using assms(4) covector_inf_comp_3 by presburger
also have "... = top * v\<^sup>T * e"
by (simp add: conv_dist_comp)
also have "... = top * r\<^sup>T * t\<^sup>\<star> * e"
by (simp add: assms(5) comp_associative)
also have "... \<le> top * r\<^sup>T * (prim_W w v e)\<^sup>\<star> * e"
by (metis assms(4,6,7) mst_extends_old_tree star_isotone mult_left_isotone mult_right_isotone)
finally have 11: "top * e \<le> top * r\<^sup>T * (prim_W w v e)\<^sup>\<star> * e"
.
have "r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_EP w v e) \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (top * e * (prim_P w v e)\<^sup>T\<^sup>\<star>)"
using 10 mult_right_isotone by blast
also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\<star> * top * e * (prim_P w v e)\<^sup>T\<^sup>\<star>"
by (simp add: mult_assoc)
also have "... \<le> top * e * (prim_P w v e)\<^sup>T\<^sup>\<star>"
by (metis comp_associative comp_inf_covector inf.idem inf.sup_right_divisibility)
also have "... \<le> top * r\<^sup>T * (prim_W w v e)\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star>"
using 11 by (simp add: mult_left_isotone)
also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\<star> * e * (prim_P w v e)\<^sup>T\<^sup>\<star>"
using assms(2) vector_conv_covector by auto
also have "... \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_W w v e) * (prim_P w v e)\<^sup>T\<^sup>\<star>"
by (simp add: mult_left_isotone mult_right_isotone)
also have "... \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_W w v e) * (prim_W w v e)\<^sup>\<star>"
by (meson dual_order.trans mult_right_isotone star_isotone sup_ge1 sup_ge2)
also have "... \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star>"
by (metis mult_assoc mult_right_isotone star.circ_transitive_equal star.left_plus_below_circ)
finally have 12: "r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_EP w v e) \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star>"
.
have "r\<^sup>T * (prim_W w v e)\<^sup>\<star> * w \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (w \<squnion> prim_EP w v e)"
by (simp add: inf_assoc)
also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\<star> * ((w \<squnion> prim_EP w v e) \<sqinter> (-(prim_EP w v e) \<squnion> prim_EP w v e))"
by (metis assms(1) inf_top_right stone)
also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\<star> * ((w \<sqinter> -(prim_EP w v e)) \<squnion> prim_EP w v e)"
by (simp add: sup_inf_distrib2)
also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (w \<sqinter> -(prim_EP w v e)) \<squnion> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_EP w v e)"
by (simp add: comp_left_dist_sup)
also have "... \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_W w v e) \<squnion> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_EP w v e)"
using mult_right_isotone sup_left_isotone by auto
also have "... \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star> \<squnion> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * (prim_EP w v e)"
using mult_assoc mult_right_isotone star.circ_plus_same star.left_plus_below_circ sup_left_isotone by auto
also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\<star>"
using 12 sup.absorb1 by blast
finally have "r\<^sup>T \<squnion> r\<^sup>T * (prim_W w v e)\<^sup>\<star> * w \<le> r\<^sup>T * (prim_W w v e)\<^sup>\<star>"
using 1 by simp
thus ?thesis
by (simp add: star_right_induct)
qed
text \<open>
Some of the following lemmas already hold in pseudocomplemented distributive Kleene allegories.
\<close>
subsubsection \<open>Exchange gives Minimum Spanning Trees\<close>
text \<open>
The lemmas in this section are used to show that the after exchange we obtain a minimum spanning tree.
The following lemmas show various interactions between the three constituents of the tree after exchange.
\<close>
lemma epm_1:
"vector v \<Longrightarrow> prim_E w v e \<squnion> prim_P w v e = prim_EP w v e"
by (metis inf_commute inf_sup_distrib1 mult_assoc mult_right_dist_sup regular_closed_p regular_complement_top vector_conv_compl)
lemma epm_2:
assumes "regular (prim_EP w v e)"
and "vector v"
shows "(w \<sqinter> -(prim_EP w v e)) \<squnion> prim_P w v e \<squnion> prim_E w v e = w"
proof -
have "(w \<sqinter> -(prim_EP w v e)) \<squnion> prim_P w v e \<squnion> prim_E w v e = (w \<sqinter> -(prim_EP w v e)) \<squnion> prim_EP w v e"
using epm_1 sup_assoc sup_commute assms(2) by (simp add: inf_sup_distrib1)
also have "... = w \<squnion> prim_EP w v e"
by (metis assms(1) inf_top.right_neutral regular_complement_top sup_inf_distrib2)
also have "... = w"
by (simp add: sup_inf_distrib1)
finally show ?thesis
.
qed
lemma epm_4:
assumes "e \<le> w"
and "injective w"
and "w * v \<le> v"
and "e \<le> v * -v\<^sup>T"
shows "top * e * w\<^sup>T\<^sup>+ \<le> top * v\<^sup>T"
proof -
have "w\<^sup>\<star> * v \<le> v"
by (simp add: assms(3) star_left_induct_mult)
hence 1: "v\<^sup>T * w\<^sup>T\<^sup>\<star> \<le> v\<^sup>T"
using conv_star_commute conv_dist_comp conv_isotone by fastforce
have "e * w\<^sup>T \<le> w * w\<^sup>T \<sqinter> e * w\<^sup>T"
by (simp add: assms(1) mult_left_isotone)
also have "... \<le> 1 \<sqinter> e * w\<^sup>T"
using assms(2) inf.sup_left_isotone by auto
also have "... = 1 \<sqinter> w * e\<^sup>T"
using calculation conv_dist_comp conv_involutive coreflexive_symmetric by fastforce
also have "... \<le> w * e\<^sup>T"
by simp
also have "... \<le> w * -v * v\<^sup>T"
by (metis assms(4) conv_complement conv_dist_comp conv_involutive conv_order mult_assoc mult_right_isotone)
also have "... \<le> top * v\<^sup>T"
by (simp add: mult_left_isotone)
finally have "top * e * w\<^sup>T\<^sup>+ \<le> top * v\<^sup>T * w\<^sup>T\<^sup>\<star>"
- by (metis antisym comp_associative comp_isotone dense_top_closed mult_left_isotone transitive_top_closed)
+ by (metis order.antisym comp_associative comp_isotone dense_top_closed mult_left_isotone transitive_top_closed)
also have "... \<le> top * v\<^sup>T"
using 1 by (simp add: mult_assoc mult_right_isotone)
finally show ?thesis
.
qed
lemma epm_5:
assumes "e \<le> w"
and "injective w"
and "w * v \<le> v"
and "e \<le> v * -v\<^sup>T"
and "vector v"
shows "prim_P w v e = bot"
proof -
have 1: "e = w \<sqinter> top * e"
by (simp add: assms(1,2) epm_3)
have 2: "top * e * w\<^sup>T\<^sup>+ \<le> top * v\<^sup>T"
by (simp add: assms(1-4) epm_4)
have 3: "-v * -v\<^sup>T \<sqinter> top * v\<^sup>T = bot"
by (simp add: assms(5) comp_associative covector_vector_comp inf.sup_monoid.add_commute schroeder_2)
have "prim_P w v e = (w \<sqinter> -v * -v\<^sup>T \<sqinter> top * e) \<squnion> (w \<sqinter> -v * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>+)"
by (metis inf_sup_distrib1 mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute)
also have "... \<le> (e \<sqinter> -v * -v\<^sup>T) \<squnion> (w \<sqinter> -v * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>+)"
using 1 by (metis comp_inf.mult_semi_associative inf.sup_monoid.add_commute semiring.add_right_mono)
also have "... \<le> (e \<sqinter> -v * -v\<^sup>T) \<squnion> (w \<sqinter> -v * -v\<^sup>T \<sqinter> top * v\<^sup>T)"
using 2 by (metis sup_right_isotone inf.sup_right_isotone)
also have "... \<le> (e \<sqinter> -v * -v\<^sup>T) \<squnion> (-v * -v\<^sup>T \<sqinter> top * v\<^sup>T)"
using inf.assoc le_infI2 by auto
also have "... \<le> v * -v\<^sup>T \<sqinter> -v * -v\<^sup>T"
using 3 assms(4) inf.sup_left_isotone by auto
also have "... \<le> v * top \<sqinter> -v * top"
using inf.sup_mono mult_right_isotone top_greatest by blast
also have "... = bot"
using assms(5) inf_compl_bot vector_complement_closed by auto
finally show ?thesis
by (simp add: le_iff_inf)
qed
lemma epm_6:
assumes "e \<le> w"
and "injective w"
and "w * v \<le> v"
and "e \<le> v * -v\<^sup>T"
and "vector v"
shows "prim_E w v e = e"
proof -
have 1: "e \<le> --v * -v\<^sup>T"
using assms(4) mult_isotone order_lesseq_imp pp_increasing by blast
have 2: "top * e * w\<^sup>T\<^sup>+ \<le> top * v\<^sup>T"
by (simp add: assms(1-4) epm_4)
have 3: "e = w \<sqinter> top * e"
by (simp add: assms(1,2) epm_3)
hence "e \<le> top * e * w\<^sup>T\<^sup>\<star>"
by (metis le_infI2 star.circ_back_loop_fixpoint sup.commute sup_ge1)
hence 4: "e \<le> prim_E w v e"
using 1 by (simp add: assms(1))
have 5: "--v * -v\<^sup>T \<sqinter> top * v\<^sup>T = bot"
by (simp add: assms(5) comp_associative covector_vector_comp inf.sup_monoid.add_commute schroeder_2)
have "prim_E w v e = (w \<sqinter> --v * -v\<^sup>T \<sqinter> top * e) \<squnion> (w \<sqinter> --v * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>+)"
by (metis inf_sup_distrib1 mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute)
also have "... \<le> (e \<sqinter> --v * -v\<^sup>T) \<squnion> (w \<sqinter> --v * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>+)"
using 3 by (metis comp_inf.mult_semi_associative inf.sup_monoid.add_commute semiring.add_right_mono)
also have "... \<le> (e \<sqinter> --v * -v\<^sup>T) \<squnion> (w \<sqinter> --v * -v\<^sup>T \<sqinter> top * v\<^sup>T)"
using 2 by (metis sup_right_isotone inf.sup_right_isotone)
also have "... \<le> (e \<sqinter> --v * -v\<^sup>T) \<squnion> (--v * -v\<^sup>T \<sqinter> top * v\<^sup>T)"
using inf.assoc le_infI2 by auto
also have "... \<le> e"
by (simp add: "5")
finally show ?thesis
- using 4 by (simp add: antisym)
+ using 4 by (simp add: order.antisym)
qed
lemma epm_7:
"regular (prim_EP w v e) \<Longrightarrow> e \<le> w \<Longrightarrow> injective w \<Longrightarrow> w * v \<le> v \<Longrightarrow> e \<le> v * -v\<^sup>T \<Longrightarrow> vector v \<Longrightarrow> prim_W w v e = w"
by (metis conv_bot epm_2 epm_5 epm_6)
lemma epm_8:
assumes "acyclic w"
shows "(w \<sqinter> -(prim_EP w v e)) \<sqinter> (prim_P w v e)\<^sup>T = bot"
proof -
have "(w \<sqinter> -(prim_EP w v e)) \<sqinter> (prim_P w v e)\<^sup>T \<le> w \<sqinter> w\<^sup>T"
by (meson conv_isotone inf_le1 inf_mono order_trans)
thus ?thesis
by (metis assms acyclic_asymmetric inf.commute le_bot)
qed
lemma epm_9:
assumes "e \<le> v * -v\<^sup>T"
and "vector v"
shows "(w \<sqinter> -(prim_EP w v e)) \<sqinter> e = bot"
proof -
have 1: "e \<le> -v\<^sup>T"
by (metis assms complement_conv_sub vector_conv_covector ev p_antitone_iff p_bot)
have "(w \<sqinter> -(prim_EP w v e)) \<sqinter> e = (w \<sqinter> --v\<^sup>T \<sqinter> e) \<squnion> (w \<sqinter> -(top * e * w\<^sup>T\<^sup>\<star>) \<sqinter> e)"
by (simp add: inf_commute inf_sup_distrib1)
also have "... \<le> (--v\<^sup>T \<sqinter> e) \<squnion> (-(top * e * w\<^sup>T\<^sup>\<star>) \<sqinter> e)"
using comp_inf.mult_left_isotone inf.cobounded2 semiring.add_mono by blast
also have "... = -(top * e * w\<^sup>T\<^sup>\<star>) \<sqinter> e"
using 1 by (metis inf.sup_relative_same_increasing inf_commute inf_sup_distrib1 maddux_3_13 regular_closed_p)
also have "... = bot"
by (metis inf.sup_relative_same_increasing inf_bot_right inf_commute inf_p mult_left_isotone star_outer_increasing top_greatest)
finally show ?thesis
by (simp add: le_iff_inf)
qed
lemma epm_10:
assumes "e \<le> v * -v\<^sup>T"
and "vector v"
shows "(prim_P w v e)\<^sup>T \<sqinter> e = bot"
proof -
have "(prim_P w v e)\<^sup>T \<le> -v * -v\<^sup>T"
by (simp add: conv_complement conv_dist_comp conv_dist_inf inf.absorb_iff1 inf.left_commute inf_commute)
hence "(prim_P w v e)\<^sup>T \<sqinter> e \<le> -v * -v\<^sup>T \<sqinter> v * -v\<^sup>T"
using assms(1) inf_mono by blast
also have "... \<le> -v * top \<sqinter> v * top"
using inf.sup_mono mult_right_isotone top_greatest by blast
also have "... = bot"
using assms(2) inf_compl_bot vector_complement_closed by auto
finally show ?thesis
by (simp add: le_iff_inf)
qed
lemma epm_11:
assumes "vector v"
shows "(w \<sqinter> -(prim_EP w v e)) \<sqinter> prim_P w v e = bot"
proof -
have "prim_P w v e \<le> prim_EP w v e"
by (metis assms comp_isotone inf.sup_left_isotone inf.sup_right_isotone order.refl top_greatest vector_conv_compl)
thus ?thesis
using inf_le2 order_trans p_antitone pseudo_complement by blast
qed
lemma epm_12:
assumes "vector v"
shows "(w \<sqinter> -(prim_EP w v e)) \<sqinter> prim_E w v e = bot"
proof -
have "prim_E w v e \<le> prim_EP w v e"
by (metis assms comp_isotone inf.sup_left_isotone inf.sup_right_isotone order.refl top_greatest vector_conv_compl)
thus ?thesis
using inf_le2 order_trans p_antitone pseudo_complement by blast
qed
lemma epm_13:
assumes "vector v"
shows "prim_P w v e \<sqinter> prim_E w v e = bot"
proof -
have "prim_P w v e \<sqinter> prim_E w v e \<le> -v * -v\<^sup>T \<sqinter> --v * -v\<^sup>T"
by (meson dual_order.trans inf.cobounded1 inf.sup_mono inf_le2)
also have "... \<le> -v * top \<sqinter> --v * top"
using inf.sup_mono mult_right_isotone top_greatest by blast
also have "... = bot"
using assms inf_compl_bot vector_complement_closed by auto
finally show ?thesis
by (simp add: le_iff_inf)
qed
text \<open>
The following lemmas show that the relation characterising the edge across the cut is an arc.
\<close>
lemma arc_edge_1:
assumes "e \<le> v * -v\<^sup>T \<sqinter> g"
and "vector v"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "t \<le> g"
and "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
shows "top * e \<le> v\<^sup>T * w\<^sup>\<star>"
proof -
have "top * e \<le> top * (v * -v\<^sup>T \<sqinter> g)"
using assms(1) mult_right_isotone by auto
also have "... \<le> top * (v * top \<sqinter> g)"
using inf.sup_right_isotone inf_commute mult_right_isotone top_greatest by presburger
also have "... = v\<^sup>T * g"
by (metis assms(2) covector_inf_comp_3 inf_top.left_neutral)
also have "... = r\<^sup>T * t\<^sup>\<star> * g"
by (simp add: assms(3))
also have "... \<le> r\<^sup>T * g\<^sup>\<star> * g"
by (simp add: assms(4) mult_left_isotone mult_right_isotone star_isotone)
also have "... \<le> r\<^sup>T * g\<^sup>\<star>"
by (simp add: mult_assoc mult_right_isotone star.right_plus_below_circ)
also have "... \<le> r\<^sup>T * w\<^sup>\<star>"
by (simp add: assms(5))
also have "... \<le> v\<^sup>T * w\<^sup>\<star>"
by (metis assms(3) mult_left_isotone mult_right_isotone mult_1_right star.circ_reflexive)
finally show ?thesis
.
qed
lemma arc_edge_2:
assumes "e \<le> v * -v\<^sup>T \<sqinter> g"
and "vector v"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "t \<le> g"
and "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
and "w * v \<le> v"
and "injective w"
shows "top * e * w\<^sup>T\<^sup>\<star> \<le> v\<^sup>T * w\<^sup>\<star>"
proof -
have 1: "top * e \<le> v\<^sup>T * w\<^sup>\<star>"
using assms(1-5) arc_edge_1 by blast
have "v\<^sup>T * w\<^sup>\<star> * w\<^sup>T = v\<^sup>T * w\<^sup>T \<squnion> v\<^sup>T * w\<^sup>+ * w\<^sup>T"
by (metis mult_assoc mult_left_dist_sup star.circ_loop_fixpoint sup_commute)
also have "... \<le> v\<^sup>T \<squnion> v\<^sup>T * w\<^sup>+ * w\<^sup>T"
by (metis assms(6) conv_dist_comp conv_isotone sup_left_isotone)
also have "... = v\<^sup>T \<squnion> v\<^sup>T * w\<^sup>\<star> * (w * w\<^sup>T)"
by (metis mult_assoc star_plus)
also have "... \<le> v\<^sup>T \<squnion> v\<^sup>T * w\<^sup>\<star>"
by (metis assms(7) mult_right_isotone mult_1_right sup_right_isotone)
also have "... = v\<^sup>T * w\<^sup>\<star>"
by (metis star.circ_back_loop_fixpoint sup_absorb2 sup_ge2)
finally show ?thesis
using 1 star_right_induct by auto
qed
lemma arc_edge_3:
assumes "e \<le> v * -v\<^sup>T \<sqinter> g"
and "vector v"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "t \<le> g"
and "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
and "w * v \<le> v"
and "injective w"
and "prim_E w v e = bot"
shows "e = bot"
proof -
have "bot = prim_E w v e"
by (simp add: assms(8))
also have "... = w \<sqinter> --v * top \<sqinter> top * -v\<^sup>T \<sqinter> top * e * w\<^sup>T\<^sup>\<star>"
by (metis assms(2) comp_inf_covector inf.assoc inf_top.left_neutral vector_conv_compl)
also have "... = w \<sqinter> top * e * w\<^sup>T\<^sup>\<star> \<sqinter> -v\<^sup>T \<sqinter> --v"
using assms(2) inf.assoc inf.commute vector_conv_compl vector_complement_closed by (simp add: inf_assoc)
finally have 1: "w \<sqinter> top * e * w\<^sup>T\<^sup>\<star> \<sqinter> -v\<^sup>T \<le> -v"
using shunting_1_pp by force
have "w\<^sup>\<star> * e\<^sup>T * top = (top * e * w\<^sup>T\<^sup>\<star>)\<^sup>T"
by (simp add: conv_star_commute comp_associative conv_dist_comp)
also have "... \<le> (v\<^sup>T * w\<^sup>\<star>)\<^sup>T"
using assms(1-7) arc_edge_2 by (simp add: conv_isotone)
also have "... = w\<^sup>T\<^sup>\<star> * v"
by (simp add: conv_star_commute conv_dist_comp)
finally have 2: "w\<^sup>\<star> * e\<^sup>T * top \<le> w\<^sup>T\<^sup>\<star> * v"
.
have "(w\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top)\<^sup>T * -v = (w \<sqinter> top * e * w\<^sup>T\<^sup>\<star>) * -v"
by (simp add: conv_dist_comp conv_dist_inf conv_star_commute mult_assoc)
also have "... = (w \<sqinter> top * e * w\<^sup>T\<^sup>\<star> \<sqinter> -v\<^sup>T) * top"
by (metis assms(2) conv_complement covector_inf_comp_3 inf_top.right_neutral vector_complement_closed)
also have "... \<le> -v * top"
using 1 by (simp add: comp_isotone)
also have "... = -v"
using assms(2) vector_complement_closed by auto
finally have "(w\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top) * --v \<le> --v"
using p_antitone_iff schroeder_3_p by auto
hence "w\<^sup>\<star> * e\<^sup>T * top \<sqinter> w\<^sup>T * --v \<le> --v"
by (simp add: inf_vector_comp)
hence 3: "w\<^sup>T * --v \<le> --v \<squnion> -(w\<^sup>\<star> * e\<^sup>T * top)"
by (simp add: inf.commute shunting_p)
have "w\<^sup>T * -(w\<^sup>\<star> * e\<^sup>T * top) \<le> -(w\<^sup>\<star> * e\<^sup>T * top)"
by (metis mult_assoc p_antitone p_antitone_iff schroeder_3_p star.circ_loop_fixpoint sup_commute sup_right_divisibility)
also have "... \<le> --v \<squnion> -(w\<^sup>\<star> * e\<^sup>T * top)"
by simp
finally have "w\<^sup>T * (--v \<squnion> -(w\<^sup>\<star> * e\<^sup>T * top)) \<le> --v \<squnion> -(w\<^sup>\<star> * e\<^sup>T * top)"
using 3 by (simp add: mult_left_dist_sup)
hence "w\<^sup>T\<^sup>\<star> * (--v \<squnion> -(w\<^sup>\<star> * e\<^sup>T * top)) \<le> --v \<squnion> -(w\<^sup>\<star> * e\<^sup>T * top)"
using star_left_induct_mult_iff by blast
hence "w\<^sup>T\<^sup>\<star> * --v \<le> --v \<squnion> -(w\<^sup>\<star> * e\<^sup>T * top)"
by (simp add: semiring.distrib_left)
hence "w\<^sup>\<star> * e\<^sup>T * top \<sqinter> w\<^sup>T\<^sup>\<star> * --v \<le> --v"
by (simp add: inf_commute shunting_p)
hence "w\<^sup>\<star> * e\<^sup>T * top \<le> --v"
using 2 by (metis inf.absorb1 p_antitone_iff p_comp_pp vector_export_comp)
hence 4: "e\<^sup>T * top \<le> --v"
by (metis mult_assoc star.circ_loop_fixpoint sup.bounded_iff)
have "e\<^sup>T * top \<le> (v * -v\<^sup>T)\<^sup>T * top"
using assms(1) comp_isotone conv_isotone by auto
also have "... \<le> -v * top"
by (simp add: conv_complement conv_dist_comp mult_assoc mult_right_isotone)
also have "... = -v"
using assms(2) vector_complement_closed by auto
finally have "e\<^sup>T * top \<le> bot"
using 4 shunting_1_pp by auto
hence "e\<^sup>T = bot"
- using antisym bot_least top_right_mult_increasing by blast
+ using order.antisym bot_least top_right_mult_increasing by blast
thus ?thesis
using conv_bot by fastforce
qed
lemma arc_edge_4:
assumes "e \<le> v * -v\<^sup>T \<sqinter> g"
and "vector v"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "t \<le> g"
and "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
and "arc e"
shows "top * prim_E w v e * top = top"
proof -
have "--v\<^sup>T * w = (--v\<^sup>T * w \<sqinter> -v\<^sup>T) \<squnion> (--v\<^sup>T * w \<sqinter> --v\<^sup>T)"
by (simp add: maddux_3_11_pp)
also have "... \<le> (--v\<^sup>T * w \<sqinter> -v\<^sup>T) \<squnion> --v\<^sup>T"
using sup_right_isotone by auto
also have "... = --v\<^sup>T * (w \<sqinter> -v\<^sup>T) \<squnion> --v\<^sup>T"
using assms(2) covector_comp_inf covector_complement_closed vector_conv_covector by auto
also have "... \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> \<squnion> --v\<^sup>T"
by (metis star.circ_back_loop_fixpoint sup.cobounded2 sup_left_isotone)
finally have 1: "--v\<^sup>T * w \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> \<squnion> --v\<^sup>T"
.
have "--v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> * w \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> \<squnion> --v\<^sup>T"
by (simp add: le_supI1 mult_assoc mult_right_isotone star.circ_plus_same star.left_plus_below_circ)
hence 2: "(--v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> \<squnion> --v\<^sup>T) * w \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> \<squnion> --v\<^sup>T"
using 1 by (simp add: inf.orderE mult_right_dist_sup)
have "v\<^sup>T \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> \<squnion> --v\<^sup>T"
by (simp add: pp_increasing sup.coboundedI2)
hence "v\<^sup>T * w\<^sup>\<star> \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star> \<squnion> --v\<^sup>T"
using 2 by (simp add: star_right_induct)
hence 3: "-v\<^sup>T \<sqinter> v\<^sup>T * w\<^sup>\<star> \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star>"
by (metis inf_commute shunting_var_p)
have "top * e = top * e \<sqinter> v\<^sup>T * w\<^sup>\<star>"
by (meson assms(1-5) arc_edge_1 inf.orderE)
also have "... \<le> top * v * -v\<^sup>T \<sqinter> v\<^sup>T * w\<^sup>\<star>"
using assms(1) inf.sup_left_isotone mult_assoc mult_right_isotone by auto
also have "... \<le> top * -v\<^sup>T \<sqinter> v\<^sup>T * w\<^sup>\<star>"
using inf.sup_left_isotone mult_left_isotone top_greatest by blast
also have "... = -v\<^sup>T \<sqinter> v\<^sup>T * w\<^sup>\<star>"
by (simp add: assms(2) vector_conv_compl)
also have "... \<le> --v\<^sup>T * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star>"
using 3 by simp
also have "... = (top \<sqinter> (--v)\<^sup>T) * (w \<sqinter> -v\<^sup>T) * w\<^sup>\<star>"
by (simp add: conv_complement)
also have "... = top * (w \<sqinter> --v \<sqinter> -v\<^sup>T) * w\<^sup>\<star>"
using assms(2) covector_inf_comp_3 inf_assoc inf_left_commute vector_complement_closed by presburger
also have "... = top * (w \<sqinter> --v * -v\<^sup>T) * w\<^sup>\<star>"
by (metis assms(2) vector_complement_closed conv_complement inf_assoc vector_covector)
finally have "top * (e\<^sup>T * top)\<^sup>T \<le> top * (w \<sqinter> --v * -v\<^sup>T) * w\<^sup>\<star>"
by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top)
hence "top \<le> top * (w \<sqinter> --v * -v\<^sup>T) * w\<^sup>\<star> * (e\<^sup>T * top)"
using assms(6) shunt_bijective by blast
also have "... = top * (w \<sqinter> --v * -v\<^sup>T) * (top * e * w\<^sup>\<star>\<^sup>T)\<^sup>T"
by (simp add: conv_dist_comp mult_assoc)
also have "... = top * (w \<sqinter> --v * -v\<^sup>T \<sqinter> top * e * w\<^sup>\<star>\<^sup>T) * top"
by (simp add: comp_inf_vector_1 mult_assoc)
finally show ?thesis
by (simp add: conv_star_commute top_le)
qed
lemma arc_edge_5:
assumes "vector v"
and "w * v \<le> v"
and "injective w"
and "arc e"
shows "(prim_E w v e)\<^sup>T * top * prim_E w v e \<le> 1"
proof -
have 1: "e\<^sup>T * top * e \<le> 1"
by (simp add: assms(4) point_injective)
have "prim_E w v e \<le> --v * top"
by (simp add: inf_commute le_infI2 mult_right_isotone)
hence 2: "prim_E w v e \<le> --v"
by (simp add: assms(1) vector_complement_closed)
have 3: "w * --v \<le> --v"
by (simp add: assms(2) p_antitone p_antitone_iff)
have "w \<sqinter> top * prim_E w v e \<le> w * (prim_E w v e)\<^sup>T * prim_E w v e"
by (metis dedekind_2 inf.commute inf_top.left_neutral)
also have "... \<le> w * w\<^sup>T * prim_E w v e"
by (simp add: conv_isotone le_infI1 mult_left_isotone mult_right_isotone)
also have "... \<le> prim_E w v e"
by (metis assms(3) mult_left_isotone mult_left_one)
finally have 4: "w \<sqinter> top * prim_E w v e \<le> prim_E w v e"
.
have "w\<^sup>+ \<sqinter> top * prim_E w v e = w\<^sup>\<star> * (w \<sqinter> top * prim_E w v e)"
by (simp add: comp_inf_covector star_plus)
also have "... \<le> w\<^sup>\<star> * prim_E w v e"
using 4 by (simp add: mult_right_isotone)
also have "... \<le> --v"
using 2 3 star_left_induct sup.bounded_iff by blast
finally have 5: "w\<^sup>+ \<sqinter> top * prim_E w v e \<sqinter> -v = bot"
using shunting_1_pp by blast
hence 6: "w\<^sup>+\<^sup>T \<sqinter> (prim_E w v e)\<^sup>T * top \<sqinter> -v\<^sup>T = bot"
using conv_complement conv_dist_comp conv_dist_inf conv_top conv_bot by force
have "(prim_E w v e)\<^sup>T * top * prim_E w v e \<le> (top * e * w\<^sup>T\<^sup>\<star>)\<^sup>T * top * (top * e * w\<^sup>T\<^sup>\<star>)"
by (simp add: conv_isotone mult_isotone)
also have "... = w\<^sup>\<star> * e\<^sup>T * top * e * w\<^sup>T\<^sup>\<star>"
by (metis conv_star_commute conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top)
also have "... \<le> w\<^sup>\<star> * w\<^sup>T\<^sup>\<star>"
using 1 by (metis mult_assoc mult_1_right mult_right_isotone mult_left_isotone)
also have "... = w\<^sup>\<star> \<squnion> w\<^sup>T\<^sup>\<star>"
- by (metis assms(3) cancel_separate inf.eq_iff star.circ_sup_sub_sup_one_1 star.circ_plus_one star_involutive)
+ by (metis assms(3) cancel_separate order.eq_iff star.circ_sup_sub_sup_one_1 star.circ_plus_one star_involutive)
also have "... = w\<^sup>+ \<squnion> w\<^sup>T\<^sup>+ \<squnion> 1"
by (metis star.circ_plus_one star_left_unfold_equal sup.assoc sup.commute)
finally have 7: "(prim_E w v e)\<^sup>T * top * prim_E w v e \<le> w\<^sup>+ \<squnion> w\<^sup>T\<^sup>+ \<squnion> 1"
.
have "prim_E w v e \<le> --v * -v\<^sup>T"
by (simp add: le_infI1)
also have "... \<le> top * -v\<^sup>T"
by (simp add: mult_left_isotone)
also have "... = -v\<^sup>T"
by (simp add: assms(1) vector_conv_compl)
finally have 8: "prim_E w v e \<le> -v\<^sup>T"
.
hence 9: "(prim_E w v e)\<^sup>T \<le> -v"
by (metis conv_complement conv_involutive conv_isotone)
have "(prim_E w v e)\<^sup>T * top * prim_E w v e = (w\<^sup>+ \<squnion> w\<^sup>T\<^sup>+ \<squnion> 1) \<sqinter> (prim_E w v e)\<^sup>T * top * prim_E w v e"
using 7 by (simp add: inf.absorb_iff2)
also have "... = (1 \<sqinter> (prim_E w v e)\<^sup>T * top * prim_E w v e) \<squnion> (w\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top * prim_E w v e) \<squnion> (w\<^sup>T\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top * prim_E w v e)"
using comp_inf.mult_right_dist_sup sup_assoc sup_commute by auto
also have "... \<le> 1 \<squnion> (w\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top * prim_E w v e) \<squnion> (w\<^sup>T\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top * prim_E w v e)"
using inf_le1 sup_left_isotone by blast
also have "... \<le> 1 \<squnion> (w\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top * prim_E w v e) \<squnion> (w\<^sup>T\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top * -v\<^sup>T)"
using 8 inf.sup_right_isotone mult_right_isotone sup_right_isotone by blast
also have "... \<le> 1 \<squnion> (w\<^sup>+ \<sqinter> -v * top * prim_E w v e) \<squnion> (w\<^sup>T\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top * -v\<^sup>T)"
using 9 by (metis inf.sup_right_isotone mult_left_isotone sup.commute sup_right_isotone)
also have "... = 1 \<squnion> (w\<^sup>+ \<sqinter> -v * top \<sqinter> top * prim_E w v e) \<squnion> (w\<^sup>T\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top \<sqinter> top * -v\<^sup>T)"
by (metis (no_types) vector_export_comp inf_top_right inf_assoc)
also have "... = 1 \<squnion> (w\<^sup>+ \<sqinter> -v \<sqinter> top * prim_E w v e) \<squnion> (w\<^sup>T\<^sup>+ \<sqinter> (prim_E w v e)\<^sup>T * top \<sqinter> -v\<^sup>T)"
using assms(1) vector_complement_closed vector_conv_compl by auto
also have "... = 1"
using 5 6 by (simp add: conv_star_commute conv_dist_comp inf.commute inf_assoc star.circ_plus_same)
finally show ?thesis
.
qed
lemma arc_edge_6:
assumes "vector v"
and "w * v \<le> v"
and "injective w"
and "arc e"
shows "prim_E w v e * top * (prim_E w v e)\<^sup>T \<le> 1"
proof -
have "prim_E w v e * 1 * (prim_E w v e)\<^sup>T \<le> w * w\<^sup>T"
using comp_isotone conv_order inf.coboundedI1 mult_one_associative by auto
also have "... \<le> 1"
by (simp add: assms(3))
finally have 1: "prim_E w v e * 1 * (prim_E w v e)\<^sup>T \<le> 1"
.
have "(prim_E w v e)\<^sup>T * top * prim_E w v e \<le> 1"
by (simp add: assms arc_edge_5)
also have "... \<le> --1"
by (simp add: pp_increasing)
finally have 2: "prim_E w v e * -1 * (prim_E w v e)\<^sup>T \<le> bot"
by (metis conv_involutive regular_closed_bot regular_dense_top triple_schroeder_p)
have "prim_E w v e * top * (prim_E w v e)\<^sup>T = prim_E w v e * 1 * (prim_E w v e)\<^sup>T \<squnion> prim_E w v e * -1 * (prim_E w v e)\<^sup>T"
by (metis mult_left_dist_sup mult_right_dist_sup regular_complement_top regular_one_closed)
also have "... \<le> 1"
using 1 2 by (simp add: bot_unique)
finally show ?thesis
.
qed
lemma arc_edge:
assumes "e \<le> v * -v\<^sup>T \<sqinter> g"
and "vector v"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "t \<le> g"
and "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
and "w * v \<le> v"
and "injective w"
and "arc e"
shows "arc (prim_E w v e)"
proof (intro conjI)
have "prim_E w v e * top * (prim_E w v e)\<^sup>T \<le> 1"
using assms(2,6-8) arc_edge_6 by simp
thus "injective (prim_E w v e * top)"
by (metis conv_dist_comp conv_top mult_assoc top_mult_top)
next
show "surjective (prim_E w v e * top)"
using assms(1-5,8) arc_edge_4 mult_assoc by simp
next
have "(prim_E w v e)\<^sup>T * top * prim_E w v e \<le> 1"
using assms(2,6-8) arc_edge_5 by simp
thus "injective ((prim_E w v e)\<^sup>T * top)"
by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top)
next
have "top * prim_E w v e * top = top"
using assms(1-5,8) arc_edge_4 by simp
thus "surjective ((prim_E w v e)\<^sup>T * top)"
by (metis mult_assoc conv_dist_comp conv_top)
qed
subsubsection \<open>Invariant implies Postcondition\<close>
text \<open>
The lemmas in this section are used to show that the invariant implies the postcondition at the end of the algorithm.
The following lemma shows that the nodes reachable in the graph are the same as those reachable in the constructed tree.
\<close>
lemma span_post:
assumes "regular v"
and "vector v"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "v * -v\<^sup>T \<sqinter> g = bot"
and "t \<le> v * v\<^sup>T \<sqinter> g"
and "r\<^sup>T * (v * v\<^sup>T \<sqinter> g)\<^sup>\<star> \<le> r\<^sup>T * t\<^sup>\<star>"
shows "v\<^sup>T = r\<^sup>T * g\<^sup>\<star>"
proof -
let ?vv = "v * v\<^sup>T \<sqinter> g"
have 1: "r\<^sup>T \<le> v\<^sup>T"
using assms(3) mult_right_isotone mult_1_right star.circ_reflexive by fastforce
have "v * top \<sqinter> g = (v * v\<^sup>T \<squnion> v * -v\<^sup>T) \<sqinter> g"
by (metis assms(1) conv_complement mult_left_dist_sup regular_complement_top)
also have "... = ?vv \<squnion> (v * -v\<^sup>T \<sqinter> g)"
by (simp add: inf_sup_distrib2)
also have "... = ?vv"
by (simp add: assms(4))
finally have 2: "v * top \<sqinter> g = ?vv"
by simp
have "r\<^sup>T * ?vv\<^sup>\<star> \<le> v\<^sup>T * ?vv\<^sup>\<star>"
using 1 by (simp add: comp_left_isotone)
also have "... \<le> v\<^sup>T * (v * v\<^sup>T)\<^sup>\<star>"
by (simp add: comp_right_isotone star.circ_isotone)
also have "... \<le> v\<^sup>T"
by (simp add: assms(2) vector_star_1)
finally have "r\<^sup>T * ?vv\<^sup>\<star> \<le> v\<^sup>T"
by simp
hence "r\<^sup>T * ?vv\<^sup>\<star> * g = (r\<^sup>T * ?vv\<^sup>\<star> \<sqinter> v\<^sup>T) * g"
by (simp add: inf.absorb1)
also have "... = r\<^sup>T * ?vv\<^sup>\<star> * (v * top \<sqinter> g)"
by (simp add: assms(2) covector_inf_comp_3)
also have "... = r\<^sup>T * ?vv\<^sup>\<star> * ?vv"
using 2 by simp
also have "... \<le> r\<^sup>T * ?vv\<^sup>\<star>"
by (simp add: comp_associative comp_right_isotone star.left_plus_below_circ star_plus)
finally have "r\<^sup>T \<squnion> r\<^sup>T * ?vv\<^sup>\<star> * g \<le> r\<^sup>T * ?vv\<^sup>\<star>"
using star.circ_back_loop_prefixpoint by auto
hence "r\<^sup>T * g\<^sup>\<star> \<le> r\<^sup>T * ?vv\<^sup>\<star>"
using star_right_induct by blast
hence "r\<^sup>T * g\<^sup>\<star> = r\<^sup>T * ?vv\<^sup>\<star>"
- by (simp add: antisym mult_right_isotone star_isotone)
+ by (simp add: order.antisym mult_right_isotone star_isotone)
also have "... = r\<^sup>T * t\<^sup>\<star>"
- using assms(5,6) antisym mult_right_isotone star_isotone by auto
+ using assms(5,6) order.antisym mult_right_isotone star_isotone by auto
also have "... = v\<^sup>T"
by (simp add: assms(3))
finally show ?thesis
by simp
qed
text \<open>
The following lemma shows that the minimum spanning tree extending a tree is the same as the tree at the end of the algorithm.
\<close>
lemma mst_post:
assumes "vector r"
and "injective r"
and "v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
and "forest w"
and "t \<le> w"
and "w \<le> v * v\<^sup>T"
shows "w = t"
proof -
have 1: "vector v"
using assms(1,3) covector_mult_closed vector_conv_covector by auto
have "w * v \<le> v * v\<^sup>T * v"
by (simp add: assms(6) mult_left_isotone)
also have "... \<le> v"
using 1 by (metis mult_assoc mult_right_isotone top_greatest)
finally have 2: "w * v \<le> v"
.
have 3: "r \<le> v"
by (metis assms(3) conv_order mult_right_isotone mult_1_right star.circ_reflexive)
have 4: "v \<sqinter> -r = t\<^sup>T\<^sup>\<star> * r \<sqinter> -r"
by (metis assms(3) conv_dist_comp conv_involutive conv_star_commute)
also have "... = (r \<squnion> t\<^sup>T\<^sup>+ * r) \<sqinter> -r"
using mult_assoc star.circ_loop_fixpoint sup_commute by auto
also have "... \<le> t\<^sup>T\<^sup>+ * r"
by (simp add: shunting)
also have "... \<le> t\<^sup>T * top"
by (simp add: comp_isotone mult_assoc)
finally have "1 \<sqinter> (v \<sqinter> -r) * (v \<sqinter> -r)\<^sup>T \<le> 1 \<sqinter> t\<^sup>T * top * (t\<^sup>T * top)\<^sup>T"
using conv_order inf.sup_right_isotone mult_isotone by auto
also have "... = 1 \<sqinter> t\<^sup>T * top * t"
by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top)
also have "... \<le> t\<^sup>T * (top * t \<sqinter> t * 1)"
by (metis conv_involutive dedekind_1 inf.commute mult_assoc)
also have "... \<le> t\<^sup>T * t"
by (simp add: mult_right_isotone)
finally have 5: "1 \<sqinter> (v \<sqinter> -r) * (v \<sqinter> -r)\<^sup>T \<le> t\<^sup>T * t"
.
have "w * w\<^sup>+ \<le> -1"
by (metis assms(4) mult_right_isotone order_trans star.circ_increasing star.left_plus_circ)
hence 6: "w\<^sup>T\<^sup>+ \<le> -w"
by (metis conv_star_commute mult_assoc mult_1_left triple_schroeder_p)
have "w * r \<sqinter> w\<^sup>T\<^sup>+ * r = (w \<sqinter> w\<^sup>T\<^sup>+) * r"
using assms(2) by (simp add: injective_comp_right_dist_inf)
also have "... = bot"
using 6 p_antitone pseudo_complement_pp semiring.mult_not_zero by blast
finally have 7: "w * r \<sqinter> w\<^sup>T\<^sup>+ * r = bot"
.
have "-1 * r \<le> -r"
using assms(2) dual_order.trans pp_increasing schroeder_4_p by blast
hence "-1 * r * top \<le> -r"
by (simp add: assms(1) comp_associative)
hence 8: "r\<^sup>T * -1 * r \<le> bot"
by (simp add: mult_assoc schroeder_6_p)
have "r\<^sup>T * w * r \<le> r\<^sup>T * w\<^sup>+ * r"
by (simp add: mult_left_isotone mult_right_isotone star.circ_mult_increasing)
also have "... \<le> r\<^sup>T * -1 * r"
by (simp add: assms(4) comp_isotone)
finally have "r\<^sup>T * w * r \<le> bot"
using 8 by simp
hence "w * r * top \<le> -r"
by (simp add: mult_assoc schroeder_6_p)
hence "w * r \<le> -r"
by (simp add: assms(1) comp_associative)
hence "w * r \<le> -r \<sqinter> w * v"
using 3 by (simp add: mult_right_isotone)
also have "... \<le> -r \<sqinter> v"
using 2 by (simp add: le_infI2)
also have "... = -r \<sqinter> t\<^sup>T\<^sup>\<star> * r"
using 4 by (simp add: inf_commute)
also have "... \<le> -r \<sqinter> w\<^sup>T\<^sup>\<star> * r"
using assms(5) comp_inf.mult_right_isotone conv_isotone mult_left_isotone star_isotone by auto
also have "... = -r \<sqinter> (r \<squnion> w\<^sup>T\<^sup>+ * r)"
using mult_assoc star.circ_loop_fixpoint sup_commute by auto
also have "... \<le> w\<^sup>T\<^sup>+ * r"
using inf.commute maddux_3_13 by auto
finally have "w * r = bot"
using 7 by (simp add: le_iff_inf)
hence "w = w \<sqinter> top * -r\<^sup>T"
by (metis complement_conv_sub conv_dist_comp conv_involutive conv_bot inf.assoc inf.orderE regular_closed_bot regular_dense_top top_left_mult_increasing)
also have "... = w \<sqinter> v * v\<^sup>T \<sqinter> top * -r\<^sup>T"
by (simp add: assms(6) inf_absorb1)
also have "... \<le> w \<sqinter> top * v\<^sup>T \<sqinter> top * -r\<^sup>T"
using comp_inf.mult_left_isotone comp_inf.mult_right_isotone mult_left_isotone by auto
also have "... = w \<sqinter> top * (v\<^sup>T \<sqinter> -r\<^sup>T)"
using 1 assms(1) covector_inf_closed inf_assoc vector_conv_compl vector_conv_covector by auto
also have "... = w * (1 \<sqinter> (v \<sqinter> -r) * top)"
by (simp add: comp_inf_vector conv_complement conv_dist_inf)
also have "... = w * (1 \<sqinter> (v \<sqinter> -r) * (v \<sqinter> -r)\<^sup>T)"
by (metis conv_top dedekind_eq inf_commute inf_top_left mult_1_left mult_1_right)
also have "... \<le> w * t\<^sup>T * t"
using 5 by (simp add: comp_isotone mult_assoc)
also have "... \<le> w * w\<^sup>T * t"
by (simp add: assms(5) comp_isotone conv_isotone)
also have "... \<le> t"
using assms(4) mult_left_isotone mult_1_left by fastforce
finally show ?thesis
- by (simp add: assms(5) antisym)
+ by (simp add: assms(5) order.antisym)
qed
subsection \<open>Kruskal's Algorithm\<close>
text \<open>
The following results are used for proving the correctness of Kruskal's minimum spanning tree algorithm.
\<close>
subsubsection \<open>Preservation of Invariant\<close>
text \<open>
We first treat the preservation of the invariant.
The following lemmas show conditions necessary for preserving that \<open>f\<close> is a forest.
\<close>
lemma kruskal_injective_inv_2:
assumes "arc e"
and "acyclic f"
shows "top * e * f\<^sup>T\<^sup>\<star> * f\<^sup>T \<le> -e"
proof -
have "f \<le> -f\<^sup>T\<^sup>\<star>"
using assms(2) acyclic_star_below_complement p_antitone_iff by simp
hence "e * f \<le> top * e * -f\<^sup>T\<^sup>\<star>"
by (simp add: comp_isotone top_left_mult_increasing)
also have "... = -(top * e * f\<^sup>T\<^sup>\<star>)"
by (metis assms(1) comp_mapping_complement conv_dist_comp conv_involutive conv_top)
finally show ?thesis
using schroeder_4_p by simp
qed
lemma kruskal_injective_inv_3:
assumes "arc e"
and "forest f"
shows "(top * e * f\<^sup>T\<^sup>\<star>)\<^sup>T * (top * e * f\<^sup>T\<^sup>\<star>) \<sqinter> f\<^sup>T * f \<le> 1"
proof -
have "(top * e * f\<^sup>T\<^sup>\<star>)\<^sup>T * (top * e * f\<^sup>T\<^sup>\<star>) = f\<^sup>\<star> * e\<^sup>T * top * e * f\<^sup>T\<^sup>\<star>"
by (metis conv_dist_comp conv_involutive conv_star_commute conv_top vector_top_closed mult_assoc)
also have "... \<le> f\<^sup>\<star> * f\<^sup>T\<^sup>\<star>"
by (metis assms(1) arc_expanded mult_left_isotone mult_right_isotone mult_1_left mult_assoc)
finally have "(top * e * f\<^sup>T\<^sup>\<star>)\<^sup>T * (top * e * f\<^sup>T\<^sup>\<star>) \<sqinter> f\<^sup>T * f \<le> f\<^sup>\<star> * f\<^sup>T\<^sup>\<star> \<sqinter> f\<^sup>T * f"
using inf.sup_left_isotone by simp
also have "... \<le> 1"
using assms(2) forest_separate by simp
finally show ?thesis
by simp
qed
lemma kruskal_acyclic_inv:
assumes "acyclic f"
and "covector q"
and "(f \<sqinter> q)\<^sup>T * f\<^sup>\<star> * e = bot"
and "e * f\<^sup>\<star> * e = bot"
and "f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> \<le> -e"
shows "acyclic ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e)"
proof -
have "(f \<sqinter> -q) * (f \<sqinter> q)\<^sup>T = (f \<sqinter> -q) * (f\<^sup>T \<sqinter> q\<^sup>T)"
by (simp add: conv_dist_inf)
hence 1: "(f \<sqinter> -q) * (f \<sqinter> q)\<^sup>T = bot"
by (metis assms(2) comp_inf.semiring.mult_zero_right comp_inf_vector_1 conv_bot covector_bot_closed inf.sup_monoid.add_assoc p_inf)
hence 2: "(f \<sqinter> -q)\<^sup>\<star> * (f \<sqinter> q)\<^sup>T = (f \<sqinter> q)\<^sup>T"
using mult_right_zero star_absorb star_simulation_right_equal by fastforce
hence 3: "((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T)\<^sup>+ = (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+"
by (simp add: plus_sup)
have 4: "((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T)\<^sup>\<star> = (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>\<star>"
using 2 by (simp add: star.circ_sup_9)
have "(f \<sqinter> q)\<^sup>T * (f \<sqinter> -q)\<^sup>\<star> * e \<le> (f \<sqinter> q)\<^sup>T * f\<^sup>\<star> * e"
by (simp add: mult_left_isotone mult_right_isotone star_isotone)
hence "(f \<sqinter> q)\<^sup>T * (f \<sqinter> -q)\<^sup>\<star> * e = bot"
using assms(3) le_bot by simp
hence 5: "(f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>\<star> * e = (f \<sqinter> -q)\<^sup>\<star> * e"
by (metis comp_associative conv_bot conv_dist_comp conv_involutive conv_star_commute star_absorb)
have "e * (f \<sqinter> -q)\<^sup>\<star> * e \<le> e * f\<^sup>\<star> * e"
by (simp add: mult_left_isotone mult_right_isotone star_isotone)
hence "e * (f \<sqinter> -q)\<^sup>\<star> * e = bot"
using assms(4) le_bot by simp
hence 6: "((f \<sqinter> -q)\<^sup>\<star> * e)\<^sup>+ = (f \<sqinter> -q)\<^sup>\<star> * e"
by (simp add: comp_associative star_absorb)
have "f\<^sup>T\<^sup>\<star> * 1 * f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> \<le> -e"
by (simp add: assms(5) star.circ_transitive_equal)
hence 7: "f\<^sup>\<star> * e * f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> \<le> -1"
by (metis comp_right_one conv_involutive conv_one conv_star_commute triple_schroeder_p)
have "(f \<sqinter> -q)\<^sup>+ * (f \<sqinter> q)\<^sup>T\<^sup>+ \<le> -1"
using 1 2 by (metis forest_bot mult_left_zero mult_assoc)
hence 8: "(f \<sqinter> q)\<^sup>T\<^sup>+ * (f \<sqinter> -q)\<^sup>+ \<le> -1"
using comp_commute_below_diversity by simp
have 9: "f\<^sup>T\<^sup>+ \<le> -1"
using assms(1) acyclic_star_below_complement schroeder_5_p by force
have "((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e)\<^sup>+ = (((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T)\<^sup>\<star> * e)\<^sup>\<star> * ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T)\<^sup>+ \<squnion> (((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T)\<^sup>\<star> * e)\<^sup>+"
by (simp add: plus_sup)
also have "... = ((f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>\<star> * e)\<^sup>\<star> * ((f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+) \<squnion> ((f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>\<star> * e)\<^sup>+"
using 3 4 by simp
also have "... = ((f \<sqinter> -q)\<^sup>\<star> * e)\<^sup>\<star> * ((f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+) \<squnion> ((f \<sqinter> -q)\<^sup>\<star> * e)\<^sup>+"
using 5 by simp
also have "... = ((f \<sqinter> -q)\<^sup>\<star> * e \<squnion> 1) * ((f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+) \<squnion> (f \<sqinter> -q)\<^sup>\<star> * e"
using 6 by (metis star_left_unfold_equal sup_monoid.add_commute)
also have "... = (f \<sqinter> -q)\<^sup>\<star> * e \<squnion> (f \<sqinter> -q)\<^sup>\<star> * e * (f \<sqinter> q)\<^sup>T\<^sup>+ \<squnion> (f \<sqinter> -q)\<^sup>\<star> * e * (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+"
using comp_associative mult_left_dist_sup mult_right_dist_sup sup_assoc sup_commute by simp
also have "... = (f \<sqinter> -q)\<^sup>\<star> * e * (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>\<star> \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+"
by (metis star.circ_back_loop_fixpoint star_plus sup_monoid.add_commute mult_assoc)
also have "... \<le> f\<^sup>\<star> * e * f\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>\<star> \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+"
using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone conv_isotone order_trans inf_le1 by meson
also have "... \<le> f\<^sup>\<star> * e * f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>\<star> * (f \<sqinter> -q)\<^sup>+ \<squnion> f\<^sup>T\<^sup>+"
using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone sup_right_isotone conv_isotone order_trans inf_le1 by meson
also have "... = f\<^sup>\<star> * e * f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+ * (f \<sqinter> -q)\<^sup>+ \<squnion> (f \<sqinter> -q)\<^sup>+ \<squnion> f\<^sup>T\<^sup>+"
by (simp add: star.circ_loop_fixpoint sup_monoid.add_assoc mult_assoc)
also have "... \<le> f\<^sup>\<star> * e * f\<^sup>T\<^sup>\<star> * f\<^sup>\<star> \<squnion> (f \<sqinter> q)\<^sup>T\<^sup>+ * (f \<sqinter> -q)\<^sup>+ \<squnion> f\<^sup>+ \<squnion> f\<^sup>T\<^sup>+"
using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone sup_right_isotone order_trans inf_le1 by meson
also have "... \<le> -1"
using 7 8 9 assms(1) by simp
finally show ?thesis
by simp
qed
lemma kruskal_exchange_acyclic_inv_1:
assumes "acyclic f"
and "covector q"
shows "acyclic ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T)"
using kruskal_acyclic_inv[where e=bot] by (simp add: assms)
lemma kruskal_exchange_acyclic_inv_2:
assumes "acyclic w"
and "injective w"
and "d \<le> w"
and "bijective (d\<^sup>T * top)"
and "bijective (e * top)"
and "d \<le> top * e\<^sup>T * w\<^sup>T\<^sup>\<star>"
and "w * e\<^sup>T * top = bot"
shows "acyclic ((w \<sqinter> -d) \<squnion> e)"
proof -
let ?v = "w \<sqinter> -d"
let ?w = "?v \<squnion> e"
have "d\<^sup>T * top \<le> w\<^sup>\<star> * e * top"
by (metis assms(6) comp_associative comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_isotone conv_dist_comp conv_involutive conv_order conv_star_commute conv_top inf.cobounded1 vector_top_closed)
hence 1: "e * top \<le> w\<^sup>T\<^sup>\<star> * d\<^sup>T * top"
by (metis assms(4,5) bijective_reverse comp_associative conv_star_commute)
have 2: "?v * d\<^sup>T * top = bot"
by (simp add: assms(2,3) kruskal_exchange_acyclic_inv_3)
have "?v * w\<^sup>T\<^sup>+ * d\<^sup>T * top \<le> w * w\<^sup>T\<^sup>+ * d\<^sup>T * top"
by (simp add: mult_left_isotone)
also have "... \<le> w\<^sup>T\<^sup>\<star> * d\<^sup>T * top"
by (metis assms(2) mult_left_isotone mult_1_left mult_assoc)
finally have "?v * w\<^sup>T\<^sup>\<star> * d\<^sup>T * top \<le> w\<^sup>T\<^sup>\<star> * d\<^sup>T * top"
using 2 by (metis bot_least comp_associative mult_right_dist_sup star.circ_back_loop_fixpoint star.circ_plus_same sup_least)
hence 3: "?v\<^sup>\<star> * e * top \<le> w\<^sup>T\<^sup>\<star> * d\<^sup>T * top"
using 1 by (simp add: comp_associative star_left_induct sup_least)
have "d * e\<^sup>T \<le> bot"
by (metis assms(3,7) conv_bot conv_dist_comp conv_involutive conv_top order.trans inf.absorb2 inf.cobounded2 inf_commute le_bot p_antitone_iff p_top schroeder_4_p top_left_mult_increasing)
hence 4: "e\<^sup>T * top \<le> -(d\<^sup>T * top)"
by (metis (no_types) comp_associative inf.cobounded2 le_bot p_antitone_iff schroeder_3_p semiring.mult_zero_left)
have "?v\<^sup>T * -(d\<^sup>T * top) \<le> -(d\<^sup>T * top)"
using schroeder_3_p mult_assoc 2 by simp
hence "?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top \<le> -(d\<^sup>T * top)"
using 4 by (simp add: comp_associative star_left_induct sup_least)
hence 5: "d\<^sup>T * top \<le> -(?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top)"
by (simp add: p_antitone_iff)
have "w * ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top = w * e\<^sup>T * top \<squnion> w * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top"
by (metis star_left_unfold_equal mult_right_dist_sup mult_left_dist_sup mult_1_right mult_assoc)
also have "... = w * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top"
using assms(7) by simp
also have "... \<le> w * w\<^sup>T * ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
by (simp add: comp_associative conv_isotone mult_left_isotone mult_right_isotone)
also have "... \<le> ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
by (metis assms(2) mult_1_left mult_left_isotone)
finally have "w * ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top \<le> --(?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top)"
by (simp add: p_antitone p_antitone_iff)
hence "w\<^sup>T * -(?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top) \<le> -(?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top)"
using comp_associative schroeder_3_p by simp
hence 6: "w\<^sup>T\<^sup>\<star> * d\<^sup>T * top \<le> -(?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top)"
using 5 by (simp add: comp_associative star_left_induct sup_least)
have "e * ?v\<^sup>\<star> * e \<le> e * ?v\<^sup>\<star> * e * top"
by (simp add: top_right_mult_increasing)
also have "... \<le> e * w\<^sup>T\<^sup>\<star> * d\<^sup>T * top"
using 3 by (simp add: comp_associative mult_right_isotone)
also have "... \<le> e * -(?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top)"
using 6 by (simp add: comp_associative mult_right_isotone)
also have "... \<le> bot"
by (metis conv_complement_sub_leq conv_dist_comp conv_involutive conv_star_commute le_bot mult_right_sub_dist_sup_right p_bot regular_closed_bot star.circ_back_loop_fixpoint)
finally have 7: "e * ?v\<^sup>\<star> * e = bot"
- by (simp add: antisym)
+ by (simp add: order.antisym)
hence "?v\<^sup>\<star> * e \<le> -1"
by (metis bot_least comp_associative comp_commute_below_diversity ex231d order_lesseq_imp semiring.mult_zero_left star.circ_left_top)
hence 8: "?v\<^sup>\<star> * e * ?v\<^sup>\<star> \<le> -1"
by (metis comp_associative comp_commute_below_diversity star.circ_transitive_equal)
have "1 \<sqinter> ?w\<^sup>+ = 1 \<sqinter> ?w * ?v\<^sup>\<star> * (e * ?v\<^sup>\<star>)\<^sup>\<star>"
by (simp add: star_sup_1 mult_assoc)
also have "... = 1 \<sqinter> ?w * ?v\<^sup>\<star> * (e * ?v\<^sup>\<star> \<squnion> 1)"
using 7 by (metis star.circ_mult_1 star_absorb sup_monoid.add_commute mult_assoc)
also have "... = 1 \<sqinter> (?v\<^sup>+ * e * ?v\<^sup>\<star> \<squnion> ?v\<^sup>+ \<squnion> e * ?v\<^sup>\<star> * e * ?v\<^sup>\<star> \<squnion> e * ?v\<^sup>\<star>)"
by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup sup_assoc sup_commute sup_left_commute)
also have "... = 1 \<sqinter> (?v\<^sup>+ * e * ?v\<^sup>\<star> \<squnion> ?v\<^sup>+ \<squnion> e * ?v\<^sup>\<star>)"
using 7 by simp
also have "... = 1 \<sqinter> (?v\<^sup>\<star> * e * ?v\<^sup>\<star> \<squnion> ?v\<^sup>+)"
by (metis (mono_tags, hide_lams) comp_associative star.circ_loop_fixpoint sup_assoc sup_commute)
also have "... \<le> 1 \<sqinter> (?v\<^sup>\<star> * e * ?v\<^sup>\<star> \<squnion> w\<^sup>+)"
using comp_inf.mult_right_isotone comp_isotone semiring.add_right_mono star_isotone sup_commute by simp
also have "... = (1 \<sqinter> ?v\<^sup>\<star> * e * ?v\<^sup>\<star>) \<squnion> (1 \<sqinter> w\<^sup>+)"
by (simp add: inf_sup_distrib1)
also have "... = 1 \<sqinter> ?v\<^sup>\<star> * e * ?v\<^sup>\<star>"
by (metis assms(1) inf_commute pseudo_complement sup_bot_right)
also have "... = bot"
using 8 p_antitone_iff pseudo_complement by simp
finally show ?thesis
using le_bot p_antitone_iff pseudo_complement by auto
qed
subsubsection \<open>Exchange gives Spanning Trees\<close>
text \<open>
The lemmas in this section are used to show that the relation after exchange represents a spanning tree.
\<close>
lemma inf_star_import:
assumes "x \<le> z"
and "univalent z"
and "reflexive y"
and "regular z"
shows "x\<^sup>\<star> * y \<sqinter> z\<^sup>\<star> \<le> x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>)"
proof -
have 1: "y \<le> x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>) \<squnion> -z\<^sup>\<star>"
by (metis assms(4) pp_dist_star shunting_var_p star.circ_loop_fixpoint sup.cobounded2)
have "x * -z\<^sup>\<star> \<sqinter> z\<^sup>+ \<le> x * (-z\<^sup>\<star> \<sqinter> x\<^sup>T * z\<^sup>+)"
by (simp add: dedekind_1)
also have "... \<le> x * (-z\<^sup>\<star> \<sqinter> z\<^sup>T * z\<^sup>+)"
using assms(1) comp_inf.mult_right_isotone conv_isotone mult_left_isotone mult_right_isotone by simp
also have "... \<le> x * (-z\<^sup>\<star> \<sqinter> 1 * z\<^sup>\<star>)"
by (metis assms(2) comp_associative comp_inf.mult_right_isotone mult_left_isotone mult_right_isotone)
finally have 2: "x * -z\<^sup>\<star> \<sqinter> z\<^sup>+ = bot"
- by (simp add: antisym)
+ by (simp add: order.antisym)
have "x * -z\<^sup>\<star> \<sqinter> z\<^sup>\<star> = (x * -z\<^sup>\<star> \<sqinter> z\<^sup>+) \<squnion> (x * -z\<^sup>\<star> \<sqinter> 1)"
by (metis comp_inf.semiring.distrib_left star_left_unfold_equal sup_commute)
also have "... \<le> x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>)"
using 2 by (simp add: assms(3) inf.coboundedI2 reflexive_mult_closed star.circ_reflexive)
finally have "x * -z\<^sup>\<star> \<le> x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>) \<squnion> -z\<^sup>\<star>"
by (metis assms(4) pp_dist_star shunting_var_p)
hence "x * (x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>) \<squnion> -z\<^sup>\<star>) \<le> x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>) \<squnion> -z\<^sup>\<star>"
by (metis le_supE le_supI mult_left_dist_sup star.circ_loop_fixpoint sup.cobounded1)
hence "x\<^sup>\<star> * y \<le> x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>) \<squnion> -z\<^sup>\<star>"
using 1 by (simp add: star_left_induct)
hence "x\<^sup>\<star> * y \<sqinter> --z\<^sup>\<star> \<le> x\<^sup>\<star> * (y \<sqinter> z\<^sup>\<star>)"
using shunting_var_p by simp
thus ?thesis
using order.trans inf.sup_right_isotone pp_increasing by blast
qed
lemma kruskal_exchange_forest_components_inv:
assumes "injective ((w \<sqinter> -d) \<squnion> e)"
and "regular d"
and "e * top * e = e"
and "d \<le> top * e\<^sup>T * w\<^sup>T\<^sup>\<star>"
and "w * e\<^sup>T * top = bot"
and "injective w"
and "d \<le> w"
and "d \<le> (w \<sqinter> -d)\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
shows "forest_components w \<le> forest_components ((w \<sqinter> -d) \<squnion> e)"
proof -
let ?v = "w \<sqinter> -d"
let ?w = "?v \<squnion> e"
let ?f = "forest_components ?w"
have 1: "?v * d\<^sup>T * top = bot"
by (simp add: assms(6,7) kruskal_exchange_acyclic_inv_3)
have 2: "d * e\<^sup>T \<le> bot"
by (metis assms(5,7) conv_bot conv_dist_comp conv_involutive conv_top order.trans inf.absorb2 inf.cobounded2 inf_commute le_bot p_antitone_iff p_top schroeder_4_p top_left_mult_increasing)
have "w\<^sup>\<star> * e\<^sup>T * top = e\<^sup>T * top"
by (metis assms(5) conv_bot conv_dist_comp conv_involutive conv_star_commute star.circ_top star_absorb)
hence "w\<^sup>\<star> * e\<^sup>T * top \<le> -(d\<^sup>T * top)"
using 2 by (metis (no_types) comp_associative inf.cobounded2 le_bot p_antitone_iff schroeder_3_p semiring.mult_zero_left)
hence 3: "e\<^sup>T * top \<le> -(w\<^sup>T\<^sup>\<star> * d\<^sup>T * top)"
by (metis conv_star_commute p_antitone_iff schroeder_3_p mult_assoc)
have "?v * w\<^sup>T\<^sup>\<star> * d\<^sup>T * top = ?v * d\<^sup>T * top \<squnion> ?v * w\<^sup>T\<^sup>+ * d\<^sup>T * top"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute)
also have "... \<le> w * w\<^sup>T\<^sup>+ * d\<^sup>T * top"
using 1 by (simp add: mult_left_isotone)
also have "... \<le> w\<^sup>T\<^sup>\<star> * d\<^sup>T * top"
by (metis assms(6) mult_assoc mult_1_left mult_left_isotone)
finally have "?v * w\<^sup>T\<^sup>\<star> * d\<^sup>T * top \<le> --(w\<^sup>T\<^sup>\<star> * d\<^sup>T * top)"
using p_antitone p_antitone_iff by auto
hence 4: "?v\<^sup>T * -(w\<^sup>T\<^sup>\<star> * d\<^sup>T * top) \<le> -(w\<^sup>T\<^sup>\<star> * d\<^sup>T * top)"
using comp_associative schroeder_3_p by simp
have 5: "injective ?v"
using assms(1) conv_dist_sup mult_left_dist_sup mult_right_dist_sup by simp
have "?v * ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top = ?v * e\<^sup>T * top \<squnion> ?v * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute)
also have "... \<le> w * e\<^sup>T * top \<squnion> ?v * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top"
using mult_left_isotone sup_left_isotone by simp
also have "... \<le> w * e\<^sup>T * top \<squnion> ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
using 5 by (metis mult_assoc mult_1_left mult_left_isotone sup_right_isotone)
finally have "?v * ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top \<le> ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
by (simp add: assms(5))
hence "?v\<^sup>\<star> * d * top \<le> ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
by (metis assms(8) star_left_induct sup_least comp_associative mult_right_sub_dist_sup_right sup.orderE vector_top_closed)
also have "... \<le> -(w\<^sup>T\<^sup>\<star> * d\<^sup>T * top)"
using 3 4 by (simp add: comp_associative star_left_induct)
also have "... \<le> -(d\<^sup>T * top)"
by (metis p_antitone star.circ_left_top star_outer_increasing mult_assoc)
finally have 6: "?v\<^sup>\<star> * d * top \<le> -(d\<^sup>T * top)"
by simp
have "d\<^sup>T * top \<le> w\<^sup>\<star> * e * top"
by (metis assms(4) comp_associative comp_inf.star.circ_sup_2 comp_isotone conv_dist_comp conv_involutive conv_order conv_star_commute conv_top vector_top_closed)
also have "... \<le> (?v \<squnion> d)\<^sup>\<star> * e * top"
by (metis assms(2) comp_inf.semiring.distrib_left maddux_3_11_pp mult_left_isotone star_isotone sup.cobounded2 sup_commute sup_inf_distrib1)
also have "... = ?v\<^sup>\<star> * (d * ?v\<^sup>\<star>)\<^sup>\<star> * e * top"
by (simp add: star_sup_1)
also have "... = ?v\<^sup>\<star> * e * top \<squnion> ?v\<^sup>\<star> * d * ?v\<^sup>\<star> * (d * ?v\<^sup>\<star>)\<^sup>\<star> * e * top"
by (metis semiring.distrib_right star.circ_unfold_sum star_decompose_1 star_decompose_3 mult_assoc)
also have "... \<le> ?v\<^sup>\<star> * e * top \<squnion> ?v\<^sup>\<star> * d * top"
by (metis comp_associative comp_isotone le_supI mult_left_dist_sup mult_right_dist_sup mult_right_isotone star.circ_decompose_5 star_decompose_3 sup.cobounded1 sup_commute top.extremum)
finally have "d\<^sup>T * top \<le> ?v\<^sup>\<star> * e * top \<squnion> (d\<^sup>T * top \<sqinter> ?v\<^sup>\<star> * d * top)"
using sup_inf_distrib2 sup_monoid.add_commute by simp
hence "d\<^sup>T * top \<le> ?v\<^sup>\<star> * e * top"
using 6 by (metis inf_commute pseudo_complement sup_monoid.add_0_right)
hence 7: "d \<le> top * e\<^sup>T * ?v\<^sup>T\<^sup>\<star>"
by (metis comp_associative conv_dist_comp conv_involutive conv_isotone conv_star_commute conv_top order.trans top_right_mult_increasing)
have 8: "?v \<le> ?f"
using forest_components_increasing le_supE by blast
have "d \<le> ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top \<sqinter> top * e\<^sup>T * ?v\<^sup>T\<^sup>\<star>"
using 7 assms(8) by simp
also have "... = ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * top * e\<^sup>T * ?v\<^sup>T\<^sup>\<star>"
by (metis inf_top_right vector_inf_comp vector_top_closed mult_assoc)
also have "... = ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * ?v\<^sup>T\<^sup>\<star>"
by (metis assms(3) comp_associative conv_dist_comp conv_top)
also have "... \<le> ?v\<^sup>T\<^sup>\<star> * e\<^sup>T * ?f"
using 8 by (metis assms(1) forest_components_equivalence cancel_separate_1 conv_dist_comp conv_order mult_left_isotone star_involutive star_isotone)
also have "... \<le> ?v\<^sup>T\<^sup>\<star> * ?f * ?f"
by (metis assms(1) forest_components_equivalence forest_components_increasing conv_isotone le_supE mult_left_isotone mult_right_isotone)
also have "... \<le> ?f * ?f * ?f"
by (metis comp_associative comp_isotone conv_dist_sup star.circ_loop_fixpoint star_isotone sup.cobounded1 sup.cobounded2)
also have "... = ?f"
by (simp add: assms(1) forest_components_equivalence preorder_idempotent)
finally have "w \<le> ?f"
using 8 by (metis assms(2) shunting_var_p sup.orderE)
thus ?thesis
using assms(1) forest_components_idempotent forest_components_isotone by fastforce
qed
lemma kruskal_spanning_inv:
assumes "injective ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e)"
and "regular q"
and "regular e"
and "(-h \<sqinter> --g)\<^sup>\<star> \<le> forest_components f"
shows "components (-(h \<sqinter> -e \<sqinter> -e\<^sup>T) \<sqinter> g) \<le> forest_components ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e)"
proof -
let ?f = "(f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e"
let ?h = "h \<sqinter> -e \<sqinter> -e\<^sup>T"
let ?F = "forest_components f"
let ?FF = "forest_components ?f"
have 1: "equivalence ?FF"
using assms(1) forest_components_equivalence by simp
hence 2: "?f * ?FF \<le> ?FF"
using order.trans forest_components_increasing mult_left_isotone by blast
have 3: "?f\<^sup>T * ?FF \<le> ?FF"
using 1 by (metis forest_components_increasing mult_left_isotone conv_isotone preorder_idempotent)
have "(f \<sqinter> q) * ?FF \<le> ?f\<^sup>T * ?FF"
using conv_dist_sup conv_involutive sup_assoc sup_left_commute mult_left_isotone by simp
hence 4: "(f \<sqinter> q) * ?FF \<le> ?FF"
using 3 order.trans by blast
have "(f \<sqinter> -q) * ?FF \<le> ?f * ?FF"
using le_supI1 mult_left_isotone by simp
hence "(f \<sqinter> -q) * ?FF \<le> ?FF"
using 2 order.trans by blast
hence "((f \<sqinter> q) \<squnion> (f \<sqinter> -q)) * ?FF \<le> ?FF"
using 4 mult_right_dist_sup by simp
hence "f * ?FF \<le> ?FF"
by (metis assms(2) maddux_3_11_pp)
hence 5: "f\<^sup>\<star> * ?FF \<le> ?FF"
using star_left_induct_mult_iff by simp
have "(f \<sqinter> -q)\<^sup>T * ?FF \<le> ?f\<^sup>T * ?FF"
by (meson conv_isotone order.trans mult_left_isotone sup.cobounded1)
hence 6: "(f \<sqinter> -q)\<^sup>T * ?FF \<le> ?FF"
using 3 order.trans by blast
have "(f \<sqinter> q)\<^sup>T * ?FF \<le> ?f * ?FF"
by (simp add: mult_left_isotone sup.left_commute sup_assoc)
hence "(f \<sqinter> q)\<^sup>T * ?FF \<le> ?FF"
using 2 order.trans by blast
hence "((f \<sqinter> -q)\<^sup>T \<squnion> (f \<sqinter> q)\<^sup>T) * ?FF \<le> ?FF"
using 6 mult_right_dist_sup by simp
hence "f\<^sup>T * ?FF \<le> ?FF"
by (metis assms(2) conv_dist_sup maddux_3_11_pp)
hence 7: "?F * ?FF \<le> ?FF"
using 5 star_left_induct mult_assoc by simp
have 8: "e * ?FF \<le> ?FF"
using 2 by (simp add: mult_right_dist_sup mult_left_isotone)
have "e\<^sup>T * ?FF \<le> ?f\<^sup>T * ?FF"
by (simp add: mult_left_isotone conv_isotone)
also have "... \<le> ?FF * ?FF"
using 1 by (metis forest_components_increasing mult_left_isotone conv_isotone)
finally have "e\<^sup>T * ?FF \<le> ?FF"
using 1 preorder_idempotent by auto
hence 9: "(?F \<squnion> e \<squnion> e\<^sup>T) * ?FF \<le> ?FF"
using 7 8 mult_right_dist_sup by simp
have "components (-?h \<sqinter> g) \<le> ((-h \<sqinter> --g) \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
by (metis assms(3) comp_inf.mult_left_sub_dist_sup_left conv_complement p_dist_inf pp_dist_inf regular_closed_p star_isotone sup_inf_distrib2 sup_monoid.add_assoc)
also have "... \<le> ((-h \<sqinter> --g)\<^sup>\<star> \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
using star.circ_increasing star_isotone sup_left_isotone by simp
also have "... \<le> (?F \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
using assms(4) sup_left_isotone star_isotone by simp
also have "... \<le> ?FF"
using 1 9 star_left_induct by force
finally show ?thesis
by simp
qed
lemma kruskal_exchange_spanning_inv_1:
assumes "injective ((w \<sqinter> -q) \<squnion> (w \<sqinter> q)\<^sup>T)"
and "regular (w \<sqinter> q)"
and "components g \<le> forest_components w"
shows "components g \<le> forest_components ((w \<sqinter> -q) \<squnion> (w \<sqinter> q)\<^sup>T)"
proof -
let ?p = "w \<sqinter> q"
let ?w = "(w \<sqinter> -q) \<squnion> ?p\<^sup>T"
have 1: "w \<sqinter> -?p \<le> forest_components ?w"
by (metis forest_components_increasing inf_import_p le_supE)
have "w \<sqinter> ?p \<le> ?w\<^sup>T"
by (simp add: conv_dist_sup)
also have "... \<le> forest_components ?w"
by (metis assms(1) conv_isotone forest_components_equivalence forest_components_increasing)
finally have "w \<sqinter> (?p \<squnion> -?p) \<le> forest_components ?w"
using 1 inf_sup_distrib1 by simp
hence "w \<le> forest_components ?w"
by (metis assms(2) inf_top_right stone)
hence 2: "w\<^sup>\<star> \<le> forest_components ?w"
using assms(1) star_isotone forest_components_star by force
hence 3: "w\<^sup>T\<^sup>\<star> \<le> forest_components ?w"
using assms(1) conv_isotone conv_star_commute forest_components_equivalence by force
have "components g \<le> forest_components w"
using assms(3) by simp
also have "... \<le> forest_components ?w * forest_components ?w"
using 2 3 mult_isotone by simp
also have "... = forest_components ?w"
using assms(1) forest_components_equivalence preorder_idempotent by simp
finally show ?thesis
by simp
qed
lemma kruskal_exchange_spanning_inv_2:
assumes "injective w"
and "w\<^sup>\<star> * e\<^sup>T = e\<^sup>T"
and "f \<squnion> f\<^sup>T \<le> (w \<sqinter> -d \<sqinter> -d\<^sup>T) \<squnion> (w\<^sup>T \<sqinter> -d \<sqinter> -d\<^sup>T)"
and "d \<le> forest_components f * e\<^sup>T * top"
shows "d \<le> (w \<sqinter> -d)\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
proof -
have 1: "(w \<sqinter> -d \<sqinter> -d\<^sup>T) * (w\<^sup>T \<sqinter> -d \<sqinter> -d\<^sup>T) \<le> 1"
using assms(1) comp_isotone order.trans inf.cobounded1 by blast
have "d \<le> forest_components f * e\<^sup>T * top"
using assms(4) by simp
also have "... \<le> (f \<squnion> f\<^sup>T)\<^sup>\<star> * (f \<squnion> f\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
by (simp add: comp_isotone star_isotone)
also have "... = (f \<squnion> f\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
by (simp add: star.circ_transitive_equal)
also have "... \<le> ((w \<sqinter> -d \<sqinter> -d\<^sup>T) \<squnion> (w\<^sup>T \<sqinter> -d \<sqinter> -d\<^sup>T))\<^sup>\<star> * e\<^sup>T * top"
using assms(3) by (simp add: comp_isotone star_isotone)
also have "... = (w\<^sup>T \<sqinter> -d \<sqinter> -d\<^sup>T)\<^sup>\<star> * (w \<sqinter> -d \<sqinter> -d\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
using 1 cancel_separate_1 by simp
also have "... \<le> (w\<^sup>T \<sqinter> -d \<sqinter> -d\<^sup>T)\<^sup>\<star> * w\<^sup>\<star> * e\<^sup>T * top"
by (simp add: inf_assoc mult_left_isotone mult_right_isotone star_isotone)
also have "... = (w\<^sup>T \<sqinter> -d \<sqinter> -d\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
using assms(2) mult_assoc by simp
also have "... \<le> (w\<^sup>T \<sqinter> -d\<^sup>T)\<^sup>\<star> * e\<^sup>T * top"
using mult_left_isotone conv_isotone star_isotone comp_inf.mult_right_isotone inf.cobounded2 inf.left_commute inf.sup_monoid.add_commute by presburger
also have "... = (w \<sqinter> -d)\<^sup>T\<^sup>\<star> * e\<^sup>T * top"
using conv_complement conv_dist_inf by presburger
finally show ?thesis
by simp
qed
lemma kruskal_spanning_inv_1:
assumes "e \<le> F"
and "regular e"
and "components (-h \<sqinter> g) \<le> F"
and "equivalence F"
shows "components (-(h \<sqinter> -e \<sqinter> -e\<^sup>T) \<sqinter> g) \<le> F"
proof -
have 1: "F * F \<le> F"
using assms(4) by simp
hence 2: "e * F \<le> F"
using assms(1) mult_left_isotone order_lesseq_imp by blast
have "e\<^sup>T * F \<le> F"
by (metis assms(1,4) conv_isotone mult_left_isotone preorder_idempotent)
hence 3: "(F \<squnion> e \<squnion> e\<^sup>T) * F \<le> F"
using 1 2 mult_right_dist_sup by simp
have "components (-(h \<sqinter> -e \<sqinter> -e\<^sup>T) \<sqinter> g) \<le> ((-h \<sqinter> --g) \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
by (metis assms(2) comp_inf.mult_left_sub_dist_sup_left conv_complement p_dist_inf pp_dist_inf regular_closed_p star_isotone sup_inf_distrib2 sup_monoid.add_assoc)
also have "... \<le> ((-h \<sqinter> --g)\<^sup>\<star> \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
using sup_left_isotone star.circ_increasing star_isotone by simp
also have "... \<le> (F \<squnion> e \<squnion> e\<^sup>T)\<^sup>\<star>"
using assms(3) sup_left_isotone star_isotone by simp
also have "... \<le> F"
using 3 assms(4) star_left_induct by force
finally show ?thesis
by simp
qed
lemma kruskal_reroot_edge:
assumes "injective (e\<^sup>T * top)"
and "acyclic w"
shows "((w \<sqinter> -(top * e * w\<^sup>T\<^sup>\<star>)) \<squnion> (w \<sqinter> top * e * w\<^sup>T\<^sup>\<star>)\<^sup>T) * e\<^sup>T = bot"
proof -
let ?q = "top * e * w\<^sup>T\<^sup>\<star>"
let ?p = "w \<sqinter> ?q"
let ?w = "(w \<sqinter> -?q) \<squnion> ?p\<^sup>T"
have "(w \<sqinter> -?q) * e\<^sup>T * top = w * (e\<^sup>T * top \<sqinter> -?q\<^sup>T)"
by (metis comp_associative comp_inf_vector_1 conv_complement covector_complement_closed vector_top_closed)
also have "... = w * (e\<^sup>T * top \<sqinter> -(w\<^sup>\<star> * e\<^sup>T * top))"
by (simp add: conv_dist_comp conv_star_commute mult_assoc)
also have "... = bot"
by (metis comp_associative comp_inf.semiring.mult_not_zero inf.sup_relative_same_increasing inf_p mult_right_zero star.circ_loop_fixpoint sup_commute sup_left_divisibility)
finally have 1: "(w \<sqinter> -?q) * e\<^sup>T * top = bot"
by simp
have "?p\<^sup>T * e\<^sup>T * top = (w\<^sup>T \<sqinter> w\<^sup>\<star> * e\<^sup>T * top) * e\<^sup>T * top"
by (simp add: conv_dist_comp conv_star_commute mult_assoc conv_dist_inf)
also have "... = w\<^sup>\<star> * e\<^sup>T * top \<sqinter> w\<^sup>T * e\<^sup>T * top"
by (simp add: inf_vector_comp vector_export_comp)
also have "... = (w\<^sup>\<star> \<sqinter> w\<^sup>T) * e\<^sup>T * top"
using assms(1) injective_comp_right_dist_inf mult_assoc by simp
also have "... = bot"
using assms(2) acyclic_star_below_complement_1 semiring.mult_not_zero by blast
finally have "?w * e\<^sup>T * top = bot"
using 1 mult_right_dist_sup by simp
thus ?thesis
by (metis star.circ_top star_absorb)
qed
subsubsection \<open>Exchange gives Minimum Spanning Trees\<close>
text \<open>
The lemmas in this section are used to show that the after exchange we obtain a minimum spanning tree.
The following lemmas show that the relation characterising the edge across the cut is an arc.
\<close>
lemma kruskal_edge_arc:
assumes "equivalence F"
and "forest w"
and "arc e"
and "regular F"
and "F \<le> forest_components (F \<sqinter> w)"
and "regular w"
and "w * e\<^sup>T = bot"
and "e * F * e = bot"
and "e\<^sup>T \<le> w\<^sup>\<star>"
shows "arc (w \<sqinter> top * e\<^sup>T * w\<^sup>T\<^sup>\<star> \<sqinter> F * e\<^sup>T * top \<sqinter> top * e * -F)"
proof (unfold arc_expanded, intro conjI)
let ?E = "top * e\<^sup>T * w\<^sup>T\<^sup>\<star>"
let ?F = "F * e\<^sup>T * top"
let ?G = "top * e * -F"
let ?FF = "F * e\<^sup>T * e * F"
let ?GG = "-F * e\<^sup>T * e * -F"
let ?w = "forest_components (F \<sqinter> w)"
have "F \<sqinter> w\<^sup>T\<^sup>\<star> \<le> forest_components (F \<sqinter> w) \<sqinter> w\<^sup>T\<^sup>\<star>"
by (simp add: assms(5) inf.coboundedI1)
also have "... \<le> (F \<sqinter> w)\<^sup>T\<^sup>\<star> * ((F \<sqinter> w)\<^sup>\<star> \<sqinter> w\<^sup>T\<^sup>\<star>)"
apply (rule inf_star_import)
apply (simp add: conv_isotone)
apply (simp add: assms(2))
apply (simp add: star.circ_reflexive)
by (metis assms(6) conv_complement)
also have "... \<le> (F \<sqinter> w)\<^sup>T\<^sup>\<star> * (w\<^sup>\<star> \<sqinter> w\<^sup>T\<^sup>\<star>)"
using comp_inf.mult_left_isotone mult_right_isotone star_isotone by simp
also have "... = (F \<sqinter> w)\<^sup>T\<^sup>\<star>"
by (simp add: assms(2) acyclic_star_inf_conv)
finally have "w * (F \<sqinter> w\<^sup>T\<^sup>\<star>) * e\<^sup>T * e \<le> w * (F \<sqinter> w)\<^sup>T\<^sup>\<star> * e\<^sup>T * e"
by (simp add: mult_left_isotone mult_right_isotone)
also have "... = w * e\<^sup>T * e \<squnion> w * (F \<sqinter> w)\<^sup>T\<^sup>+ * e\<^sup>T * e"
by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute)
also have "... = w * (F \<sqinter> w)\<^sup>T\<^sup>+ * e\<^sup>T * e"
by (simp add: assms(7))
also have "... \<le> w * (F \<sqinter> w)\<^sup>T\<^sup>+"
by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone)
also have "... \<le> w * w\<^sup>T * (F \<sqinter> w)\<^sup>T\<^sup>\<star>"
by (simp add: comp_associative conv_isotone mult_left_isotone mult_right_isotone)
also have "... \<le> (F \<sqinter> w)\<^sup>T\<^sup>\<star>"
using assms(2) coreflexive_comp_top_inf inf.sup_right_divisibility by auto
also have "... \<le> F\<^sup>T\<^sup>\<star>"
by (simp add: conv_dist_inf star_isotone)
finally have 1: "w * (F \<sqinter> w\<^sup>T\<^sup>\<star>) * e\<^sup>T * e \<le> F"
- by (metis assms(1) antisym mult_1_left mult_left_isotone star.circ_plus_same star.circ_reflexive star.left_plus_below_circ star_left_induct_mult_iff)
+ by (metis assms(1) order.antisym mult_1_left mult_left_isotone star.circ_plus_same star.circ_reflexive star.left_plus_below_circ star_left_induct_mult_iff)
have "F * e\<^sup>T * e \<le> forest_components (F \<sqinter> w) * e\<^sup>T * e"
by (simp add: assms(5) mult_left_isotone)
also have "... \<le> forest_components w * e\<^sup>T * e"
by (simp add: comp_isotone conv_dist_inf star_isotone)
also have "... = w\<^sup>T\<^sup>\<star> * e\<^sup>T * e"
by (metis (no_types) assms(7) comp_associative conv_bot conv_dist_comp conv_involutive conv_star_commute star_absorb)
also have "... \<le> w\<^sup>T\<^sup>\<star>"
by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone)
finally have 2: "F * e\<^sup>T * e \<le> w\<^sup>T\<^sup>\<star>"
by simp
have "w * F * e\<^sup>T * e \<le> w * F * e\<^sup>T * e * e\<^sup>T * e"
using comp_associative ex231c mult_right_isotone by simp
also have "... = w * (F * e\<^sup>T * e \<sqinter> w\<^sup>T\<^sup>\<star>) * e\<^sup>T * e"
using 2 by (simp add: comp_associative inf.absorb1)
also have "... \<le> w * (F \<sqinter> w\<^sup>T\<^sup>\<star>) * e\<^sup>T * e"
by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone mult_left_isotone inf.sup_left_isotone)
also have "... \<le> F"
using 1 by simp
finally have 3: "w * F * e\<^sup>T * e \<le> F"
by simp
hence "e\<^sup>T * e * F * w\<^sup>T \<le> F"
by (metis assms(1) conv_dist_comp conv_dist_inf conv_involutive inf.absorb_iff1 mult_assoc)
hence "e\<^sup>T * e * F * w\<^sup>T \<le> e\<^sup>T * top \<sqinter> F"
by (simp add: comp_associative mult_right_isotone)
also have "... \<le> e\<^sup>T * e * F"
by (metis conv_involutive dedekind_1 inf_top_left mult_assoc)
finally have 4: "e\<^sup>T * e * F * w\<^sup>T \<le> e\<^sup>T * e * F"
by simp
have "(top * e)\<^sup>T * (?F \<sqinter> w\<^sup>T\<^sup>\<star>) = e\<^sup>T * top * e * F * w\<^sup>T\<^sup>\<star>"
by (metis assms(1) comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb conv_dist_comp conv_involutive conv_top covector_inf_comp_3 vector_top_closed mult_assoc)
also have "... = e\<^sup>T * e * F * w\<^sup>T\<^sup>\<star>"
by (simp add: assms(3) arc_top_edge)
also have "... \<le> e\<^sup>T * e * F"
using 4 star_right_induct_mult by simp
also have "... \<le> F"
by (metis assms(3) arc_injective conv_involutive mult_1_left mult_left_isotone)
finally have 5: "(top * e)\<^sup>T * (?F \<sqinter> w\<^sup>T\<^sup>\<star>) \<le> F"
by simp
have "(?F \<sqinter> w) * w\<^sup>T\<^sup>+ = ?F \<sqinter> w * w\<^sup>T\<^sup>+"
by (simp add: vector_export_comp)
also have "... \<le> ?F \<sqinter> w\<^sup>T\<^sup>\<star>"
by (metis assms(2) comp_associative inf.sup_right_isotone mult_left_isotone star.circ_transitive_equal star_left_unfold_equal sup.absorb_iff2 sup_monoid.add_assoc)
also have 6: "... \<le> top * e * F"
using 5 by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top)
finally have 7: "(?F \<sqinter> w) * w\<^sup>T\<^sup>+ \<le> top * e * F"
by simp
have "e\<^sup>T * top * e \<le> 1"
by (simp add: assms(3) point_injective)
also have "... \<le> F"
by (simp add: assms(1))
finally have 8: "e * -F * e\<^sup>T \<le> bot"
by (metis p_antitone p_antitone_iff p_bot regular_closed_bot schroeder_3_p schroeder_4_p mult_assoc)
have "?FF \<sqinter> w * (w\<^sup>T\<^sup>+ \<sqinter> ?GG) * w\<^sup>T \<le> ?F \<sqinter> w * (w\<^sup>T\<^sup>+ \<sqinter> ?GG) * w\<^sup>T"
using comp_inf.mult_left_isotone mult_isotone mult_assoc by simp
also have "... \<le> ?F \<sqinter> w * (w\<^sup>T\<^sup>+ \<sqinter> ?G) * w\<^sup>T"
by (metis assms(3) arc_top_edge comp_inf.star.circ_decompose_9 comp_inf_covector inf.sup_right_isotone inf_le2 mult_left_isotone mult_right_isotone vector_top_closed mult_assoc)
also have "... = (?F \<sqinter> w) * (w\<^sup>T\<^sup>+ \<sqinter> ?G) * w\<^sup>T"
by (simp add: vector_export_comp)
also have "... = (?F \<sqinter> w) * w\<^sup>T\<^sup>+ * (?G\<^sup>T \<sqinter> w\<^sup>T)"
by (simp add: covector_comp_inf covector_comp_inf_1 covector_mult_closed)
also have "... \<le> top * e * F * (?G\<^sup>T \<sqinter> w\<^sup>T)"
using 7 mult_left_isotone by simp
also have "... \<le> top * e * F * ?G\<^sup>T"
by (simp add: mult_right_isotone)
also have "... = top * e * -F * e\<^sup>T * top"
by (metis assms(1) conv_complement conv_dist_comp conv_top equivalence_comp_left_complement mult_assoc)
finally have 9: "?FF \<sqinter> w * (w\<^sup>T\<^sup>+ \<sqinter> ?GG) * w\<^sup>T = bot"
using 8 by (metis comp_associative covector_bot_closed le_bot vector_bot_closed)
hence 10: "?FF \<sqinter> w * (w\<^sup>+ \<sqinter> ?GG) * w\<^sup>T = bot"
using assms(1) comp_associative conv_bot conv_complement conv_dist_comp conv_dist_inf conv_star_commute star.circ_plus_same by fastforce
have "(w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G) * top * (w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G)\<^sup>T = (?F \<sqinter> (w \<sqinter> ?E \<sqinter> ?G)) * top * ((w \<sqinter> ?E \<sqinter> ?G)\<^sup>T \<sqinter> ?F\<^sup>T)"
by (simp add: conv_dist_inf inf_commute inf_left_commute)
also have "... = (?F \<sqinter> (w \<sqinter> ?E \<sqinter> ?G)) * top * (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T \<sqinter> ?F\<^sup>T"
using covector_comp_inf vector_conv_covector vector_mult_closed vector_top_closed by simp
also have "... = ?F \<sqinter> (w \<sqinter> ?E \<sqinter> ?G) * top * (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T \<sqinter> ?F\<^sup>T"
by (simp add: vector_export_comp)
also have "... = ?F \<sqinter> top * e * F \<sqinter> (w \<sqinter> ?E \<sqinter> ?G) * top * (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T"
by (simp add: assms(1) conv_dist_comp inf_assoc inf_commute mult_assoc)
also have "... = ?F * e * F \<sqinter> (w \<sqinter> ?E \<sqinter> ?G) * top * (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T"
by (metis comp_associative comp_inf_covector inf_top.left_neutral)
also have "... = ?FF \<sqinter> (w \<sqinter> ?E \<sqinter> ?G) * (top * (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T)"
using assms(3) arc_top_edge comp_associative by simp
also have "... = ?FF \<sqinter> (w \<sqinter> ?E \<sqinter> ?G) * (top * (?G\<^sup>T \<sqinter> (?E\<^sup>T \<sqinter> w\<^sup>T)))"
by (simp add: conv_dist_inf inf_assoc inf_commute inf_left_commute)
also have "... = ?FF \<sqinter> (w \<sqinter> ?E \<sqinter> ?G) * (?G * (?E\<^sup>T \<sqinter> w\<^sup>T))"
by (metis covector_comp_inf_1 covector_top_closed covector_mult_closed inf_top_left)
also have "... = ?FF \<sqinter> (w \<sqinter> ?E \<sqinter> ?G) * (?G \<sqinter> ?E) * w\<^sup>T"
by (metis covector_comp_inf_1 covector_top_closed mult_assoc)
also have "... = ?FF \<sqinter> (w \<sqinter> ?E) * (?G\<^sup>T \<sqinter> ?G \<sqinter> ?E) * w\<^sup>T"
by (metis covector_comp_inf_1 covector_mult_closed inf.sup_monoid.add_assoc vector_top_closed)
also have "... = ?FF \<sqinter> w * (?E\<^sup>T \<sqinter> ?G\<^sup>T \<sqinter> ?G \<sqinter> ?E) * w\<^sup>T"
by (metis covector_comp_inf_1 covector_mult_closed inf.sup_monoid.add_assoc vector_top_closed)
also have "... = ?FF \<sqinter> w * (?E\<^sup>T \<sqinter> ?E \<sqinter> (?G\<^sup>T \<sqinter> ?G)) * w\<^sup>T"
by (simp add: inf_commute inf_left_commute)
also have "... = ?FF \<sqinter> w * (?E\<^sup>T \<sqinter> ?E \<sqinter> (-F * e\<^sup>T * top \<sqinter> ?G)) * w\<^sup>T"
by (simp add: assms(1) conv_complement conv_dist_comp mult_assoc)
also have "... = ?FF \<sqinter> w * (?E\<^sup>T \<sqinter> ?E \<sqinter> (-F * e\<^sup>T * ?G)) * w\<^sup>T"
by (metis comp_associative comp_inf_covector inf_top.left_neutral)
also have "... = ?FF \<sqinter> w * (?E\<^sup>T \<sqinter> ?E \<sqinter> ?GG) * w\<^sup>T"
by (metis assms(3) arc_top_edge comp_associative)
also have "... = ?FF \<sqinter> w * (w\<^sup>\<star> * e * top \<sqinter> ?E \<sqinter> ?GG) * w\<^sup>T"
by (simp add: comp_associative conv_dist_comp conv_star_commute)
also have "... = ?FF \<sqinter> w * (w\<^sup>\<star> * e * ?E \<sqinter> ?GG) * w\<^sup>T"
by (metis comp_associative comp_inf_covector inf_top.left_neutral)
also have "... \<le> ?FF \<sqinter> w * (w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> ?GG) * w\<^sup>T"
by (metis assms(3) mult_assoc mult_1_right mult_left_isotone mult_right_isotone inf.sup_left_isotone inf.sup_right_isotone arc_expanded)
also have "... = ?FF \<sqinter> w * ((w\<^sup>+ \<squnion> 1 \<squnion> w\<^sup>T\<^sup>\<star>) \<sqinter> ?GG) * w\<^sup>T"
by (simp add: assms(2) cancel_separate_eq star_left_unfold_equal sup_monoid.add_commute)
also have "... = ?FF \<sqinter> w * ((w\<^sup>+ \<squnion> 1 \<squnion> w\<^sup>T\<^sup>+) \<sqinter> ?GG) * w\<^sup>T"
using star.circ_plus_one star_left_unfold_equal sup_assoc by presburger
also have "... = (?FF \<sqinter> w * (w\<^sup>+ \<sqinter> ?GG) * w\<^sup>T) \<squnion> (?FF \<sqinter> w * (1 \<sqinter> ?GG) * w\<^sup>T) \<squnion> (?FF \<sqinter> w * (w\<^sup>T\<^sup>+ \<sqinter> ?GG) * w\<^sup>T)"
by (simp add: inf_sup_distrib1 inf_sup_distrib2 semiring.distrib_left semiring.distrib_right)
also have "... \<le> w * (1 \<sqinter> ?GG) * w\<^sup>T"
using 9 10 by simp
also have "... \<le> w * w\<^sup>T"
by (metis inf.cobounded1 mult_1_right mult_left_isotone mult_right_isotone)
also have "... \<le> 1"
by (simp add: assms(2))
finally show "(w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G) * top * (w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G)\<^sup>T \<le> 1"
by simp
have "w\<^sup>T\<^sup>+ \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w \<le> w\<^sup>T\<^sup>+ \<sqinter> ?G \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w"
using top_greatest inf.sup_left_isotone inf.sup_right_isotone mult_left_isotone by simp
also have "... \<le> w\<^sup>T\<^sup>+ \<sqinter> ?G \<sqinter> w\<^sup>T * ?F"
using comp_associative inf.sup_right_isotone mult_right_isotone top.extremum by presburger
also have "... = w\<^sup>T * (w\<^sup>T\<^sup>\<star> \<sqinter> ?F) \<sqinter> ?G"
using assms(2) inf_assoc inf_commute inf_left_commute univalent_comp_left_dist_inf by simp
also have "... \<le> w\<^sup>T * (top * e * F) \<sqinter> ?G"
using 6 by (metis inf.sup_monoid.add_commute inf.sup_right_isotone mult_right_isotone)
also have "... \<le> top * e * F \<sqinter> ?G"
by (metis comp_associative comp_inf_covector mult_left_isotone top.extremum)
also have "... = bot"
by (metis assms(3) conv_dist_comp conv_involutive conv_top inf_p mult_right_zero univalent_comp_left_dist_inf)
finally have 11: "w\<^sup>T\<^sup>+ \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w = bot"
- by (simp add: antisym)
+ by (simp add: order.antisym)
hence 12: "w\<^sup>+ \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w = bot"
using assms(1) comp_associative conv_bot conv_complement conv_dist_comp conv_dist_inf conv_star_commute star.circ_plus_same by fastforce
have "(w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G)\<^sup>T * top * (w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G) = ((w \<sqinter> ?E \<sqinter> ?G)\<^sup>T \<sqinter> ?F\<^sup>T) * top * (?F \<sqinter> (w \<sqinter> ?E \<sqinter> ?G))"
by (simp add: conv_dist_inf inf_commute inf_left_commute)
also have "... = (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T * ?F * (?F \<sqinter> (w \<sqinter> ?E \<sqinter> ?G))"
by (simp add: covector_inf_comp_3 vector_mult_closed)
also have "... = (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * (w \<sqinter> ?E \<sqinter> ?G)"
using covector_comp_inf covector_inf_comp_3 vector_conv_covector vector_mult_closed by simp
also have "... = (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * (w \<sqinter> ?E) \<sqinter> ?G"
by (simp add: comp_associative comp_inf_covector)
also have "... = (w \<sqinter> ?E \<sqinter> ?G)\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w \<sqinter> ?E \<sqinter> ?G"
by (simp add: comp_associative comp_inf_covector)
also have "... = (?G\<^sup>T \<sqinter> (?E\<^sup>T \<sqinter> w\<^sup>T)) * (?F \<sqinter> ?F\<^sup>T) * w \<sqinter> ?E \<sqinter> ?G"
by (simp add: conv_dist_inf inf.left_commute inf.sup_monoid.add_commute)
also have "... = ?G\<^sup>T \<sqinter> (?E\<^sup>T \<sqinter> w\<^sup>T) * (?F \<sqinter> ?F\<^sup>T) * w \<sqinter> ?E \<sqinter> ?G"
by (metis (no_types) comp_associative conv_dist_comp conv_top vector_export_comp)
also have "... = ?G\<^sup>T \<sqinter> ?E\<^sup>T \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w \<sqinter> ?E \<sqinter> ?G"
by (metis (no_types) comp_associative inf_assoc conv_dist_comp conv_top vector_export_comp)
also have "... = ?E\<^sup>T \<sqinter> ?E \<sqinter> (?G\<^sup>T \<sqinter> ?G) \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w"
by (simp add: inf_assoc inf.left_commute inf.sup_monoid.add_commute)
also have "... = w\<^sup>\<star> * e * top \<sqinter> ?E \<sqinter> (?G\<^sup>T \<sqinter> ?G) \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w"
by (simp add: comp_associative conv_dist_comp conv_star_commute)
also have "... = w\<^sup>\<star> * e * ?E \<sqinter> (?G\<^sup>T \<sqinter> ?G) \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w"
by (metis comp_associative comp_inf_covector inf_top.left_neutral)
also have "... \<le> w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> (?G\<^sup>T \<sqinter> ?G) \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w"
by (metis assms(3) mult_assoc mult_1_right mult_left_isotone mult_right_isotone inf.sup_left_isotone arc_expanded)
also have "... = w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> (-F * e\<^sup>T * top \<sqinter> ?G) \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w"
by (simp add: assms(1) conv_complement conv_dist_comp mult_assoc)
also have "... = w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> -F * e\<^sup>T * ?G \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w"
by (metis comp_associative comp_inf_covector inf_top.left_neutral)
also have "... = w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * (?F \<sqinter> ?F\<^sup>T) * w"
by (metis assms(3) arc_top_edge mult_assoc)
also have "... = w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * (?F \<sqinter> top * e * F) * w"
by (simp add: assms(1) conv_dist_comp mult_assoc)
also have "... = w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * (?F * e * F) * w"
by (metis comp_associative comp_inf_covector inf_top.left_neutral)
also have "... = w\<^sup>\<star> * w\<^sup>T\<^sup>\<star> \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w"
by (metis assms(3) arc_top_edge mult_assoc)
also have "... = (w\<^sup>+ \<squnion> 1 \<squnion> w\<^sup>T\<^sup>\<star>) \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w"
by (simp add: assms(2) cancel_separate_eq star_left_unfold_equal sup_monoid.add_commute)
also have "... = (w\<^sup>+ \<squnion> 1 \<squnion> w\<^sup>T\<^sup>+) \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w"
using star.circ_plus_one star_left_unfold_equal sup_assoc by presburger
also have "... = (w\<^sup>+ \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w) \<squnion> (1 \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w) \<squnion> (w\<^sup>T\<^sup>+ \<sqinter> -F * e\<^sup>T * e * -F \<sqinter> w\<^sup>T * F * e\<^sup>T * e * F * w)"
by (simp add: inf_sup_distrib2)
also have "... \<le> 1"
using 11 12 by (simp add: inf.coboundedI1)
finally show "(w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G)\<^sup>T * top * (w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G) \<le> 1"
by simp
have "(w \<sqinter> -F) * (F \<sqinter> w\<^sup>T) \<le> w * w\<^sup>T \<sqinter> -F * F"
by (simp add: mult_isotone)
also have "... \<le> 1 \<sqinter> -F"
using assms(1,2) comp_inf.comp_isotone equivalence_comp_right_complement by auto
also have "... = bot"
using assms(1) bot_unique pp_isotone pseudo_complement_pp by blast
finally have 13: "(w \<sqinter> -F) * (F \<sqinter> w\<^sup>T) = bot"
- by (simp add: antisym)
+ by (simp add: order.antisym)
have "w \<sqinter> ?G \<le> F * (w \<sqinter> ?G)"
by (metis assms(1) mult_1_left mult_right_dist_sup sup.absorb_iff2)
also have "... \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by (metis eq_refl le_supE star.circ_back_loop_fixpoint)
finally have 14: "w \<sqinter> ?G \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by simp
have "w \<sqinter> top * e * F \<le> w * (e * F)\<^sup>T * e * F"
by (metis (no_types) comp_inf.star_slide dedekind_2 inf_left_commute inf_top_right mult_assoc)
also have "... \<le> F"
using 3 assms(1) by (metis comp_associative conv_dist_comp mult_left_isotone preorder_idempotent)
finally have "w \<sqinter> -F \<le> -(top * e * F)"
using order.trans p_shunting_swap pp_increasing by blast
also have "... = ?G"
by (metis assms(3) comp_mapping_complement conv_dist_comp conv_involutive conv_top)
finally have "(w \<sqinter> -F) * F * (w \<sqinter> ?G) = (w \<sqinter> -F \<sqinter> ?G) * F * (w \<sqinter> ?G)"
by (simp add: inf.absorb1)
also have "... \<le> (w \<sqinter> -F \<sqinter> ?G) * F * w"
by (simp add: comp_isotone)
also have "... \<le> (w \<sqinter> -F \<sqinter> ?G) * forest_components (F \<sqinter> w) * w"
by (simp add: assms(5) mult_left_isotone mult_right_isotone)
also have "... \<le> (w \<sqinter> -F \<sqinter> ?G) * (F \<sqinter> w)\<^sup>T\<^sup>\<star> * w\<^sup>\<star> * w"
by (simp add: mult_left_isotone mult_right_isotone star_isotone mult_assoc)
also have "... \<le> (w \<sqinter> -F \<sqinter> ?G) * (F \<sqinter> w)\<^sup>T\<^sup>\<star> * w\<^sup>\<star>"
by (simp add: comp_associative mult_right_isotone star.circ_plus_same star.left_plus_below_circ)
also have "... = (w \<sqinter> -F \<sqinter> ?G) * w\<^sup>\<star> \<squnion> (w \<sqinter> -F \<sqinter> ?G) * (F \<sqinter> w)\<^sup>T\<^sup>+ * w\<^sup>\<star>"
by (metis comp_associative inf.sup_monoid.add_assoc mult_left_dist_sup star.circ_loop_fixpoint sup_commute)
also have "... \<le> (w \<sqinter> -F \<sqinter> ?G) * w\<^sup>\<star> \<squnion> (w \<sqinter> -F \<sqinter> ?G) * (F \<sqinter> w)\<^sup>T * top"
by (metis mult_assoc top_greatest mult_right_isotone sup_right_isotone)
also have "... \<le> (w \<sqinter> -F \<sqinter> ?G) * w\<^sup>\<star> \<squnion> (w \<sqinter> -F) * (F \<sqinter> w)\<^sup>T * top"
using inf.cobounded1 mult_left_isotone sup_right_isotone by blast
also have "... \<le> (w \<sqinter> ?G) * w\<^sup>\<star> \<squnion> (w \<sqinter> -F) * (F \<sqinter> w)\<^sup>T * top"
using inf.sup_monoid.add_assoc inf.sup_right_isotone mult_left_isotone sup_commute sup_right_isotone by simp
also have "... = (w \<sqinter> ?G) * w\<^sup>\<star> \<squnion> (w \<sqinter> -F) * (F \<sqinter> w\<^sup>T) * top"
by (simp add: assms(1) conv_dist_inf)
also have "... \<le> 1 * (w \<sqinter> ?G) * w\<^sup>\<star>"
using 13 by simp
also have "... \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
using assms(1) mult_left_isotone by blast
finally have 15: "(w \<sqinter> -F) * F * (w \<sqinter> ?G) \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by simp
have "(w \<sqinter> F) * F * (w \<sqinter> ?G) \<le> F * F * (w \<sqinter> ?G)"
by (simp add: mult_left_isotone)
also have "... = F * (w \<sqinter> ?G)"
by (simp add: assms(1) preorder_idempotent)
also have "... \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by (metis eq_refl le_supE star.circ_back_loop_fixpoint)
finally have "(w \<sqinter> F) * F * (w \<sqinter> ?G) \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by simp
hence "((w \<sqinter> F) \<squnion> (w \<sqinter> -F)) * F * (w \<sqinter> ?G) \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
using 15 by (simp add: semiring.distrib_right)
hence "w * F * (w \<sqinter> ?G) \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by (metis assms(4) maddux_3_11_pp)
hence "w * F * (w \<sqinter> ?G) * w\<^sup>\<star> \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by (metis (full_types) comp_associative mult_left_isotone star.circ_transitive_equal)
hence "w\<^sup>\<star> * (w \<sqinter> ?G) \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
using 14 by (simp add: mult_assoc star_left_induct)
hence 16: "w\<^sup>+ \<sqinter> ?G \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
by (simp add: covector_comp_inf covector_mult_closed star.circ_plus_same)
have 17: "e\<^sup>T * top * e\<^sup>T \<le> -F"
using assms(8) le_bot triple_schroeder_p by simp
hence "(top * e)\<^sup>T * e\<^sup>T \<le> -F"
by (simp add: conv_dist_comp)
hence 18: "e\<^sup>T \<le> ?G"
by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top)
have "e\<^sup>T \<le> -F"
using 17 by (simp add: assms(3) arc_top_arc)
also have "... \<le> -1"
by (simp add: assms(1) p_antitone)
finally have "e\<^sup>T \<le> w\<^sup>\<star> \<sqinter> -1"
using assms(9) by simp
also have "... \<le> w\<^sup>+"
using shunting_var_p star_left_unfold_equal sup_commute by simp
finally have "e\<^sup>T \<le> w\<^sup>+ \<sqinter> ?G"
using 18 by simp
hence "e\<^sup>T \<le> F * (w \<sqinter> ?G) * w\<^sup>\<star>"
using 16 order_trans by blast
also have "... = (F * w \<sqinter> ?G) * w\<^sup>\<star>"
by (simp add: comp_associative comp_inf_covector)
finally have "e\<^sup>T * top * e\<^sup>T \<le> (F * w \<sqinter> ?G) * w\<^sup>\<star>"
by (simp add: assms(3) arc_top_arc)
hence "e\<^sup>T * top * (e * top)\<^sup>T \<le> (F * w \<sqinter> ?G) * w\<^sup>\<star>"
by (metis conv_dist_comp conv_top vector_top_closed mult_assoc)
hence "e\<^sup>T * top \<le> (F * w \<sqinter> ?G) * w\<^sup>\<star> * e * top"
by (metis assms(3) shunt_bijective mult_assoc)
hence "(top * e)\<^sup>T * top \<le> (F * w \<sqinter> ?G) * w\<^sup>\<star> * e * top"
by (simp add: conv_dist_comp mult_assoc)
hence "top \<le> top * e * (F * w \<sqinter> ?G) * w\<^sup>\<star> * e * top"
by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top mult_assoc)
also have "... = top * e * F * w * (w\<^sup>\<star> * e * top \<sqinter> ?G\<^sup>T)"
by (metis comp_associative comp_inf_vector_1)
also have "... = top * (w \<sqinter> (top * e * F)\<^sup>T) * (w\<^sup>\<star> * e * top \<sqinter> ?G\<^sup>T)"
by (metis comp_inf_vector_1 inf_top.left_neutral)
also have "... = top * (w \<sqinter> ?F) * (w\<^sup>\<star> * e * top \<sqinter> ?G\<^sup>T)"
by (simp add: assms(1) conv_dist_comp mult_assoc)
also have "... = top * (w \<sqinter> ?F) * (?E\<^sup>T \<sqinter> ?G\<^sup>T)"
by (simp add: comp_associative conv_dist_comp conv_star_commute)
also have "... = top * (w \<sqinter> ?F \<sqinter> ?G) * ?E\<^sup>T"
by (simp add: comp_associative comp_inf_vector_1)
also have "... = top * (w \<sqinter> ?F \<sqinter> ?G \<sqinter> ?E) * top"
using comp_inf_vector_1 mult_assoc by simp
finally show "top * (w \<sqinter> ?E \<sqinter> ?F \<sqinter> ?G) * top = top"
by (simp add: inf_commute inf_left_commute top_le)
qed
lemma kruskal_edge_arc_1:
assumes "e \<le> --h"
and "h \<le> g"
and "symmetric g"
and "components g \<le> forest_components w"
and "w * e\<^sup>T = bot"
shows "e\<^sup>T \<le> w\<^sup>\<star>"
proof -
have "w\<^sup>T * top \<le> -(e\<^sup>T * top)"
using assms(5) schroeder_3_p vector_bot_closed mult_assoc by fastforce
hence 1: "w\<^sup>T * top \<sqinter> e\<^sup>T * top = bot"
using pseudo_complement by simp
have "e\<^sup>T \<le> e\<^sup>T * top \<sqinter> --h\<^sup>T"
using assms(1) conv_complement conv_isotone top_right_mult_increasing by fastforce
also have "... \<le> e\<^sup>T * top \<sqinter> --g"
by (metis assms(2,3) inf.sup_right_isotone pp_isotone conv_isotone)
also have "... \<le> e\<^sup>T * top \<sqinter> components g"
using inf.sup_right_isotone star.circ_increasing by simp
also have "... \<le> e\<^sup>T * top \<sqinter> forest_components w"
using assms(4) comp_inf.mult_right_isotone by simp
also have "... = (e\<^sup>T * top \<sqinter> w\<^sup>T\<^sup>\<star>) * w\<^sup>\<star>"
by (simp add: inf_assoc vector_export_comp)
also have "... = (e\<^sup>T * top \<sqinter> 1) * w\<^sup>\<star> \<squnion> (e\<^sup>T * top \<sqinter> w\<^sup>T\<^sup>+) * w\<^sup>\<star>"
by (metis inf_sup_distrib1 semiring.distrib_right star_left_unfold_equal)
also have "... \<le> w\<^sup>\<star> \<squnion> (e\<^sup>T * top \<sqinter> w\<^sup>T\<^sup>+) * w\<^sup>\<star>"
by (metis inf_le2 mult_1_left mult_left_isotone sup_left_isotone)
also have "... \<le> w\<^sup>\<star> \<squnion> (e\<^sup>T * top \<sqinter> w\<^sup>T) * top"
using comp_associative comp_inf.mult_right_isotone sup_right_isotone mult_right_isotone top.extremum vector_export_comp by presburger
also have "... = w\<^sup>\<star>"
using 1 inf.sup_monoid.add_commute inf_vector_comp by simp
finally show ?thesis
by simp
qed
lemma kruskal_edge_between_components_1:
assumes "equivalence F"
and "mapping (top * e)"
shows "F \<le> -(w \<sqinter> top * e\<^sup>T * w\<^sup>T\<^sup>\<star> \<sqinter> F * e\<^sup>T * top \<sqinter> top * e * -F)"
proof -
let ?d = "w \<sqinter> top * e\<^sup>T * w\<^sup>T\<^sup>\<star> \<sqinter> F * e\<^sup>T * top \<sqinter> top * e * -F"
have "?d \<sqinter> F \<le> F * e\<^sup>T * top \<sqinter> F"
by (meson inf_le1 inf_le2 le_infI order_trans)
also have "... \<le> (F * e\<^sup>T * top)\<^sup>T * F"
by (simp add: mult_assoc vector_restrict_comp_conv)
also have "... = top * e * F * F"
by (simp add: assms(1) comp_associative conv_dist_comp conv_star_commute)
also have "... = top * e * F"
using assms(1) preorder_idempotent mult_assoc by fastforce
finally have "?d \<sqinter> F \<le> top * e * F \<sqinter> top * e * -F"
by (simp add: le_infI1)
also have "... = top * e * F \<sqinter> -(top * e * F)"
using assms(2) conv_dist_comp total_conv_surjective comp_mapping_complement by simp
finally show ?thesis
by (metis inf_p le_bot p_antitone_iff pseudo_complement)
qed
lemma kruskal_edge_between_components_2:
assumes "forest_components f \<le> -d"
and "injective f"
and "f \<squnion> f\<^sup>T \<le> w \<squnion> w\<^sup>T"
shows "f \<squnion> f\<^sup>T \<le> (w \<sqinter> -d \<sqinter> -d\<^sup>T) \<squnion> (w\<^sup>T \<sqinter> -d \<sqinter> -d\<^sup>T)"
proof -
let ?F = "forest_components f"
have "?F\<^sup>T \<le> -d\<^sup>T"
using assms(1) conv_complement conv_order by fastforce
hence 1: "?F \<le> -d\<^sup>T"
by (simp add: conv_dist_comp conv_star_commute)
have "equivalence ?F"
using assms(2) forest_components_equivalence by simp
hence "f \<squnion> f\<^sup>T \<le> ?F"
by (metis conv_dist_inf forest_components_increasing inf.absorb_iff2 sup.boundedI)
also have "... \<le> -d \<sqinter> -d\<^sup>T"
using 1 assms(1) by simp
finally have "f \<squnion> f\<^sup>T \<le> -d \<sqinter> -d\<^sup>T"
by simp
thus ?thesis
by (metis assms(3) inf_sup_distrib2 le_inf_iff)
qed
end
subsection \<open>Related Structures\<close>
text \<open>
Stone algebras can be expanded to Stone-Kleene relation algebras by reusing some operations.
\<close>
sublocale stone_algebra < comp_inf: stone_kleene_relation_algebra where star = "\<lambda>x . top" and one = top and times = inf and conv = id
apply unfold_locales
by simp
text \<open>
Every bounded linear order can be expanded to a Stone algebra, which can be expanded to a Stone relation algebra, which can be expanded to a Stone-Kleene relation algebra.
\<close>
class linorder_stone_kleene_relation_algebra_expansion = linorder_stone_relation_algebra_expansion + star +
assumes star_def [simp]: "x\<^sup>\<star> = top"
begin
subclass kleene_algebra
apply unfold_locales
apply simp
apply (simp add: min.coboundedI1 min.commute)
by (simp add: min.coboundedI1)
subclass stone_kleene_relation_algebra
apply unfold_locales
by simp
end
text \<open>
A Kleene relation algebra is based on a relation algebra.
\<close>
class kleene_relation_algebra = relation_algebra + stone_kleene_relation_algebra
end
diff --git a/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy
--- a/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy
+++ b/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy
@@ -1,1313 +1,1313 @@
(* Title: Matrix Kleene Algebras
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Matrix Kleene Algebras\<close>
text \<open>
This theory gives a matrix model of Stone-Kleene relation algebras.
The main result is that matrices over Kleene algebras form Kleene algebras.
The automata-based construction is due to Conway \cite{Conway1971}.
An implementation of the construction in Isabelle/HOL that extends \cite{ArmstrongGomesStruthWeber2016} was given in \cite{Asplund2014} without a correctness proof.
For specifying the size of matrices, Isabelle/HOL's type system requires the use of types, not sets.
This creates two issues when trying to implement Conway's recursive construction directly.
First, the matrix size changes for recursive calls, which requires dependent types.
Second, some submatrices used in the construction are not square, which requires typed Kleene algebras \cite{Kozen1998}, that is, categories of Kleene algebras.
Because these instruments are not available in Isabelle/HOL, we use square matrices with a constant size given by the argument of the Kleene star operation.
Smaller, possibly rectangular submatrices are identified by two lists of indices: one for the rows to include and one for the columns to include.
Lists are used to make recursive calls deterministic; otherwise sets would be sufficient.
\<close>
theory Matrix_Kleene_Algebras
imports Stone_Relation_Algebras.Matrix_Relation_Algebras Kleene_Relation_Algebras
begin
subsection \<open>Matrix Restrictions\<close>
text \<open>
In this section we develop a calculus of matrix restrictions.
The restriction of a matrix to specific row and column indices is implemented by the following function, which keeps the size of the matrix and sets all unused entries to \<open>bot\<close>.
\<close>
definition restrict_matrix :: "'a list \<Rightarrow> ('a,'b::bot) square \<Rightarrow> 'a list \<Rightarrow> ('a,'b) square" ("_ \<langle>_\<rangle> _" [90,41,90] 91)
where "restrict_matrix as f bs = (\<lambda>(i,j) . if List.member as i \<and> List.member bs j then f (i,j) else bot)"
text \<open>
The following function captures Conway's automata-based construction of the Kleene star of a matrix.
An index \<open>k\<close> is chosen and \<open>s\<close> contains all other indices.
The matrix is split into four submatrices \<open>a\<close>, \<open>b\<close>, \<open>c\<close>, \<open>d\<close> including/not including row/column \<open>k\<close>.
Four matrices are computed containing the entries given by Conway's construction.
These four matrices are added to obtain the result.
All matrices involved in the function have the same size, but matrix restriction is used to set irrelevant entries to \<open>bot\<close>.
\<close>
primrec star_matrix' :: "'a list \<Rightarrow> ('a,'b::{star,times,bounded_semilattice_sup_bot}) square \<Rightarrow> ('a,'b) square" where
"star_matrix' Nil g = mbot" |
"star_matrix' (k#s) g = (
let r = [k] in
let a = r\<langle>g\<rangle>r in
let b = r\<langle>g\<rangle>s in
let c = s\<langle>g\<rangle>r in
let d = s\<langle>g\<rangle>s in
let as = r\<langle>star o a\<rangle>r in
let ds = star_matrix' s d in
let e = a \<oplus> b \<odot> ds \<odot> c in
let es = r\<langle>star o e\<rangle>r in
let f = d \<oplus> c \<odot> as \<odot> b in
let fs = star_matrix' s f in
es \<oplus> as \<odot> b \<odot> fs \<oplus> ds \<odot> c \<odot> es \<oplus> fs
)"
text \<open>
The Kleene star of the whole matrix is obtained by taking as indices all elements of the underlying type \<open>'a\<close>.
This is conveniently supplied by the \<open>enum\<close> class.
\<close>
fun star_matrix :: "('a::enum,'b::{star,times,bounded_semilattice_sup_bot}) square \<Rightarrow> ('a,'b) square" ("_\<^sup>\<odot>" [100] 100) where "star_matrix f = star_matrix' (enum_class.enum::'a list) f"
text \<open>
The following lemmas deconstruct matrices with non-empty restrictions.
\<close>
lemma restrict_empty_left:
"[]\<langle>f\<rangle>ls = mbot"
by (unfold restrict_matrix_def List.member_def bot_matrix_def) auto
lemma restrict_empty_right:
"ks\<langle>f\<rangle>[] = mbot"
by (unfold restrict_matrix_def List.member_def bot_matrix_def) auto
lemma restrict_nonempty_left:
fixes f :: "('a,'b::bounded_semilattice_sup_bot) square"
shows "(k#ks)\<langle>f\<rangle>ls = [k]\<langle>f\<rangle>ls \<oplus> ks\<langle>f\<rangle>ls"
by (unfold restrict_matrix_def List.member_def sup_matrix_def) auto
lemma restrict_nonempty_right:
fixes f :: "('a,'b::bounded_semilattice_sup_bot) square"
shows "ks\<langle>f\<rangle>(l#ls) = ks\<langle>f\<rangle>[l] \<oplus> ks\<langle>f\<rangle>ls"
by (unfold restrict_matrix_def List.member_def sup_matrix_def) auto
lemma restrict_nonempty:
fixes f :: "('a,'b::bounded_semilattice_sup_bot) square"
shows "(k#ks)\<langle>f\<rangle>(l#ls) = [k]\<langle>f\<rangle>[l] \<oplus> [k]\<langle>f\<rangle>ls \<oplus> ks\<langle>f\<rangle>[l] \<oplus> ks\<langle>f\<rangle>ls"
by (unfold restrict_matrix_def List.member_def sup_matrix_def) auto
text \<open>
The following predicate captures that two index sets are disjoint.
This has consequences for composition and the unit matrix.
\<close>
abbreviation "disjoint ks ls \<equiv> \<not>(\<exists>x . List.member ks x \<and> List.member ls x)"
lemma times_disjoint:
fixes f g :: "('a,'b::idempotent_semiring) square"
assumes "disjoint ls ms"
shows "ks\<langle>f\<rangle>ls \<odot> ms\<langle>g\<rangle>ns = mbot"
proof (rule ext, rule prod_cases)
fix i j
have "(ks\<langle>f\<rangle>ls \<odot> ms\<langle>g\<rangle>ns) (i,j) = (\<Squnion>\<^sub>k (ks\<langle>f\<rangle>ls) (i,k) * (ms\<langle>g\<rangle>ns) (k,j))"
by (simp add: times_matrix_def)
also have "... = (\<Squnion>\<^sub>k (if List.member ks i \<and> List.member ls k then f (i,k) else bot) * (if List.member ms k \<and> List.member ns j then g (k,j) else bot))"
by (simp add: restrict_matrix_def)
also have "... = (\<Squnion>\<^sub>k if List.member ms k \<and> List.member ns j then bot * g (k,j) else (if List.member ks i \<and> List.member ls k then f (i,k) else bot) * bot)"
using assms by (auto intro: sup_monoid.sum.cong)
also have "... = (\<Squnion>\<^sub>(k::'a) bot)"
by (simp add: sup_monoid.sum.neutral)
also have "... = bot"
by (simp add: eq_iff le_funI)
also have "... = mbot (i,j)"
by (simp add: bot_matrix_def)
finally show "(ks\<langle>f\<rangle>ls \<odot> ms\<langle>g\<rangle>ns) (i,j) = mbot (i,j)"
.
qed
lemma one_disjoint:
assumes "disjoint ks ls"
shows "ks\<langle>(mone::('a,'b::idempotent_semiring) square)\<rangle>ls = mbot"
proof (rule ext, rule prod_cases)
let ?o = "mone::('a,'b) square"
fix i j
have "(ks\<langle>?o\<rangle>ls) (i,j) = (if List.member ks i \<and> List.member ls j then if i = j then 1 else bot else bot)"
by (simp add: restrict_matrix_def one_matrix_def)
also have "... = bot"
using assms by auto
also have "... = mbot (i,j)"
by (simp add: bot_matrix_def)
finally show "(ks\<langle>?o\<rangle>ls) (i,j) = mbot (i,j)"
.
qed
text \<open>
The following predicate captures that an index set is a subset of another index set.
This has consequences for repeated restrictions.
\<close>
abbreviation "is_sublist ks ls \<equiv> \<forall>x . List.member ks x \<longrightarrow> List.member ls x"
lemma restrict_sublist:
assumes "is_sublist ls ks"
and "is_sublist ms ns"
shows "ls\<langle>ks\<langle>f\<rangle>ns\<rangle>ms = ls\<langle>f\<rangle>ms"
proof (rule ext, rule prod_cases)
fix i j
show "(ls\<langle>ks\<langle>f\<rangle>ns\<rangle>ms) (i,j) = (ls\<langle>f\<rangle>ms) (i,j)"
proof (cases "List.member ls i \<and> List.member ms j")
case True thus ?thesis
by (simp add: assms restrict_matrix_def)
next
case False thus ?thesis
by (unfold restrict_matrix_def) auto
qed
qed
lemma restrict_superlist:
assumes "is_sublist ls ks"
and "is_sublist ms ns"
shows "ks\<langle>ls\<langle>f\<rangle>ms\<rangle>ns = ls\<langle>f\<rangle>ms"
proof (rule ext, rule prod_cases)
fix i j
show "(ks\<langle>ls\<langle>f\<rangle>ms\<rangle>ns) (i,j) = (ls\<langle>f\<rangle>ms) (i,j)"
proof (cases "List.member ls i \<and> List.member ms j")
case True thus ?thesis
by (simp add: assms restrict_matrix_def)
next
case False thus ?thesis
by (unfold restrict_matrix_def) auto
qed
qed
text \<open>
The following lemmas give the sizes of the results of some matrix operations.
\<close>
lemma restrict_sup:
fixes f g :: "('a,'b::bounded_semilattice_sup_bot) square"
shows "ks\<langle>f \<oplus> g\<rangle>ls = ks\<langle>f\<rangle>ls \<oplus> ks\<langle>g\<rangle>ls"
by (unfold restrict_matrix_def sup_matrix_def) auto
lemma restrict_times:
fixes f g :: "('a,'b::idempotent_semiring) square"
shows "ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms = ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
proof (rule ext, rule prod_cases)
fix i j
have "(ks\<langle>(ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)\<rangle>ms) (i,j) = (if List.member ks i \<and> List.member ms j then (\<Squnion>\<^sub>k (ks\<langle>f\<rangle>ls) (i,k) * (ls\<langle>g\<rangle>ms) (k,j)) else bot)"
by (simp add: times_matrix_def restrict_matrix_def)
also have "... = (if List.member ks i \<and> List.member ms j then (\<Squnion>\<^sub>k (if List.member ks i \<and> List.member ls k then f (i,k) else bot) * (if List.member ls k \<and> List.member ms j then g (k,j) else bot)) else bot)"
by (simp add: restrict_matrix_def)
also have "... = (if List.member ks i \<and> List.member ms j then (\<Squnion>\<^sub>k if List.member ls k then f (i,k) * g (k,j) else bot) else bot)"
by (auto intro: sup_monoid.sum.cong)
also have "... = (\<Squnion>\<^sub>k if List.member ks i \<and> List.member ms j then (if List.member ls k then f (i,k) * g (k,j) else bot) else bot)"
by auto
also have "... = (\<Squnion>\<^sub>k (if List.member ks i \<and> List.member ls k then f (i,k) else bot) * (if List.member ls k \<and> List.member ms j then g (k,j) else bot))"
by (auto intro: sup_monoid.sum.cong)
also have "... = (\<Squnion>\<^sub>k (ks\<langle>f\<rangle>ls) (i,k) * (ls\<langle>g\<rangle>ms) (k,j))"
by (simp add: restrict_matrix_def)
also have "... = (ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) (i,j)"
by (simp add: times_matrix_def)
finally show "(ks\<langle>(ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)\<rangle>ms) (i,j) = (ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) (i,j)"
.
qed
lemma restrict_star:
fixes g :: "('a,'b::kleene_algebra) square"
shows "t\<langle>star_matrix' t g\<rangle>t = star_matrix' t g"
proof (induct arbitrary: g rule: list.induct)
case Nil show ?case
by (simp add: restrict_empty_left)
next
case (Cons k s)
let ?t = "k#s"
assume "\<And>g::('a,'b) square . s\<langle>star_matrix' s g\<rangle>s = star_matrix' s g"
hence 1: "\<And>g::('a,'b) square . ?t\<langle>star_matrix' s g\<rangle>?t = star_matrix' s g"
by (metis member_rec(1) restrict_superlist)
show "?t\<langle>star_matrix' ?t g\<rangle>?t = star_matrix' ?t g"
proof -
let ?r = "[k]"
let ?a = "?r\<langle>g\<rangle>?r"
let ?b = "?r\<langle>g\<rangle>s"
let ?c = "s\<langle>g\<rangle>?r"
let ?d = "s\<langle>g\<rangle>s"
let ?as = "?r\<langle>star o ?a\<rangle>?r"
let ?ds = "star_matrix' s ?d"
let ?e = "?a \<oplus> ?b \<odot> ?ds \<odot> ?c"
let ?es = "?r\<langle>star o ?e\<rangle>?r"
let ?f = "?d \<oplus> ?c \<odot> ?as \<odot> ?b"
let ?fs = "star_matrix' s ?f"
have 2: "?t\<langle>?as\<rangle>?t = ?as \<and> ?t\<langle>?b\<rangle>?t = ?b \<and> ?t\<langle>?c\<rangle>?t = ?c \<and> ?t\<langle>?es\<rangle>?t = ?es"
by (simp add: restrict_superlist member_def)
have 3: "?t\<langle>?ds\<rangle>?t = ?ds \<and> ?t\<langle>?fs\<rangle>?t = ?fs"
using 1 by simp
have 4: "?t\<langle>?t\<langle>?as\<rangle>?t \<odot> ?t\<langle>?b\<rangle>?t \<odot> ?t\<langle>?fs\<rangle>?t\<rangle>?t = ?t\<langle>?as\<rangle>?t \<odot> ?t\<langle>?b\<rangle>?t \<odot> ?t\<langle>?fs\<rangle>?t"
by (metis (no_types) restrict_times)
have 5: "?t\<langle>?t\<langle>?ds\<rangle>?t \<odot> ?t\<langle>?c\<rangle>?t \<odot> ?t\<langle>?es\<rangle>?t\<rangle>?t = ?t\<langle>?ds\<rangle>?t \<odot> ?t\<langle>?c\<rangle>?t \<odot> ?t\<langle>?es\<rangle>?t"
by (metis (no_types) restrict_times)
have "?t\<langle>star_matrix' ?t g\<rangle>?t = ?t\<langle>?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs\<rangle>?t"
by (metis star_matrix'.simps(2))
also have "... = ?t\<langle>?es\<rangle>?t \<oplus> ?t\<langle>?as \<odot> ?b \<odot> ?fs\<rangle>?t \<oplus> ?t\<langle>?ds \<odot> ?c \<odot> ?es\<rangle>?t \<oplus> ?t\<langle>?fs\<rangle>?t"
by (simp add: restrict_sup)
also have "... = ?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs"
using 2 3 4 5 by simp
also have "... = star_matrix' ?t g"
by (metis star_matrix'.simps(2))
finally show ?thesis
.
qed
qed
lemma restrict_one:
assumes "\<not> List.member ks k"
shows "(k#ks)\<langle>(mone::('a,'b::idempotent_semiring) square)\<rangle>(k#ks) = [k]\<langle>mone\<rangle>[k] \<oplus> ks\<langle>mone\<rangle>ks"
by (subst restrict_nonempty) (simp add: assms member_rec one_disjoint)
lemma restrict_one_left_unit:
"ks\<langle>(mone::('a::finite,'b::idempotent_semiring) square)\<rangle>ks \<odot> ks\<langle>f\<rangle>ls = ks\<langle>f\<rangle>ls"
proof (rule ext, rule prod_cases)
let ?o = "mone::('a,'b::idempotent_semiring) square"
fix i j
have "(ks\<langle>?o\<rangle>ks \<odot> ks\<langle>f\<rangle>ls) (i,j) = (\<Squnion>\<^sub>k (ks\<langle>?o\<rangle>ks) (i,k) * (ks\<langle>f\<rangle>ls) (k,j))"
by (simp add: times_matrix_def)
also have "... = (\<Squnion>\<^sub>k (if List.member ks i \<and> List.member ks k then ?o (i,k) else bot) * (if List.member ks k \<and> List.member ls j then f (k,j) else bot))"
by (simp add: restrict_matrix_def)
also have "... = (\<Squnion>\<^sub>k (if List.member ks i \<and> List.member ks k then (if i = k then 1 else bot) else bot) * (if List.member ks k \<and> List.member ls j then f (k,j) else bot))"
by (unfold one_matrix_def) auto
also have "... = (\<Squnion>\<^sub>k (if i = k then (if List.member ks i then 1 else bot) else bot) * (if List.member ks k \<and> List.member ls j then f (k,j) else bot))"
by (auto intro: sup_monoid.sum.cong)
also have "... = (\<Squnion>\<^sub>k if i = k then (if List.member ks i then 1 else bot) * (if List.member ks i \<and> List.member ls j then f (i,j) else bot) else bot)"
by (rule sup_monoid.sum.cong) simp_all
also have "... = (if List.member ks i then 1 else bot) * (if List.member ks i \<and> List.member ls j then f (i,j) else bot)"
by simp
also have "... = (if List.member ks i \<and> List.member ls j then f (i,j) else bot)"
by simp
also have "... = (ks\<langle>f\<rangle>ls) (i,j)"
by (simp add: restrict_matrix_def)
finally show "(ks\<langle>?o\<rangle>ks \<odot> ks\<langle>f\<rangle>ls) (i,j) = (ks\<langle>f\<rangle>ls) (i,j)"
.
qed
text \<open>
The following lemmas consider restrictions to singleton index sets.
\<close>
lemma restrict_singleton:
"([k]\<langle>f\<rangle>[l]) (i,j) = (if i = k \<and> j = l then f (i,j) else bot)"
by (simp add: restrict_matrix_def List.member_def)
lemma restrict_singleton_list:
"([k]\<langle>f\<rangle>ls) (i,j) = (if i = k \<and> List.member ls j then f (i,j) else bot)"
by (simp add: restrict_matrix_def List.member_def)
lemma restrict_list_singleton:
"(ks\<langle>f\<rangle>[l]) (i,j) = (if List.member ks i \<and> j = l then f (i,j) else bot)"
by (simp add: restrict_matrix_def List.member_def)
lemma restrict_singleton_product:
fixes f g :: "('a::finite,'b::kleene_algebra) square"
shows "([k]\<langle>f\<rangle>[l] \<odot> [m]\<langle>g\<rangle>[n]) (i,j) = (if i = k \<and> l = m \<and> j = n then f (i,l) * g (m,j) else bot)"
proof -
have "([k]\<langle>f\<rangle>[l] \<odot> [m]\<langle>g\<rangle>[n]) (i,j) = (\<Squnion>\<^sub>h ([k]\<langle>f\<rangle>[l]) (i,h) * ([m]\<langle>g\<rangle>[n]) (h,j))"
by (simp add: times_matrix_def)
also have "... = (\<Squnion>\<^sub>h (if i = k \<and> h = l then f (i,h) else bot) * (if h = m \<and> j = n then g (h,j) else bot))"
by (simp add: restrict_singleton)
also have "... = (\<Squnion>\<^sub>h if h = l then (if i = k then f (i,h) else bot) * (if h = m \<and> j = n then g (h,j) else bot) else bot)"
by (rule sup_monoid.sum.cong) auto
also have "... = (if i = k then f (i,l) else bot) * (if l = m \<and> j = n then g (l,j) else bot)"
by simp
also have "... = (if i = k \<and> l = m \<and> j = n then f (i,l) * g (m,j) else bot)"
by simp
finally show ?thesis
.
qed
text \<open>
The Kleene star unfold law holds for matrices with a single entry on the diagonal.
\<close>
lemma restrict_star_unfold:
"[l]\<langle>(mone::('a::finite,'b::kleene_algebra) square)\<rangle>[l] \<oplus> [l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l] = [l]\<langle>star o f\<rangle>[l]"
proof (rule ext, rule prod_cases)
let ?o = "mone::('a,'b::kleene_algebra) square"
fix i j
have "([l]\<langle>?o\<rangle>[l] \<oplus> [l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l]) (i,j) = ([l]\<langle>?o\<rangle>[l]) (i,j) \<squnion> ([l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l]) (i,j)"
by (simp add: sup_matrix_def)
also have "... = ([l]\<langle>?o\<rangle>[l]) (i,j) \<squnion> (\<Squnion>\<^sub>k ([l]\<langle>f\<rangle>[l]) (i,k) * ([l]\<langle>star o f\<rangle>[l]) (k,j))"
by (simp add: times_matrix_def)
also have "... = ([l]\<langle>?o\<rangle>[l]) (i,j) \<squnion> (\<Squnion>\<^sub>k (if i = l \<and> k = l then f (i,k) else bot) * (if k = l \<and> j = l then (f (k,j))\<^sup>\<star> else bot))"
by (simp add: restrict_singleton o_def)
also have "... = ([l]\<langle>?o\<rangle>[l]) (i,j) \<squnion> (\<Squnion>\<^sub>k if k = l then (if i = l then f (i,k) else bot) * (if j = l then (f (k,j))\<^sup>\<star> else bot) else bot)"
apply (rule arg_cong2[where f=sup])
apply simp
by (rule sup_monoid.sum.cong) auto
also have "... = ([l]\<langle>?o\<rangle>[l]) (i,j) \<squnion> (if i = l then f (i,l) else bot) * (if j = l then (f (l,j))\<^sup>\<star> else bot)"
by simp
also have "... = (if i = l \<and> j = l then 1 \<squnion> f (l,l) * (f (l,l))\<^sup>\<star> else bot)"
by (simp add: restrict_singleton one_matrix_def)
also have "... = (if i = l \<and> j = l then (f (l,l))\<^sup>\<star> else bot)"
by (simp add: star_left_unfold_equal)
also have "... = ([l]\<langle>star o f\<rangle>[l]) (i,j)"
by (simp add: restrict_singleton o_def)
finally show "([l]\<langle>?o\<rangle>[l] \<oplus> [l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l]) (i,j) = ([l]\<langle>star o f\<rangle>[l]) (i,j)"
.
qed
lemma restrict_all:
"enum_class.enum\<langle>f\<rangle>enum_class.enum = f"
by (simp add: restrict_matrix_def List.member_def enum_UNIV)
text \<open>
The following shows the various components of a matrix product.
It is essentially a recursive implementation of the product.
\<close>
lemma restrict_nonempty_product:
fixes f g :: "('a::finite,'b::idempotent_semiring) square"
assumes "\<not> List.member ls l"
shows "(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms) = ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)"
proof -
have "(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms) = ([k]\<langle>f\<rangle>[l] \<oplus> [k]\<langle>f\<rangle>ls \<oplus> ks\<langle>f\<rangle>[l] \<oplus> ks\<langle>f\<rangle>ls) \<odot> ([l]\<langle>g\<rangle>[m] \<oplus> [l]\<langle>g\<rangle>ms \<oplus> ls\<langle>g\<rangle>[m] \<oplus> ls\<langle>g\<rangle>ms)"
by (metis restrict_nonempty)
also have "... = [k]\<langle>f\<rangle>[l] \<odot> ([l]\<langle>g\<rangle>[m] \<oplus> [l]\<langle>g\<rangle>ms \<oplus> ls\<langle>g\<rangle>[m] \<oplus> ls\<langle>g\<rangle>ms) \<oplus> [k]\<langle>f\<rangle>ls \<odot> ([l]\<langle>g\<rangle>[m] \<oplus> [l]\<langle>g\<rangle>ms \<oplus> ls\<langle>g\<rangle>[m] \<oplus> ls\<langle>g\<rangle>ms) \<oplus> ks\<langle>f\<rangle>[l] \<odot> ([l]\<langle>g\<rangle>[m] \<oplus> [l]\<langle>g\<rangle>ms \<oplus> ls\<langle>g\<rangle>[m] \<oplus> ls\<langle>g\<rangle>ms) \<oplus> ks\<langle>f\<rangle>ls \<odot> ([l]\<langle>g\<rangle>[m] \<oplus> [l]\<langle>g\<rangle>ms \<oplus> ls\<langle>g\<rangle>[m] \<oplus> ls\<langle>g\<rangle>ms)"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>[l] \<odot> ls\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>[l] \<odot> ls\<langle>g\<rangle>ms) \<oplus> ([k]\<langle>f\<rangle>ls \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>[l] \<odot> ls\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>[l] \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>ls \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)"
by (simp add: matrix_idempotent_semiring.mult_left_dist_sup)
also have "... = ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms) \<oplus> ([k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)"
using assms by (simp add: List.member_def times_disjoint)
also have "... = ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)"
by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc matrix_semilattice_sup.sup_left_commute)
finally show ?thesis
.
qed
text \<open>
Equality of matrices is componentwise.
\<close>
lemma restrict_nonempty_eq:
"(k#ks)\<langle>f\<rangle>(l#ls) = (k#ks)\<langle>g\<rangle>(l#ls) \<longleftrightarrow> [k]\<langle>f\<rangle>[l] = [k]\<langle>g\<rangle>[l] \<and> [k]\<langle>f\<rangle>ls = [k]\<langle>g\<rangle>ls \<and> ks\<langle>f\<rangle>[l] = ks\<langle>g\<rangle>[l] \<and> ks\<langle>f\<rangle>ls = ks\<langle>g\<rangle>ls"
proof
assume 1: "(k#ks)\<langle>f\<rangle>(l#ls) = (k#ks)\<langle>g\<rangle>(l#ls)"
have 2: "is_sublist [k] (k#ks) \<and> is_sublist ks (k#ks) \<and> is_sublist [l] (l#ls) \<and> is_sublist ls (l#ls)"
by (simp add: member_rec)
hence "[k]\<langle>f\<rangle>[l] = [k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls)\<rangle>[l] \<and> [k]\<langle>f\<rangle>ls = [k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls)\<rangle>ls \<and> ks\<langle>f\<rangle>[l] = ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls)\<rangle>[l] \<and> ks\<langle>f\<rangle>ls = ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls)\<rangle>ls"
by (simp add: restrict_sublist)
thus "[k]\<langle>f\<rangle>[l] = [k]\<langle>g\<rangle>[l] \<and> [k]\<langle>f\<rangle>ls = [k]\<langle>g\<rangle>ls \<and> ks\<langle>f\<rangle>[l] = ks\<langle>g\<rangle>[l] \<and> ks\<langle>f\<rangle>ls = ks\<langle>g\<rangle>ls"
using 1 2 by (simp add: restrict_sublist)
next
assume 3: "[k]\<langle>f\<rangle>[l] = [k]\<langle>g\<rangle>[l] \<and> [k]\<langle>f\<rangle>ls = [k]\<langle>g\<rangle>ls \<and> ks\<langle>f\<rangle>[l] = ks\<langle>g\<rangle>[l] \<and> ks\<langle>f\<rangle>ls = ks\<langle>g\<rangle>ls"
show "(k#ks)\<langle>f\<rangle>(l#ls) = (k#ks)\<langle>g\<rangle>(l#ls)"
proof (rule ext, rule prod_cases)
fix i j
have 4: "f (k,l) = g (k,l)"
using 3 by (metis restrict_singleton)
have 5: "List.member ls j \<Longrightarrow> f (k,j) = g (k,j)"
using 3 by (metis restrict_singleton_list)
have 6: "List.member ks i \<Longrightarrow> f (i,l) = g (i,l)"
using 3 by (metis restrict_list_singleton)
have "(ks\<langle>f\<rangle>ls) (i,j) = (ks\<langle>g\<rangle>ls) (i,j)"
using 3 by simp
hence 7: "List.member ks i \<Longrightarrow> List.member ls j \<Longrightarrow> f (i,j) = g (i,j)"
by (simp add: restrict_matrix_def)
have "((k#ks)\<langle>f\<rangle>(l#ls)) (i,j) = (if (i = k \<or> List.member ks i) \<and> (j = l \<or> List.member ls j) then f (i,j) else bot)"
by (simp add: restrict_matrix_def List.member_def)
also have "... = (if i = k \<and> j = l then f (i,j) else if i = k \<and> List.member ls j then f (i,j) else if List.member ks i \<and> j = l then f (i,j) else if List.member ks i \<and> List.member ls j then f (i,j) else bot)"
by auto
also have "... = (if i = k \<and> j = l then g (i,j) else if i = k \<and> List.member ls j then g (i,j) else if List.member ks i \<and> j = l then g (i,j) else if List.member ks i \<and> List.member ls j then g (i,j) else bot)"
using 4 5 6 7 by simp
also have "... = (if (i = k \<or> List.member ks i) \<and> (j = l \<or> List.member ls j) then g (i,j) else bot)"
by auto
also have "... = ((k#ks)\<langle>g\<rangle>(l#ls)) (i,j)"
by (simp add: restrict_matrix_def List.member_def)
finally show "((k#ks)\<langle>f\<rangle>(l#ls)) (i,j) = ((k#ks)\<langle>g\<rangle>(l#ls)) (i,j)"
.
qed
qed
text \<open>
Inequality of matrices is componentwise.
\<close>
lemma restrict_nonempty_less_eq:
fixes f g :: "('a,'b::idempotent_semiring) square"
shows "(k#ks)\<langle>f\<rangle>(l#ls) \<preceq> (k#ks)\<langle>g\<rangle>(l#ls) \<longleftrightarrow> [k]\<langle>f\<rangle>[l] \<preceq> [k]\<langle>g\<rangle>[l] \<and> [k]\<langle>f\<rangle>ls \<preceq> [k]\<langle>g\<rangle>ls \<and> ks\<langle>f\<rangle>[l] \<preceq> ks\<langle>g\<rangle>[l] \<and> ks\<langle>f\<rangle>ls \<preceq> ks\<langle>g\<rangle>ls"
by (unfold matrix_semilattice_sup.sup.order_iff) (metis (no_types, lifting) restrict_nonempty_eq restrict_sup)
text \<open>
The following lemmas treat repeated restrictions to disjoint index sets.
\<close>
lemma restrict_disjoint_left:
assumes "disjoint ks ms"
shows "ms\<langle>ks\<langle>f\<rangle>ls\<rangle>ns = mbot"
proof (rule ext, rule prod_cases)
fix i j
have "(ms\<langle>ks\<langle>f\<rangle>ls\<rangle>ns) (i,j) = (if List.member ms i \<and> List.member ns j then if List.member ks i \<and> List.member ls j then f (i,j) else bot else bot)"
by (simp add: restrict_matrix_def)
thus "(ms\<langle>ks\<langle>f\<rangle>ls\<rangle>ns) (i,j) = mbot (i,j)"
using assms by (simp add: bot_matrix_def)
qed
lemma restrict_disjoint_right:
assumes "disjoint ls ns"
shows "ms\<langle>ks\<langle>f\<rangle>ls\<rangle>ns = mbot"
proof (rule ext, rule prod_cases)
fix i j
have "(ms\<langle>ks\<langle>f\<rangle>ls\<rangle>ns) (i,j) = (if List.member ms i \<and> List.member ns j then if List.member ks i \<and> List.member ls j then f (i,j) else bot else bot)"
by (simp add: restrict_matrix_def)
thus "(ms\<langle>ks\<langle>f\<rangle>ls\<rangle>ns) (i,j) = mbot (i,j)"
using assms by (simp add: bot_matrix_def)
qed
text \<open>
The following lemma expresses the equality of a matrix and a product of two matrices componentwise.
\<close>
lemma restrict_nonempty_product_eq:
fixes f g h :: "('a::finite,'b::idempotent_semiring) square"
assumes "\<not> List.member ks k"
and "\<not> List.member ls l"
and "\<not> List.member ms m"
shows "(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms) = (k#ks)\<langle>h\<rangle>(m#ms) \<longleftrightarrow> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] = [k]\<langle>h\<rangle>[m] \<and> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms = [k]\<langle>h\<rangle>ms \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] = ks\<langle>h\<rangle>[m] \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms = ks\<langle>h\<rangle>ms"
proof -
have 1: "disjoint [k] ks \<and> disjoint [m] ms"
by (simp add: assms(1,3) member_rec)
have 2: "[k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]"
proof -
have "[k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = [k]\<langle>([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)\<rangle>[m]"
by (simp add: assms(2) restrict_nonempty_product)
also have "... = [k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> [k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> [k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>[m] \<oplus> [k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>[m]"
by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup)
also have "... = [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<oplus> [k]\<langle>[k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m] \<oplus> [k]\<langle>[k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m] \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m]"
by (simp add: restrict_times)
also have "... = [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]"
using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right)
finally show ?thesis
.
qed
have 3: "[k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
proof -
have "[k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = [k]\<langle>([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)\<rangle>ms"
by (simp add: assms(2) restrict_nonempty_product)
also have "... = [k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms \<oplus> [k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms \<oplus> [k]\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms \<oplus> [k]\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms"
by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup)
also have "... = [k]\<langle>[k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>[k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms\<rangle>ms \<oplus> [k]\<langle>ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms\<rangle>ms"
by (simp add: restrict_times)
also have "... = [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left)
finally show ?thesis
.
qed
have 4: "ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]"
proof -
have "ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = ks\<langle>([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)\<rangle>[m]"
by (simp add: assms(2) restrict_nonempty_product)
also have "... = ks\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> ks\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> ks\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>[m] \<oplus> ks\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>[m] \<oplus> ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m] \<oplus> ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>[m] \<oplus> ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>[m]"
by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup)
also have "... = ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>[m] \<oplus> ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>[m] \<oplus> ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m] \<oplus> ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m] \<oplus> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<oplus> ks\<langle>ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m] \<oplus> ks\<langle>ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms\<rangle>[m]"
by (simp add: restrict_times)
also have "... = ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]"
using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left)
finally show ?thesis
.
qed
have 5: "ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
proof -
have "ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = ks\<langle>([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> ([k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]) \<oplus> (ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms)\<rangle>ms"
by (simp add: assms(2) restrict_nonempty_product)
also have "... = ks\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms \<oplus> ks\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms \<oplus> ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms \<oplus> ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms"
by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup)
also have "... = ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms\<rangle>ms\<rangle>ms \<oplus> ks\<langle>[k]\<langle>[k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms\<rangle>ms\<rangle>ms \<oplus> ks\<langle>ks\<langle>ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>ks\<langle>ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]\<rangle>[m]\<rangle>ms \<oplus> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
by (simp add: restrict_times)
also have "... = ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left)
finally show ?thesis
.
qed
have "(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms) = (k#ks)\<langle>h\<rangle>(m#ms) \<longleftrightarrow> (k#ks)\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>(m#ms) = (k#ks)\<langle>h\<rangle>(m#ms)"
by (simp add: restrict_times)
also have "... \<longleftrightarrow> [k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = [k]\<langle>h\<rangle>[m] \<and> [k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = [k]\<langle>h\<rangle>ms \<and> ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = ks\<langle>h\<rangle>[m] \<and> ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = ks\<langle>h\<rangle>ms"
by (meson restrict_nonempty_eq)
also have "... \<longleftrightarrow> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] = [k]\<langle>h\<rangle>[m] \<and> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms = [k]\<langle>h\<rangle>ms \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] = ks\<langle>h\<rangle>[m] \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms = ks\<langle>h\<rangle>ms"
using 2 3 4 5 by simp
finally show ?thesis
by simp
qed
text \<open>
The following lemma gives a componentwise characterisation of the inequality of a matrix and a product of two matrices.
\<close>
lemma restrict_nonempty_product_less_eq:
fixes f g h :: "('a::finite,'b::idempotent_semiring) square"
assumes "\<not> List.member ks k"
and "\<not> List.member ls l"
and "\<not> List.member ms m"
shows "(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms) \<preceq> (k#ks)\<langle>h\<rangle>(m#ms) \<longleftrightarrow> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<preceq> [k]\<langle>h\<rangle>[m] \<and> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms \<preceq> [k]\<langle>h\<rangle>ms \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<preceq> ks\<langle>h\<rangle>[m] \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms \<preceq> ks\<langle>h\<rangle>ms"
proof -
have 1: "[k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]"
by (metis assms restrict_nonempty_product_eq restrict_times)
have 2: "[k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
by (metis assms restrict_nonempty_product_eq restrict_times)
have 3: "ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] = ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m]"
by (metis assms restrict_nonempty_product_eq restrict_times)
have 4: "ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms = ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms"
by (metis assms restrict_nonempty_product_eq restrict_times)
have "(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms) \<preceq> (k#ks)\<langle>h\<rangle>(m#ms) \<longleftrightarrow> (k#ks)\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>(m#ms) \<preceq> (k#ks)\<langle>h\<rangle>(m#ms)"
by (simp add: restrict_times)
also have "... \<longleftrightarrow> [k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] \<preceq> [k]\<langle>h\<rangle>[m] \<and> [k]\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms \<preceq> [k]\<langle>h\<rangle>ms \<and> ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>[m] \<preceq> ks\<langle>h\<rangle>[m] \<and> ks\<langle>(k#ks)\<langle>f\<rangle>(l#ls) \<odot> (l#ls)\<langle>g\<rangle>(m#ms)\<rangle>ms \<preceq> ks\<langle>h\<rangle>ms"
by (meson restrict_nonempty_less_eq)
also have "... \<longleftrightarrow> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<preceq> [k]\<langle>h\<rangle>[m] \<and> [k]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> [k]\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms \<preceq> [k]\<langle>h\<rangle>ms \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>[m] \<preceq> ks\<langle>h\<rangle>[m] \<and> ks\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<oplus> ks\<langle>f\<rangle>ls \<odot> ls\<langle>g\<rangle>ms \<preceq> ks\<langle>h\<rangle>ms"
using 1 2 3 4 by simp
finally show ?thesis
by simp
qed
text \<open>
The Kleene star induction laws hold for matrices with a single entry on the diagonal.
The matrix \<open>g\<close> can actually contain a whole row/colum at the appropriate index.
\<close>
lemma restrict_star_left_induct:
fixes f g :: "('a::finite,'b::kleene_algebra) square"
shows "distinct ms \<Longrightarrow> [l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<preceq> [l]\<langle>g\<rangle>ms \<Longrightarrow> [l]\<langle>star o f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<preceq> [l]\<langle>g\<rangle>ms"
proof (induct ms)
case Nil thus ?case
by (simp add: restrict_empty_right)
next
case (Cons m ms)
assume 1: "distinct ms \<Longrightarrow> [l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<preceq> [l]\<langle>g\<rangle>ms \<Longrightarrow> [l]\<langle>star o f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<preceq> [l]\<langle>g\<rangle>ms"
assume 2: "distinct (m#ms)"
assume 3: "[l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>(m#ms) \<preceq> [l]\<langle>g\<rangle>(m#ms)"
have 4: "[l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<preceq> [l]\<langle>g\<rangle>[m] \<and> [l]\<langle>f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<preceq> [l]\<langle>g\<rangle>ms"
using 2 3 by (metis distinct.simps(2) matrix_semilattice_sup.sup.bounded_iff member_def member_rec(2) restrict_nonempty_product_less_eq)
hence 5: "[l]\<langle>star o f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>ms \<preceq> [l]\<langle>g\<rangle>ms"
using 1 2 by simp
have "f (l,l) * g (l,m) \<le> g (l,m)"
using 4 by (metis restrict_singleton_product restrict_singleton less_eq_matrix_def)
hence 6: "(f (l,l))\<^sup>\<star> * g (l,m) \<le> g (l,m)"
by (simp add: star_left_induct_mult)
have "[l]\<langle>star o f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m] \<preceq> [l]\<langle>g\<rangle>[m]"
proof (unfold less_eq_matrix_def, rule allI, rule prod_cases)
fix i j
have "([l]\<langle>star o f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]) (i,j) = (\<Squnion>\<^sub>k ([l]\<langle>star o f\<rangle>[l]) (i,k) * ([l]\<langle>g\<rangle>[m]) (k,j))"
by (simp add: times_matrix_def)
also have "... = (\<Squnion>\<^sub>k (if i = l \<and> k = l then (f (i,k))\<^sup>\<star> else bot) * (if k = l \<and> j = m then g (k,j) else bot))"
by (simp add: restrict_singleton o_def)
also have "... = (\<Squnion>\<^sub>k if k = l then (if i = l then (f (i,k))\<^sup>\<star> else bot) * (if j = m then g (k,j) else bot) else bot)"
by (rule sup_monoid.sum.cong) auto
also have "... = (if i = l then (f (i,l))\<^sup>\<star> else bot) * (if j = m then g (l,j) else bot)"
by simp
also have "... = (if i = l \<and> j = m then (f (l,l))\<^sup>\<star> * g (l,m) else bot)"
by simp
also have "... \<le> ([l]\<langle>g\<rangle>[m]) (i,j)"
using 6 by (simp add: restrict_singleton)
finally show "([l]\<langle>star o f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>[m]) (i,j) \<le> ([l]\<langle>g\<rangle>[m]) (i,j)"
.
qed
thus "[l]\<langle>star o f\<rangle>[l] \<odot> [l]\<langle>g\<rangle>(m#ms) \<preceq> [l]\<langle>g\<rangle>(m#ms)"
using 2 5 by (metis (no_types, hide_lams) matrix_idempotent_semiring.mult_left_dist_sup matrix_semilattice_sup.sup.mono restrict_nonempty_right)
qed
lemma restrict_star_right_induct:
fixes f g :: "('a::finite,'b::kleene_algebra) square"
shows "distinct ms \<Longrightarrow> ms\<langle>g\<rangle>[l] \<odot> [l]\<langle>f\<rangle>[l] \<preceq> ms\<langle>g\<rangle>[l] \<Longrightarrow> ms\<langle>g\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l] \<preceq> ms\<langle>g\<rangle>[l]"
proof (induct ms)
case Nil thus ?case
by (simp add: restrict_empty_left)
next
case (Cons m ms)
assume 1: "distinct ms \<Longrightarrow> ms\<langle>g\<rangle>[l] \<odot> [l]\<langle>f\<rangle>[l] \<preceq> ms\<langle>g\<rangle>[l] \<Longrightarrow> ms\<langle>g\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l] \<preceq> ms\<langle>g\<rangle>[l]"
assume 2: "distinct (m#ms)"
assume 3: "(m#ms)\<langle>g\<rangle>[l] \<odot> [l]\<langle>f\<rangle>[l] \<preceq> (m#ms)\<langle>g\<rangle>[l]"
have 4: "[m]\<langle>g\<rangle>[l] \<odot> [l]\<langle>f\<rangle>[l] \<preceq> [m]\<langle>g\<rangle>[l] \<and> ms\<langle>g\<rangle>[l] \<odot> [l]\<langle>f\<rangle>[l] \<preceq> ms\<langle>g\<rangle>[l]"
using 2 3 by (metis distinct.simps(2) matrix_semilattice_sup.sup.bounded_iff member_def member_rec(2) restrict_nonempty_product_less_eq)
hence 5: "ms\<langle>g\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l] \<preceq> ms\<langle>g\<rangle>[l]"
using 1 2 by simp
have "g (m,l) * f (l,l) \<le> g (m,l)"
using 4 by (metis restrict_singleton_product restrict_singleton less_eq_matrix_def)
hence 6: "g (m,l) * (f (l,l))\<^sup>\<star> \<le> g (m,l)"
by (simp add: star_right_induct_mult)
have "[m]\<langle>g\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l] \<preceq> [m]\<langle>g\<rangle>[l]"
proof (unfold less_eq_matrix_def, rule allI, rule prod_cases)
fix i j
have "([m]\<langle>g\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l]) (i,j) = (\<Squnion>\<^sub>k ([m]\<langle>g\<rangle>[l]) (i,k) * ([l]\<langle>star o f\<rangle>[l]) (k,j))"
by (simp add: times_matrix_def)
also have "... = (\<Squnion>\<^sub>k (if i = m \<and> k = l then g (i,k) else bot) * (if k = l \<and> j = l then (f (k,j))\<^sup>\<star> else bot))"
by (simp add: restrict_singleton o_def)
also have "... = (\<Squnion>\<^sub>k if k = l then (if i = m then g (i,k) else bot) * (if j = l then (f (k,j))\<^sup>\<star> else bot) else bot)"
by (rule sup_monoid.sum.cong) auto
also have "... = (if i = m then g (i,l) else bot) * (if j = l then (f (l,j))\<^sup>\<star> else bot)"
by simp
also have "... = (if i = m \<and> j = l then g (m,l) * (f (l,l))\<^sup>\<star> else bot)"
by simp
also have "... \<le> ([m]\<langle>g\<rangle>[l]) (i,j)"
using 6 by (simp add: restrict_singleton)
finally show "([m]\<langle>g\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l]) (i,j) \<le> ([m]\<langle>g\<rangle>[l]) (i,j)"
.
qed
thus "(m#ms)\<langle>g\<rangle>[l] \<odot> [l]\<langle>star o f\<rangle>[l] \<preceq> (m#ms)\<langle>g\<rangle>[l]"
using 2 5 by (metis (no_types, hide_lams) matrix_idempotent_semiring.mult_right_dist_sup matrix_semilattice_sup.sup.mono restrict_nonempty_left)
qed
lemma restrict_pp:
fixes f :: "('a,'b::p_algebra) square"
shows "ks\<langle>\<ominus>\<ominus>f\<rangle>ls = \<ominus>\<ominus>(ks\<langle>f\<rangle>ls)"
by (unfold restrict_matrix_def uminus_matrix_def) auto
lemma pp_star_commute:
fixes f :: "('a,'b::stone_kleene_relation_algebra) square"
shows "\<ominus>\<ominus>(star o f) = star o \<ominus>\<ominus>f"
by (simp add: uminus_matrix_def o_def pp_dist_star)
subsection \<open>Matrices form a Kleene Algebra\<close>
text \<open>
Matrices over Kleene algebras form a Kleene algebra using Conway's construction.
It remains to prove one unfold and two induction axioms of the Kleene star.
Each proof is by induction over the size of the matrix represented by an index list.
\<close>
interpretation matrix_kleene_algebra: kleene_algebra_var where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::kleene_algebra) square" and one = one_matrix and times = times_matrix and star = star_matrix
proof
fix y :: "('a,'b) square"
let ?e = "enum_class.enum::'a list"
let ?o = "mone :: ('a,'b) square"
have "\<forall>g :: ('a,'b) square . distinct ?e \<longrightarrow> (?e\<langle>?o\<rangle>?e \<oplus> ?e\<langle>g\<rangle>?e \<odot> star_matrix' ?e g) = (star_matrix' ?e g)"
proof (induct rule: list.induct)
case Nil thus ?case
by (simp add: restrict_empty_left)
next
case (Cons k s)
let ?t = "k#s"
assume 1: "\<forall>g :: ('a,'b) square . distinct s \<longrightarrow> (s\<langle>?o\<rangle>s \<oplus> s\<langle>g\<rangle>s \<odot> star_matrix' s g) = (star_matrix' s g)"
show "\<forall>g :: ('a,'b) square . distinct ?t \<longrightarrow> (?t\<langle>?o\<rangle>?t \<oplus> ?t\<langle>g\<rangle>?t \<odot> star_matrix' ?t g) = (star_matrix' ?t g)"
proof (rule allI, rule impI)
fix g :: "('a,'b) square"
assume 2: "distinct ?t"
let ?r = "[k]"
let ?a = "?r\<langle>g\<rangle>?r"
let ?b = "?r\<langle>g\<rangle>s"
let ?c = "s\<langle>g\<rangle>?r"
let ?d = "s\<langle>g\<rangle>s"
let ?as = "?r\<langle>star o ?a\<rangle>?r"
let ?ds = "star_matrix' s ?d"
let ?e = "?a \<oplus> ?b \<odot> ?ds \<odot> ?c"
let ?es = "?r\<langle>star o ?e\<rangle>?r"
let ?f = "?d \<oplus> ?c \<odot> ?as \<odot> ?b"
let ?fs = "star_matrix' s ?f"
have "s\<langle>?ds\<rangle>s = ?ds \<and> s\<langle>?fs\<rangle>s = ?fs"
by (simp add: restrict_star)
hence 3: "?r\<langle>?e\<rangle>?r = ?e \<and> s\<langle>?f\<rangle>s = ?f"
by (metis (no_types, lifting) restrict_one_left_unit restrict_sup restrict_times)
have 4: "disjoint s ?r \<and> disjoint ?r s"
using 2 by (simp add: in_set_member member_rec)
hence 5: "?t\<langle>?o\<rangle>?t = ?r\<langle>?o\<rangle>?r \<oplus> s\<langle>?o\<rangle>s"
by (meson member_rec(1) restrict_one)
have 6: "?t\<langle>g\<rangle>?t \<odot> ?es = ?a \<odot> ?es \<oplus> ?c \<odot> ?es"
proof -
have "?t\<langle>g\<rangle>?t \<odot> ?es = (?a \<oplus> ?b \<oplus> ?c \<oplus> ?d) \<odot> ?es"
by (metis restrict_nonempty)
also have "... = ?a \<odot> ?es \<oplus> ?b \<odot> ?es \<oplus> ?c \<odot> ?es \<oplus> ?d \<odot> ?es"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?a \<odot> ?es \<oplus> ?c \<odot> ?es"
using 4 by (simp add: times_disjoint)
finally show ?thesis
.
qed
have 7: "?t\<langle>g\<rangle>?t \<odot> ?as \<odot> ?b \<odot> ?fs = ?a \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?c \<odot> ?as \<odot> ?b \<odot> ?fs"
proof -
have "?t\<langle>g\<rangle>?t \<odot> ?as \<odot> ?b \<odot> ?fs = (?a \<oplus> ?b \<oplus> ?c \<oplus> ?d) \<odot> ?as \<odot> ?b \<odot> ?fs"
by (metis restrict_nonempty)
also have "... = ?a \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?b \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?c \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?d \<odot> ?as \<odot> ?b \<odot> ?fs"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?a \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?c \<odot> ?as \<odot> ?b \<odot> ?fs"
using 4 by (simp add: times_disjoint)
finally show ?thesis
.
qed
have 8: "?t\<langle>g\<rangle>?t \<odot> ?ds \<odot> ?c \<odot> ?es = ?b \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?d \<odot> ?ds \<odot> ?c \<odot> ?es"
proof -
have "?t\<langle>g\<rangle>?t \<odot> ?ds \<odot> ?c \<odot> ?es = (?a \<oplus> ?b \<oplus> ?c \<oplus> ?d) \<odot> ?ds \<odot> ?c \<odot> ?es"
by (metis restrict_nonempty)
also have "... = ?a \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?b \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?c \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?d \<odot> ?ds \<odot> ?c \<odot> ?es"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?b \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?d \<odot> ?ds \<odot> ?c \<odot> ?es"
using 4 by (metis (no_types, lifting) times_disjoint matrix_idempotent_semiring.mult_left_zero restrict_star matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left)
finally show ?thesis
.
qed
have 9: "?t\<langle>g\<rangle>?t \<odot> ?fs = ?b \<odot> ?fs \<oplus> ?d \<odot> ?fs"
proof -
have "?t\<langle>g\<rangle>?t \<odot> ?fs = (?a \<oplus> ?b \<oplus> ?c \<oplus> ?d) \<odot> ?fs"
by (metis restrict_nonempty)
also have "... = ?a \<odot> ?fs \<oplus> ?b \<odot> ?fs \<oplus> ?c \<odot> ?fs \<oplus> ?d \<odot> ?fs"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?b \<odot> ?fs \<oplus> ?d \<odot> ?fs"
using 4 by (metis (no_types, lifting) times_disjoint restrict_star matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left)
finally show ?thesis
.
qed
have "?t\<langle>?o\<rangle>?t \<oplus> ?t\<langle>g\<rangle>?t \<odot> star_matrix' ?t g = ?t\<langle>?o\<rangle>?t \<oplus> ?t\<langle>g\<rangle>?t \<odot> (?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs)"
by (metis star_matrix'.simps(2))
also have "... = ?t\<langle>?o\<rangle>?t \<oplus> ?t\<langle>g\<rangle>?t \<odot> ?es \<oplus> ?t\<langle>g\<rangle>?t \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?t\<langle>g\<rangle>?t \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?t\<langle>g\<rangle>?t \<odot> ?fs"
by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc matrix_semilattice_sup.sup_assoc)
also have "... = ?r\<langle>?o\<rangle>?r \<oplus> s\<langle>?o\<rangle>s \<oplus> ?a \<odot> ?es \<oplus> ?c \<odot> ?es \<oplus> ?a \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?c \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?b \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?d \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?b \<odot> ?fs \<oplus> ?d \<odot> ?fs"
using 5 6 7 8 9 by (simp add: matrix_semilattice_sup.sup.assoc)
also have "... = (?r\<langle>?o\<rangle>?r \<oplus> (?a \<odot> ?es \<oplus> ?b \<odot> ?ds \<odot> ?c \<odot> ?es)) \<oplus> (?b \<odot> ?fs \<oplus> ?a \<odot> ?as \<odot> ?b \<odot> ?fs) \<oplus> (?c \<odot> ?es \<oplus> ?d \<odot> ?ds \<odot> ?c \<odot> ?es) \<oplus> (s\<langle>?o\<rangle>s \<oplus> (?d \<odot> ?fs \<oplus> ?c \<odot> ?as \<odot> ?b \<odot> ?fs))"
by (simp only: matrix_semilattice_sup.sup_assoc matrix_semilattice_sup.sup_commute matrix_semilattice_sup.sup_left_commute)
also have "... = (?r\<langle>?o\<rangle>?r \<oplus> (?a \<odot> ?es \<oplus> ?b \<odot> ?ds \<odot> ?c \<odot> ?es)) \<oplus> (?r\<langle>?o\<rangle>?r \<odot> ?b \<odot> ?fs \<oplus> ?a \<odot> ?as \<odot> ?b \<odot> ?fs) \<oplus> (s\<langle>?o\<rangle>s \<odot> ?c \<odot> ?es \<oplus> ?d \<odot> ?ds \<odot> ?c \<odot> ?es) \<oplus> (s\<langle>?o\<rangle>s \<oplus> (?d \<odot> ?fs \<oplus> ?c \<odot> ?as \<odot> ?b \<odot> ?fs))"
by (simp add: restrict_one_left_unit)
also have "... = (?r\<langle>?o\<rangle>?r \<oplus> ?e \<odot> ?es) \<oplus> ((?r\<langle>?o\<rangle>?r \<oplus> ?a \<odot> ?as) \<odot> ?b \<odot> ?fs) \<oplus> ((s\<langle>?o\<rangle>s \<oplus> ?d \<odot> ?ds) \<odot> ?c \<odot> ?es) \<oplus> (s\<langle>?o\<rangle>s \<oplus> ?f \<odot> ?fs)"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = (?r\<langle>?o\<rangle>?r \<oplus> ?e \<odot> ?es) \<oplus> ((?r\<langle>?o\<rangle>?r \<oplus> ?a \<odot> ?as) \<odot> ?b \<odot> ?fs) \<oplus> ((s\<langle>?o\<rangle>s \<oplus> ?d \<odot> ?ds) \<odot> ?c \<odot> ?es) \<oplus> ?fs"
using 1 2 3 by (metis distinct.simps(2))
also have "... = (?r\<langle>?o\<rangle>?r \<oplus> ?e \<odot> ?es) \<oplus> ((?r\<langle>?o\<rangle>?r \<oplus> ?a \<odot> ?as) \<odot> ?b \<odot> ?fs) \<oplus> (?ds \<odot> ?c \<odot> ?es) \<oplus> ?fs"
using 1 2 by (metis (no_types, lifting) distinct.simps(2) restrict_superlist)
also have "... = ?es \<oplus> ((?r\<langle>?o\<rangle>?r \<oplus> ?a \<odot> ?as) \<odot> ?b \<odot> ?fs) \<oplus> (?ds \<odot> ?c \<odot> ?es) \<oplus> ?fs"
using 3 by (metis restrict_star_unfold)
also have "... = ?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs"
by (metis (no_types, lifting) restrict_one_left_unit restrict_star_unfold restrict_times)
also have "... = star_matrix' ?t g"
by (metis star_matrix'.simps(2))
finally show "?t\<langle>?o\<rangle>?t \<oplus> ?t\<langle>g\<rangle>?t \<odot> star_matrix' ?t g = star_matrix' ?t g"
.
qed
qed
thus "?o \<oplus> y \<odot> y\<^sup>\<odot> \<preceq> y\<^sup>\<odot>"
by (simp add: enum_distinct restrict_all)
next
fix x y z :: "('a,'b) square"
let ?e = "enum_class.enum::'a list"
have "\<forall>g h :: ('a,'b) square . \<forall>zs . distinct ?e \<and> distinct zs \<longrightarrow> (?e\<langle>g\<rangle>?e \<odot> ?e\<langle>h\<rangle>zs \<preceq> ?e\<langle>h\<rangle>zs \<longrightarrow> star_matrix' ?e g \<odot> ?e\<langle>h\<rangle>zs \<preceq> ?e\<langle>h\<rangle>zs)"
proof (induct rule: list.induct)
case Nil thus ?case
by (simp add: restrict_empty_left)
case (Cons k s)
let ?t = "k#s"
assume 1: "\<forall>g h :: ('a,'b) square . \<forall>zs . distinct s \<and> distinct zs \<longrightarrow> (s\<langle>g\<rangle>s \<odot> s\<langle>h\<rangle>zs \<preceq> s\<langle>h\<rangle>zs \<longrightarrow> star_matrix' s g \<odot> s\<langle>h\<rangle>zs \<preceq> s\<langle>h\<rangle>zs)"
show "\<forall>g h :: ('a,'b) square . \<forall>zs . distinct ?t \<and> distinct zs \<longrightarrow> (?t\<langle>g\<rangle>?t \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs \<longrightarrow> star_matrix' ?t g \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs)"
proof (intro allI)
fix g h :: "('a,'b) square"
fix zs :: "'a list"
show "distinct ?t \<and> distinct zs \<longrightarrow> (?t\<langle>g\<rangle>?t \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs \<longrightarrow> star_matrix' ?t g \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs)"
proof (cases zs)
case Nil thus ?thesis
by (metis restrict_empty_right restrict_star restrict_times)
next
case (Cons y ys)
assume 2: "zs = y#ys"
show "distinct ?t \<and> distinct zs \<longrightarrow> (?t\<langle>g\<rangle>?t \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs \<longrightarrow> star_matrix' ?t g \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs)"
proof (intro impI)
let ?y = "[y]"
assume 3: "distinct ?t \<and> distinct zs"
hence 4: "distinct s \<and> distinct ys \<and> \<not> List.member s k \<and> \<not> List.member ys y"
using 2 by (simp add: List.member_def)
let ?r = "[k]"
let ?a = "?r\<langle>g\<rangle>?r"
let ?b = "?r\<langle>g\<rangle>s"
let ?c = "s\<langle>g\<rangle>?r"
let ?d = "s\<langle>g\<rangle>s"
let ?as = "?r\<langle>star o ?a\<rangle>?r"
let ?ds = "star_matrix' s ?d"
let ?e = "?a \<oplus> ?b \<odot> ?ds \<odot> ?c"
let ?es = "?r\<langle>star o ?e\<rangle>?r"
let ?f = "?d \<oplus> ?c \<odot> ?as \<odot> ?b"
let ?fs = "star_matrix' s ?f"
let ?ha = "?r\<langle>h\<rangle>?y"
let ?hb = "?r\<langle>h\<rangle>ys"
let ?hc = "s\<langle>h\<rangle>?y"
let ?hd = "s\<langle>h\<rangle>ys"
assume "?t\<langle>g\<rangle>?t \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs"
hence 5: "?a \<odot> ?ha \<oplus> ?b \<odot> ?hc \<preceq> ?ha \<and> ?a \<odot> ?hb \<oplus> ?b \<odot> ?hd \<preceq> ?hb \<and> ?c \<odot> ?ha \<oplus> ?d \<odot> ?hc \<preceq> ?hc \<and> ?c \<odot> ?hb \<oplus> ?d \<odot> ?hd \<preceq> ?hd"
using 2 3 4 by (simp add: restrict_nonempty_product_less_eq)
have 6: "s\<langle>?ds\<rangle>s = ?ds \<and> s\<langle>?fs\<rangle>s = ?fs"
by (simp add: restrict_star)
hence 7: "?r\<langle>?e\<rangle>?r = ?e \<and> s\<langle>?f\<rangle>s = ?f"
by (metis (no_types, lifting) restrict_one_left_unit restrict_sup restrict_times)
have 8: "disjoint s ?r \<and> disjoint ?r s"
using 3 by (simp add: in_set_member member_rec(1) member_rec(2))
have 9: "?es \<odot> ?t\<langle>h\<rangle>zs = ?es \<odot> ?ha \<oplus> ?es \<odot> ?hb"
proof -
have "?es \<odot> ?t\<langle>h\<rangle>zs = ?es \<odot> (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd)"
using 2 by (metis restrict_nonempty)
also have "... = ?es \<odot> ?ha \<oplus> ?es \<odot> ?hb \<oplus> ?es \<odot> ?hc \<oplus> ?es \<odot> ?hd"
by (simp add: matrix_idempotent_semiring.mult_left_dist_sup)
also have "... = ?es \<odot> ?ha \<oplus> ?es \<odot> ?hb"
using 8 by (simp add: times_disjoint)
finally show ?thesis
.
qed
have 10: "?as \<odot> ?b \<odot> ?fs \<odot> ?t\<langle>h\<rangle>zs = ?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd"
proof -
have "?as \<odot> ?b \<odot> ?fs \<odot> ?t\<langle>h\<rangle>zs = ?as \<odot> ?b \<odot> ?fs \<odot> (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd)"
using 2 by (metis restrict_nonempty)
also have "... = ?as \<odot> ?b \<odot> ?fs \<odot> ?ha \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hb \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd"
by (simp add: matrix_idempotent_semiring.mult_left_dist_sup)
also have "... = ?as \<odot> ?b \<odot> (?fs \<odot> ?ha) \<oplus> ?as \<odot> ?b \<odot> (?fs \<odot> ?hb) \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd"
by (simp add: matrix_monoid.mult_assoc)
also have "... = ?as \<odot> ?b \<odot> mbot \<oplus> ?as \<odot> ?b \<odot> mbot \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd"
using 6 8 by (metis (no_types) times_disjoint)
also have "... = ?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd"
by simp
finally show ?thesis
.
qed
have 11: "?ds \<odot> ?c \<odot> ?es \<odot> ?t\<langle>h\<rangle>zs = ?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hb"
proof -
have "?ds \<odot> ?c \<odot> ?es \<odot> ?t\<langle>h\<rangle>zs = ?ds \<odot> ?c \<odot> ?es \<odot> (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd)"
using 2 by (metis restrict_nonempty)
also have "... = ?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hc \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hd"
by (simp add: matrix_idempotent_semiring.mult_left_dist_sup)
also have "... = ?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<oplus> ?ds \<odot> ?c \<odot> (?es \<odot> ?hc) \<oplus> ?ds \<odot> ?c \<odot> (?es \<odot> ?hd)"
by (simp add: matrix_monoid.mult_assoc)
also have "... = ?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<oplus> ?ds \<odot> ?c \<odot> mbot \<oplus> ?ds \<odot> ?c \<odot> mbot"
using 8 by (metis times_disjoint)
also have "... = ?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hb"
by simp
finally show ?thesis
.
qed
have 12: "?fs \<odot> ?t\<langle>h\<rangle>zs = ?fs \<odot> ?hc \<oplus> ?fs \<odot> ?hd"
proof -
have "?fs \<odot> ?t\<langle>h\<rangle>zs = ?fs \<odot> (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd)"
using 2 by (metis restrict_nonempty)
also have "... = ?fs \<odot> ?ha \<oplus> ?fs \<odot> ?hb \<oplus> ?fs \<odot> ?hc \<oplus> ?fs \<odot> ?hd"
by (simp add: matrix_idempotent_semiring.mult_left_dist_sup)
also have "... = ?fs \<odot> ?hc \<oplus> ?fs \<odot> ?hd"
using 6 8 by (metis (no_types) times_disjoint matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left)
finally show ?thesis
.
qed
have 13: "?es \<odot> ?ha \<preceq> ?ha"
proof -
have "?b \<odot> ?ds \<odot> ?c \<odot> ?ha \<preceq> ?b \<odot> ?ds \<odot> ?hc"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?b \<odot> ?hc"
using 1 3 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc member_rec(2) restrict_sublist)
also have "... \<preceq> ?ha"
using 5 by simp
finally have "?e \<odot> ?ha \<preceq> ?ha"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
thus ?thesis
using 7 by (simp add: restrict_star_left_induct)
qed
have 14: "?es \<odot> ?hb \<preceq> ?hb"
proof -
have "?b \<odot> ?ds \<odot> ?c \<odot> ?hb \<preceq> ?b \<odot> ?ds \<odot> ?hd"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?b \<odot> ?hd"
using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc restrict_sublist)
also have "... \<preceq> ?hb"
using 5 by simp
finally have "?e \<odot> ?hb \<preceq> ?hb"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
thus ?thesis
using 4 7 by (simp add: restrict_star_left_induct)
qed
have 15: "?fs \<odot> ?hc \<preceq> ?hc"
proof -
have "?c \<odot> ?as \<odot> ?b \<odot> ?hc \<preceq> ?c \<odot> ?as \<odot> ?ha"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?c \<odot> ?ha"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc restrict_star_left_induct restrict_sublist)
also have "... \<preceq> ?hc"
using 5 by simp
finally have "?f \<odot> ?hc \<preceq> ?hc"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
thus ?thesis
using 1 3 7 by simp
qed
have 16: "?fs \<odot> ?hd \<preceq> ?hd"
proof -
have "?c \<odot> ?as \<odot> ?b \<odot> ?hd \<preceq> ?c \<odot> ?as \<odot> ?hb"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?c \<odot> ?hb"
using 4 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc restrict_star_left_induct restrict_sublist)
also have "... \<preceq> ?hd"
using 5 by simp
finally have "?f \<odot> ?hd \<preceq> ?hd"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
thus ?thesis
using 1 4 7 by simp
qed
have 17: "?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<preceq> ?ha"
proof -
have "?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<preceq> ?as \<odot> ?b \<odot> ?hc"
using 15 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?as \<odot> ?ha"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?ha"
using 5 by (simp add: restrict_star_left_induct restrict_sublist)
finally show ?thesis
.
qed
have 18: "?as \<odot> ?b \<odot> ?fs \<odot> ?hd \<preceq> ?hb"
proof -
have "?as \<odot> ?b \<odot> ?fs \<odot> ?hd \<preceq> ?as \<odot> ?b \<odot> ?hd"
using 16 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?as \<odot> ?hb"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?hb"
using 4 5 by (simp add: restrict_star_left_induct restrict_sublist)
finally show ?thesis
.
qed
have 19: "?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<preceq> ?hc"
proof -
have "?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<preceq> ?ds \<odot> ?c \<odot> ?ha"
using 13 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?ds \<odot> ?hc"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?hc"
using 1 3 5 by (simp add: restrict_sublist)
finally show ?thesis
.
qed
have 20: "?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<preceq> ?hd"
proof -
have "?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<preceq> ?ds \<odot> ?c \<odot> ?hb"
using 14 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?ds \<odot> ?hd"
using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc)
also have "... \<preceq> ?hd"
using 1 4 5 by (simp add: restrict_sublist)
finally show ?thesis
.
qed
have 21: "?es \<odot> ?ha \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<preceq> ?ha"
using 13 17 matrix_semilattice_sup.le_supI by blast
have 22: "?es \<odot> ?hb \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd \<preceq> ?hb"
using 14 18 matrix_semilattice_sup.le_supI by blast
have 23: "?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?fs \<odot> ?hc \<preceq> ?hc"
using 15 19 matrix_semilattice_sup.le_supI by blast
have 24: "?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<oplus> ?fs \<odot> ?hd \<preceq> ?hd"
using 16 20 matrix_semilattice_sup.le_supI by blast
have "star_matrix' ?t g \<odot> ?t\<langle>h\<rangle>zs = (?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs) \<odot> ?t\<langle>h\<rangle>zs"
by (metis star_matrix'.simps(2))
also have "... = ?es \<odot> ?t\<langle>h\<rangle>zs \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?t\<langle>h\<rangle>zs \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?t\<langle>h\<rangle>zs \<oplus> ?fs \<odot> ?t\<langle>h\<rangle>zs"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?es \<odot> ?ha \<oplus> ?es \<odot> ?hb \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hc \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<oplus> ?fs \<odot> ?hc \<oplus> ?fs \<odot> ?hd"
using 9 10 11 12 by (simp only: matrix_semilattice_sup.sup_assoc)
also have "... = (?es \<odot> ?ha \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hc) \<oplus> (?es \<odot> ?hb \<oplus> ?as \<odot> ?b \<odot> ?fs \<odot> ?hd) \<oplus> (?ds \<odot> ?c \<odot> ?es \<odot> ?ha \<oplus> ?fs \<odot> ?hc) \<oplus> (?ds \<odot> ?c \<odot> ?es \<odot> ?hb \<oplus> ?fs \<odot> ?hd)"
by (simp only: matrix_semilattice_sup.sup_assoc matrix_semilattice_sup.sup_commute matrix_semilattice_sup.sup_left_commute)
also have "... \<preceq> ?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd"
using 21 22 23 24 matrix_semilattice_sup.sup.mono by blast
also have "... = ?t\<langle>h\<rangle>zs"
using 2 by (metis restrict_nonempty)
finally show "star_matrix' ?t g \<odot> ?t\<langle>h\<rangle>zs \<preceq> ?t\<langle>h\<rangle>zs"
.
qed
qed
qed
qed
hence "\<forall>zs . distinct zs \<longrightarrow> (y \<odot> ?e\<langle>x\<rangle>zs \<preceq> ?e\<langle>x\<rangle>zs \<longrightarrow> y\<^sup>\<odot> \<odot> ?e\<langle>x\<rangle>zs \<preceq> ?e\<langle>x\<rangle>zs)"
by (simp add: enum_distinct restrict_all)
thus "y \<odot> x \<preceq> x \<longrightarrow> y\<^sup>\<odot> \<odot> x \<preceq> x"
by (metis restrict_all enum_distinct)
next
fix x y z :: "('a,'b) square"
let ?e = "enum_class.enum::'a list"
have "\<forall>g h :: ('a,'b) square . \<forall>zs . distinct ?e \<and> distinct zs \<longrightarrow> (zs\<langle>h\<rangle>?e \<odot> ?e\<langle>g\<rangle>?e \<preceq> zs\<langle>h\<rangle>?e \<longrightarrow> zs\<langle>h\<rangle>?e \<odot> star_matrix' ?e g \<preceq> zs\<langle>h\<rangle>?e)"
proof (induct rule:list.induct)
case Nil thus ?case
by (simp add: restrict_empty_left)
case (Cons k s)
let ?t = "k#s"
assume 1: "\<forall>g h :: ('a,'b) square . \<forall>zs . distinct s \<and> distinct zs \<longrightarrow> (zs\<langle>h\<rangle>s \<odot> s\<langle>g\<rangle>s \<preceq> zs\<langle>h\<rangle>s \<longrightarrow> zs\<langle>h\<rangle>s \<odot> star_matrix' s g \<preceq> zs\<langle>h\<rangle>s)"
show "\<forall>g h :: ('a,'b) square . \<forall>zs . distinct ?t \<and> distinct zs \<longrightarrow> (zs\<langle>h\<rangle>?t \<odot> ?t\<langle>g\<rangle>?t \<preceq> zs\<langle>h\<rangle>?t \<longrightarrow> zs\<langle>h\<rangle>?t \<odot> star_matrix' ?t g \<preceq> zs\<langle>h\<rangle>?t)"
proof (intro allI)
fix g h :: "('a,'b) square"
fix zs :: "'a list"
show "distinct ?t \<and> distinct zs \<longrightarrow> (zs\<langle>h\<rangle>?t \<odot> ?t\<langle>g\<rangle>?t \<preceq> zs\<langle>h\<rangle>?t \<longrightarrow> zs\<langle>h\<rangle>?t \<odot> star_matrix' ?t g \<preceq> zs\<langle>h\<rangle>?t)"
proof (cases zs)
case Nil thus ?thesis
by (metis restrict_empty_left restrict_star restrict_times)
next
case (Cons y ys)
assume 2: "zs = y#ys"
show "distinct ?t \<and> distinct zs \<longrightarrow> (zs\<langle>h\<rangle>?t \<odot> ?t\<langle>g\<rangle>?t \<preceq> zs\<langle>h\<rangle>?t \<longrightarrow> zs\<langle>h\<rangle>?t \<odot> star_matrix' ?t g \<preceq> zs\<langle>h\<rangle>?t)"
proof (intro impI)
let ?y = "[y]"
assume 3: "distinct ?t \<and> distinct zs"
hence 4: "distinct s \<and> distinct ys \<and> \<not> List.member s k \<and> \<not> List.member ys y"
using 2 by (simp add: List.member_def)
let ?r = "[k]"
let ?a = "?r\<langle>g\<rangle>?r"
let ?b = "?r\<langle>g\<rangle>s"
let ?c = "s\<langle>g\<rangle>?r"
let ?d = "s\<langle>g\<rangle>s"
let ?as = "?r\<langle>star o ?a\<rangle>?r"
let ?ds = "star_matrix' s ?d"
let ?e = "?a \<oplus> ?b \<odot> ?ds \<odot> ?c"
let ?es = "?r\<langle>star o ?e\<rangle>?r"
let ?f = "?d \<oplus> ?c \<odot> ?as \<odot> ?b"
let ?fs = "star_matrix' s ?f"
let ?ha = "?y\<langle>h\<rangle>?r"
let ?hb = "?y\<langle>h\<rangle>s"
let ?hc = "ys\<langle>h\<rangle>?r"
let ?hd = "ys\<langle>h\<rangle>s"
assume "zs\<langle>h\<rangle>?t \<odot> ?t\<langle>g\<rangle>?t \<preceq> zs\<langle>h\<rangle>?t"
hence 5: "?ha \<odot> ?a \<oplus> ?hb \<odot> ?c \<preceq> ?ha \<and> ?ha \<odot> ?b \<oplus> ?hb \<odot> ?d \<preceq> ?hb \<and> ?hc \<odot> ?a \<oplus> ?hd \<odot> ?c \<preceq> ?hc \<and> ?hc \<odot> ?b \<oplus> ?hd \<odot> ?d \<preceq> ?hd"
using 2 3 4 by (simp add: restrict_nonempty_product_less_eq)
have 6: "s\<langle>?ds\<rangle>s = ?ds \<and> s\<langle>?fs\<rangle>s = ?fs"
by (simp add: restrict_star)
hence 7: "?r\<langle>?e\<rangle>?r = ?e \<and> s\<langle>?f\<rangle>s = ?f"
by (metis (no_types, lifting) restrict_one_left_unit restrict_sup restrict_times)
have 8: "disjoint s ?r \<and> disjoint ?r s"
using 3 by (simp add: in_set_member member_rec)
have 9: "zs\<langle>h\<rangle>?t \<odot> ?es = ?ha \<odot> ?es \<oplus> ?hc \<odot> ?es"
proof -
have "zs\<langle>h\<rangle>?t \<odot> ?es = (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd) \<odot> ?es"
using 2 by (metis restrict_nonempty)
also have "... = ?ha \<odot> ?es \<oplus> ?hb \<odot> ?es \<oplus> ?hc \<odot> ?es \<oplus> ?hd \<odot> ?es"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?ha \<odot> ?es \<oplus> ?hc \<odot> ?es"
using 8 by (simp add: times_disjoint)
finally show ?thesis
.
qed
have 10: "zs\<langle>h\<rangle>?t \<odot> ?as \<odot> ?b \<odot> ?fs = ?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hc \<odot> ?as \<odot> ?b \<odot> ?fs"
proof -
have "zs\<langle>h\<rangle>?t \<odot> ?as \<odot> ?b \<odot> ?fs = (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd) \<odot> ?as \<odot> ?b \<odot> ?fs"
using 2 by (metis restrict_nonempty)
also have "... = ?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hb \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hc \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hd \<odot> ?as \<odot> ?b \<odot> ?fs"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> mbot \<odot> ?b \<odot> ?fs \<oplus> ?hc \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> mbot \<odot> ?b \<odot> ?fs"
using 8 by (metis (no_types) times_disjoint)
also have "... = ?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hc \<odot> ?as \<odot> ?b \<odot> ?fs"
by simp
finally show ?thesis
.
qed
have 11: "zs\<langle>h\<rangle>?t \<odot> ?ds \<odot> ?c \<odot> ?es = ?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?hd \<odot> ?ds \<odot> ?c \<odot> ?es"
proof -
have "zs\<langle>h\<rangle>?t \<odot> ?ds \<odot> ?c \<odot> ?es = (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd) \<odot> ?ds \<odot> ?c \<odot> ?es"
using 2 by (metis restrict_nonempty)
also have "... = ?ha \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?hc \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?hd \<odot> ?ds \<odot> ?c \<odot> ?es"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = mbot \<odot> ?c \<odot> ?es \<oplus> ?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> mbot \<odot> ?c \<odot> ?es \<oplus> ?hd \<odot> ?ds \<odot> ?c \<odot> ?es"
using 6 8 by (metis (no_types) times_disjoint)
also have "... = ?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?hd \<odot> ?ds \<odot> ?c \<odot> ?es"
by simp
finally show ?thesis
.
qed
have 12: "zs\<langle>h\<rangle>?t \<odot> ?fs = ?hb \<odot> ?fs \<oplus> ?hd \<odot> ?fs"
proof -
have "zs\<langle>h\<rangle>?t \<odot> ?fs = (?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd) \<odot> ?fs"
using 2 by (metis restrict_nonempty)
also have "... = ?ha \<odot> ?fs \<oplus> ?hb \<odot> ?fs \<oplus> ?hc \<odot> ?fs \<oplus> ?hd \<odot> ?fs"
by (simp add: matrix_idempotent_semiring.mult_right_dist_sup)
also have "... = ?hb \<odot> ?fs \<oplus> ?hd \<odot> ?fs"
using 6 8 by (metis (no_types) times_disjoint matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left)
finally show ?thesis
.
qed
have 13: "?ha \<odot> ?es \<preceq> ?ha"
proof -
have "?ha \<odot> ?b \<odot> ?ds \<odot> ?c \<preceq> ?hb \<odot> ?ds \<odot> ?c"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?hb \<odot> ?c"
using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist)
also have "... \<preceq> ?ha"
using 5 by simp
finally have "?ha \<odot> ?e \<preceq> ?ha"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc)
thus ?thesis
using 7 by (simp add: restrict_star_right_induct)
qed
have 14: "?hb \<odot> ?fs \<preceq> ?hb"
proof -
have "?hb \<odot> ?c \<odot> ?as \<odot> ?b \<preceq> ?ha \<odot> ?as \<odot> ?b"
using 5 by (metis matrix_semilattice_sup.le_supE matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?ha \<odot> ?b"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist)
also have "... \<preceq> ?hb"
using 5 by simp
finally have "?hb \<odot> ?f \<preceq> ?hb"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc)
thus ?thesis
using 1 3 7 by simp
qed
have 15: "?hc \<odot> ?es \<preceq> ?hc"
proof -
have "?hc \<odot> ?b \<odot> ?ds \<odot> ?c \<preceq> ?hd \<odot> ?ds \<odot> ?c"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?hd \<odot> ?c"
using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist)
also have "... \<preceq> ?hc"
using 5 by simp
finally have "?hc \<odot> ?e \<preceq> ?hc"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc)
thus ?thesis
using 4 7 by (simp add: restrict_star_right_induct)
qed
have 16: "?hd \<odot> ?fs \<preceq> ?hd"
proof -
have "?hd \<odot> ?c \<odot> ?as \<odot> ?b \<preceq> ?hc \<odot> ?as \<odot> ?b"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?hc \<odot> ?b"
using 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist)
also have "... \<preceq> ?hd"
using 5 by simp
finally have "?hd \<odot> ?f \<preceq> ?hd"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc)
thus ?thesis
using 1 4 7 by simp
qed
have 17: "?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<preceq> ?ha"
proof -
have "?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<preceq> ?hb \<odot> ?c \<odot> ?es"
using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist)
also have "... \<preceq> ?ha \<odot> ?es"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?ha"
using 13 by simp
finally show ?thesis
.
qed
have 18: "?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<preceq> ?hb"
proof -
have "?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<preceq> ?ha \<odot> ?b \<odot> ?fs"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist)
also have "... \<preceq> ?hb \<odot> ?fs"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?hb"
using 14 by simp
finally show ?thesis
by simp
qed
have 19: "?hd \<odot> ?ds \<odot> ?c \<odot> ?es \<preceq> ?hc"
proof -
have "?hd \<odot> ?ds \<odot> ?c \<odot> ?es \<preceq> ?hd \<odot> ?c \<odot> ?es"
using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist)
also have "... \<preceq> ?hc \<odot> ?es"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?hc"
using 15 by simp
finally show ?thesis
by simp
qed
have 20: "?hc \<odot> ?as \<odot> ?b \<odot> ?fs \<preceq> ?hd"
proof -
have "?hc \<odot> ?as \<odot> ?b \<odot> ?fs \<preceq> ?hc \<odot> ?b \<odot> ?fs"
using 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist)
also have "... \<preceq> ?hd \<odot> ?fs"
using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone)
also have "... \<preceq> ?hd"
using 16 by simp
finally show ?thesis
by simp
qed
have 21: "?ha \<odot> ?es \<oplus> ?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<preceq> ?ha"
using 13 17 matrix_semilattice_sup.le_supI by blast
have 22: "?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hb \<odot> ?fs \<preceq> ?hb"
using 14 18 matrix_semilattice_sup.le_supI by blast
have 23: "?hc \<odot> ?es \<oplus> ?hd \<odot> ?ds \<odot> ?c \<odot> ?es \<preceq> ?hc"
using 15 19 matrix_semilattice_sup.le_supI by blast
have 24: "?hc \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hd \<odot> ?fs \<preceq> ?hd"
using 16 20 matrix_semilattice_sup.le_supI by blast
have "zs\<langle>h\<rangle>?t \<odot> star_matrix' ?t g = zs\<langle>h\<rangle>?t \<odot> (?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs)"
by (metis star_matrix'.simps(2))
also have "... = zs\<langle>h\<rangle>?t \<odot> ?es \<oplus> zs\<langle>h\<rangle>?t \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> zs\<langle>h\<rangle>?t \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> zs\<langle>h\<rangle>?t \<odot> ?fs"
by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc)
also have "... = ?ha \<odot> ?es \<oplus> ?hc \<odot> ?es \<oplus> ?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hc \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hb \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?hd \<odot> ?ds \<odot> ?c \<odot> ?es \<oplus> ?hb \<odot> ?fs \<oplus> ?hd \<odot> ?fs"
using 9 10 11 12 by (simp add: matrix_semilattice_sup.sup_assoc)
also have "... = (?ha \<odot> ?es \<oplus> ?hb \<odot> ?ds \<odot> ?c \<odot> ?es) \<oplus> (?ha \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hb \<odot> ?fs) \<oplus> (?hc \<odot> ?es \<oplus> ?hd \<odot> ?ds \<odot> ?c \<odot> ?es) \<oplus> (?hc \<odot> ?as \<odot> ?b \<odot> ?fs \<oplus> ?hd \<odot> ?fs)"
using 9 10 11 12 by (simp only: matrix_semilattice_sup.sup_assoc matrix_semilattice_sup.sup_commute matrix_semilattice_sup.sup_left_commute)
also have "... \<preceq> ?ha \<oplus> ?hb \<oplus> ?hc \<oplus> ?hd"
using 21 22 23 24 matrix_semilattice_sup.sup.mono by blast
also have "... = zs\<langle>h\<rangle>?t"
using 2 by (metis restrict_nonempty)
finally show "zs\<langle>h\<rangle>?t \<odot> star_matrix' ?t g \<preceq> zs\<langle>h\<rangle>?t"
.
qed
qed
qed
qed
hence "\<forall>zs . distinct zs \<longrightarrow> (zs\<langle>x\<rangle>?e \<odot> y \<preceq> zs\<langle>x\<rangle>?e \<longrightarrow> zs\<langle>x\<rangle>?e \<odot> y\<^sup>\<odot> \<preceq> zs\<langle>x\<rangle>?e)"
by (simp add: enum_distinct restrict_all)
thus "x \<odot> y \<preceq> x \<longrightarrow> x \<odot> y\<^sup>\<odot> \<preceq> x"
by (metis restrict_all enum_distinct)
qed
subsection \<open>Matrices form a Stone-Kleene Relation Algebra\<close>
text \<open>
Matrices over Stone-Kleene relation algebras form a Stone-Kleene relation algebra.
It remains to prove the axiom about the interaction of Kleene star and double complement.
\<close>
interpretation matrix_stone_kleene_relation_algebra: stone_kleene_relation_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::stone_kleene_relation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix
proof
fix x :: "('a,'b) square"
let ?e = "enum_class.enum::'a list"
let ?o = "mone :: ('a,'b) square"
show "\<ominus>\<ominus>(x\<^sup>\<odot>) = (\<ominus>\<ominus>x)\<^sup>\<odot>"
- proof (rule matrix_order.antisym)
+ proof (rule matrix_order.order_antisym)
have "\<forall>g :: ('a,'b) square . distinct ?e \<longrightarrow> \<ominus>\<ominus>(star_matrix' ?e (\<ominus>\<ominus>g)) = star_matrix' ?e (\<ominus>\<ominus>g)"
proof (induct rule: list.induct)
case Nil thus ?case
by simp
next
case (Cons k s)
let ?t = "k#s"
assume 1: "\<forall>g :: ('a,'b) square . distinct s \<longrightarrow> \<ominus>\<ominus>(star_matrix' s (\<ominus>\<ominus>g)) = star_matrix' s (\<ominus>\<ominus>g)"
show "\<forall>g :: ('a,'b) square . distinct ?t \<longrightarrow> \<ominus>\<ominus>(star_matrix' ?t (\<ominus>\<ominus>g)) = star_matrix' ?t (\<ominus>\<ominus>g)"
proof (rule allI, rule impI)
fix g :: "('a,'b) square"
assume 2: "distinct ?t"
let ?r = "[k]"
let ?a = "?r\<langle>\<ominus>\<ominus>g\<rangle>?r"
let ?b = "?r\<langle>\<ominus>\<ominus>g\<rangle>s"
let ?c = "s\<langle>\<ominus>\<ominus>g\<rangle>?r"
let ?d = "s\<langle>\<ominus>\<ominus>g\<rangle>s"
let ?as = "?r\<langle>star o ?a\<rangle>?r"
let ?ds = "star_matrix' s ?d"
let ?e = "?a \<oplus> ?b \<odot> ?ds \<odot> ?c"
let ?es = "?r\<langle>star o ?e\<rangle>?r"
let ?f = "?d \<oplus> ?c \<odot> ?as \<odot> ?b"
let ?fs = "star_matrix' s ?f"
have "s\<langle>?ds\<rangle>s = ?ds \<and> s\<langle>?fs\<rangle>s = ?fs"
by (simp add: restrict_star)
have 3: "\<ominus>\<ominus>?a = ?a \<and> \<ominus>\<ominus>?b = ?b \<and> \<ominus>\<ominus>?c = ?c \<and> \<ominus>\<ominus>?d = ?d"
by (metis matrix_p_algebra.regular_closed_p restrict_pp)
hence 4: "\<ominus>\<ominus>?as = ?as"
by (metis pp_star_commute restrict_pp)
hence "\<ominus>\<ominus>?f = ?f"
using 3 by (metis matrix_stone_algebra.regular_closed_sup matrix_stone_relation_algebra.regular_mult_closed)
hence 5: "\<ominus>\<ominus>?fs = ?fs"
using 1 2 by (metis distinct.simps(2))
have 6: "\<ominus>\<ominus>?ds = ?ds"
using 1 2 by (simp add: restrict_pp)
hence "\<ominus>\<ominus>?e = ?e"
using 3 by (metis matrix_stone_algebra.regular_closed_sup matrix_stone_relation_algebra.regular_mult_closed)
hence 7: "\<ominus>\<ominus>?es = ?es"
by (metis pp_star_commute restrict_pp)
have "\<ominus>\<ominus>(star_matrix' ?t (\<ominus>\<ominus>g)) = \<ominus>\<ominus>(?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs)"
by (metis star_matrix'.simps(2))
also have "... = \<ominus>\<ominus>?es \<oplus> \<ominus>\<ominus>?as \<odot> \<ominus>\<ominus>?b \<odot> \<ominus>\<ominus>?fs \<oplus> \<ominus>\<ominus>?ds \<odot> \<ominus>\<ominus>?c \<odot> \<ominus>\<ominus>?es \<oplus> \<ominus>\<ominus>?fs"
by (simp add: matrix_stone_relation_algebra.pp_dist_comp)
also have "... = ?es \<oplus> ?as \<odot> ?b \<odot> ?fs \<oplus> ?ds \<odot> ?c \<odot> ?es \<oplus> ?fs"
using 3 4 5 6 7 by simp
finally show "\<ominus>\<ominus>(star_matrix' ?t (\<ominus>\<ominus>g)) = star_matrix' ?t (\<ominus>\<ominus>g)"
by (metis star_matrix'.simps(2))
qed
qed
hence "(\<ominus>\<ominus>x)\<^sup>\<odot> = \<ominus>\<ominus>((\<ominus>\<ominus>x)\<^sup>\<odot>)"
by (simp add: enum_distinct restrict_all)
thus "\<ominus>\<ominus>(x\<^sup>\<odot>) \<preceq> (\<ominus>\<ominus>x)\<^sup>\<odot>"
by (metis matrix_kleene_algebra.star.circ_isotone matrix_p_algebra.pp_increasing matrix_p_algebra.pp_isotone)
next
have "?o \<oplus> \<ominus>\<ominus>x \<odot> \<ominus>\<ominus>(x\<^sup>\<odot>) \<preceq> \<ominus>\<ominus>(x\<^sup>\<odot>)"
by (metis matrix_kleene_algebra.star_left_unfold_equal matrix_p_algebra.sup_pp_semi_commute matrix_stone_relation_algebra.pp_dist_comp)
thus "(\<ominus>\<ominus>x)\<^sup>\<odot> \<preceq> \<ominus>\<ominus>(x\<^sup>\<odot>)"
using matrix_kleene_algebra.star_left_induct by fastforce
qed
qed
end
diff --git a/thys/Stone_Relation_Algebras/Fixpoints.thy b/thys/Stone_Relation_Algebras/Fixpoints.thy
--- a/thys/Stone_Relation_Algebras/Fixpoints.thy
+++ b/thys/Stone_Relation_Algebras/Fixpoints.thy
@@ -1,624 +1,624 @@
(* Title: Fixpoints
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Fixpoints\<close>
text \<open>
This theory develops a fixpoint calculus based on partial orders.
Besides fixpoints we consider least prefixpoints and greatest postfixpoints of functions on a partial order.
We do not assume that the underlying structure is complete or that all functions are continuous or isotone.
Assumptions about the existence of fixpoints and necessary properties of the involved functions will be stated explicitly in each theorem.
This way, the results can be instantiated by various structures, such as complete lattices and Kleene algebras, which impose different kinds of restriction.
See, for example, \cite{AartsBackhouseBoitenDoornbosGasterenGeldropHoogendijkVoermansWoude1995,DaveyPriestley2002} for fixpoint calculi in complete lattices.
Our fixpoint calculus contains similar rules, in particular:
\begin{itemize}
\item unfold rule,
\item fixpoint operators preserve isotonicity,
\item square rule,
\item rolling rule,
\item various fusion rules,
\item exchange rule and
\item diagonal rule.
\end{itemize}
All of our rules are based on existence rather than completeness of the underlying structure.
We have applied results from this theory in \cite{Guttmann2012c} and subsequent papers for unifying and reasoning about the semantics of recursion in various relational and matrix-based computation models.
\<close>
theory Fixpoints
imports Stone_Algebras.Lattice_Basics
begin
text \<open>
The whole calculus is based on partial orders only.
\<close>
context order
begin
text \<open>
We first define when an element $x$ is a least/greatest (pre/post)fixpoint of a given function $f$.
\<close>
definition is_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where "is_fixpoint f x \<equiv> f x = x"
definition is_prefixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where "is_prefixpoint f x \<equiv> f x \<le> x"
definition is_postfixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where "is_postfixpoint f x \<equiv> f x \<ge> x"
definition is_least_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where "is_least_fixpoint f x \<equiv> f x = x \<and> (\<forall>y . f y = y \<longrightarrow> x \<le> y)"
definition is_greatest_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where "is_greatest_fixpoint f x \<equiv> f x = x \<and> (\<forall>y . f y = y \<longrightarrow> x \<ge> y)"
definition is_least_prefixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where "is_least_prefixpoint f x \<equiv> f x \<le> x \<and> (\<forall>y . f y \<le> y \<longrightarrow> x \<le> y)"
definition is_greatest_postfixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where "is_greatest_postfixpoint f x \<equiv> f x \<ge> x \<and> (\<forall>y . f y \<ge> y \<longrightarrow> x \<ge> y)"
text \<open>
Next follows the existence of the corresponding fixpoints for a given function $f$.
\<close>
definition has_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where "has_fixpoint f \<equiv> \<exists>x . is_fixpoint f x"
definition has_prefixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where "has_prefixpoint f \<equiv> \<exists>x . is_prefixpoint f x"
definition has_postfixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where "has_postfixpoint f \<equiv> \<exists>x . is_postfixpoint f x"
definition has_least_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where "has_least_fixpoint f \<equiv> \<exists>x . is_least_fixpoint f x"
definition has_greatest_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where "has_greatest_fixpoint f \<equiv> \<exists>x . is_greatest_fixpoint f x"
definition has_least_prefixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where "has_least_prefixpoint f \<equiv> \<exists>x . is_least_prefixpoint f x"
definition has_greatest_postfixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> bool" where "has_greatest_postfixpoint f \<equiv> \<exists>x . is_greatest_postfixpoint f x"
text \<open>
The actual least/greatest (pre/post)fixpoints of a given function $f$ are extracted by the following operators.
\<close>
definition the_least_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" ("\<mu> _" [201] 200) where "\<mu> f = (THE x . is_least_fixpoint f x)"
definition the_greatest_fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" ("\<nu> _" [201] 200) where "\<nu> f = (THE x . is_greatest_fixpoint f x)"
definition the_least_prefixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" ("p\<mu> _" [201] 200) where "p\<mu> f = (THE x . is_least_prefixpoint f x)"
definition the_greatest_postfixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" ("p\<nu> _" [201] 200) where "p\<nu> f = (THE x . is_greatest_postfixpoint f x)"
text \<open>
We start with basic consequences of the above definitions.
\<close>
lemma least_fixpoint_unique:
"has_least_fixpoint f \<Longrightarrow> \<exists>!x . is_least_fixpoint f x"
- using has_least_fixpoint_def is_least_fixpoint_def antisym by auto
+ using has_least_fixpoint_def is_least_fixpoint_def order.antisym by auto
lemma greatest_fixpoint_unique:
"has_greatest_fixpoint f \<Longrightarrow> \<exists>!x . is_greatest_fixpoint f x"
- using has_greatest_fixpoint_def is_greatest_fixpoint_def antisym by auto
+ using has_greatest_fixpoint_def is_greatest_fixpoint_def order.antisym by auto
lemma least_prefixpoint_unique:
"has_least_prefixpoint f \<Longrightarrow> \<exists>!x . is_least_prefixpoint f x"
- using has_least_prefixpoint_def is_least_prefixpoint_def antisym by auto
+ using has_least_prefixpoint_def is_least_prefixpoint_def order.antisym by auto
lemma greatest_postfixpoint_unique:
"has_greatest_postfixpoint f \<Longrightarrow> \<exists>!x . is_greatest_postfixpoint f x"
- using has_greatest_postfixpoint_def is_greatest_postfixpoint_def antisym by auto
+ using has_greatest_postfixpoint_def is_greatest_postfixpoint_def order.antisym by auto
lemma least_fixpoint:
"has_least_fixpoint f \<Longrightarrow> is_least_fixpoint f (\<mu> f)"
by (simp add: least_fixpoint_unique theI' the_least_fixpoint_def)
lemma greatest_fixpoint:
"has_greatest_fixpoint f \<Longrightarrow> is_greatest_fixpoint f (\<nu> f)"
by (simp add: greatest_fixpoint_unique theI' the_greatest_fixpoint_def)
lemma least_prefixpoint:
"has_least_prefixpoint f \<Longrightarrow> is_least_prefixpoint f (p\<mu> f)"
by (simp add: least_prefixpoint_unique theI' the_least_prefixpoint_def)
lemma greatest_postfixpoint:
"has_greatest_postfixpoint f \<Longrightarrow> is_greatest_postfixpoint f (p\<nu> f)"
by (simp add: greatest_postfixpoint_unique theI' the_greatest_postfixpoint_def)
lemma least_fixpoint_same:
"is_least_fixpoint f x \<Longrightarrow> x = \<mu> f"
- by (simp add: is_least_fixpoint_def antisym the_equality the_least_fixpoint_def)
+ by (simp add: is_least_fixpoint_def order.antisym the_equality the_least_fixpoint_def)
lemma greatest_fixpoint_same:
"is_greatest_fixpoint f x \<Longrightarrow> x = \<nu> f"
using greatest_fixpoint greatest_fixpoint_unique has_greatest_fixpoint_def by auto
lemma least_prefixpoint_same:
"is_least_prefixpoint f x \<Longrightarrow> x = p\<mu> f"
using has_least_prefixpoint_def least_prefixpoint least_prefixpoint_unique by blast
lemma greatest_postfixpoint_same:
"is_greatest_postfixpoint f x \<Longrightarrow> x = p\<nu> f"
using greatest_postfixpoint greatest_postfixpoint_unique has_greatest_postfixpoint_def by auto
lemma least_fixpoint_char:
"is_least_fixpoint f x \<longleftrightarrow> has_least_fixpoint f \<and> x = \<mu> f"
using has_least_fixpoint_def least_fixpoint_same by auto
lemma least_prefixpoint_char:
"is_least_prefixpoint f x \<longleftrightarrow> has_least_prefixpoint f \<and> x = p\<mu> f"
using has_least_prefixpoint_def least_prefixpoint_same by auto
lemma greatest_fixpoint_char:
"is_greatest_fixpoint f x \<longleftrightarrow> has_greatest_fixpoint f \<and> x = \<nu> f"
using greatest_fixpoint_same has_greatest_fixpoint_def by auto
lemma greatest_postfixpoint_char:
"is_greatest_postfixpoint f x \<longleftrightarrow> has_greatest_postfixpoint f \<and> x = p\<nu> f"
using greatest_postfixpoint_same has_greatest_postfixpoint_def by auto
text \<open>
Next come the unfold rules for least/greatest (pre/post)fixpoints.
\<close>
lemma mu_unfold:
"has_least_fixpoint f \<Longrightarrow> f (\<mu> f) = \<mu> f"
using is_least_fixpoint_def least_fixpoint by auto
lemma pmu_unfold:
"has_least_prefixpoint f \<Longrightarrow> f (p\<mu> f) \<le> p\<mu> f"
using is_least_prefixpoint_def least_prefixpoint by blast
lemma nu_unfold:
"has_greatest_fixpoint f \<Longrightarrow> \<nu> f = f (\<nu> f)"
by (metis is_greatest_fixpoint_def greatest_fixpoint)
lemma pnu_unfold:
"has_greatest_postfixpoint f \<Longrightarrow> p\<nu> f \<le> f (p\<nu> f)"
using greatest_postfixpoint is_greatest_postfixpoint_def by auto
text \<open>
Pre-/postfixpoints of isotone functions are fixpoints.
\<close>
lemma least_prefixpoint_fixpoint:
"has_least_prefixpoint f \<Longrightarrow> isotone f \<Longrightarrow> is_least_fixpoint f (p\<mu> f)"
- using is_least_fixpoint_def is_least_prefixpoint_def least_prefixpoint antisym isotone_def by auto
+ using is_least_fixpoint_def is_least_prefixpoint_def least_prefixpoint order.antisym isotone_def by auto
lemma pmu_mu:
"has_least_prefixpoint f \<Longrightarrow> isotone f \<Longrightarrow> p\<mu> f = \<mu> f"
by (simp add: least_fixpoint_same least_prefixpoint_fixpoint)
lemma greatest_postfixpoint_fixpoint:
"has_greatest_postfixpoint f \<Longrightarrow> isotone f \<Longrightarrow> is_greatest_fixpoint f (p\<nu> f)"
- using greatest_postfixpoint is_greatest_fixpoint_def is_greatest_postfixpoint_def antisym isotone_def by auto
+ using greatest_postfixpoint is_greatest_fixpoint_def is_greatest_postfixpoint_def order.antisym isotone_def by auto
lemma pnu_nu:
"has_greatest_postfixpoint f \<Longrightarrow> isotone f \<Longrightarrow> p\<nu> f = \<nu> f"
by (simp add: greatest_fixpoint_same greatest_postfixpoint_fixpoint)
text \<open>
The fixpoint operators preserve isotonicity.
\<close>
lemma pmu_isotone:
"has_least_prefixpoint f \<Longrightarrow> has_least_prefixpoint g \<Longrightarrow> f \<le>\<le> g \<Longrightarrow> p\<mu> f \<le> p\<mu> g"
by (metis is_least_prefixpoint_def least_prefixpoint order_trans lifted_less_eq_def)
lemma mu_isotone:
"has_least_prefixpoint f \<Longrightarrow> has_least_prefixpoint g \<Longrightarrow> isotone f \<Longrightarrow> isotone g \<Longrightarrow> f \<le>\<le> g \<Longrightarrow> \<mu> f \<le> \<mu> g"
using pmu_isotone pmu_mu by fastforce
lemma pnu_isotone:
"has_greatest_postfixpoint f \<Longrightarrow> has_greatest_postfixpoint g \<Longrightarrow> f \<le>\<le> g \<Longrightarrow> p\<nu> f \<le> p\<nu> g"
by (metis greatest_postfixpoint is_greatest_postfixpoint_def order_trans lifted_less_eq_def)
lemma nu_isotone:
"has_greatest_postfixpoint f \<Longrightarrow> has_greatest_postfixpoint g \<Longrightarrow> isotone f \<Longrightarrow> isotone g \<Longrightarrow> f \<le>\<le> g \<Longrightarrow> \<nu> f \<le> \<nu> g"
using pnu_isotone pnu_nu by fastforce
text \<open>
The square rule for fixpoints of a function applied twice.
\<close>
lemma mu_square:
"isotone f \<Longrightarrow> has_least_fixpoint f \<Longrightarrow> has_least_fixpoint (f \<circ> f) \<Longrightarrow> \<mu> f = \<mu> (f \<circ> f)"
- by (metis (no_types, hide_lams) antisym is_least_fixpoint_def isotone_def least_fixpoint_char least_fixpoint_unique o_apply)
+ by (metis (no_types, hide_lams) order.antisym is_least_fixpoint_def isotone_def least_fixpoint_char least_fixpoint_unique o_apply)
lemma nu_square:
"isotone f \<Longrightarrow> has_greatest_fixpoint f \<Longrightarrow> has_greatest_fixpoint (f \<circ> f) \<Longrightarrow> \<nu> f = \<nu> (f \<circ> f)"
- by (metis (no_types, hide_lams) antisym is_greatest_fixpoint_def isotone_def greatest_fixpoint_char greatest_fixpoint_unique o_apply)
+ by (metis (no_types, hide_lams) order.antisym is_greatest_fixpoint_def isotone_def greatest_fixpoint_char greatest_fixpoint_unique o_apply)
text \<open>
The rolling rule for fixpoints of the composition of two functions.
\<close>
lemma mu_roll:
assumes "isotone g"
and "has_least_fixpoint (f \<circ> g)"
and "has_least_fixpoint (g \<circ> f)"
shows "\<mu> (g \<circ> f) = g (\<mu> (f \<circ> g))"
-proof (rule antisym)
+proof (rule order.antisym)
show "\<mu> (g \<circ> f) \<le> g (\<mu> (f \<circ> g))"
by (metis assms(2-3) comp_apply is_least_fixpoint_def least_fixpoint)
next
have "is_least_fixpoint (g \<circ> f) (\<mu> (g \<circ> f))"
by (simp add: assms(3) least_fixpoint)
thus "g (\<mu> (f \<circ> g)) \<le> \<mu> (g \<circ> f)"
by (metis (no_types) assms(1-2) comp_def is_least_fixpoint_def least_fixpoint isotone_def)
qed
lemma nu_roll:
assumes "isotone g"
and "has_greatest_fixpoint (f \<circ> g)"
and "has_greatest_fixpoint (g \<circ> f)"
shows "\<nu> (g \<circ> f) = g (\<nu> (f \<circ> g))"
-proof (rule antisym)
+proof (rule order.antisym)
have 1: "is_greatest_fixpoint (f \<circ> g) (\<nu> (f \<circ> g))"
by (simp add: assms(2) greatest_fixpoint)
have "is_greatest_fixpoint (g \<circ> f) (\<nu> (g \<circ> f))"
by (simp add: assms(3) greatest_fixpoint)
thus "\<nu> (g \<circ> f) \<le> g (\<nu> (f \<circ> g))"
using 1 by (metis (no_types) assms(1) comp_def is_greatest_fixpoint_def isotone_def)
next
show "g (\<nu> (f \<circ> g)) \<le> \<nu> (g \<circ> f)"
by (metis assms(2-3) comp_apply greatest_fixpoint is_greatest_fixpoint_def)
qed
text \<open>
Least (pre)fixpoints are below greatest (post)fixpoints.
\<close>
lemma mu_below_nu:
"has_least_fixpoint f \<Longrightarrow> has_greatest_fixpoint f \<Longrightarrow> \<mu> f \<le> \<nu> f"
using greatest_fixpoint is_greatest_fixpoint_def mu_unfold by auto
lemma pmu_below_pnu_fix:
"has_fixpoint f \<Longrightarrow> has_least_prefixpoint f \<Longrightarrow> has_greatest_postfixpoint f \<Longrightarrow> p\<mu> f \<le> p\<nu> f"
by (metis greatest_postfixpoint has_fixpoint_def is_fixpoint_def is_greatest_postfixpoint_def is_least_prefixpoint_def least_prefixpoint order_refl order_trans)
lemma pmu_below_pnu_iso:
"isotone f \<Longrightarrow> has_least_prefixpoint f \<Longrightarrow> has_greatest_postfixpoint f \<Longrightarrow> p\<mu> f \<le> p\<nu> f"
using greatest_postfixpoint_fixpoint is_greatest_fixpoint_def is_least_fixpoint_def least_prefixpoint_fixpoint by auto
text \<open>
Several variants of the fusion rule for fixpoints follow.
\<close>
lemma mu_fusion_1:
assumes "galois l u"
and "isotone h"
and "has_least_prefixpoint g"
and "has_least_fixpoint h"
and "l (g (u (\<mu> h))) \<le> h (l (u (\<mu> h)))"
shows "l (p\<mu> g) \<le> \<mu> h"
proof -
have "l (g (u (\<mu> h))) \<le> \<mu> h"
by (metis assms(1,2,4,5) galois_char isotone_def order_lesseq_imp mu_unfold)
thus "l (p\<mu> g) \<le> \<mu> h"
using assms(1,3) is_least_prefixpoint_def least_prefixpoint galois_def by auto
qed
lemma mu_fusion_2:
"galois l u \<Longrightarrow> isotone h \<Longrightarrow> has_least_prefixpoint g \<Longrightarrow> has_least_fixpoint h \<Longrightarrow> l \<circ> g \<le>\<le> h \<circ> l \<Longrightarrow> l (p\<mu> g) \<le> \<mu> h"
by (simp add: mu_fusion_1 lifted_less_eq_def)
lemma mu_fusion_equal_1:
"galois l u \<Longrightarrow> isotone g \<Longrightarrow> isotone h \<Longrightarrow> has_least_prefixpoint g \<Longrightarrow> has_least_fixpoint h \<Longrightarrow> l (g (u (\<mu> h))) \<le> h(l(u(\<mu> h))) \<Longrightarrow> l (g (p\<mu> g)) = h (l (p\<mu> g)) \<Longrightarrow> \<mu> h = l (p\<mu> g) \<and> \<mu> h = l (\<mu> g)"
- by (metis antisym least_fixpoint least_prefixpoint_fixpoint is_least_fixpoint_def mu_fusion_1 pmu_mu)
+ by (metis order.antisym least_fixpoint least_prefixpoint_fixpoint is_least_fixpoint_def mu_fusion_1 pmu_mu)
lemma mu_fusion_equal_2:
"galois l u \<Longrightarrow> isotone h \<Longrightarrow> has_least_prefixpoint g \<Longrightarrow> has_least_prefixpoint h \<Longrightarrow> l (g (u (\<mu> h))) \<le> h (l (u (\<mu> h))) \<and> l (g (p\<mu> g)) = h (l (p\<mu> g)) \<longrightarrow> p\<mu> h = l (p\<mu> g) \<and> \<mu> h = l (p\<mu> g)"
- by (metis is_least_prefixpoint_def least_fixpoint_char least_prefixpoint least_prefixpoint_fixpoint antisym galois_char isotone_def mu_fusion_1)
+ by (metis is_least_prefixpoint_def least_fixpoint_char least_prefixpoint least_prefixpoint_fixpoint order.antisym galois_char isotone_def mu_fusion_1)
lemma mu_fusion_equal_3:
assumes "galois l u"
and "isotone g"
and "isotone h"
and "has_least_prefixpoint g"
and "has_least_fixpoint h"
and "l \<circ> g = h \<circ> l"
shows "\<mu> h = l (p\<mu> g)"
and "\<mu> h = l (\<mu> g)"
proof -
have "\<forall>x . l (g x) = h (l x)"
using assms(6) comp_eq_elim by blast
thus "\<mu> h = l (p\<mu> g)"
using assms(1-5) mu_fusion_equal_1 by auto
thus "\<mu> h = l (\<mu> g)"
by (simp add: assms(2,4) pmu_mu)
qed
lemma mu_fusion_equal_4:
assumes "galois l u"
and "isotone h"
and "has_least_prefixpoint g"
and "has_least_prefixpoint h"
and "l \<circ> g = h \<circ> l"
shows "p\<mu> h = l (p\<mu> g)"
and "\<mu> h = l (p\<mu> g)"
proof -
have "\<forall>x . l (g x) = h (l x)"
using assms(5) comp_eq_elim by blast
thus "p\<mu> h = l (p\<mu> g)"
using assms(1-4) mu_fusion_equal_2 by auto
thus "\<mu> h = l (p\<mu> g)"
by (simp add: assms(2,4) pmu_mu)
qed
lemma nu_fusion_1:
assumes "galois l u"
and "isotone h"
and "has_greatest_postfixpoint g"
and "has_greatest_fixpoint h"
and "h (u (l (\<nu> h))) \<le> u (g (l (\<nu> h)))"
shows "\<nu> h \<le> u (p\<nu> g)"
proof -
have "\<nu> h \<le> u (g (l (\<nu> h)))"
by (metis assms(1,2,4,5) order_trans galois_char isotone_def nu_unfold)
thus "\<nu> h \<le> u (p\<nu> g)"
by (metis assms(1,3) greatest_postfixpoint is_greatest_postfixpoint_def ord.galois_def)
qed
lemma nu_fusion_2:
"galois l u \<Longrightarrow> isotone h \<Longrightarrow> has_greatest_postfixpoint g \<Longrightarrow> has_greatest_fixpoint h \<Longrightarrow> h \<circ> u \<le>\<le> u \<circ> g \<Longrightarrow> \<nu> h \<le> u (p\<nu> g)"
by (simp add: nu_fusion_1 lifted_less_eq_def)
lemma nu_fusion_equal_1:
"galois l u \<Longrightarrow> isotone g \<Longrightarrow> isotone h \<Longrightarrow> has_greatest_postfixpoint g \<Longrightarrow> has_greatest_fixpoint h \<Longrightarrow> h (u (l (\<nu> h))) \<le> u (g (l (\<nu> h))) \<Longrightarrow> h (u (p\<nu> g)) = u (g (p\<nu> g)) \<Longrightarrow> \<nu> h = u (p\<nu> g) \<and> \<nu> h = u (\<nu> g)"
- by (metis greatest_fixpoint_char greatest_postfixpoint_fixpoint is_greatest_fixpoint_def antisym nu_fusion_1)
+ by (metis greatest_fixpoint_char greatest_postfixpoint_fixpoint is_greatest_fixpoint_def order.antisym nu_fusion_1)
lemma nu_fusion_equal_2:
"galois l u \<Longrightarrow> isotone h \<Longrightarrow> has_greatest_postfixpoint g \<Longrightarrow> has_greatest_postfixpoint h \<Longrightarrow> h (u (l (\<nu> h))) \<le> u (g (l (\<nu> h))) \<and> h (u (p\<nu> g)) = u (g (p\<nu> g)) \<Longrightarrow> p\<nu> h = u (p\<nu> g) \<and> \<nu> h = u (p\<nu> g)"
- by (metis greatest_fixpoint_char greatest_postfixpoint greatest_postfixpoint_fixpoint is_greatest_postfixpoint_def antisym galois_char nu_fusion_1 isotone_def)
+ by (metis greatest_fixpoint_char greatest_postfixpoint greatest_postfixpoint_fixpoint is_greatest_postfixpoint_def order.antisym galois_char nu_fusion_1 isotone_def)
lemma nu_fusion_equal_3:
assumes "galois l u"
and "isotone g"
and "isotone h"
and "has_greatest_postfixpoint g"
and "has_greatest_fixpoint h"
and "h \<circ> u = u \<circ> g"
shows "\<nu> h = u (p\<nu> g)"
and "\<nu> h = u (\<nu> g)"
proof -
have "\<forall>x . u (g x) = h (u x)"
using assms(6) comp_eq_dest by fastforce
thus "\<nu> h = u (p\<nu> g)"
using assms(1-5) nu_fusion_equal_1 by auto
thus "\<nu> h = u (\<nu> g)"
by (simp add: assms(2-4) pnu_nu)
qed
lemma nu_fusion_equal_4:
assumes "galois l u"
and "isotone h"
and "has_greatest_postfixpoint g"
and "has_greatest_postfixpoint h"
and "h \<circ> u = u \<circ> g"
shows "p\<nu> h = u (p\<nu> g)"
and "\<nu> h = u (p\<nu> g)"
proof -
have "\<forall>x . u (g x) = h (u x)"
using assms(5) comp_eq_dest by fastforce
thus "p\<nu> h = u (p\<nu> g)"
using assms(1-4) nu_fusion_equal_2 by auto
thus "\<nu> h = u (p\<nu> g)"
by (simp add: assms(2,4) pnu_nu)
qed
text \<open>
Next come the exchange rules for replacing the first/second function in a composition.
\<close>
lemma mu_exchange_1:
assumes "galois l u"
and "isotone g"
and "isotone h"
and "has_least_prefixpoint (l \<circ> h)"
and "has_least_prefixpoint (h \<circ> g)"
and "has_least_fixpoint (g \<circ> h)"
and "l \<circ> h \<circ> g \<le>\<le> g \<circ> h \<circ> l"
shows "\<mu> (l \<circ> h) \<le> \<mu> (g \<circ> h)"
proof -
have 1: "l \<circ> (h \<circ> g) \<le>\<le> (g \<circ> h) \<circ> l"
by (simp add: assms(7) rewriteL_comp_comp)
have "(l \<circ> h) (\<mu> (g \<circ> h)) = l (\<mu> (h \<circ> g))"
by (metis assms(2,3,5,6) comp_apply least_fixpoint_char least_prefixpoint_fixpoint isotone_def mu_roll)
also have "... \<le> \<mu> (g \<circ> h)"
using 1 by (metis assms(1-3,5,6) comp_apply least_fixpoint_char least_prefixpoint_fixpoint isotone_def mu_fusion_2)
finally have "p\<mu> (l \<circ> h) \<le> \<mu> (g \<circ> h)"
using assms(4) is_least_prefixpoint_def least_prefixpoint by blast
thus "\<mu> (l \<circ> h) \<le> \<mu> (g \<circ> h)"
by (metis assms(1,3,4) galois_char isotone_def least_fixpoint_char least_prefixpoint_fixpoint o_apply)
qed
lemma mu_exchange_2:
assumes "galois l u"
and "isotone g"
and "isotone h"
and "has_least_prefixpoint (l \<circ> h)"
and "has_least_prefixpoint (h \<circ> l)"
and "has_least_prefixpoint (h \<circ> g)"
and "has_least_fixpoint (g \<circ> h)"
and "has_least_fixpoint (h \<circ> g)"
and "l \<circ> h \<circ> g \<le>\<le> g \<circ> h \<circ> l"
shows "\<mu> (h \<circ> l) \<le> \<mu> (h \<circ> g)"
proof -
have "\<mu> (h \<circ> l) = h (\<mu> (l \<circ> h))"
by (metis (no_types, lifting) assms(1,3-5) galois_char isotone_def least_fixpoint_char least_prefixpoint_fixpoint mu_roll o_apply)
also have "... \<le> h (\<mu> (g \<circ> h))"
using assms(1-4,6,7,9) isotone_def mu_exchange_1 by blast
also have "... = \<mu> (h \<circ> g)"
by (simp add: assms(3,7,8) mu_roll)
finally show ?thesis
.
qed
lemma mu_exchange_equal:
assumes "galois l u"
and "galois k t"
and "isotone h"
and "has_least_prefixpoint (l \<circ> h)"
and "has_least_prefixpoint (h \<circ> l)"
and "has_least_prefixpoint (k \<circ> h)"
and "has_least_prefixpoint (h \<circ> k)"
and "l \<circ> h \<circ> k = k \<circ> h \<circ> l"
shows "\<mu> (l \<circ> h) = \<mu> (k \<circ> h)"
and "\<mu> (h \<circ> l) = \<mu> (h \<circ> k)"
proof -
have 1: "has_least_fixpoint (k \<circ> h)"
using assms(2,3,6) least_fixpoint_char least_prefixpoint_fixpoint galois_char isotone_def by auto
have 2: "has_least_fixpoint (h \<circ> k)"
using assms(2,3,7) least_fixpoint_char least_prefixpoint_fixpoint galois_char isotone_def by auto
have 3: "has_least_fixpoint (l \<circ> h)"
using assms(1,3,4) least_fixpoint_char least_prefixpoint_fixpoint galois_char isotone_def by auto
have 4: "has_least_fixpoint (h \<circ> l)"
using assms(1,3,5) least_fixpoint_char least_prefixpoint_fixpoint galois_char isotone_def by auto
show "\<mu> (h \<circ> l) = \<mu> (h \<circ> k)"
- using 1 2 3 4 assms antisym galois_char lifted_reflexive mu_exchange_2 by auto
+ using 1 2 3 4 assms order.antisym galois_char lifted_reflexive mu_exchange_2 by auto
show "\<mu> (l \<circ> h) = \<mu> (k \<circ> h)"
- using 1 2 3 4 assms antisym galois_char lifted_reflexive mu_exchange_1 by auto
+ using 1 2 3 4 assms order.antisym galois_char lifted_reflexive mu_exchange_1 by auto
qed
lemma nu_exchange_1:
assumes "galois l u"
and "isotone g"
and "isotone h"
and "has_greatest_postfixpoint (u \<circ> h)"
and "has_greatest_postfixpoint (h \<circ> g)"
and "has_greatest_fixpoint (g \<circ> h)"
and "g \<circ> h \<circ> u \<le>\<le> u \<circ> h \<circ> g"
shows "\<nu> (g \<circ> h) \<le> \<nu> (u \<circ> h)"
proof -
have "(g \<circ> h) \<circ> u \<le>\<le> u \<circ> (h \<circ> g)"
by (simp add: assms(7) o_assoc)
hence "\<nu> (g \<circ> h) \<le> u (\<nu> (h \<circ> g))"
by (metis assms(1-3,5,6) greatest_fixpoint_char greatest_postfixpoint_fixpoint isotone_def nu_fusion_2 o_apply)
also have "... = (u \<circ> h) (\<nu> (g \<circ> h))"
by (metis assms(2,3,5,6) greatest_fixpoint_char greatest_postfixpoint_fixpoint isotone_def nu_roll o_apply)
finally have "\<nu> (g \<circ> h) \<le> p\<nu> (u \<circ> h)"
using assms(4) greatest_postfixpoint is_greatest_postfixpoint_def by blast
thus "\<nu> (g \<circ> h) \<le> \<nu> (u \<circ> h)"
using assms(1,3,4) galois_char greatest_fixpoint_char greatest_postfixpoint_fixpoint isotone_def by auto
qed
lemma nu_exchange_2:
assumes "galois l u"
and "isotone g"
and "isotone h"
and "has_greatest_postfixpoint (u \<circ> h)"
and "has_greatest_postfixpoint (h \<circ> u)"
and "has_greatest_postfixpoint (h \<circ> g)"
and "has_greatest_fixpoint (g \<circ> h)"
and "has_greatest_fixpoint (h \<circ> g)"
and "g \<circ> h \<circ> u \<le>\<le> u \<circ> h \<circ> g"
shows "\<nu> (h \<circ> g) \<le> \<nu> (h \<circ> u)"
proof -
have "\<nu> (h \<circ> g) = h (\<nu> (g \<circ> h))"
by (simp add: assms(3,7,8) nu_roll)
also have "... \<le> h (\<nu> (u \<circ> h))"
using assms(1-4,6,7,9) isotone_def nu_exchange_1 by blast
also have "... = \<nu> (h \<circ> u)"
by (metis (no_types, lifting) assms(1,3-5) galois_char greatest_fixpoint_char greatest_postfixpoint_fixpoint isotone_def nu_roll o_apply)
finally show "\<nu> (h \<circ> g) \<le> \<nu> (h \<circ> u)"
.
qed
lemma nu_exchange_equal:
assumes "galois l u"
and "galois k t"
and "isotone h"
and "has_greatest_postfixpoint (u \<circ> h)"
and "has_greatest_postfixpoint (h \<circ> u)"
and "has_greatest_postfixpoint (t \<circ> h)"
and "has_greatest_postfixpoint (h \<circ> t)"
and "u \<circ> h \<circ> t = t \<circ> h \<circ> u"
shows "\<nu> (u \<circ> h) = \<nu> (t \<circ> h)"
and "\<nu> (h \<circ> u) = \<nu> (h \<circ> t)"
proof -
have 1: "has_greatest_fixpoint (u \<circ> h)"
using assms(1,3,4) greatest_fixpoint_char greatest_postfixpoint_fixpoint galois_char isotone_def by auto
have 2: "has_greatest_fixpoint (h \<circ> u)"
using assms(1,3,5) greatest_fixpoint_char greatest_postfixpoint_fixpoint galois_char isotone_def by auto
have 3: "has_greatest_fixpoint (t \<circ> h)"
using assms(2,3,6) greatest_fixpoint_char greatest_postfixpoint_fixpoint galois_char isotone_def by auto
have 4: "has_greatest_fixpoint (h \<circ> t)"
using assms(2,3,7) greatest_fixpoint_char greatest_postfixpoint_fixpoint galois_char isotone_def by auto
show "\<nu> (u \<circ> h) = \<nu> (t \<circ> h)"
- using 1 2 3 4 assms antisym galois_char lifted_reflexive nu_exchange_1 by auto
+ using 1 2 3 4 assms order.antisym galois_char lifted_reflexive nu_exchange_1 by auto
show "\<nu> (h \<circ> u) = \<nu> (h \<circ> t)"
- using 1 2 3 4 assms antisym galois_char lifted_reflexive nu_exchange_2 by auto
+ using 1 2 3 4 assms order.antisym galois_char lifted_reflexive nu_exchange_2 by auto
qed
text \<open>
The following results generalise parts of \cite[Exercise 8.27]{DaveyPriestley2002} from continuous functions on complete partial orders to the present setting.
\<close>
lemma mu_commute_fixpoint_1:
"isotone f \<Longrightarrow> has_least_fixpoint (f \<circ> g) \<Longrightarrow> f \<circ> g = g \<circ> f \<Longrightarrow> is_fixpoint f (\<mu> (f \<circ> g))"
by (metis is_fixpoint_def mu_roll)
lemma mu_commute_fixpoint_2:
"isotone g \<Longrightarrow> has_least_fixpoint (f \<circ> g) \<Longrightarrow> f \<circ> g = g \<circ> f \<Longrightarrow> is_fixpoint g (\<mu> (f \<circ> g))"
by (simp add: mu_commute_fixpoint_1)
lemma mu_commute_least_fixpoint:
"isotone f \<Longrightarrow> isotone g \<Longrightarrow> has_least_fixpoint f \<Longrightarrow> has_least_fixpoint g \<Longrightarrow> has_least_fixpoint (f \<circ> g) \<Longrightarrow> f \<circ> g = g \<circ> f \<Longrightarrow> \<mu> (f \<circ> g) = \<mu> f \<Longrightarrow> \<mu> g \<le> \<mu> f"
by (metis is_least_fixpoint_def least_fixpoint mu_roll)
text \<open>
The converse of the preceding result is claimed for continuous $f$, $g$ on a complete partial order; it is unknown whether it holds without these additional assumptions.
\<close>
lemma nu_commute_fixpoint_1:
"isotone f \<Longrightarrow> has_greatest_fixpoint (f \<circ> g) \<Longrightarrow> f \<circ> g = g \<circ> f \<Longrightarrow> is_fixpoint f (\<nu>(f \<circ> g))"
by (metis is_fixpoint_def nu_roll)
lemma nu_commute_fixpoint_2:
"isotone g \<Longrightarrow> has_greatest_fixpoint (f \<circ> g) \<Longrightarrow> f \<circ> g = g \<circ> f \<Longrightarrow> is_fixpoint g (\<nu>(f \<circ> g))"
by (simp add: nu_commute_fixpoint_1)
lemma nu_commute_greatest_fixpoint:
"isotone f \<Longrightarrow> isotone g \<Longrightarrow> has_greatest_fixpoint f \<Longrightarrow> has_greatest_fixpoint g \<Longrightarrow> has_greatest_fixpoint (f \<circ> g) \<Longrightarrow> f \<circ> g = g \<circ> f \<Longrightarrow> \<nu> (f \<circ> g) = \<nu> f \<Longrightarrow> \<nu> f \<le> \<nu> g"
by (metis greatest_fixpoint is_greatest_fixpoint_def nu_roll)
text \<open>
Finally, we show a number of versions of the diagonal rule for functions with two arguments.
\<close>
lemma mu_diagonal_1:
assumes "isotone (\<lambda>x . \<mu> (\<lambda>y . f x y))"
and "\<forall>x . has_least_fixpoint (\<lambda>y . f x y)"
and "has_least_prefixpoint (\<lambda>x . \<mu> (\<lambda>y . f x y))"
shows "\<mu> (\<lambda>x . f x x) = \<mu> (\<lambda>x . \<mu> (\<lambda>y . f x y))"
proof -
let ?g = "\<lambda>x . \<mu> (\<lambda>y . f x y)"
have 1: "is_least_prefixpoint ?g (\<mu> ?g)"
using assms(1,3) least_prefixpoint pmu_mu by fastforce
have "f (\<mu> ?g) (\<mu> ?g) = \<mu> ?g"
by (metis (no_types, lifting) assms is_least_fixpoint_def least_fixpoint_char least_prefixpoint_fixpoint)
hence "is_least_fixpoint (\<lambda>x . f x x) (\<mu> ?g)"
using 1 assms(2) is_least_fixpoint_def is_least_prefixpoint_def least_fixpoint by auto
thus ?thesis
using least_fixpoint_same by simp
qed
lemma mu_diagonal_2:
"\<forall>x . isotone (\<lambda>y . f x y) \<and> isotone (\<lambda>y . f y x) \<and> has_least_prefixpoint (\<lambda>y . f x y) \<Longrightarrow> has_least_prefixpoint (\<lambda>x . \<mu> (\<lambda>y . f x y)) \<Longrightarrow> \<mu> (\<lambda>x . f x x) = \<mu> (\<lambda>x . \<mu> (\<lambda>y . f x y))"
apply (rule mu_diagonal_1)
using isotone_def lifted_less_eq_def mu_isotone apply simp
using has_least_fixpoint_def least_prefixpoint_fixpoint apply blast
by simp
lemma nu_diagonal_1:
assumes "isotone (\<lambda>x . \<nu> (\<lambda>y . f x y))"
and "\<forall>x . has_greatest_fixpoint (\<lambda>y . f x y)"
and "has_greatest_postfixpoint (\<lambda>x . \<nu> (\<lambda>y . f x y))"
shows "\<nu> (\<lambda>x . f x x) = \<nu> (\<lambda>x . \<nu> (\<lambda>y . f x y))"
proof -
let ?g = "\<lambda>x . \<nu> (\<lambda>y . f x y)"
have 1: "is_greatest_postfixpoint ?g (\<nu> ?g)"
using assms(1,3) greatest_postfixpoint pnu_nu by fastforce
have "f (\<nu> ?g) (\<nu> ?g) = \<nu> ?g"
by (metis (no_types, lifting) assms is_greatest_fixpoint_def greatest_fixpoint_char greatest_postfixpoint_fixpoint)
hence "is_greatest_fixpoint (\<lambda>x . f x x) (\<nu> ?g)"
using 1 assms(2) is_greatest_fixpoint_def is_greatest_postfixpoint_def greatest_fixpoint by auto
thus ?thesis
using greatest_fixpoint_same by simp
qed
lemma nu_diagonal_2:
"\<forall>x . isotone (\<lambda>y . f x y) \<and> isotone (\<lambda>y . f y x) \<and> has_greatest_postfixpoint (\<lambda>y . f x y) \<Longrightarrow> has_greatest_postfixpoint (\<lambda>x . \<nu> (\<lambda>y . f x y)) \<Longrightarrow> \<nu> (\<lambda>x . f x x) = \<nu> (\<lambda>x . \<nu> (\<lambda>y . f x y))"
apply (rule nu_diagonal_1)
using isotone_def lifted_less_eq_def nu_isotone apply simp
using has_greatest_fixpoint_def greatest_postfixpoint_fixpoint apply blast
by simp
end
end
diff --git a/thys/Stone_Relation_Algebras/Relation_Algebras.thy b/thys/Stone_Relation_Algebras/Relation_Algebras.thy
--- a/thys/Stone_Relation_Algebras/Relation_Algebras.thy
+++ b/thys/Stone_Relation_Algebras/Relation_Algebras.thy
@@ -1,1973 +1,1973 @@
(* Title: Relation Algebras
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Relation Algebras\<close>
text \<open>
The main structures introduced in this theory are Stone relation algebras.
They generalise Tarski's relation algebras \cite{Tarski1941} by weakening the Boolean algebra lattice structure to a Stone algebra.
Our motivation is to generalise relation-algebraic methods from unweighted graphs to weighted graphs.
Unlike unweighted graphs, weighted graphs do not form a Boolean algebra because there is no complement operation on the edge weights.
However, edge weights form a Stone algebra, and matrices over edge weights (that is, weighted graphs) form a Stone relation algebra.
The development in this theory is described in our papers \cite{Guttmann2016c,Guttmann2017b}.
Our main application there is the verification of Prim's minimum spanning tree algorithm.
Related work about fuzzy relations \cite{Goguen1967,Winter2001b}, Dedekind categories \cite{KawaharaFurusawa2001} and rough relations \cite{Comer1993,Pawlak1996} is also discussed in these papers.
In particular, Stone relation algebras do not assume that the underlying lattice is complete or a Heyting algebra, and they do not assume that composition has residuals.
We proceed in two steps.
First, we study the positive fragment in the form of single-object bounded distributive allegories \cite{FreydScedrov1990}.
Second, we extend these structures by a pseudocomplement operation with additional axioms to obtain Stone relation algebras.
Tarski's relation algebras are then obtained by a simple extension that imposes a Boolean algebra.
See, for example, \cite{BirdMoor1997,HirschHodkinson2002,Maddux1996,Maddux2006,Schmidt2011,SchmidtStroehlein1993} for further details about relations and relation algebras, and \cite{AndrekaMikulas2011,BredihinSchein1978} for algebras of relations with a smaller signature.
\<close>
theory Relation_Algebras
imports Stone_Algebras.P_Algebras Semirings
begin
subsection \<open>Single-Object Bounded Distributive Allegories\<close>
text \<open>
We start with developing bounded distributive allegories.
The following definitions concern properties of relations that require converse in addition to lattice and semiring operations.
\<close>
class conv =
fixes conv :: "'a \<Rightarrow> 'a" ("_\<^sup>T" [100] 100)
class bounded_distrib_allegory_signature = inf + sup + times + conv + bot + top + one + ord
begin
subclass times_one_ord .
subclass times_top .
abbreviation total_var :: "'a \<Rightarrow> bool" where "total_var x \<equiv> 1 \<le> x * x\<^sup>T"
abbreviation surjective_var :: "'a \<Rightarrow> bool" where "surjective_var x \<equiv> 1 \<le> x\<^sup>T * x"
abbreviation univalent :: "'a \<Rightarrow> bool" where "univalent x \<equiv> x\<^sup>T * x \<le> 1"
abbreviation injective :: "'a \<Rightarrow> bool" where "injective x \<equiv> x * x\<^sup>T \<le> 1"
abbreviation mapping :: "'a \<Rightarrow> bool" where "mapping x \<equiv> univalent x \<and> total x"
abbreviation bijective :: "'a \<Rightarrow> bool" where "bijective x \<equiv> injective x \<and> surjective x"
abbreviation point :: "'a \<Rightarrow> bool" where "point x \<equiv> vector x \<and> bijective x"
abbreviation arc :: "'a \<Rightarrow> bool" where "arc x \<equiv> bijective (x * top) \<and> bijective (x\<^sup>T * top)" (* earlier name: atom *)
abbreviation symmetric :: "'a \<Rightarrow> bool" where "symmetric x \<equiv> x\<^sup>T = x"
abbreviation antisymmetric :: "'a \<Rightarrow> bool" where "antisymmetric x \<equiv> x \<sqinter> x\<^sup>T \<le> 1"
abbreviation asymmetric :: "'a \<Rightarrow> bool" where "asymmetric x \<equiv> x \<sqinter> x\<^sup>T = bot"
abbreviation linear :: "'a \<Rightarrow> bool" where "linear x \<equiv> x \<squnion> x\<^sup>T = top"
abbreviation equivalence :: "'a \<Rightarrow> bool" where "equivalence x \<equiv> preorder x \<and> symmetric x"
abbreviation order :: "'a \<Rightarrow> bool" where "order x \<equiv> preorder x \<and> antisymmetric x"
abbreviation linear_order :: "'a \<Rightarrow> bool" where "linear_order x \<equiv> order x \<and> linear x"
end
text \<open>
We reuse the relation algebra axioms given in \cite{Maddux1996} except for one -- see lemma \<open>conv_complement_sub\<close> below -- which we replace with the Dedekind rule (or modular law) \<open>dedekind_1\<close>.
The Dedekind rule or variants of it are known from \cite{BirdMoor1997,FreydScedrov1990,KawaharaFurusawaMori1999,SchmidtStroehlein1993}.
We add \<open>comp_left_zero\<close>, which follows in relation algebras but not in the present setting.
The main change is that only a bounded distributive lattice is required, not a Boolean algebra.
\<close>
class bounded_distrib_allegory = bounded_distrib_lattice + times + one + conv +
assumes comp_associative : "(x * y) * z = x * (y * z)"
assumes comp_right_dist_sup : "(x \<squnion> y) * z = (x * z) \<squnion> (y * z)"
assumes comp_left_zero [simp]: "bot * x = bot"
assumes comp_left_one [simp]: "1 * x = x"
assumes conv_involutive [simp]: "x\<^sup>T\<^sup>T = x"
assumes conv_dist_sup : "(x \<squnion> y)\<^sup>T = x\<^sup>T \<squnion> y\<^sup>T"
assumes conv_dist_comp : "(x * y)\<^sup>T = y\<^sup>T * x\<^sup>T"
assumes dedekind_1 : "x * y \<sqinter> z \<le> x * (y \<sqinter> (x\<^sup>T * z))"
begin
subclass bounded_distrib_allegory_signature .
text \<open>
Many properties of relation algebras already follow in bounded distributive allegories.
\<close>
lemma conv_isotone:
"x \<le> y \<Longrightarrow> x\<^sup>T \<le> y\<^sup>T"
by (metis conv_dist_sup le_iff_sup)
lemma conv_order:
"x \<le> y \<longleftrightarrow> x\<^sup>T \<le> y\<^sup>T"
using conv_isotone by fastforce
lemma conv_bot [simp]:
"bot\<^sup>T = bot"
using conv_order bot_unique by force
lemma conv_top [simp]:
"top\<^sup>T = top"
- by (metis conv_involutive conv_order eq_iff top_greatest)
+ by (metis conv_involutive conv_order order.eq_iff top_greatest)
lemma conv_dist_inf:
"(x \<sqinter> y)\<^sup>T = x\<^sup>T \<sqinter> y\<^sup>T"
- apply (rule antisym)
+ apply (rule order.antisym)
using conv_order apply simp
by (metis conv_order conv_involutive inf.boundedI inf.cobounded1 inf.cobounded2)
lemma conv_inf_bot_iff:
"bot = x\<^sup>T \<sqinter> y \<longleftrightarrow> bot = x \<sqinter> y\<^sup>T"
using conv_dist_inf conv_bot by fastforce
lemma conv_one [simp]:
"1\<^sup>T = 1"
by (metis comp_left_one conv_dist_comp conv_involutive)
lemma comp_left_dist_sup:
"(x * y) \<squnion> (x * z) = x * (y \<squnion> z)"
by (metis comp_right_dist_sup conv_involutive conv_dist_sup conv_dist_comp)
lemma comp_right_isotone:
"x \<le> y \<Longrightarrow> z * x \<le> z * y"
by (simp add: comp_left_dist_sup sup.absorb_iff1)
lemma comp_left_isotone:
"x \<le> y \<Longrightarrow> x * z \<le> y * z"
by (metis comp_right_dist_sup le_iff_sup)
lemma comp_isotone:
"x \<le> y \<Longrightarrow> w \<le> z \<Longrightarrow> x * w \<le> y * z"
using comp_left_isotone comp_right_isotone order.trans by blast
lemma comp_left_subdist_inf:
"(x \<sqinter> y) * z \<le> x * z \<sqinter> y * z"
by (simp add: comp_left_isotone)
lemma comp_left_increasing_sup:
"x * y \<le> (x \<squnion> z) * y"
by (simp add: comp_left_isotone)
lemma comp_right_subdist_inf:
"x * (y \<sqinter> z) \<le> x * y \<sqinter> x * z"
by (simp add: comp_right_isotone)
lemma comp_right_increasing_sup:
"x * y \<le> x * (y \<squnion> z)"
by (simp add: comp_right_isotone)
lemma comp_right_zero [simp]:
"x * bot = bot"
by (metis comp_left_zero conv_dist_comp conv_involutive)
lemma comp_right_one [simp]:
"x * 1 = x"
by (metis comp_left_one conv_dist_comp conv_involutive)
lemma comp_left_conjugate:
"conjugate (\<lambda>y . x * y) (\<lambda>y . x\<^sup>T * y)"
apply (unfold conjugate_def, intro allI)
by (metis comp_right_zero bot.extremum_unique conv_involutive dedekind_1 inf.commute)
lemma comp_right_conjugate:
"conjugate (\<lambda>y . y * x) (\<lambda>y . y * x\<^sup>T)"
apply (unfold conjugate_def, intro allI)
by (metis comp_left_conjugate[unfolded conjugate_def] conv_inf_bot_iff conv_dist_comp conv_involutive)
text \<open>
We still obtain a semiring structure.
\<close>
subclass bounded_idempotent_semiring
by (unfold_locales)
(auto simp: comp_right_isotone comp_right_dist_sup comp_associative comp_left_dist_sup)
sublocale inf: semiring_0 sup bot inf
by (unfold_locales, auto simp: inf_sup_distrib2 inf_sup_distrib1 inf_assoc)
lemma schroeder_1:
"x * y \<sqinter> z = bot \<longleftrightarrow> x\<^sup>T * z \<sqinter> y = bot"
using abel_semigroup.commute comp_left_conjugate conjugate_def inf.abel_semigroup_axioms by fastforce
lemma schroeder_2:
"x * y \<sqinter> z = bot \<longleftrightarrow> z * y\<^sup>T \<sqinter> x = bot"
by (metis comp_right_conjugate conjugate_def inf_commute)
lemma comp_additive:
"additive (\<lambda>y . x * y) \<and> additive (\<lambda>y . x\<^sup>T * y) \<and> additive (\<lambda>y . y * x) \<and> additive (\<lambda>y . y * x\<^sup>T)"
by (simp add: comp_left_dist_sup additive_def comp_right_dist_sup)
lemma dedekind_2:
"y * x \<sqinter> z \<le> (y \<sqinter> (z * x\<^sup>T)) * x"
by (metis conv_dist_inf conv_order conv_dist_comp dedekind_1)
text \<open>
The intersection with a vector can still be exported from the first argument of a composition, and many other properties of vectors and covectors continue to hold.
\<close>
lemma vector_inf_comp:
"vector x \<Longrightarrow> (x \<sqinter> y) * z = x \<sqinter> (y * z)"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis comp_left_subdist_inf comp_right_isotone inf.sup_left_isotone order_lesseq_imp top_greatest)
by (metis comp_left_isotone comp_right_isotone dedekind_2 inf_commute inf_mono order_refl order_trans top_greatest)
lemma vector_inf_closed:
"vector x \<Longrightarrow> vector y \<Longrightarrow> vector (x \<sqinter> y)"
by (simp add: vector_inf_comp)
lemma vector_inf_one_comp:
"vector x \<Longrightarrow> (x \<sqinter> 1) * y = x \<sqinter> y"
by (simp add: vector_inf_comp)
lemma covector_inf_comp_1:
assumes "vector x"
shows "(y \<sqinter> x\<^sup>T) * z = (y \<sqinter> x\<^sup>T) * (x \<sqinter> z)"
proof -
have "(y \<sqinter> x\<^sup>T) * z \<le> (y \<sqinter> x\<^sup>T) * (z \<sqinter> ((y\<^sup>T \<sqinter> x) * top))"
by (metis inf_top_right dedekind_1 conv_dist_inf conv_involutive)
also have "... \<le> (y \<sqinter> x\<^sup>T) * (x \<sqinter> z)"
by (metis assms comp_left_isotone comp_right_isotone inf_le2 inf_mono order_refl inf_commute)
finally show ?thesis
- by (simp add: comp_right_isotone antisym)
+ by (simp add: comp_right_isotone order.antisym)
qed
lemma covector_inf_comp_2:
assumes "vector x"
shows "y * (x \<sqinter> z) = (y \<sqinter> x\<^sup>T) * (x \<sqinter> z)"
proof -
have "y * (x \<sqinter> z) \<le> (y \<sqinter> (top * (x \<sqinter> z)\<^sup>T)) * (x \<sqinter> z)"
by (metis dedekind_2 inf_top_right)
also have "... \<le> (y \<sqinter> x\<^sup>T) * (x \<sqinter> z)"
by (metis assms comp_left_isotone conv_dist_comp conv_order conv_top eq_refl inf_le1 inf_mono)
finally show ?thesis
- using comp_left_subdist_inf antisym by auto
+ using comp_left_subdist_inf order.antisym by auto
qed
lemma covector_inf_comp_3:
"vector x \<Longrightarrow> (y \<sqinter> x\<^sup>T) * z = y * (x \<sqinter> z)"
by (metis covector_inf_comp_1 covector_inf_comp_2)
lemma covector_inf_closed:
"covector x \<Longrightarrow> covector y \<Longrightarrow> covector (x \<sqinter> y)"
- by (metis comp_right_subdist_inf inf.antisym top_left_mult_increasing)
+ by (metis comp_right_subdist_inf order.antisym top_left_mult_increasing)
lemma vector_conv_covector:
"vector v \<longleftrightarrow> covector (v\<^sup>T)"
by (metis conv_dist_comp conv_involutive conv_top)
lemma covector_conv_vector:
"covector v \<longleftrightarrow> vector (v\<^sup>T)"
by (simp add: vector_conv_covector)
lemma covector_comp_inf:
"covector z \<Longrightarrow> x * (y \<sqinter> z) = x * y \<sqinter> z"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis comp_isotone comp_right_subdist_inf inf.boundedE inf.boundedI inf.cobounded2 top.extremum)
by (metis comp_left_isotone comp_right_isotone dedekind_1 inf_commute inf_mono order_refl order_trans top_greatest)
lemma vector_restrict_comp_conv:
"vector x \<Longrightarrow> x \<sqinter> y \<le> x\<^sup>T * y"
by (metis covector_inf_comp_3 eq_refl inf.sup_monoid.add_commute inf_top_right le_supE sup.orderE top_left_mult_increasing)
lemma covector_restrict_comp_conv:
"covector x \<Longrightarrow> y \<sqinter> x \<le> y * x\<^sup>T"
by (metis conv_dist_comp conv_dist_inf conv_order conv_top inf.sup_monoid.add_commute vector_restrict_comp_conv)
lemma covector_comp_inf_1:
"covector x \<Longrightarrow> (y \<sqinter> x) * z = y * (x\<^sup>T \<sqinter> z)"
using covector_conv_vector covector_inf_comp_3 by fastforce
text \<open>
We still have two ways to represent surjectivity and totality.
\<close>
lemma surjective_var:
"surjective x \<longleftrightarrow> surjective_var x"
proof
assume "surjective x"
thus "surjective_var x"
by (metis dedekind_2 comp_left_one inf_absorb2 top_greatest)
next
assume "surjective_var x"
hence "x\<^sup>T * (x * top) = top"
by (metis comp_left_isotone comp_associative comp_left_one top_le)
thus "surjective x"
by (metis comp_right_isotone conv_top conv_dist_comp conv_involutive top_greatest top_le)
qed
lemma total_var:
"total x \<longleftrightarrow> total_var x"
by (metis conv_top conv_dist_comp conv_involutive surjective_var)
lemma surjective_conv_total:
"surjective x \<longleftrightarrow> total (x\<^sup>T)"
by (metis conv_top conv_dist_comp conv_involutive)
lemma total_conv_surjective:
"total x \<longleftrightarrow> surjective (x\<^sup>T)"
by (simp add: surjective_conv_total)
lemma injective_conv_univalent:
"injective x \<longleftrightarrow> univalent (x\<^sup>T)"
by simp
lemma univalent_conv_injective:
"univalent x \<longleftrightarrow> injective (x\<^sup>T)"
by simp
text \<open>
We continue with studying further closure properties.
\<close>
lemma univalent_bot_closed:
"univalent bot"
by simp
lemma univalent_one_closed:
"univalent 1"
by simp
lemma univalent_inf_closed:
"univalent x \<Longrightarrow> univalent (x \<sqinter> y)"
by (metis comp_left_subdist_inf comp_right_subdist_inf conv_dist_inf inf.cobounded1 order_lesseq_imp)
lemma univalent_mult_closed:
assumes "univalent x"
and "univalent y"
shows "univalent (x * y)"
proof -
have "(x * y)\<^sup>T * x \<le> y\<^sup>T"
by (metis assms(1) comp_left_isotone comp_right_one conv_one conv_order comp_associative conv_dist_comp conv_involutive)
thus ?thesis
by (metis assms(2) comp_left_isotone comp_associative dual_order.trans)
qed
lemma injective_bot_closed:
"injective bot"
by simp
lemma injective_one_closed:
"injective 1"
by simp
lemma injective_inf_closed:
"injective x \<Longrightarrow> injective (x \<sqinter> y)"
by (metis conv_dist_inf injective_conv_univalent univalent_inf_closed)
lemma injective_mult_closed:
"injective x \<Longrightarrow> injective y \<Longrightarrow> injective (x * y)"
by (metis injective_conv_univalent conv_dist_comp univalent_mult_closed)
lemma mapping_one_closed:
"mapping 1"
by simp
lemma mapping_mult_closed:
"mapping x \<Longrightarrow> mapping y \<Longrightarrow> mapping (x * y)"
by (simp add: comp_associative univalent_mult_closed)
lemma bijective_one_closed:
"bijective 1"
by simp
lemma bijective_mult_closed:
"bijective x \<Longrightarrow> bijective y \<Longrightarrow> bijective (x * y)"
by (metis injective_mult_closed comp_associative)
lemma bijective_conv_mapping:
"bijective x \<longleftrightarrow> mapping (x\<^sup>T)"
by (simp add: surjective_conv_total)
lemma mapping_conv_bijective:
"mapping x \<longleftrightarrow> bijective (x\<^sup>T)"
by (simp add: total_conv_surjective)
lemma reflexive_inf_closed:
"reflexive x \<Longrightarrow> reflexive y \<Longrightarrow> reflexive (x \<sqinter> y)"
by simp
lemma reflexive_conv_closed:
"reflexive x \<Longrightarrow> reflexive (x\<^sup>T)"
using conv_isotone by force
lemma coreflexive_inf_closed:
"coreflexive x \<Longrightarrow> coreflexive (x \<sqinter> y)"
by (simp add: le_infI1)
lemma coreflexive_conv_closed:
"coreflexive x \<Longrightarrow> coreflexive (x\<^sup>T)"
using conv_order by force
lemma coreflexive_symmetric:
"coreflexive x \<Longrightarrow> symmetric x"
by (metis comp_right_one comp_right_subdist_inf conv_dist_inf conv_dist_comp conv_involutive dedekind_1 inf.absorb1 inf_absorb2)
lemma transitive_inf_closed:
"transitive x \<Longrightarrow> transitive y \<Longrightarrow> transitive (x \<sqinter> y)"
by (meson comp_left_subdist_inf inf.cobounded1 inf.sup_mono inf_le2 mult_right_isotone order.trans)
lemma transitive_conv_closed:
"transitive x \<Longrightarrow> transitive (x\<^sup>T)"
using conv_order conv_dist_comp by fastforce
lemma dense_conv_closed:
"dense_rel x \<Longrightarrow> dense_rel (x\<^sup>T)"
using conv_order conv_dist_comp by fastforce
lemma idempotent_conv_closed:
"idempotent x \<Longrightarrow> idempotent (x\<^sup>T)"
by (metis conv_dist_comp)
lemma preorder_inf_closed:
"preorder x \<Longrightarrow> preorder y \<Longrightarrow> preorder (x \<sqinter> y)"
using transitive_inf_closed by auto
lemma preorder_conv_closed:
"preorder x \<Longrightarrow> preorder (x\<^sup>T)"
by (simp add: reflexive_conv_closed transitive_conv_closed)
lemma symmetric_bot_closed:
"symmetric bot"
by simp
lemma symmetric_one_closed:
"symmetric 1"
by simp
lemma symmetric_top_closed:
"symmetric top"
by simp
lemma symmetric_inf_closed:
"symmetric x \<Longrightarrow> symmetric y \<Longrightarrow> symmetric (x \<sqinter> y)"
by (simp add: conv_dist_inf)
lemma symmetric_sup_closed:
"symmetric x \<Longrightarrow> symmetric y \<Longrightarrow> symmetric (x \<squnion> y)"
by (simp add: conv_dist_sup)
lemma symmetric_conv_closed:
"symmetric x \<Longrightarrow> symmetric (x\<^sup>T)"
by simp
lemma one_inf_conv:
"1 \<sqinter> x = 1 \<sqinter> x\<^sup>T"
by (metis conv_dist_inf coreflexive_symmetric inf.cobounded1 symmetric_one_closed)
lemma antisymmetric_bot_closed:
"antisymmetric bot"
by simp
lemma antisymmetric_one_closed:
"antisymmetric 1"
by simp
lemma antisymmetric_inf_closed:
"antisymmetric x \<Longrightarrow> antisymmetric (x \<sqinter> y)"
by (rule order_trans[where y="x \<sqinter> x\<^sup>T"]) (simp_all add: conv_isotone inf.coboundedI2 inf.sup_assoc)
lemma antisymmetric_conv_closed:
"antisymmetric x \<Longrightarrow> antisymmetric (x\<^sup>T)"
by (simp add: inf_commute)
lemma asymmetric_bot_closed:
"asymmetric bot"
by simp
lemma asymmetric_inf_closed:
"asymmetric x \<Longrightarrow> asymmetric (x \<sqinter> y)"
by (metis conv_dist_inf inf.mult_zero_left inf.left_commute inf_assoc)
lemma asymmetric_conv_closed:
"asymmetric x \<Longrightarrow> asymmetric (x\<^sup>T)"
by (simp add: inf_commute)
lemma linear_top_closed:
"linear top"
by simp
lemma linear_sup_closed:
"linear x \<Longrightarrow> linear (x \<squnion> y)"
by (metis conv_dist_sup sup_assoc sup_commute sup_top_right)
lemma linear_reflexive:
"linear x \<Longrightarrow> reflexive x"
by (metis one_inf_conv inf.distrib_left inf.cobounded2 inf.orderE reflexive_top_closed sup.idem)
lemma linear_conv_closed:
"linear x \<Longrightarrow> linear (x\<^sup>T)"
by (simp add: sup_commute)
lemma linear_comp_closed:
assumes "linear x"
and "linear y"
shows "linear (x * y)"
proof -
have "reflexive y"
by (simp add: assms(2) linear_reflexive)
hence "x \<squnion> x\<^sup>T \<le> x * y \<squnion> y\<^sup>T * x\<^sup>T"
by (metis case_split_left case_split_right le_supI sup.cobounded1 sup.cobounded2 sup.idem reflexive_conv_closed)
thus ?thesis
by (simp add: assms(1) conv_dist_comp top_le)
qed
lemma equivalence_one_closed:
"equivalence 1"
by simp
lemma equivalence_top_closed:
"equivalence top"
by simp
lemma equivalence_inf_closed:
"equivalence x \<Longrightarrow> equivalence y \<Longrightarrow> equivalence (x \<sqinter> y)"
using conv_dist_inf preorder_inf_closed by auto
lemma equivalence_conv_closed:
"equivalence x \<Longrightarrow> equivalence (x\<^sup>T)"
by simp
lemma order_one_closed:
"order 1"
by simp
lemma order_inf_closed:
"order x \<Longrightarrow> order y \<Longrightarrow> order (x \<sqinter> y)"
using antisymmetric_inf_closed transitive_inf_closed by auto
lemma order_conv_closed:
"order x \<Longrightarrow> order (x\<^sup>T)"
by (simp add: inf_commute reflexive_conv_closed transitive_conv_closed)
lemma linear_order_conv_closed:
"linear_order x \<Longrightarrow> linear_order (x\<^sup>T)"
using equivalence_top_closed conv_dist_sup inf_commute reflexive_conv_closed transitive_conv_closed by force
text \<open>
We show a fact about equivalences.
\<close>
lemma equivalence_comp_dist_inf:
"equivalence x \<Longrightarrow> x * y \<sqinter> x * z = x * (y \<sqinter> x * z)"
- by (metis antisym comp_right_subdist_inf dedekind_1 eq_iff inf.absorb1 inf.absorb2 mult_1_right mult_assoc)
+ by (metis order.antisym comp_right_subdist_inf dedekind_1 order.eq_iff inf.absorb1 inf.absorb2 mult_1_right mult_assoc)
text \<open>
The following result generalises the fact that composition with a test amounts to intersection with the corresponding vector.
Both tests and vectors can be used to represent sets as relations.
\<close>
lemma coreflexive_comp_top_inf:
"coreflexive x \<Longrightarrow> x * top \<sqinter> y = x * y"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis comp_left_isotone comp_left_one coreflexive_symmetric dedekind_1 inf_top_left order_trans)
using comp_left_isotone comp_right_isotone by fastforce
lemma coreflexive_comp_top_inf_one:
"coreflexive x \<Longrightarrow> x * top \<sqinter> 1 = x"
by (simp add: coreflexive_comp_top_inf)
lemma coreflexive_comp_inf:
"coreflexive x \<Longrightarrow> coreflexive y \<Longrightarrow> x * y = x \<sqinter> y"
by (metis (full_types) coreflexive_comp_top_inf coreflexive_comp_top_inf_one inf.mult_assoc inf.absorb2)
lemma coreflexive_comp_inf_comp:
assumes "coreflexive x"
and "coreflexive y"
shows "(x * z) \<sqinter> (y * z) = (x \<sqinter> y) * z"
proof -
have "(x * z) \<sqinter> (y * z) = x * top \<sqinter> z \<sqinter> y * top \<sqinter> z"
using assms coreflexive_comp_top_inf inf_assoc by auto
also have "... = x * top \<sqinter> y * top \<sqinter> z"
by (simp add: inf.commute inf.left_commute)
also have "... = (x \<sqinter> y) * top \<sqinter> z"
by (metis assms coreflexive_comp_inf coreflexive_comp_top_inf mult_assoc)
also have "... = (x \<sqinter> y) * z"
by (simp add: assms(1) coreflexive_comp_top_inf coreflexive_inf_closed)
finally show ?thesis
.
qed
lemma test_comp_test_inf:
"(x \<sqinter> 1) * y * (z \<sqinter> 1) = (x \<sqinter> 1) * y \<sqinter> y * (z \<sqinter> 1)"
by (smt comp_right_one comp_right_subdist_inf coreflexive_comp_top_inf inf.left_commute inf.orderE inf_le2 mult_assoc)
lemma test_comp_test_top:
"y \<sqinter> (x \<sqinter> 1) * top * (z \<sqinter> 1) = (x \<sqinter> 1) * y * (z \<sqinter> 1)"
proof -
have "\<forall>u v . (v \<sqinter> u\<^sup>T)\<^sup>T = v\<^sup>T \<sqinter> u"
using conv_dist_inf by auto
thus ?thesis
by (smt conv_dist_comp conv_involutive coreflexive_comp_top_inf inf.cobounded2 inf.left_commute inf.sup_monoid.add_commute symmetric_one_closed mult_assoc symmetric_top_closed)
qed
lemma coreflexive_idempotent:
"coreflexive x \<Longrightarrow> idempotent x"
by (simp add: coreflexive_comp_inf)
lemma coreflexive_univalent:
"coreflexive x \<Longrightarrow> univalent x"
by (simp add: coreflexive_idempotent coreflexive_symmetric)
lemma coreflexive_injective:
"coreflexive x \<Longrightarrow> injective x"
by (simp add: coreflexive_idempotent coreflexive_symmetric)
lemma coreflexive_commutative:
"coreflexive x \<Longrightarrow> coreflexive y \<Longrightarrow> x * y = y * x"
by (simp add: coreflexive_comp_inf inf.commute)
lemma coreflexive_dedekind:
"coreflexive x \<Longrightarrow> coreflexive y \<Longrightarrow> coreflexive z \<Longrightarrow> x * y \<sqinter> z \<le> x * (y \<sqinter> x * z)"
by (simp add: coreflexive_comp_inf inf.coboundedI1 inf.left_commute)
text \<open>
Also the equational version of the Dedekind rule continues to hold.
\<close>
lemma dedekind_eq:
"x * y \<sqinter> z = (x \<sqinter> (z * y\<^sup>T)) * (y \<sqinter> (x\<^sup>T * z)) \<sqinter> z"
-proof (rule antisym)
+proof (rule order.antisym)
have "x * y \<sqinter> z \<le> x * (y \<sqinter> (x\<^sup>T * z)) \<sqinter> z"
by (simp add: dedekind_1)
also have "... \<le> (x \<sqinter> (z * (y \<sqinter> (x\<^sup>T * z))\<^sup>T)) * (y \<sqinter> (x\<^sup>T * z)) \<sqinter> z"
by (simp add: dedekind_2)
also have "... \<le> (x \<sqinter> (z * y\<^sup>T)) * (y \<sqinter> (x\<^sup>T * z)) \<sqinter> z"
by (metis comp_left_isotone comp_right_isotone inf_mono conv_order inf.cobounded1 order_refl)
finally show "x * y \<sqinter> z \<le> (x \<sqinter> (z * y\<^sup>T)) * (y \<sqinter> (x\<^sup>T * z)) \<sqinter> z"
.
next
show "(x \<sqinter> (z * y\<^sup>T)) * (y \<sqinter> (x\<^sup>T * z)) \<sqinter> z \<le> x * y \<sqinter> z"
using comp_isotone inf.sup_left_isotone by auto
qed
lemma dedekind:
"x * y \<sqinter> z \<le> (x \<sqinter> (z * y\<^sup>T)) * (y \<sqinter> (x\<^sup>T * z))"
by (metis dedekind_eq inf.cobounded1)
lemma vector_export_comp:
"(x * top \<sqinter> y) * z = x * top \<sqinter> y * z"
proof -
have "vector (x * top)"
by (simp add: comp_associative)
thus ?thesis
by (simp add: vector_inf_comp)
qed
lemma vector_export_comp_unit:
"(x * top \<sqinter> 1) * y = x * top \<sqinter> y"
by (simp add: vector_export_comp)
text \<open>
We solve a few exercises from \cite{SchmidtStroehlein1993}.
\<close>
lemma ex231a [simp]:
"(1 \<sqinter> x * x\<^sup>T) * x = x"
by (metis inf.cobounded1 inf.idem inf_right_idem comp_left_one conv_one coreflexive_comp_top_inf dedekind_eq)
lemma ex231b [simp]:
"x * (1 \<sqinter> x\<^sup>T * x) = x"
by (metis conv_dist_comp conv_dist_inf conv_involutive conv_one ex231a)
lemma ex231c:
"x \<le> x * x\<^sup>T * x"
by (metis comp_left_isotone ex231a inf_le2)
lemma ex231d:
"x \<le> x * top * x"
by (metis comp_left_isotone comp_right_isotone top_greatest order_trans ex231c)
lemma ex231e [simp]:
"x * top * x * top = x * top"
- by (metis ex231d antisym comp_associative mult_right_isotone top.extremum)
+ by (metis ex231d order.antisym comp_associative mult_right_isotone top.extremum)
lemma arc_injective:
"arc x \<Longrightarrow> injective x"
by (metis conv_dist_inf conv_involutive inf.absorb2 top_right_mult_increasing univalent_inf_closed)
lemma arc_conv_closed:
"arc x \<Longrightarrow> arc (x\<^sup>T)"
by simp
lemma arc_univalent:
"arc x \<Longrightarrow> univalent x"
using arc_conv_closed arc_injective univalent_conv_injective by blast
lemma injective_codomain:
assumes "injective x"
shows "x * (x \<sqinter> 1) = x \<sqinter> 1"
-proof (rule antisym)
+proof (rule order.antisym)
show "x * (x \<sqinter> 1) \<le> x \<sqinter> 1"
by (metis assms comp_right_one dual_order.trans inf.boundedI inf.cobounded1 inf.sup_monoid.add_commute mult_right_isotone one_inf_conv)
next
show "x \<sqinter> 1 \<le> x * (x \<sqinter> 1)"
by (metis coreflexive_idempotent inf.cobounded1 inf.cobounded2 mult_left_isotone)
qed
text \<open>
The following result generalises \cite[Exercise 2]{Oliveira2009}.
It is used to show that the while-loop preserves injectivity of the constructed tree.
\<close>
lemma injective_sup:
assumes "injective t"
and "e * t\<^sup>T \<le> 1"
and "injective e"
shows "injective (t \<squnion> e)"
proof -
have "(t \<squnion> e) * (t \<squnion> e)\<^sup>T = t * t\<^sup>T \<squnion> t * e\<^sup>T \<squnion> e * t\<^sup>T \<squnion> e * e\<^sup>T"
by (simp add: comp_left_dist_sup conv_dist_sup semiring.distrib_right sup.assoc)
thus ?thesis
using assms coreflexive_symmetric conv_dist_comp by fastforce
qed
lemma injective_inv:
"injective t \<Longrightarrow> e * t\<^sup>T = bot \<Longrightarrow> arc e \<Longrightarrow> injective (t \<squnion> e)"
using arc_injective injective_sup bot_least by blast
lemma univalent_sup:
"univalent t \<Longrightarrow> e\<^sup>T * t \<le> 1 \<Longrightarrow> univalent e \<Longrightarrow> univalent (t \<squnion> e)"
by (metis injective_sup conv_dist_sup conv_involutive)
lemma point_injective:
"arc x \<Longrightarrow> x\<^sup>T * top * x \<le> 1"
by (metis conv_top comp_associative conv_dist_comp conv_involutive vector_top_closed)
lemma vv_transitive:
"vector v \<Longrightarrow> (v * v\<^sup>T) * (v * v\<^sup>T) \<le> v * v\<^sup>T"
by (metis comp_associative comp_left_isotone comp_right_isotone top_greatest)
lemma epm_3:
assumes "e \<le> w"
and "injective w"
shows "e = w \<sqinter> top * e"
proof -
have "w \<sqinter> top * e \<le> w * e\<^sup>T * e"
by (metis (no_types, lifting) inf.absorb2 top.extremum dedekind_2 inf.commute)
also have "... \<le> w * w\<^sup>T * e"
by (simp add: assms(1) conv_isotone mult_left_isotone mult_right_isotone)
also have "... \<le> e"
using assms(2) coreflexive_comp_top_inf inf.sup_right_divisibility by blast
finally show ?thesis
- by (simp add: assms(1) top_left_mult_increasing antisym)
+ by (simp add: assms(1) top_left_mult_increasing order.antisym)
qed
lemma comp_inf_vector:
"x * (y \<sqinter> z * top) = (x \<sqinter> top * z\<^sup>T) * y"
by (metis conv_top covector_inf_comp_3 comp_associative conv_dist_comp inf.commute vector_top_closed)
lemma inf_vector_comp:
"(x \<sqinter> y * top) * z = y * top \<sqinter> x * z"
using inf.commute vector_export_comp by auto
lemma comp_inf_covector:
"x * (y \<sqinter> top * z) = x * y \<sqinter> top * z"
by (simp add: covector_comp_inf covector_mult_closed)
text \<open>
Well-known distributivity properties of univalent and injective relations over meet continue to hold.
\<close>
lemma univalent_comp_left_dist_inf:
assumes "univalent x"
shows "x * (y \<sqinter> z) = x * y \<sqinter> x * z"
-proof (rule antisym)
+proof (rule order.antisym)
show "x * (y \<sqinter> z) \<le> x * y \<sqinter> x * z"
by (simp add: comp_right_isotone)
next
have "x * y \<sqinter> x * z \<le> (x \<sqinter> x * z * y\<^sup>T) * (y \<sqinter> x\<^sup>T * x * z)"
by (metis comp_associative dedekind)
also have "... \<le> x * (y \<sqinter> x\<^sup>T * x * z)"
by (simp add: comp_left_isotone)
also have "... \<le> x * (y \<sqinter> 1 * z)"
using assms comp_left_isotone comp_right_isotone inf.sup_right_isotone by blast
finally show "x * y \<sqinter> x * z \<le> x * (y \<sqinter> z)"
by simp
qed
lemma injective_comp_right_dist_inf:
"injective z \<Longrightarrow> (x \<sqinter> y) * z = x * z \<sqinter> y * z"
by (metis univalent_comp_left_dist_inf conv_dist_comp conv_involutive conv_dist_inf)
lemma vector_covector:
"vector v \<Longrightarrow> vector w \<Longrightarrow> v \<sqinter> w\<^sup>T = v * w\<^sup>T"
by (metis covector_comp_inf inf_top_left vector_conv_covector)
lemma comp_inf_vector_1:
"(x \<sqinter> top * y) * z = x * (z \<sqinter> (top * y)\<^sup>T)"
by (simp add: comp_inf_vector conv_dist_comp)
text \<open>
The shunting properties for bijective relations and mappings continue to hold.
\<close>
lemma shunt_bijective:
assumes "bijective z"
shows "x \<le> y * z \<longleftrightarrow> x * z\<^sup>T \<le> y"
proof
assume "x \<le> y * z"
hence "x * z\<^sup>T \<le> y * z * z\<^sup>T"
by (simp add: mult_left_isotone)
also have "... \<le> y"
using assms comp_associative mult_right_isotone by fastforce
finally show "x * z\<^sup>T \<le> y"
.
next
assume 1: "x * z\<^sup>T \<le> y"
have "x = x \<sqinter> top * z"
by (simp add: assms)
also have "... \<le> x * z\<^sup>T * z"
by (metis dedekind_2 inf_commute inf_top.right_neutral)
also have "... \<le> y * z"
using 1 by (simp add: mult_left_isotone)
finally show "x \<le> y * z"
.
qed
lemma shunt_mapping:
"mapping z \<Longrightarrow> x \<le> z * y \<longleftrightarrow> z\<^sup>T * x \<le> y"
by (metis shunt_bijective mapping_conv_bijective conv_order conv_dist_comp conv_involutive)
lemma bijective_reverse:
assumes "bijective p"
and "bijective q"
shows "p \<le> r * q \<longleftrightarrow> q \<le> r\<^sup>T * p"
proof -
have "p \<le> r * q \<longleftrightarrow> p * q\<^sup>T \<le> r"
by (simp add: assms(2) shunt_bijective)
also have "... \<longleftrightarrow> q\<^sup>T \<le> p\<^sup>T * r"
by (metis assms(1) conv_dist_comp conv_involutive conv_order shunt_bijective)
also have "... \<longleftrightarrow> q \<le> r\<^sup>T * p"
using conv_dist_comp conv_isotone by fastforce
finally show ?thesis
by simp
qed
lemma arc_expanded:
"arc x \<longleftrightarrow> x * top * x\<^sup>T \<le> 1 \<and> x\<^sup>T * top * x \<le> 1 \<and> top * x * top = top"
by (metis conv_top comp_associative conv_dist_comp conv_involutive vector_top_closed)
lemma arc_top_arc:
assumes "arc x"
shows "x * top * x = x"
by (metis assms epm_3 top_right_mult_increasing vector_inf_comp vector_mult_closed vector_top_closed)
lemma arc_top_edge:
assumes "arc x"
shows "x\<^sup>T * top * x = x\<^sup>T * x"
proof -
have "x\<^sup>T = x\<^sup>T * top \<sqinter> top * x\<^sup>T"
using assms epm_3 top_right_mult_increasing by simp
thus ?thesis
by (metis comp_inf_vector_1 conv_dist_comp conv_involutive conv_top inf.absorb1 top_right_mult_increasing)
qed
text \<open>Lemmas \<open>arc_eq_1\<close> and \<open>arc_eq_2\<close> were contributed by Nicolas Robinson-O'Brien.\<close>
lemma arc_eq_1:
assumes "arc x"
shows "x = x * x\<^sup>T * x"
proof -
have "x * x\<^sup>T * x \<le> x * top * x"
by (simp add: mult_left_isotone mult_right_isotone)
also have "... \<le> x"
by (simp add: assms arc_top_arc)
finally have "x * x\<^sup>T * x \<le> x"
by simp
thus ?thesis
- by (simp add: antisym ex231c)
+ by (simp add: order.antisym ex231c)
qed
lemma arc_eq_2:
assumes "arc x"
shows "x\<^sup>T = x\<^sup>T * x * x\<^sup>T"
using arc_eq_1 assms conv_involutive by fastforce
lemma points_arc:
"point x \<Longrightarrow> point y \<Longrightarrow> arc (x * y\<^sup>T)"
by (metis comp_associative conv_dist_comp conv_involutive equivalence_top_closed)
lemma point_arc:
"point x \<Longrightarrow> arc (x * x\<^sup>T)"
by (simp add: points_arc)
end
subsection \<open>Single-Object Pseudocomplemented Distributive Allegories\<close>
text \<open>
We extend single-object bounded distributive allegories by a pseudocomplement operation.
The following definitions concern properties of relations that require a pseudocomplement.
\<close>
class relation_algebra_signature = bounded_distrib_allegory_signature + uminus
begin
abbreviation irreflexive :: "'a \<Rightarrow> bool" where "irreflexive x \<equiv> x \<le> -1"
abbreviation strict_linear :: "'a \<Rightarrow> bool" where "strict_linear x \<equiv> x \<squnion> x\<^sup>T = -1"
abbreviation strict_order :: "'a \<Rightarrow> bool" where "strict_order x \<equiv> irreflexive x \<and> transitive x"
abbreviation linear_strict_order :: "'a \<Rightarrow> bool" where "linear_strict_order x \<equiv> strict_order x \<and> strict_linear x"
text \<open>
The following variants are useful for the graph model.
\<close>
abbreviation pp_mapping :: "'a \<Rightarrow> bool" where "pp_mapping x \<equiv> univalent x \<and> total (--x)"
abbreviation pp_bijective :: "'a \<Rightarrow> bool" where "pp_bijective x \<equiv> injective x \<and> surjective (--x)"
abbreviation pp_point :: "'a \<Rightarrow> bool" where "pp_point x \<equiv> vector x \<and> pp_bijective x"
abbreviation pp_arc :: "'a \<Rightarrow> bool" where "pp_arc x \<equiv> pp_bijective (x * top) \<and> pp_bijective (x\<^sup>T * top)"
end
class pd_allegory = bounded_distrib_allegory + p_algebra
begin
subclass relation_algebra_signature .
subclass pd_algebra ..
lemma conv_complement_1:
"-(x\<^sup>T) \<squnion> (-x)\<^sup>T = (-x)\<^sup>T"
by (metis conv_dist_inf conv_order bot_least conv_involutive pseudo_complement sup.absorb2 sup.cobounded2)
lemma conv_complement:
"(-x)\<^sup>T = -(x\<^sup>T)"
by (metis conv_complement_1 conv_dist_sup conv_involutive sup_commute)
lemma conv_complement_sub_inf [simp]:
"x\<^sup>T * -(x * y) \<sqinter> y = bot"
by (metis comp_left_zero conv_dist_comp conv_involutive dedekind_1 inf_import_p inf_p inf_right_idem ppp pseudo_complement regular_closed_bot)
lemma conv_complement_sub_leq:
"x\<^sup>T * -(x * y) \<le> -y"
using pseudo_complement conv_complement_sub_inf by blast
lemma conv_complement_sub [simp]:
"x\<^sup>T * -(x * y) \<squnion> -y = -y"
by (simp add: conv_complement_sub_leq sup.absorb2)
lemma complement_conv_sub:
"-(y * x) * x\<^sup>T \<le> -y"
by (metis conv_complement conv_complement_sub_leq conv_order conv_dist_comp)
text \<open>
The following so-called Schr\"oder equivalences, or De Morgan's Theorem K, hold only with a pseudocomplemented element on both right-hand sides.
\<close>
lemma schroeder_3_p:
"x * y \<le> -z \<longleftrightarrow> x\<^sup>T * z \<le> -y"
using pseudo_complement schroeder_1 by auto
lemma schroeder_4_p:
"x * y \<le> -z \<longleftrightarrow> z * y\<^sup>T \<le> -x"
using pseudo_complement schroeder_2 by auto
lemma comp_pp_semi_commute:
"x * --y \<le> --(x * y)"
using conv_complement_sub_leq schroeder_3_p by fastforce
text \<open>
The following result looks similar to a property of (anti)domain.
\<close>
lemma p_comp_pp [simp]:
"-(x * --y) = -(x * y)"
- using comp_pp_semi_commute comp_right_isotone inf.eq_iff p_antitone pp_increasing by fastforce
+ using comp_pp_semi_commute comp_right_isotone order.eq_iff p_antitone pp_increasing by fastforce
lemma pp_comp_semi_commute:
"--x * y \<le> --(x * y)"
using complement_conv_sub schroeder_4_p by fastforce
lemma p_pp_comp [simp]:
"-(--x * y) = -(x * y)"
- using pp_comp_semi_commute comp_left_isotone inf.eq_iff p_antitone pp_increasing by fastforce
+ using pp_comp_semi_commute comp_left_isotone order.eq_iff p_antitone pp_increasing by fastforce
lemma pp_comp_subdist:
"--x * --y \<le> --(x * y)"
by (simp add: p_antitone_iff)
lemma theorem24xxiii:
"x * y \<sqinter> -(x * z) = x * (y \<sqinter> -z) \<sqinter> -(x * z)"
proof -
have "x * y \<sqinter> -(x * z) \<le> x * (y \<sqinter> (x\<^sup>T * -(x * z)))"
by (simp add: dedekind_1)
also have "... \<le> x * (y \<sqinter> -z)"
using comp_right_isotone conv_complement_sub_leq inf.sup_right_isotone by auto
finally show ?thesis
- using comp_right_subdist_inf antisym inf.coboundedI2 inf.commute by auto
+ using comp_right_subdist_inf order.antisym inf.coboundedI2 inf.commute by auto
qed
text \<open>
Even in Stone relation algebras, we do not obtain the backward implication in the following result.
\<close>
lemma vector_complement_closed:
"vector x \<Longrightarrow> vector (-x)"
- by (metis complement_conv_sub conv_top eq_iff top_right_mult_increasing)
+ by (metis complement_conv_sub conv_top order.eq_iff top_right_mult_increasing)
lemma covector_complement_closed:
"covector x \<Longrightarrow> covector (-x)"
- by (metis conv_complement_sub_leq conv_top eq_iff top_left_mult_increasing)
+ by (metis conv_complement_sub_leq conv_top order.eq_iff top_left_mult_increasing)
lemma covector_vector_comp:
"vector v \<Longrightarrow> -v\<^sup>T * v = bot"
by (metis conv_bot conv_complement conv_complement_sub_inf conv_dist_comp conv_involutive inf_top.right_neutral)
lemma irreflexive_bot_closed:
"irreflexive bot"
by simp
lemma irreflexive_inf_closed:
"irreflexive x \<Longrightarrow> irreflexive (x \<sqinter> y)"
by (simp add: le_infI1)
lemma irreflexive_sup_closed:
"irreflexive x \<Longrightarrow> irreflexive y \<Longrightarrow> irreflexive (x \<squnion> y)"
by simp
lemma irreflexive_conv_closed:
"irreflexive x \<Longrightarrow> irreflexive (x\<^sup>T)"
using conv_complement conv_isotone by fastforce
lemma reflexive_complement_irreflexive:
"reflexive x \<Longrightarrow> irreflexive (-x)"
by (simp add: p_antitone)
lemma irreflexive_complement_reflexive:
"irreflexive x \<longleftrightarrow> reflexive (-x)"
by (simp add: p_antitone_iff)
lemma symmetric_complement_closed:
"symmetric x \<Longrightarrow> symmetric (-x)"
by (simp add: conv_complement)
lemma asymmetric_irreflexive:
"asymmetric x \<Longrightarrow> irreflexive x"
by (metis inf.mult_not_zero inf.left_commute inf.right_idem inf.sup_monoid.add_commute pseudo_complement one_inf_conv)
lemma linear_asymmetric:
"linear x \<Longrightarrow> asymmetric (-x)"
using conv_complement p_top by force
lemma strict_linear_sup_closed:
"strict_linear x \<Longrightarrow> strict_linear y \<Longrightarrow> strict_linear (x \<squnion> y)"
by (metis (mono_tags, hide_lams) conv_dist_sup sup.right_idem sup_assoc sup_commute)
lemma strict_linear_irreflexive:
"strict_linear x \<Longrightarrow> irreflexive x"
using sup_left_divisibility by blast
lemma strict_linear_conv_closed:
"strict_linear x \<Longrightarrow> strict_linear (x\<^sup>T)"
by (simp add: sup_commute)
lemma strict_order_var:
"strict_order x \<longleftrightarrow> asymmetric x \<and> transitive x"
by (metis asymmetric_irreflexive comp_right_one irreflexive_conv_closed conv_dist_comp dual_order.trans pseudo_complement schroeder_3_p)
lemma strict_order_bot_closed:
"strict_order bot"
by simp
lemma strict_order_inf_closed:
"strict_order x \<Longrightarrow> strict_order y \<Longrightarrow> strict_order (x \<sqinter> y)"
using inf.coboundedI1 transitive_inf_closed by auto
lemma strict_order_conv_closed:
"strict_order x \<Longrightarrow> strict_order (x\<^sup>T)"
using irreflexive_conv_closed transitive_conv_closed by blast
lemma order_strict_order:
assumes "order x"
shows "strict_order (x \<sqinter> -1)"
proof (rule conjI)
show 1: "irreflexive (x \<sqinter> -1)"
by simp
have "antisymmetric (x \<sqinter> -1)"
using antisymmetric_inf_closed assms by blast
hence "(x \<sqinter> -1) * (x \<sqinter> -1) \<sqinter> 1 \<le> (x \<sqinter> -1 \<sqinter> (x \<sqinter> -1)\<^sup>T) * (x \<sqinter> -1 \<sqinter> (x \<sqinter> -1)\<^sup>T)"
using 1 by (metis (no_types) coreflexive_symmetric irreflexive_inf_closed coreflexive_transitive dedekind_1 inf_idem mult_1_right semiring.mult_not_zero strict_order_var)
also have "... = (x \<sqinter> x\<^sup>T \<sqinter> -1) * (x \<sqinter> x\<^sup>T \<sqinter> -1)"
by (simp add: conv_complement conv_dist_inf inf.absorb2 inf.sup_monoid.add_assoc)
also have "... = bot"
- using assms inf.antisym reflexive_conv_closed by fastforce
+ using assms order.antisym reflexive_conv_closed by fastforce
finally have "(x \<sqinter> -1) * (x \<sqinter> -1) \<le> -1"
using le_bot pseudo_complement by blast
thus "transitive (x \<sqinter> -1)"
by (meson assms comp_isotone inf.boundedI inf.cobounded1 inf.order_lesseq_imp)
qed
lemma strict_order_order:
"strict_order x \<Longrightarrow> order (x \<squnion> 1)"
apply (unfold strict_order_var, intro conjI)
apply simp
apply (simp add: mult_left_dist_sup mult_right_dist_sup sup.absorb2)
using conv_dist_sup coreflexive_bot_closed sup.absorb2 sup_inf_distrib2 by fastforce
lemma linear_strict_order_conv_closed:
"linear_strict_order x \<Longrightarrow> linear_strict_order (x\<^sup>T)"
by (simp add: irreflexive_conv_closed sup_monoid.add_commute transitive_conv_closed)
lemma linear_order_strict_order:
"linear_order x \<Longrightarrow> linear_strict_order (x \<sqinter> -1)"
apply (rule conjI)
using order_strict_order apply simp
- by (metis conv_complement conv_dist_inf coreflexive_symmetric eq_iff inf.absorb2 inf.distrib_left inf.sup_monoid.add_commute top.extremum)
+ by (metis conv_complement conv_dist_inf coreflexive_symmetric order.eq_iff inf.absorb2 inf.distrib_left inf.sup_monoid.add_commute top.extremum)
lemma regular_conv_closed:
"regular x \<Longrightarrow> regular (x\<^sup>T)"
by (metis conv_complement)
text \<open>
We show a number of facts about equivalences.
\<close>
lemma equivalence_comp_left_complement:
"equivalence x \<Longrightarrow> x * -x = -x"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis conv_complement_sub_leq preorder_idempotent)
using mult_left_isotone by fastforce
lemma equivalence_comp_right_complement:
"equivalence x \<Longrightarrow> -x * x = -x"
by (metis equivalence_comp_left_complement conv_complement conv_dist_comp)
text \<open>
The pseudocomplement of tests is given by the following operation.
\<close>
abbreviation coreflexive_complement :: "'a \<Rightarrow> 'a" ("_ ''" [80] 80)
where "x ' \<equiv> -x \<sqinter> 1"
lemma coreflexive_comp_top_coreflexive_complement:
"coreflexive x \<Longrightarrow> (x * top)' = x '"
by (metis coreflexive_comp_top_inf_one inf.commute inf_import_p)
lemma coreflexive_comp_inf_complement:
"coreflexive x \<Longrightarrow> (x * y) \<sqinter> -z = (x * y) \<sqinter> -(x * z)"
by (metis coreflexive_comp_top_inf inf.sup_relative_same_increasing inf_import_p inf_le1)
lemma double_coreflexive_complement:
"x '' = (-x)'"
using inf.sup_monoid.add_commute inf_import_p by auto
lemma coreflexive_pp_dist_comp:
assumes "coreflexive x"
and "coreflexive y"
shows "(x * y)'' = x '' * y ''"
proof -
have "(x * y)'' = --(x * y) \<sqinter> 1"
by (simp add: double_coreflexive_complement)
also have "... = --x \<sqinter> --y \<sqinter> 1"
by (simp add: assms coreflexive_comp_inf)
also have "... = (--x \<sqinter> 1) * (--y \<sqinter> 1)"
by (simp add: coreflexive_comp_inf inf.left_commute inf.sup_monoid.add_assoc)
also have "... = x '' * y ''"
by (simp add: double_coreflexive_complement)
finally show ?thesis
.
qed
lemma coreflexive_pseudo_complement:
"coreflexive x \<Longrightarrow> x \<sqinter> y = bot \<longleftrightarrow> x \<le> y '"
by (simp add: pseudo_complement)
lemma pp_bijective_conv_mapping:
"pp_bijective x \<longleftrightarrow> pp_mapping (x\<^sup>T)"
by (simp add: conv_complement surjective_conv_total)
lemma pp_arc_expanded:
"pp_arc x \<longleftrightarrow> x * top * x\<^sup>T \<le> 1 \<and> x\<^sup>T * top * x \<le> 1 \<and> top * --x * top = top"
proof
assume 1: "pp_arc x"
have 2: "x * top * x\<^sup>T \<le> 1"
using 1 by (metis comp_associative conv_dist_comp equivalence_top_closed vector_top_closed)
have 3: "x\<^sup>T * top * x \<le> 1"
using 1 by (metis conv_dist_comp conv_involutive equivalence_top_closed vector_top_closed mult_assoc)
have 4: "x\<^sup>T \<le> x\<^sup>T * x * x\<^sup>T"
by (metis conv_involutive ex231c)
have "top = --(top * x) * top"
using 1 by (metis conv_complement conv_dist_comp conv_involutive equivalence_top_closed)
also have "... \<le> --(top * x\<^sup>T * top * x) * top"
using 1 by (metis eq_refl mult_assoc p_comp_pp p_pp_comp)
also have "... = (top * --(x * top) \<sqinter> --(top * x\<^sup>T * top * x)) * top"
using 1 by simp
also have "... = top * (--(x * top) \<sqinter> --(top * x\<^sup>T * top * x)) * top"
by (simp add: covector_complement_closed covector_comp_inf covector_mult_closed)
also have "... = top * --(x * top \<sqinter> top * x\<^sup>T * top * x) * top"
by simp
also have "... = top * --(x * top * x\<^sup>T * top * x) * top"
by (metis comp_associative comp_inf_covector inf_top.left_neutral)
also have "... \<le> top * --(x * top * x\<^sup>T * x * x\<^sup>T * top * x) * top"
using 4 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone)
also have "... \<le> top * --(x * x\<^sup>T * top * x) * top"
using 2 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone comp_left_one)
also have "... \<le> top * --x * top"
using 3 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone comp_right_one)
finally show "x * top * x\<^sup>T \<le> 1 \<and> x\<^sup>T * top * x \<le> 1 \<and> top * --x * top = top"
using 2 3 top_le by blast
next
assume "x * top * x\<^sup>T \<le> 1 \<and> x\<^sup>T * top * x \<le> 1 \<and> top * --x * top = top"
thus "pp_arc x"
apply (intro conjI)
apply (metis comp_associative conv_dist_comp equivalence_top_closed vector_top_closed)
apply (metis comp_associative mult_right_isotone top_le pp_comp_semi_commute)
apply (metis conv_dist_comp coreflexive_symmetric vector_conv_covector vector_top_closed mult_assoc)
by (metis conv_complement conv_dist_comp equivalence_top_closed inf.orderE inf_top.left_neutral mult_right_isotone pp_comp_semi_commute)
qed
text \<open>
The following operation represents states with infinite executions of non-strict computations.
\<close>
abbreviation N :: "'a \<Rightarrow> 'a"
where "N x \<equiv> -(-x * top) \<sqinter> 1"
lemma N_comp:
"N x * y = -(-x * top) \<sqinter> y"
by (simp add: vector_mult_closed vector_complement_closed vector_inf_one_comp)
lemma N_comp_top [simp]:
"N x * top = -(-x * top)"
by (simp add: N_comp)
lemma vector_N_pp:
"vector x \<Longrightarrow> N x = --x \<sqinter> 1"
by (simp add: vector_complement_closed)
lemma N_vector_pp [simp]:
"N (x * top) = --(x * top) \<sqinter> 1"
by (simp add: comp_associative vector_complement_closed)
lemma N_vector_top_pp [simp]:
"N (x * top) * top = --(x * top)"
by (metis N_comp_top comp_associative vector_top_closed vector_complement_closed)
lemma N_below_inf_one_pp:
"N x \<le> --x \<sqinter> 1"
using inf.sup_left_isotone p_antitone top_right_mult_increasing by auto
lemma N_below_pp:
"N x \<le> --x"
using N_below_inf_one_pp by auto
lemma N_comp_N:
"N x * N y = -(-x * top) \<sqinter> -(-y * top) \<sqinter> 1"
by (simp add: N_comp inf.mult_assoc)
lemma N_bot [simp]:
"N bot = bot"
by simp
lemma N_top [simp]:
"N top = 1"
by simp
lemma n_split_omega_mult_pp:
"xs * --xo = xo \<Longrightarrow> vector xo \<Longrightarrow> N top * xo = xs * N xo * top"
by (metis N_top N_vector_top_pp comp_associative comp_left_one)
text \<open>
Many of the following results have been derived for verifying Prim's minimum spanning tree algorithm.
\<close>
lemma ee:
assumes "vector v"
and "e \<le> v * -v\<^sup>T"
shows "e * e = bot"
proof -
have "e * v \<le> bot"
by (metis assms covector_vector_comp comp_associative mult_left_isotone mult_right_zero)
thus ?thesis
by (metis assms(2) bot_unique comp_associative mult_right_isotone semiring.mult_not_zero)
qed
lemma et:
assumes "vector v"
and "e \<le> v * -v\<^sup>T"
and "t \<le> v * v\<^sup>T"
shows "e * t = bot"
and "e * t\<^sup>T = bot"
proof -
have "e * t \<le> v * -v\<^sup>T * v * v\<^sup>T"
using assms(2-3) comp_isotone mult_assoc by fastforce
thus "e * t = bot"
by (simp add: assms(1) covector_vector_comp le_bot mult_assoc)
next
have "t\<^sup>T \<le> v * v\<^sup>T"
using assms(3) conv_order conv_dist_comp by fastforce
hence "e * t\<^sup>T \<le> v * -v\<^sup>T * v * v\<^sup>T"
by (metis assms(2) comp_associative comp_isotone)
thus "e * t\<^sup>T = bot"
by (simp add: assms(1) covector_vector_comp le_bot mult_assoc)
qed
lemma ve_dist:
assumes "e \<le> v * -v\<^sup>T"
and "vector v"
and "arc e"
shows "(v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T = v * v\<^sup>T \<squnion> v * v\<^sup>T * e \<squnion> e\<^sup>T * v * v\<^sup>T \<squnion> e\<^sup>T * e"
proof -
have "e \<le> v * top"
using assms(1) comp_right_isotone dual_order.trans top_greatest by blast
hence "v * top * e = v * top * (v * top \<sqinter> e)"
by (simp add: inf.absorb2)
also have "... = (v * top \<sqinter> top * v\<^sup>T) * e"
using assms(2) covector_inf_comp_3 vector_conv_covector by force
also have "... = v * top * v\<^sup>T * e"
by (metis assms(2) inf_top_right vector_inf_comp)
also have "... = v * v\<^sup>T * e"
by (simp add: assms(2))
finally have 1: "v * top * e = v * v\<^sup>T * e"
.
have "e\<^sup>T * top * e \<le> e\<^sup>T * top * e * e\<^sup>T * e"
using ex231c comp_associative mult_right_isotone by auto
also have "... \<le> e\<^sup>T * e"
by (metis assms(3) coreflexive_comp_top_inf le_infE mult_semi_associative point_injective)
finally have 2: "e\<^sup>T * top * e = e\<^sup>T * e"
- by (simp add: inf.antisym mult_left_isotone top_right_mult_increasing)
+ by (simp add: order.antisym mult_left_isotone top_right_mult_increasing)
have "(v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T = (v \<squnion> e\<^sup>T * top) * (v\<^sup>T \<squnion> top * e)"
by (simp add: conv_dist_comp conv_dist_sup)
also have "... = v * v\<^sup>T \<squnion> v * top * e \<squnion> e\<^sup>T * top * v\<^sup>T \<squnion> e\<^sup>T * top * top * e"
by (metis semiring.distrib_left semiring.distrib_right sup_assoc mult_assoc)
also have "... = v * v\<^sup>T \<squnion> v * top * e \<squnion> (v * top * e)\<^sup>T \<squnion> e\<^sup>T * top * e"
by (simp add: comp_associative conv_dist_comp)
also have "... = v * v\<^sup>T \<squnion> v * v\<^sup>T * e \<squnion> (v * v\<^sup>T * e)\<^sup>T \<squnion> e\<^sup>T * e"
using 1 2 by simp
finally show ?thesis
by (simp add: comp_associative conv_dist_comp)
qed
lemma ev:
"vector v \<Longrightarrow> e \<le> v * -v\<^sup>T \<Longrightarrow> e * v = bot"
- by (metis covector_vector_comp antisym bot_least comp_associative mult_left_isotone mult_right_zero)
+ by (metis covector_vector_comp order.antisym bot_least comp_associative mult_left_isotone mult_right_zero)
lemma vTeT:
"vector v \<Longrightarrow> e \<le> v * -v\<^sup>T \<Longrightarrow> v\<^sup>T * e\<^sup>T = bot"
using conv_bot ev conv_dist_comp by fastforce
text \<open>
The following result is used to show that the while-loop of Prim's algorithm preserves that the constructed tree is a subgraph of g.
\<close>
lemma prim_subgraph_inv:
assumes "e \<le> v * -v\<^sup>T \<sqinter> g"
and "t \<le> v * v\<^sup>T \<sqinter> g"
shows "t \<squnion> e \<le> ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T) \<sqinter> g"
proof (rule sup_least)
have "t \<le> ((v \<squnion> e\<^sup>T * top) * v\<^sup>T) \<sqinter> g"
using assms(2) le_supI1 mult_right_dist_sup by auto
also have "... \<le> ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T) \<sqinter> g"
using comp_right_isotone conv_dist_sup inf.sup_left_isotone by auto
finally show "t \<le> ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T) \<sqinter> g"
.
next
have "e \<le> v * top"
by (meson assms(1) inf.boundedE mult_right_isotone order.trans top.extremum)
hence "e \<le> v * top \<sqinter> top * e"
by (simp add: top_left_mult_increasing)
also have "... = v * top * e"
by (metis inf_top_right vector_export_comp)
finally have "e \<le> v * top * e \<sqinter> g"
using assms(1) by auto
also have "... = v * (e\<^sup>T * top)\<^sup>T \<sqinter> g"
by (simp add: comp_associative conv_dist_comp)
also have "... \<le> v * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g"
by (simp add: conv_dist_sup mult_left_dist_sup sup.assoc sup.orderI)
also have "... \<le> (v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T \<sqinter> g"
using inf.sup_left_isotone mult_right_sub_dist_sup_left by auto
finally show "e \<le> ((v \<squnion> e\<^sup>T * top) * (v \<squnion> e\<^sup>T * top)\<^sup>T) \<sqinter> g"
.
qed
text \<open>
The following result shows how to apply the Schr\"oder equivalence to the middle factor in a composition of three relations.
Again the elements on the right-hand side need to be pseudocomplemented.
\<close>
lemma triple_schroeder_p:
"x * y * z \<le> -w \<longleftrightarrow> x\<^sup>T * w * z\<^sup>T \<le> -y"
using mult_assoc p_antitone_iff schroeder_3_p schroeder_4_p by auto
text \<open>
The rotation versions of the Schr\"oder equivalences continue to hold, again with pseudocomplemented elements on the right-hand side.
\<close>
lemma schroeder_5_p:
"x * y \<le> -z \<longleftrightarrow> y * z\<^sup>T \<le> -x\<^sup>T"
using schroeder_3_p schroeder_4_p by auto
lemma schroeder_6_p:
"x * y \<le> -z \<longleftrightarrow> z\<^sup>T * x \<le> -y\<^sup>T"
using schroeder_3_p schroeder_4_p by auto
lemma vector_conv_compl:
"vector v \<Longrightarrow> top * -v\<^sup>T = -v\<^sup>T"
by (simp add: covector_complement_closed vector_conv_covector)
text \<open>
Composition commutes, relative to the diversity relation.
\<close>
lemma comp_commute_below_diversity:
"x * y \<le> -1 \<longleftrightarrow> y * x \<le> -1"
by (metis comp_right_one conv_dist_comp conv_one schroeder_3_p schroeder_4_p)
lemma comp_injective_below_complement:
"injective y \<Longrightarrow> -x * y \<le> -(x * y)"
by (metis p_antitone_iff comp_associative comp_right_isotone comp_right_one schroeder_4_p)
lemma comp_univalent_below_complement:
"univalent x \<Longrightarrow> x * -y \<le> -(x * y)"
by (metis p_inf pseudo_complement semiring.mult_zero_right univalent_comp_left_dist_inf)
text \<open>
Bijective relations and mappings can be exported from a pseudocomplement.
\<close>
lemma comp_bijective_complement:
"bijective y \<Longrightarrow> -x * y = -(x * y)"
- using comp_injective_below_complement complement_conv_sub antisym shunt_bijective by blast
+ using comp_injective_below_complement complement_conv_sub order.antisym shunt_bijective by blast
lemma comp_mapping_complement:
"mapping x \<Longrightarrow> x * -y = -(x * y)"
by (metis (full_types) comp_bijective_complement conv_complement conv_dist_comp conv_involutive total_conv_surjective)
text \<open>
The following facts are used in the correctness proof of Kruskal's minimum spanning tree algorithm.
\<close>
lemma kruskal_injective_inv:
assumes "injective f"
and "covector q"
and "q * f\<^sup>T \<le> q"
and "e \<le> q"
and "q * f\<^sup>T \<le> -e"
and "injective e"
and "q\<^sup>T * q \<sqinter> f\<^sup>T * f \<le> 1"
shows "injective ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e)"
proof -
have 1: "(f \<sqinter> -q) * (f \<sqinter> -q)\<^sup>T \<le> 1"
by (simp add: assms(1) injective_inf_closed)
have 2: "(f \<sqinter> -q) * (f \<sqinter> q) \<le> 1"
proof -
have 21: "bot = q * f\<^sup>T \<sqinter> - q"
by (metis assms(3) inf.sup_monoid.add_assoc inf.sup_right_divisibility inf_import_p inf_p)
have "(f \<sqinter> -q) * (f \<sqinter> q) \<le> -q * f \<sqinter> q"
by (metis assms(2) comp_inf_covector comp_isotone inf.cobounded2 inf.left_idem)
also have "... = bot"
using 21 schroeder_2 by auto
finally show ?thesis
by (simp add: bot_unique)
qed
have 3: "(f \<sqinter> -q) * e\<^sup>T \<le> 1"
proof -
have "(f \<sqinter> -q) * e\<^sup>T \<le> -q * e\<^sup>T"
by (simp add: mult_left_isotone)
also have "... = bot"
by (metis assms(2,4) bot_unique conv_bot conv_complement covector_complement_closed p_antitone p_bot regular_closed_bot schroeder_5_p)
finally show ?thesis
by (simp add: bot_unique)
qed
have 4: "(f \<sqinter> q)\<^sup>T * (f \<sqinter> -q)\<^sup>T \<le> 1"
using 2 conv_dist_comp conv_isotone by force
have 5: "(f \<sqinter> q)\<^sup>T * (f \<sqinter> q) \<le> 1"
proof -
have "(f \<sqinter> q)\<^sup>T * (f \<sqinter> q) \<le> q\<^sup>T * q \<sqinter> f\<^sup>T * f"
by (simp add: conv_isotone mult_isotone)
also have "... \<le> 1"
by (simp add: assms(7))
finally show ?thesis
by simp
qed
have 6: "(f \<sqinter> q)\<^sup>T * e\<^sup>T \<le> 1"
proof -
have "f\<^sup>T * e\<^sup>T \<le> -q\<^sup>T"
using assms(5) schroeder_5_p by simp
hence "(f \<sqinter> q)\<^sup>T * e\<^sup>T = bot"
by (metis assms(2,5) conv_bot conv_dist_comp covector_comp_inf inf.absorb1 inf.cobounded2 inf.sup_monoid.add_commute inf_left_commute inf_p schroeder_4_p)
thus ?thesis
by (simp add: bot_unique)
qed
have 7: "e * (f \<sqinter> -q)\<^sup>T \<le> 1"
using 3 conv_dist_comp coreflexive_symmetric by fastforce
have 8: "e * (f \<sqinter> q) \<le> 1"
using 6 conv_dist_comp coreflexive_symmetric by fastforce
have 9: "e * e\<^sup>T \<le> 1"
by (simp add: assms(6))
have "((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e) * ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e)\<^sup>T = (f \<sqinter> -q) * (f \<sqinter> -q)\<^sup>T \<squnion> (f \<sqinter> -q) * (f \<sqinter> q) \<squnion> (f \<sqinter> -q) * e\<^sup>T \<squnion> (f \<sqinter> q)\<^sup>T * (f \<sqinter> -q)\<^sup>T \<squnion> (f \<sqinter> q)\<^sup>T * (f \<sqinter> q) \<squnion> (f \<sqinter> q)\<^sup>T * e\<^sup>T \<squnion> e * (f \<sqinter> -q)\<^sup>T \<squnion> e * (f \<sqinter> q) \<squnion> e * e\<^sup>T"
using comp_left_dist_sup comp_right_dist_sup conv_dist_sup sup.assoc by simp
also have "... \<le> 1"
using 1 2 3 4 5 6 7 8 9 by simp
finally show ?thesis
by simp
qed
lemma kruskal_exchange_injective_inv_1:
assumes "injective f"
and "covector q"
and "q * f\<^sup>T \<le> q"
and "q\<^sup>T * q \<sqinter> f\<^sup>T * f \<le> 1"
shows "injective ((f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T)"
using kruskal_injective_inv[where e=bot] by (simp add: assms)
lemma kruskal_exchange_acyclic_inv_3:
assumes "injective w"
and "d \<le> w"
shows "(w \<sqinter> -d) * d\<^sup>T * top = bot"
proof -
have "(w \<sqinter> -d) * d\<^sup>T * top = (w \<sqinter> -d \<sqinter> (d\<^sup>T * top)\<^sup>T) * top"
by (simp add: comp_associative comp_inf_vector_1 conv_dist_comp)
also have "... = (w \<sqinter> top * d \<sqinter> -d) * top"
by (simp add: conv_dist_comp inf_commute inf_left_commute)
finally show ?thesis
using assms epm_3 by simp
qed
lemma kruskal_subgraph_inv:
assumes "f \<le> --(-h \<sqinter> g)"
and "e \<le> --g"
and "symmetric h"
and "symmetric g"
shows "(f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e \<le> --(-(h \<sqinter> -e \<sqinter> -e\<^sup>T) \<sqinter> g)"
proof -
let ?f = "(f \<sqinter> -q) \<squnion> (f \<sqinter> q)\<^sup>T \<squnion> e"
let ?h = "h \<sqinter> -e \<sqinter> -e\<^sup>T"
have 1: "f \<sqinter> -q \<le> -h \<sqinter> --g"
using assms(1) inf.coboundedI1 by simp
have "(f \<sqinter> q)\<^sup>T \<le> (-h \<sqinter> --g)\<^sup>T"
using assms(1) inf.coboundedI1 conv_isotone by simp
also have "... = -h \<sqinter> --g"
using assms(3,4) conv_complement conv_dist_inf by simp
finally have "?f \<le> (-h \<sqinter> --g) \<squnion> (e \<sqinter> --g)"
using 1 assms(2) inf.absorb1 semiring.add_right_mono by simp
also have "... \<le> (-h \<squnion> --e) \<sqinter> --g"
by (simp add: inf.coboundedI1 le_supI2 pp_increasing)
also have "... \<le> -?h \<sqinter> --g"
using inf.sup_left_isotone order_trans p_antitone_inf p_supdist_inf by blast
finally show "?f \<le> --(-?h \<sqinter> g)"
using inf_pp_semi_commute order_lesseq_imp by blast
qed
end
subsection \<open>Stone Relation Algebras\<close>
text \<open>
We add \<open>pp_dist_comp\<close> and \<open>pp_one\<close>, which follow in relation algebras but not in the present setting.
The main change is that only a Stone algebra is required, not a Boolean algebra.
\<close>
class stone_relation_algebra = pd_allegory + stone_algebra +
assumes pp_dist_comp : "--(x * y) = --x * --y"
assumes pp_one [simp]: "--1 = 1"
begin
text \<open>
The following property is a simple consequence of the Stone axiom.
We cannot hope to remove the double complement in it.
\<close>
lemma conv_complement_0_p [simp]:
"(-x)\<^sup>T \<squnion> (--x)\<^sup>T = top"
by (metis conv_top conv_dist_sup stone)
lemma theorem24xxiv_pp:
"-(x * y) \<squnion> --(x * z) = -(x * (y \<sqinter> -z)) \<squnion> --(x * z)"
by (metis p_dist_inf theorem24xxiii)
lemma asymmetric_linear:
"asymmetric x \<longleftrightarrow> linear (-x)"
by (metis conv_complement inf.distrib_left inf_p maddux_3_11_pp p_bot p_dist_inf)
lemma strict_linear_asymmetric:
"strict_linear x \<Longrightarrow> antisymmetric (-x)"
by (metis conv_complement eq_refl p_dist_sup pp_one)
lemma regular_complement_top:
"regular x \<Longrightarrow> x \<squnion> -x = top"
by (metis stone)
lemma regular_mult_closed:
"regular x \<Longrightarrow> regular y \<Longrightarrow> regular (x * y)"
by (simp add: pp_dist_comp)
lemma regular_one_closed:
"regular 1"
by simp
text \<open>
The following variants of total and surjective are useful for graphs.
\<close>
lemma pp_total:
"total (--x) \<longleftrightarrow> -(x*top) = bot"
by (simp add: dense_pp pp_dist_comp)
lemma pp_surjective:
"surjective (--x) \<longleftrightarrow> -(top*x) = bot"
by (metis p_bot p_comp_pp p_top pp_dist_comp)
text \<open>
Bijective elements and mappings are necessarily regular, that is, invariant under double-complement.
This implies that points are regular.
Moreover, also arcs are regular.
\<close>
lemma bijective_regular:
"bijective x \<Longrightarrow> regular x"
by (metis comp_bijective_complement mult_left_one regular_one_closed)
lemma mapping_regular:
"mapping x \<Longrightarrow> regular x"
by (metis bijective_regular conv_complement conv_involutive total_conv_surjective)
lemma arc_regular:
assumes "arc x"
shows "regular x"
proof -
have "--x \<le> --(x * top \<sqinter> top * x)"
by (simp add: pp_isotone top_left_mult_increasing top_right_mult_increasing)
also have "... = --(x * top) \<sqinter> --(top * x)"
by simp
also have "... = x * top \<sqinter> top * x"
by (metis assms bijective_regular conv_top conv_dist_comp conv_involutive mapping_regular)
also have "... \<le> x * x\<^sup>T * top * x"
by (metis comp_associative dedekind_1 inf.commute inf_top.right_neutral)
also have "... \<le> x"
by (metis assms comp_right_one conv_top comp_associative conv_dist_comp conv_involutive mult_right_isotone vector_top_closed)
finally show ?thesis
- by (simp add: antisym pp_increasing)
+ by (simp add: order.antisym pp_increasing)
qed
(*
lemma conv_complement_0 [simp]: "x\<^sup>T \<squnion> (-x)\<^sup>T = top" nitpick [expect=genuine] oops
lemma schroeder_3: "x * y \<le> z \<longleftrightarrow> x\<^sup>T * -z \<le> -y" nitpick [expect=genuine] oops
lemma schroeder_4: "x * y \<le> z \<longleftrightarrow> -z * y\<^sup>T \<le> -x" nitpick [expect=genuine] oops
lemma theorem24xxiv: "-(x * y) \<squnion> (x * z) = -(x * (y \<sqinter> -z)) \<squnion> (x * z)" nitpick [expect=genuine] oops
lemma vector_N: "x = x * top \<longrightarrow> N(x) = x \<sqinter> 1" nitpick [expect=genuine] oops
lemma N_vector [simp]: "N(x * top) = x * top \<sqinter> 1" nitpick [expect=genuine] oops
lemma N_vector_top [simp]: "N(x * top) * top = x * top" nitpick [expect=genuine] oops
lemma N_below_inf_one: "N(x) \<le> x \<sqinter> 1" nitpick [expect=genuine] oops
lemma N_below: "N(x) \<le> x" nitpick [expect=genuine] oops
lemma n_split_omega_mult: "xs * xo = xo \<and> xo * top = xo \<longrightarrow> N(top) * xo = xs * N(xo) * top" nitpick [expect=genuine] oops
lemma complement_vector: "vector v \<longleftrightarrow> vector (-v)" nitpick [expect=genuine] oops
lemma complement_covector: "covector v \<longleftrightarrow> covector (-v)" nitpick [expect=genuine] oops
lemma triple_schroeder: "x * y * z \<le> w \<longleftrightarrow> x\<^sup>T * -w * z\<^sup>T \<le> -y" nitpick [expect=genuine] oops
lemma schroeder_5: "x * y \<le> z \<longleftrightarrow> y * -z\<^sup>T \<le> -x\<^sup>T" nitpick [expect=genuine] oops
lemma schroeder_6: "x * y \<le> z \<longleftrightarrow> -z\<^sup>T * x \<le> -y\<^sup>T" nitpick [expect=genuine] oops
*)
end
text \<open>
Every Stone algebra can be expanded to a Stone relation algebra by identifying the semiring and lattice structures and taking identity as converse.
\<close>
sublocale stone_algebra < comp_inf: stone_relation_algebra where one = top and times = inf and conv = id
proof (unfold_locales, goal_cases)
case 7
show ?case by (simp add: inf_commute)
qed (auto simp: inf.assoc inf_sup_distrib2 inf_left_commute)
text \<open>
Every bounded linear order can be expanded to a Stone algebra, which can be expanded to a Stone relation algebra by reusing some of the operations.
In particular, composition is meet, its identity is \<open>top\<close> and converse is the identity function.
\<close>
class linorder_stone_relation_algebra_expansion = linorder_stone_algebra_expansion + times + conv + one +
assumes times_def [simp]: "x * y = min x y"
assumes conv_def [simp]: "x\<^sup>T = x"
assumes one_def [simp]: "1 = top"
begin
lemma times_inf [simp]:
"x * y = x \<sqinter> y"
by simp
subclass stone_relation_algebra
apply unfold_locales
using comp_inf.mult_right_dist_sup inf_commute inf_assoc inf_left_commute pp_dist_inf min_def by simp_all
lemma times_dense:
"x \<noteq> bot \<Longrightarrow> y \<noteq> bot \<Longrightarrow> x * y \<noteq> bot"
using inf_dense min_inf times_def by presburger
end
subsection \<open>Relation Algebras\<close>
text \<open>
For a relation algebra, we only require that the underlying lattice is a Boolean algebra.
In fact, the only missing axiom is that double-complement is the identity.
\<close>
class relation_algebra = boolean_algebra + stone_relation_algebra
begin
lemma conv_complement_0 [simp]:
"x\<^sup>T \<squnion> (-x)\<^sup>T = top"
by (simp add: conv_complement)
text \<open>
We now obtain the original formulations of the Schr\"oder equivalences.
\<close>
lemma schroeder_3:
"x * y \<le> z \<longleftrightarrow> x\<^sup>T * -z \<le> -y"
by (simp add: schroeder_3_p)
lemma schroeder_4:
"x * y \<le> z \<longleftrightarrow> -z * y\<^sup>T \<le> -x"
by (simp add: schroeder_4_p)
lemma theorem24xxiv:
"-(x * y) \<squnion> (x * z) = -(x * (y \<sqinter> -z)) \<squnion> (x * z)"
using theorem24xxiv_pp by auto
lemma vector_N:
"vector x \<Longrightarrow> N(x) = x \<sqinter> 1"
by (simp add: vector_N_pp)
lemma N_vector [simp]:
"N(x * top) = x * top \<sqinter> 1"
by simp
lemma N_vector_top [simp]:
"N(x * top) * top = x * top"
using N_vector_top_pp by simp
lemma N_below_inf_one:
"N(x) \<le> x \<sqinter> 1"
using N_below_inf_one_pp by simp
lemma N_below:
"N(x) \<le> x"
using N_below_pp by simp
lemma n_split_omega_mult:
"xs * xo = xo \<Longrightarrow> xo * top = xo \<Longrightarrow> N(top) * xo = xs * N(xo) * top"
using n_split_omega_mult_pp by simp
lemma complement_vector:
"vector v \<longleftrightarrow> vector (-v)"
using vector_complement_closed by fastforce
lemma complement_covector:
"covector v \<longleftrightarrow> covector (-v)"
using covector_complement_closed by force
lemma triple_schroeder:
"x * y * z \<le> w \<longleftrightarrow> x\<^sup>T * -w * z\<^sup>T \<le> -y"
by (simp add: triple_schroeder_p)
lemma schroeder_5:
"x * y \<le> z \<longleftrightarrow> y * -z\<^sup>T \<le> -x\<^sup>T"
by (simp add: conv_complement schroeder_5_p)
lemma schroeder_6:
"x * y \<le> z \<longleftrightarrow> -z\<^sup>T * x \<le> -y\<^sup>T"
by (simp add: conv_complement schroeder_5_p)
end
text \<open>
We briefly look at the so-called Tarski rule.
In some models of Stone relation algebras it only holds for regular elements, so we add this as an assumption.
\<close>
class stone_relation_algebra_tarski = stone_relation_algebra +
assumes tarski: "regular x \<Longrightarrow> x \<noteq> bot \<Longrightarrow> top * x * top = top"
begin
text \<open>
We can then show, for example, that every arc is contained in a pseudocomplemented relation or its pseudocomplement.
\<close>
lemma arc_in_partition:
assumes "arc x"
shows "x \<le> -y \<or> x \<le> --y"
proof -
have 1: "x * top * x\<^sup>T \<le> 1 \<and> x\<^sup>T * top * x \<le> 1"
using assms arc_expanded by auto
have "\<not> x \<le> --y \<longrightarrow> x \<le> -y"
proof
assume "\<not> x \<le> --y"
hence "x \<sqinter> -y \<noteq> bot"
using pseudo_complement by simp
hence "top * (x \<sqinter> -y) * top = top"
using assms arc_regular tarski by auto
hence "x = x \<sqinter> top * (x \<sqinter> -y) * top"
by simp
also have "... \<le> x \<sqinter> x * ((x \<sqinter> -y) * top)\<^sup>T * (x \<sqinter> -y) * top"
by (metis dedekind_2 inf.cobounded1 inf.boundedI inf_commute mult_assoc inf.absorb2 top.extremum)
also have "... = x \<sqinter> x * top * (x\<^sup>T \<sqinter> -y\<^sup>T) * (x \<sqinter> -y) * top"
by (simp add: comp_associative conv_complement conv_dist_comp conv_dist_inf)
also have "... \<le> x \<sqinter> x * top * x\<^sup>T * (x \<sqinter> -y) * top"
using inf.sup_right_isotone mult_left_isotone mult_right_isotone by auto
also have "... \<le> x \<sqinter> 1 * (x \<sqinter> -y) * top"
using 1 by (metis comp_associative comp_isotone inf.sup_right_isotone mult_1_left mult_semi_associative)
also have "... = x \<sqinter> (x \<sqinter> -y) * top"
by simp
also have "... \<le> (x \<sqinter> -y) * ((x \<sqinter> -y)\<^sup>T * x)"
by (metis dedekind_1 inf_commute inf_top_right)
also have "... \<le> (x \<sqinter> -y) * (x\<^sup>T * x)"
by (simp add: conv_dist_inf mult_left_isotone mult_right_isotone)
also have "... \<le> (x \<sqinter> -y) * (x\<^sup>T * top * x)"
by (simp add: mult_assoc mult_right_isotone top_left_mult_increasing)
also have "... \<le> x \<sqinter> -y"
using 1 by (metis mult_right_isotone mult_1_right)
finally show "x \<le> -y"
by simp
qed
thus ?thesis
by auto
qed
lemma non_bot_arc_in_partition_xor:
assumes "arc x"
and "x \<noteq> bot"
shows "(x \<le> -y \<and> \<not> x \<le> --y) \<or> (\<not> x \<le> -y \<and> x \<le> --y)"
proof -
have "x \<le> -y \<and> x \<le> --y \<longrightarrow> False"
by (simp add: assms(2) inf_absorb1 shunting_1_pp)
thus ?thesis
using assms(1) arc_in_partition by auto
qed
lemma point_in_vector_or_pseudo_complement:
assumes "point p"
and "vector v"
shows "p \<le> --v \<or> p \<le> -v"
proof (rule disjCI)
assume "\<not>(p \<le> -v)"
hence "top * (p \<sqinter> --v) = top"
by (smt assms bijective_regular regular_closed_inf regular_closed_p shunting_1_pp tarski vector_complement_closed vector_inf_closed vector_mult_closed)
thus "p \<le> --v"
by (metis assms(1) epm_3 inf.absorb_iff1 inf.cobounded1 inf_top.right_neutral)
qed
lemma distinct_points:
assumes "point x"
and "point y"
and "x \<noteq> y"
shows "x \<sqinter> y = bot"
- by (metis assms antisym comp_bijective_complement inf.sup_monoid.add_commute mult_left_one pseudo_complement regular_one_closed point_in_vector_or_pseudo_complement)
+ by (metis assms order.antisym comp_bijective_complement inf.sup_monoid.add_commute mult_left_one pseudo_complement regular_one_closed point_in_vector_or_pseudo_complement)
lemma point_in_vector_or_complement:
assumes "point p"
and "vector v"
and "regular v"
shows "p \<le> v \<or> p \<le> -v"
using assms point_in_vector_or_pseudo_complement by fastforce
lemma point_in_vector_sup:
assumes "point p"
and "vector v"
and "regular v"
and "p \<le> v \<squnion> w"
shows "p \<le> v \<or> p \<le> w"
by (metis assms inf.absorb1 shunting_var_p sup_commute point_in_vector_or_complement)
lemma point_atomic_vector:
assumes "point x"
and "vector y"
and "regular y"
and "y \<le> x"
shows "y = x \<or> y = bot"
proof (cases "x \<le> -y")
case True
thus ?thesis
using assms(4) inf.absorb2 pseudo_complement by force
next
case False
thus ?thesis
using assms point_in_vector_or_pseudo_complement by fastforce
qed
lemma point_in_vector_or_complement_2:
assumes "point x"
and "vector y"
and "regular y"
and "\<not> y \<le> -x"
shows "x \<le> y"
using assms point_in_vector_or_pseudo_complement p_antitone_iff by fastforce
text \<open>The next three lemmas \<open>arc_in_arc_or_complement\<close>, \<open>arc_in_sup_arc\<close> and \<open>different_arc_in_sup_arc\<close> were contributed by Nicolas Robinson-O'Brien.\<close>
lemma arc_in_arc_or_complement:
assumes "arc x"
and "arc y"
and "\<not> x \<le> y"
shows "x \<le> -y"
using assms arc_in_partition arc_regular by force
lemma arc_in_sup_arc:
assumes "arc x"
and "arc y"
and "x \<le> z \<squnion> y"
shows "x \<le> z \<or> x \<le> y"
proof (cases "x \<le> y")
case True
thus ?thesis
by simp
next
case False
hence "x \<le> -y"
using assms(1,2) arc_in_arc_or_complement by blast
hence "x \<le> -y \<sqinter> (z \<squnion> y)"
using assms(3) by simp
hence "x \<le> z"
by (metis inf.boundedE inf.sup_monoid.add_commute maddux_3_13 sup_commute)
thus ?thesis
by simp
qed
lemma different_arc_in_sup_arc:
assumes "arc x"
and "arc y"
and "x \<le> z \<squnion> y"
and "x \<noteq> y"
shows "x \<le> z"
proof -
have "x \<le> -y"
- using arc_in_arc_or_complement assms(1,2,4) eq_iff p_antitone_iff by blast
+ using arc_in_arc_or_complement assms(1,2,4) order.eq_iff p_antitone_iff by blast
hence "x \<le> -y \<sqinter> (z \<squnion> y)"
using assms arc_in_sup_arc by simp
thus ?thesis
by (metis order_lesseq_imp p_inf_sup_below sup_commute)
qed
end
class relation_algebra_tarski = relation_algebra + stone_relation_algebra_tarski
text \<open>
Finally, the above axioms of relation algebras do not imply that they contain at least two elements.
This is necessary, for example, to show that arcs are not empty.
\<close>
class stone_relation_algebra_consistent = stone_relation_algebra +
assumes consistent: "bot \<noteq> top"
begin
lemma arc_not_bot:
"arc x \<Longrightarrow> x \<noteq> bot"
using consistent mult_right_zero by auto
end
class relation_algebra_consistent = relation_algebra + stone_relation_algebra_consistent
class stone_relation_algebra_tarski_consistent = stone_relation_algebra_tarski + stone_relation_algebra_consistent
begin
lemma arc_in_partition_xor:
"arc x \<Longrightarrow> (x \<le> -y \<and> \<not> x \<le> --y) \<or> (\<not> x \<le> -y \<and> x \<le> --y)"
by (simp add: non_bot_arc_in_partition_xor arc_not_bot)
end
end
diff --git a/thys/Stone_Relation_Algebras/Semirings.thy b/thys/Stone_Relation_Algebras/Semirings.thy
--- a/thys/Stone_Relation_Algebras/Semirings.thy
+++ b/thys/Stone_Relation_Algebras/Semirings.thy
@@ -1,689 +1,689 @@
(* Title: Semirings
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Semirings\<close>
text \<open>
This theory develops a hierarchy of idempotent semirings.
All kinds of semiring considered here are bounded semilattices, but many lack additional properties typically assumed for semirings.
In particular, we consider the variants of semirings, in which
\begin{itemize}
\item multiplication is not required to be associative;
\item a right zero and unit of multiplication need not exist;
\item multiplication has a left residual;
\item multiplication from the left is not required to distribute over addition;
\item the semilattice order has a greatest element.
\end{itemize}
We have applied results from this theory a number of papers for unifying computation models.
For example, see \cite{Guttmann2012c} for various relational and matrix-based computation models and \cite{BerghammerGuttmann2015b} for multirelational models.
The main results in this theory relate different ways of defining reflexive-transitive closures as discussed in \cite{BerghammerGuttmann2015b}.
\<close>
theory Semirings
imports Fixpoints
begin
subsection \<open>Idempotent Semirings\<close>
text \<open>
The following definitions are standard for relations.
Putting them into a general class that depends only on the signature facilitates reuse.
Coreflexives are sometimes called partial identities, subidentities, monotypes or tests.
\<close>
class times_one_ord = times + one + ord
begin
abbreviation reflexive :: "'a \<Rightarrow> bool" where "reflexive x \<equiv> 1 \<le> x"
abbreviation coreflexive :: "'a \<Rightarrow> bool" where "coreflexive x \<equiv> x \<le> 1"
abbreviation transitive :: "'a \<Rightarrow> bool" where "transitive x \<equiv> x * x \<le> x"
abbreviation dense_rel :: "'a \<Rightarrow> bool" where "dense_rel x \<equiv> x \<le> x * x"
abbreviation idempotent :: "'a \<Rightarrow> bool" where "idempotent x \<equiv> x * x = x"
abbreviation preorder :: "'a \<Rightarrow> bool" where "preorder x \<equiv> reflexive x \<and> transitive x"
abbreviation "coreflexives \<equiv> { x . coreflexive x }"
end
text \<open>
The first algebra is a very weak idempotent semiring, in which multiplication is not necessarily associative.
\<close>
class non_associative_left_semiring = bounded_semilattice_sup_bot + times + one +
assumes mult_left_sub_dist_sup: "x * y \<squnion> x * z \<le> x * (y \<squnion> z)"
assumes mult_right_dist_sup: "(x \<squnion> y) * z = x * z \<squnion> y * z"
assumes mult_left_zero [simp]: "bot * x = bot"
assumes mult_left_one [simp]: "1 * x = x"
assumes mult_sub_right_one: "x \<le> x * 1"
begin
subclass times_one_ord .
text \<open>
We first show basic isotonicity and subdistributivity properties of multiplication.
\<close>
lemma mult_left_isotone:
"x \<le> y \<Longrightarrow> x * z \<le> y * z"
using mult_right_dist_sup sup_right_divisibility by auto
lemma mult_right_isotone:
"x \<le> y \<Longrightarrow> z * x \<le> z * y"
using mult_left_sub_dist_sup sup.bounded_iff sup_right_divisibility by auto
lemma mult_isotone:
"w \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> w * x \<le> y * z"
using order_trans mult_left_isotone mult_right_isotone by blast
lemma affine_isotone:
"isotone (\<lambda>x . y * x \<squnion> z)"
using isotone_def mult_right_isotone sup_left_isotone by auto
lemma mult_left_sub_dist_sup_left:
"x * y \<le> x * (y \<squnion> z)"
by (simp add: mult_right_isotone)
lemma mult_left_sub_dist_sup_right:
"x * z \<le> x * (y \<squnion> z)"
by (simp add: mult_right_isotone)
lemma mult_right_sub_dist_sup_left:
"x * z \<le> (x \<squnion> y) * z"
by (simp add: mult_left_isotone)
lemma mult_right_sub_dist_sup_right:
"y * z \<le> (x \<squnion> y) * z"
by (simp add: mult_left_isotone)
lemma case_split_left:
assumes "1 \<le> w \<squnion> z"
and "w * x \<le> y"
and "z * x \<le> y"
shows "x \<le> y"
proof -
have "(w \<squnion> z) * x \<le> y"
by (simp add: assms(2-3) mult_right_dist_sup)
thus ?thesis
by (metis assms(1) dual_order.trans mult_left_one mult_left_isotone)
qed
lemma case_split_left_equal:
"w \<squnion> z = 1 \<Longrightarrow> w * x = w * y \<Longrightarrow> z * x = z * y \<Longrightarrow> x = y"
by (metis mult_left_one mult_right_dist_sup)
text \<open>
Next we consider under which semiring operations the above properties are closed.
\<close>
lemma reflexive_one_closed:
"reflexive 1"
by simp
lemma reflexive_sup_closed:
"reflexive x \<Longrightarrow> reflexive (x \<squnion> y)"
by (simp add: le_supI1)
lemma reflexive_mult_closed:
"reflexive x \<Longrightarrow> reflexive y \<Longrightarrow> reflexive (x * y)"
using mult_isotone by fastforce
lemma coreflexive_bot_closed:
"coreflexive bot"
by simp
lemma coreflexive_one_closed:
"coreflexive 1"
by simp
lemma coreflexive_sup_closed:
"coreflexive x \<Longrightarrow> coreflexive y \<Longrightarrow> coreflexive (x \<squnion> y)"
by simp
lemma coreflexive_mult_closed:
"coreflexive x \<Longrightarrow> coreflexive y \<Longrightarrow> coreflexive (x * y)"
using mult_isotone by fastforce
lemma transitive_bot_closed:
"transitive bot"
by simp
lemma transitive_one_closed:
"transitive 1"
by simp
lemma dense_bot_closed:
"dense_rel bot"
by simp
lemma dense_one_closed:
"dense_rel 1"
by simp
lemma dense_sup_closed:
"dense_rel x \<Longrightarrow> dense_rel y \<Longrightarrow> dense_rel (x \<squnion> y)"
by (metis mult_right_dist_sup order_lesseq_imp sup.mono mult_left_sub_dist_sup_left mult_left_sub_dist_sup_right)
lemma idempotent_bot_closed:
"idempotent bot"
by simp
lemma idempotent_one_closed:
"idempotent 1"
by simp
lemma preorder_one_closed:
"preorder 1"
by simp
lemma coreflexive_transitive:
"coreflexive x \<Longrightarrow> transitive x"
using mult_left_isotone by fastforce
lemma preorder_idempotent:
"preorder x \<Longrightarrow> idempotent x"
- using antisym mult_isotone by fastforce
+ using order.antisym mult_isotone by fastforce
text \<open>
We study the following three ways of defining reflexive-transitive closures.
Each of them is given as a least prefixpoint, but the underlying functions are different.
They implement left recursion, right recursion and symmetric recursion, respectively.
\<close>
abbreviation Lf :: "'a \<Rightarrow> ('a \<Rightarrow> 'a)" where "Lf y \<equiv> (\<lambda>x . 1 \<squnion> x * y)"
abbreviation Rf :: "'a \<Rightarrow> ('a \<Rightarrow> 'a)" where "Rf y \<equiv> (\<lambda>x . 1 \<squnion> y * x)"
abbreviation Sf :: "'a \<Rightarrow> ('a \<Rightarrow> 'a)" where "Sf y \<equiv> (\<lambda>x . 1 \<squnion> y \<squnion> x * x)"
abbreviation lstar :: "'a \<Rightarrow> 'a" where "lstar y \<equiv> p\<mu> (Lf y)"
abbreviation rstar :: "'a \<Rightarrow> 'a" where "rstar y \<equiv> p\<mu> (Rf y)"
abbreviation sstar :: "'a \<Rightarrow> 'a" where "sstar y \<equiv> p\<mu> (Sf y)"
text \<open>
All functions are isotone and, therefore, if the prefixpoints exist they are also fixpoints.
\<close>
lemma lstar_rec_isotone:
"isotone (Lf y)"
using isotone_def sup_right_divisibility sup_right_isotone mult_right_sub_dist_sup_right by auto
lemma rstar_rec_isotone:
"isotone (Rf y)"
using isotone_def sup_right_divisibility sup_right_isotone mult_left_sub_dist_sup_right by auto
lemma sstar_rec_isotone:
"isotone (Sf y)"
using isotone_def sup_right_isotone mult_isotone by auto
lemma lstar_fixpoint:
"has_least_prefixpoint (Lf y) \<Longrightarrow> lstar y = \<mu> (Lf y)"
by (simp add: pmu_mu lstar_rec_isotone)
lemma rstar_fixpoint:
"has_least_prefixpoint (Rf y) \<Longrightarrow> rstar y = \<mu> (Rf y)"
by (simp add: pmu_mu rstar_rec_isotone)
lemma sstar_fixpoint:
"has_least_prefixpoint (Sf y) \<Longrightarrow> sstar y = \<mu> (Sf y)"
by (simp add: pmu_mu sstar_rec_isotone)
lemma sstar_increasing:
"has_least_prefixpoint (Sf y) \<Longrightarrow> y \<le> sstar y"
using order_trans pmu_unfold sup_ge1 sup_ge2 by blast
text \<open>
The fixpoint given by right recursion is always below the one given by symmetric recursion.
\<close>
lemma rstar_below_sstar:
assumes "has_least_prefixpoint (Rf y)"
and "has_least_prefixpoint (Sf y)"
shows "rstar y \<le> sstar y"
proof -
have "y \<le> sstar y"
using assms(2) pmu_unfold by force
hence "Rf y (sstar y) \<le> Sf y (sstar y)"
by (meson sup.cobounded1 sup.mono mult_left_isotone)
also have "... \<le> sstar y"
using assms(2) pmu_unfold by blast
finally show ?thesis
using assms(1) is_least_prefixpoint_def least_prefixpoint by auto
qed
end
text \<open>
Our next structure adds one half of the associativity property.
This inequality holds, for example, for multirelations under the compositions defined by Parikh and Peleg \cite{Parikh1983,Peleg1987}.
The converse inequality requires up-closed multirelations for Parikh's composition.
\<close>
class pre_left_semiring = non_associative_left_semiring +
assumes mult_semi_associative: "(x * y) * z \<le> x * (y * z)"
begin
lemma mult_one_associative [simp]:
"x * 1 * y = x * y"
by (metis dual_order.antisym mult_left_isotone mult_left_one mult_semi_associative mult_sub_right_one)
lemma mult_sup_associative_one:
"(x * (y * 1)) * z \<le> x * (y * z)"
by (metis mult_semi_associative mult_one_associative)
lemma rstar_increasing:
assumes "has_least_prefixpoint (Rf y)"
shows "y \<le> rstar y"
proof -
have "Rf y (rstar y) \<le> rstar y"
using assms pmu_unfold by blast
thus ?thesis
by (metis le_supE mult_right_isotone mult_sub_right_one sup.absorb_iff2)
qed
end
text \<open>
For the next structure we add a left residual operation.
Such a residual is available, for example, for multirelations.
The operator notation for binary division is introduced in a class that requires a unary inverse.
This is appropriate for fields, but too strong in the present context of semirings.
We therefore reintroduce it without requiring a unary inverse.
\<close>
no_notation
inverse_divide (infixl "'/" 70)
notation
divide (infixl "'/" 70)
class residuated_pre_left_semiring = pre_left_semiring + divide +
assumes lres_galois: "x * y \<le> z \<longleftrightarrow> x \<le> z / y"
begin
text \<open>
We first derive basic properties of left residuals from the Galois connection.
\<close>
lemma lres_left_isotone:
"x \<le> y \<Longrightarrow> x / z \<le> y / z"
using dual_order.trans lres_galois by blast
lemma lres_right_antitone:
"x \<le> y \<Longrightarrow> z / y \<le> z / x"
using dual_order.trans lres_galois mult_right_isotone by blast
lemma lres_inverse:
"(x / y) * y \<le> x"
by (simp add: lres_galois)
lemma lres_one:
"x / 1 \<le> x"
using mult_sub_right_one order_trans lres_inverse by blast
lemma lres_mult_sub_lres_lres:
"x / (z * y) \<le> (x / y) / z"
using lres_galois mult_semi_associative order.trans by blast
lemma mult_lres_sub_assoc:
"x * (y / z) \<le> (x * y) / z"
by (meson dual_order.trans lres_galois mult_right_isotone lres_inverse lres_mult_sub_lres_lres)
text \<open>
With the help of a left residual, it follows that left recursion is below right recursion.
\<close>
lemma lstar_below_rstar:
assumes "has_least_prefixpoint (Lf y)"
and "has_least_prefixpoint (Rf y)"
shows "lstar y \<le> rstar y"
proof -
have "y * (rstar y / y) * y \<le> y * rstar y"
using lres_galois mult_lres_sub_assoc by auto
also have "... \<le> rstar y"
using assms(2) le_supE pmu_unfold by blast
finally have "y * (rstar y / y) \<le> rstar y / y"
by (simp add: lres_galois)
hence "Rf y (rstar y / y) \<le> rstar y / y"
using assms(2) lres_galois rstar_increasing by fastforce
hence "rstar y \<le> rstar y / y"
using assms(2) is_least_prefixpoint_def least_prefixpoint by auto
hence "Lf y (rstar y) \<le> rstar y"
using assms(2) lres_galois pmu_unfold by fastforce
thus ?thesis
using assms(1) is_least_prefixpoint_def least_prefixpoint by auto
qed
text \<open>
Moreover, right recursion gives the same result as symmetric recursion.
The next proof follows an argument of \cite[Satz 10.1.5]{Berghammer2012}.
\<close>
lemma rstar_sstar:
assumes "has_least_prefixpoint (Rf y)"
and "has_least_prefixpoint (Sf y)"
shows "rstar y = sstar y"
proof -
have "Rf y (rstar y / rstar y) * rstar y \<le> rstar y \<squnion> y * ((rstar y / rstar y) * rstar y)"
using mult_right_dist_sup mult_semi_associative sup_right_isotone by auto
also have "... \<le> rstar y \<squnion> y * rstar y"
using mult_right_isotone sup_right_isotone lres_inverse by blast
also have "... \<le> rstar y"
using assms(1) pmu_unfold by fastforce
finally have "Rf y (rstar y / rstar y) \<le> rstar y / rstar y"
by (simp add: lres_galois)
hence "rstar y * rstar y \<le> rstar y"
using assms(1) is_least_prefixpoint_def least_prefixpoint lres_galois by auto
hence "y \<squnion> rstar y * rstar y \<le> rstar y"
by (simp add: assms(1) rstar_increasing)
hence "Sf y (rstar y) \<le> rstar y"
using assms(1) pmu_unfold by force
hence "sstar y \<le> rstar y"
using assms(2) is_least_prefixpoint_def least_prefixpoint by auto
thus ?thesis
- by (simp add: assms antisym rstar_below_sstar)
+ by (simp add: assms order.antisym rstar_below_sstar)
qed
end
text \<open>
In the next structure we add full associativity of multiplication, as well as a right unit.
Still, multiplication does not need to have a right zero and does not need to distribute over addition from the left.
\<close>
class idempotent_left_semiring = non_associative_left_semiring + monoid_mult
begin
subclass pre_left_semiring
by unfold_locales (simp add: mult_assoc)
lemma zero_right_mult_decreasing:
"x * bot \<le> x"
by (metis bot_least mult_1_right mult_right_isotone)
text \<open>
The following result shows that for dense coreflexives there are two equivalent ways to express that a property is preserved.
In the setting of Kleene algebras, this is well known for tests, which form a Boolean subalgebra.
The point here is that only very few properties of tests are needed to show the equivalence.
\<close>
lemma test_preserves_equation:
assumes "dense_rel p"
and "coreflexive p"
shows "p * x \<le> x * p \<longleftrightarrow> p * x = p * x * p"
proof
assume 1: "p * x \<le> x * p"
have "p * x \<le> p * p * x"
by (simp add: assms(1) mult_left_isotone)
also have "... \<le> p * x * p"
using 1 by (simp add: mult_right_isotone mult_assoc)
finally show "p * x = p * x * p"
- using assms(2) antisym mult_right_isotone by fastforce
+ using assms(2) order.antisym mult_right_isotone by fastforce
next
assume "p * x = p * x * p"
thus "p * x \<le> x * p"
by (metis assms(2) mult_left_isotone mult_left_one)
qed
end
text \<open>
The next structure has both distributivity properties of multiplication.
Only a right zero is missing from full semirings.
This is important as many computation models do not have a right zero of sequential composition.
\<close>
class idempotent_left_zero_semiring = idempotent_left_semiring +
assumes mult_left_dist_sup: "x * (y \<squnion> z) = x * y \<squnion> x * z"
begin
lemma case_split_right:
assumes "1 \<le> w \<squnion> z"
and "x * w \<le> y"
and "x * z \<le> y"
shows "x \<le> y"
proof -
have "x * (w \<squnion> z) \<le> y"
by (simp add: assms(2-3) mult_left_dist_sup)
thus ?thesis
by (metis assms(1) dual_order.trans mult_1_right mult_right_isotone)
qed
lemma case_split_right_equal:
"w \<squnion> z = 1 \<Longrightarrow> x * w = y * w \<Longrightarrow> x * z = y * z \<Longrightarrow> x = y"
by (metis mult_1_right mult_left_dist_sup)
text \<open>
This is the first structure we can connect to the semirings provided by Isabelle/HOL.
\<close>
sublocale semiring: ordered_semiring sup bot less_eq less times
apply unfold_locales
using sup_right_isotone apply blast
apply (simp add: mult_right_dist_sup)
apply (simp add: mult_left_dist_sup)
apply (simp add: mult_right_isotone)
by (simp add: mult_left_isotone)
sublocale semiring: semiring_numeral 1 times sup ..
end
text \<open>
Completing this part of the hierarchy, we obtain idempotent semirings by adding a right zero of multiplication.
\<close>
class idempotent_semiring = idempotent_left_zero_semiring +
assumes mult_right_zero [simp]: "x * bot = bot"
begin
sublocale semiring: semiring_0 sup bot times
by unfold_locales simp_all
end
subsection \<open>Bounded Idempotent Semirings\<close>
text \<open>
All of the following semirings have a greatest element in the underlying semilattice order.
With this element, we can express further standard properties of relations.
We extend each class in the above hierarchy in turn.
\<close>
class times_top = times + top
begin
abbreviation vector :: "'a \<Rightarrow> bool" where "vector x \<equiv> x * top = x"
abbreviation covector :: "'a \<Rightarrow> bool" where "covector x \<equiv> top * x = x"
abbreviation total :: "'a \<Rightarrow> bool" where "total x \<equiv> x * top = top"
abbreviation surjective :: "'a \<Rightarrow> bool" where "surjective x \<equiv> top * x = top"
abbreviation "vectors \<equiv> { x . vector x }"
abbreviation "covectors \<equiv> { x . covector x }"
end
class bounded_non_associative_left_semiring = non_associative_left_semiring + top +
assumes sup_right_top [simp]: "x \<squnion> top = top"
begin
subclass times_top .
text \<open>
We first give basic properties of the greatest element.
\<close>
lemma sup_left_top [simp]:
"top \<squnion> x = top"
using sup_right_top sup.commute by fastforce
lemma top_greatest [simp]:
"x \<le> top"
by (simp add: le_iff_sup)
lemma top_left_mult_increasing:
"x \<le> top * x"
by (metis mult_left_isotone mult_left_one top_greatest)
lemma top_right_mult_increasing:
"x \<le> x * top"
using mult_right_isotone mult_sub_right_one order_trans top_greatest by blast
lemma top_mult_top [simp]:
"top * top = top"
- by (simp add: antisym top_left_mult_increasing)
+ by (simp add: order.antisym top_left_mult_increasing)
text \<open>
Closure of the above properties under the semiring operations is considered next.
\<close>
lemma vector_bot_closed:
"vector bot"
by simp
lemma vector_top_closed:
"vector top"
by simp
lemma vector_sup_closed:
"vector x \<Longrightarrow> vector y \<Longrightarrow> vector (x \<squnion> y)"
by (simp add: mult_right_dist_sup)
lemma covector_top_closed:
"covector top"
by simp
lemma total_one_closed:
"total 1"
by simp
lemma total_top_closed:
"total top"
by simp
lemma total_sup_closed:
"total x \<Longrightarrow> total (x \<squnion> y)"
by (simp add: mult_right_dist_sup)
lemma surjective_one_closed:
"surjective 1"
- by (simp add: antisym mult_sub_right_one)
+ by (simp add: order.antisym mult_sub_right_one)
lemma surjective_top_closed:
"surjective top"
by simp
lemma surjective_sup_closed:
"surjective x \<Longrightarrow> surjective (x \<squnion> y)"
by (metis le_iff_sup mult_left_sub_dist_sup_left sup_left_top)
lemma reflexive_top_closed:
"reflexive top"
by simp
lemma transitive_top_closed:
"transitive top"
by simp
lemma dense_top_closed:
"dense_rel top"
by simp
lemma idempotent_top_closed:
"idempotent top"
by simp
lemma preorder_top_closed:
"preorder top"
by simp
end
text \<open>
Some closure properties require at least half of associativity.
\<close>
class bounded_pre_left_semiring = pre_left_semiring + bounded_non_associative_left_semiring
begin
lemma vector_mult_closed:
"vector y \<Longrightarrow> vector (x * y)"
- by (metis antisym mult_semi_associative top_right_mult_increasing)
+ by (metis order.antisym mult_semi_associative top_right_mult_increasing)
lemma surjective_mult_closed:
"surjective x \<Longrightarrow> surjective y \<Longrightarrow> surjective (x * y)"
- by (metis antisym mult_semi_associative top_greatest)
+ by (metis order.antisym mult_semi_associative top_greatest)
end
text \<open>
We next consider residuals with the greatest element.
\<close>
class bounded_residuated_pre_left_semiring = residuated_pre_left_semiring + bounded_pre_left_semiring
begin
lemma lres_top_decreasing:
"x / top \<le> x"
using lres_inverse order.trans top_right_mult_increasing by blast
lemma top_lres_absorb [simp]:
"top / x = top"
- using antisym lres_galois top_greatest by blast
+ using order.antisym lres_galois top_greatest by blast
lemma covector_lres_closed:
"covector x \<Longrightarrow> covector (x / y)"
- by (metis antisym mult_lres_sub_assoc top_left_mult_increasing)
+ by (metis order.antisym mult_lres_sub_assoc top_left_mult_increasing)
end
text \<open>
Some closure properties require full associativity.
\<close>
class bounded_idempotent_left_semiring = bounded_pre_left_semiring + idempotent_left_semiring
begin
lemma covector_mult_closed:
"covector x \<Longrightarrow> covector (x * y)"
by (metis mult_assoc)
lemma total_mult_closed:
"total x \<Longrightarrow> total y \<Longrightarrow> total (x * y)"
by (simp add: mult_assoc)
end
text \<open>
Some closure properties require distributivity from the left.
\<close>
class bounded_idempotent_left_zero_semiring = bounded_idempotent_left_semiring + idempotent_left_zero_semiring
begin
lemma covector_sup_closed:
"covector x \<Longrightarrow> covector y \<Longrightarrow> covector (x \<squnion> y)"
by (simp add: mult_left_dist_sup)
end
text \<open>
Our final structure is an idempotent semiring with a greatest element.
\<close>
class bounded_idempotent_semiring = bounded_idempotent_left_zero_semiring + idempotent_semiring
begin
lemma covector_bot_closed:
"covector bot"
by simp
end
end
diff --git a/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy b/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy
--- a/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy
+++ b/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy
@@ -1,3608 +1,3608 @@
(* Title: Subset Boolean Algebras
Authors: Walter Guttmann, Bernhard Möller
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
theory Subset_Boolean_Algebras
imports Stone_Algebras.P_Algebras
begin
section \<open>Boolean Algebras\<close>
text \<open>
We show that Isabelle/HOL's \<open>boolean_algebra\<close> class is equivalent to Huntington's axioms \cite{Huntington1933}.
See \cite{WamplerDoty2016} for related results.
\<close>
subsection \<open>Huntington's Axioms\<close>
text \<open>Definition 1\<close>
class huntington = sup + uminus +
assumes associative: "x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
assumes commutative: "x \<squnion> y = y \<squnion> x"
assumes huntington: "x = -(-x \<squnion> y) \<squnion> -(-x \<squnion> -y)"
begin
lemma top_unique:
"x \<squnion> -x = y \<squnion> -y"
proof -
have "x \<squnion> -x = y \<squnion> -(--y \<squnion> -x) \<squnion> -(--y \<squnion> --x)"
by (smt associative commutative huntington)
thus ?thesis
by (metis associative huntington)
qed
end
subsection \<open>Equivalence to \<open>boolean_algebra\<close> Class\<close>
text \<open>Definition 2\<close>
class extended = sup + inf + minus + uminus + bot + top + ord +
assumes top_def: "top = (THE x . \<forall>y . x = y \<squnion> -y)" (* define without imposing uniqueness *)
assumes bot_def: "bot = -(THE x . \<forall>y . x = y \<squnion> -y)"
assumes inf_def: "x \<sqinter> y = -(-x \<squnion> -y)"
assumes minus_def: "x - y = -(-x \<squnion> y)"
assumes less_eq_def: "x \<le> y \<longleftrightarrow> x \<squnion> y = y"
assumes less_def: "x < y \<longleftrightarrow> x \<squnion> y = y \<and> \<not> (y \<squnion> x = x)"
class huntington_extended = huntington + extended
begin
lemma top_char:
"top = x \<squnion> -x"
using top_def top_unique by auto
lemma bot_char:
"bot = -top"
by (simp add: bot_def top_def)
subclass boolean_algebra
proof
show 1: "\<And>x y. (x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (simp add: less_def less_eq_def)
show 2: "\<And>x. x \<le> x"
proof -
fix x
have "x \<squnion> top = top \<squnion> --x"
by (metis (full_types) associative top_char)
thus "x \<le> x"
by (metis (no_types) associative huntington less_eq_def top_char)
qed
show 3: "\<And>x y z. x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
by (metis associative less_eq_def)
show 4: "\<And>x y. x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
by (simp add: commutative less_eq_def)
show 5: "\<And>x y. x \<sqinter> y \<le> x"
using 2 by (metis associative huntington inf_def less_eq_def)
show 6: "\<And>x y. x \<sqinter> y \<le> y"
using 5 commutative inf_def by fastforce
show 8: "\<And>x y. x \<le> x \<squnion> y"
using 2 associative less_eq_def by auto
show 9: "\<And>y x. y \<le> x \<squnion> y"
using 8 commutative by fastforce
show 10: "\<And>y x z. y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> y \<squnion> z \<le> x"
by (metis associative less_eq_def)
show 11: "\<And>x. bot \<le> x"
using 8 by (metis bot_char huntington top_char)
show 12: "\<And>x. x \<le> top"
using 6 11 by (metis huntington bot_def inf_def less_eq_def top_def)
show 13: "\<And>x y z. x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
proof -
have 2: "\<And>x y z . x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
by (simp add: associative)
have 3: "\<And>x y z . (x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
using 2 by metis
have 4: "\<And>x y . x \<squnion> y = y \<squnion> x"
by (simp add: commutative)
have 5: "\<And>x y . x = - (- x \<squnion> y) \<squnion> - (- x \<squnion> - y)"
by (simp add: huntington)
have 6: "\<And>x y . - (- x \<squnion> y) \<squnion> - (- x \<squnion> - y) = x"
using 5 by metis
have 7: "\<And>x y . x \<sqinter> y = - (- x \<squnion> - y)"
by (simp add: inf_def)
have 10: "\<And>x y z . x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
using 3 4 by metis
have 11: "\<And>x y z . - (- x \<squnion> y) \<squnion> (- (- x \<squnion> - y) \<squnion> z) = x \<squnion> z"
using 3 6 by metis
have 12: "\<And>x y . - (x \<squnion> - y) \<squnion> - (- y \<squnion> - x) = y"
using 4 6 by metis
have 13: "\<And>x y . - (- x \<squnion> y) \<squnion> - (- y \<squnion> - x) = x"
using 4 6 by metis
have 14: "\<And>x y . - x \<squnion> - (- (- x \<squnion> y) \<squnion> - - (- x \<squnion> - y)) = - x \<squnion> y"
using 6 by metis
have 18: "\<And>x y z . - (x \<squnion> - y) \<squnion> (- (- y \<squnion> - x) \<squnion> z) = y \<squnion> z"
using 3 12 by metis
have 20: "\<And>x y . - (- x \<squnion> - y) \<squnion> - (y \<squnion> - x) = x"
using 4 12 by metis
have 21: "\<And>x y . - (x \<squnion> - y) \<squnion> - (- x \<squnion> - y) = y"
using 4 12 by metis
have 22: "\<And>x y . - x \<squnion> - (- (y \<squnion> - x) \<squnion> - - (- x \<squnion> - y)) = y \<squnion> - x"
using 6 12 by metis
have 23: "\<And>x y . - x \<squnion> - (- x \<squnion> (- y \<squnion> - (y \<squnion> - x))) = y \<squnion> - x"
using 3 4 6 12 by metis
have 24: "\<And>x y . - x \<squnion> - (- (- x \<squnion> - y) \<squnion> - - (- x \<squnion> y)) = - x \<squnion> - y"
using 6 12 by metis
have 28: "\<And>x y . - (- x \<squnion> - y) \<squnion> - (- y \<squnion> x) = y"
using 4 13 by metis
have 30: "\<And>x y . - x \<squnion> - (- y \<squnion> (- x \<squnion> - (- x \<squnion> y))) = - x \<squnion> y"
using 3 4 6 13 by metis
have 32: "\<And>x y z . - (- x \<squnion> y) \<squnion> (z \<squnion> - (- y \<squnion> - x)) = z \<squnion> x"
using 10 13 by metis
have 37: "\<And>x y z . - (- x \<squnion> - y) \<squnion> (- (y \<squnion> - x) \<squnion> z) = x \<squnion> z"
using 3 20 by metis
have 39: "\<And>x y z . - (- x \<squnion> - y) \<squnion> (z \<squnion> - (y \<squnion> - x)) = z \<squnion> x"
using 10 20 by metis
have 40: "\<And>x y z . - (x \<squnion> - y) \<squnion> (- (- x \<squnion> - y) \<squnion> z) = y \<squnion> z"
using 3 21 by metis
have 43: "\<And>x y . - x \<squnion> - (- y \<squnion> (- x \<squnion> - (y \<squnion> - x))) = y \<squnion> - x"
using 3 4 6 21 by metis
have 47: "\<And>x y z . - (x \<squnion> y) \<squnion> - (- (- x \<squnion> z) \<squnion> - (- (- x \<squnion> - z) \<squnion> y)) = - x \<squnion> z"
using 6 11 by metis
have 55: "\<And>x y . x \<squnion> - (- y \<squnion> - - x) = y \<squnion> - (- x \<squnion> y)"
using 4 11 12 by metis
have 58: "\<And>x y . x \<squnion> - (- - y \<squnion> - x) = x \<squnion> - (- x \<squnion> y)"
using 4 11 13 by metis
have 63: "\<And>x y . x \<squnion> - (- - x \<squnion> - y) = y \<squnion> - (- x \<squnion> y)"
using 4 11 21 by metis
have 71: "\<And>x y . x \<squnion> - (- y \<squnion> x) = y \<squnion> - (- x \<squnion> y)"
using 4 11 28 by metis
have 75: "\<And>x y . x \<squnion> - (- y \<squnion> x) = y \<squnion> - (y \<squnion> - x)"
using 4 71 by metis
have 78: "\<And>x y . - x \<squnion> (y \<squnion> - (- x \<squnion> (y \<squnion> - - (- x \<squnion> - y)))) = - x \<squnion> - (- x \<squnion> - y)"
using 3 4 6 71 by metis
have 86: "\<And>x y . - (- x \<squnion> - (- y \<squnion> x)) \<squnion> - (y \<squnion> - (- x \<squnion> y)) = - y \<squnion> x"
using 4 20 71 by metis
have 172: "\<And>x y . - x \<squnion> - (- x \<squnion> - y) = y \<squnion> - (- - x \<squnion> y)"
using 14 75 by metis
have 201: "\<And>x y . x \<squnion> - (- y \<squnion> - - x) = y \<squnion> - (y \<squnion> - x)"
using 4 55 by metis
have 236: "\<And>x y . x \<squnion> - (- - y \<squnion> - x) = x \<squnion> - (y \<squnion> - x)"
using 4 58 by metis
have 266: "\<And>x y . - x \<squnion> - (- (- x \<squnion> - (y \<squnion> - - x)) \<squnion> - - (- x \<squnion> - - (- - x \<squnion> y))) = - x \<squnion> - (- - x \<squnion> y)"
using 14 58 236 by metis
have 678: "\<And>x y z . - (- x \<squnion> - (- y \<squnion> x)) \<squnion> (- (y \<squnion> - (- x \<squnion> y)) \<squnion> z) = - y \<squnion> (x \<squnion> z)"
using 3 4 37 71 by smt
have 745: "\<And>x y z . - (- x \<squnion> - (- y \<squnion> x)) \<squnion> (z \<squnion> - (y \<squnion> - (- x \<squnion> y))) = z \<squnion> (- y \<squnion> x)"
using 4 39 71 by metis
have 800: "\<And>x y . - - x \<squnion> (- y \<squnion> (- (y \<squnion> - - x) \<squnion> - (- x \<squnion> (- - x \<squnion> (- y \<squnion> - (y \<squnion> - - x)))))) = x \<squnion> - (y \<squnion> - - x)"
using 3 23 63 by metis
have 944: "\<And>x y . x \<squnion> - (x \<squnion> - - (- (- x \<squnion> - y) \<squnion> - - (- x \<squnion> y))) = - (- x \<squnion> - y) \<squnion> - (- (- x \<squnion> - y) \<squnion> - - (- x \<squnion> y))"
using 4 24 71 by metis
have 948: "\<And>x y . - x \<squnion> - (- (y \<squnion> - (y \<squnion> - - x)) \<squnion> - - (- x \<squnion> (- y \<squnion> - x))) = - x \<squnion> - (- y \<squnion> - x)"
using 24 75 by metis
have 950: "\<And>x y . - x \<squnion> - (- (y \<squnion> - (- - x \<squnion> y)) \<squnion> - - (- x \<squnion> (- x \<squnion> - y))) = - x \<squnion> - (- x \<squnion> - y)"
using 24 75 by metis
have 961: "\<And>x y . - x \<squnion> - (- (y \<squnion> - (- - x \<squnion> y)) \<squnion> - - (- x \<squnion> (- - - x \<squnion> - y))) = y \<squnion> - (- - x \<squnion> y)"
using 24 63 by metis
have 966: "\<And>x y . - x \<squnion> - (- (y \<squnion> - (y \<squnion> - - x)) \<squnion> - - (- x \<squnion> (- y \<squnion> - - - x))) = y \<squnion> - (y \<squnion> - - x)"
using 24 201 by metis
have 969: "\<And>x y . - x \<squnion> - (- (- x \<squnion> - (y \<squnion> - - x)) \<squnion> - - (- x \<squnion> (- - y \<squnion> - - x))) = - x \<squnion> - (y \<squnion> - - x)"
using 24 236 by metis
have 1096: "\<And>x y z . - x \<squnion> (- (- x \<squnion> - y) \<squnion> z) = y \<squnion> (- (- - x \<squnion> y) \<squnion> z)"
using 3 172 by metis
have 1098: "\<And>x y z . - x \<squnion> (y \<squnion> - (- x \<squnion> - z)) = y \<squnion> (z \<squnion> - (- - x \<squnion> z))"
using 10 172 by metis
have 1105: "\<And>x y . x \<squnion> - x = y \<squnion> - y"
using 4 10 12 32 172 by metis
have 1109: "\<And>x y z . x \<squnion> (- x \<squnion> y) = z \<squnion> (- z \<squnion> y)"
using 3 1105 by metis
have 1110: "\<And>x y z . x \<squnion> - x = y \<squnion> (z \<squnion> - (y \<squnion> z))"
using 3 1105 by metis
have 1114: "\<And>x y . - (- x \<squnion> - - x) = - (y \<squnion> - y)"
using 7 1105 by metis
have 1115: "\<And>x y z . x \<squnion> (y \<squnion> - y) = z \<squnion> (x \<squnion> - z)"
using 10 1105 by metis
have 1117: "\<And>x y . - (x \<squnion> - - x) \<squnion> - (y \<squnion> - y) = - x"
using 4 13 1105 by metis
have 1121: "\<And>x y . - (x \<squnion> - x) \<squnion> - (y \<squnion> - - y) = - y"
using 4 28 1105 by metis
have 1122: "\<And>x . - - x = x"
using 4 28 1105 1117 by metis
have 1134: "\<And>x y z . - (x \<squnion> - y) \<squnion> (z \<squnion> - z) = y \<squnion> (- y \<squnion> - x)"
using 18 1105 1122 by metis
have 1140: "\<And>x . - x \<squnion> - (x \<squnion> (x \<squnion> - x)) = - x \<squnion> - x"
using 4 22 1105 1122 1134 by metis
have 1143: "\<And>x y . x \<squnion> (- x \<squnion> y) = y \<squnion> (x \<squnion> - y)"
using 37 1105 1122 1134 by metis
have 1155: "\<And>x y . - (x \<squnion> - x) \<squnion> - (y \<squnion> y) = - y"
using 1121 1122 by metis
have 1156: "\<And>x y . - (x \<squnion> x) \<squnion> - (y \<squnion> - y) = - x"
using 1117 1122 by metis
have 1157: "\<And>x y . - (x \<squnion> - x) = - (y \<squnion> - y)"
using 4 1114 1122 by metis
have 1167: "\<And>x y z . - x \<squnion> (y \<squnion> - (- x \<squnion> - z)) = y \<squnion> (z \<squnion> - (x \<squnion> z))"
using 1098 1122 by metis
have 1169: "\<And>x y z . - x \<squnion> (- (- x \<squnion> - y) \<squnion> z) = y \<squnion> (- (x \<squnion> y) \<squnion> z)"
using 1096 1122 by metis
have 1227: "\<And>x y . - x \<squnion> - (- x \<squnion> (y \<squnion> (x \<squnion> - (- x \<squnion> - (y \<squnion> x))))) = - x \<squnion> - (y \<squnion> x)"
using 3 4 969 1122 by smt
have 1230: "\<And>x y . - x \<squnion> - (- x \<squnion> (- y \<squnion> (- x \<squnion> - (y \<squnion> - (y \<squnion> x))))) = y \<squnion> - (y \<squnion> x)"
using 3 4 966 1122 by smt
have 1234: "\<And>x y . - x \<squnion> - (- x \<squnion> (- x \<squnion> (- y \<squnion> - (y \<squnion> - (x \<squnion> y))))) = y \<squnion> - (x \<squnion> y)"
using 3 4 961 1122 by metis
have 1239: "\<And>x y . - x \<squnion> - (- x \<squnion> - y) = y \<squnion> - (x \<squnion> y)"
using 3 4 950 1122 1234 by metis
have 1240: "\<And>x y . - x \<squnion> - (- y \<squnion> - x) = y \<squnion> - (y \<squnion> x)"
using 3 4 948 1122 1230 by metis
have 1244: "\<And>x y . x \<squnion> - (x \<squnion> (y \<squnion> (y \<squnion> - (x \<squnion> y)))) = - (- x \<squnion> - y) \<squnion> - (y \<squnion> (y \<squnion> - (x \<squnion> y)))"
using 3 4 944 1122 1167 by metis
have 1275: "\<And>x y . x \<squnion> (- y \<squnion> (- (y \<squnion> x) \<squnion> - (x \<squnion> (- x \<squnion> (- y \<squnion> - (y \<squnion> x)))))) = x \<squnion> - (y \<squnion> x)"
using 10 800 1122 by metis
have 1346: "\<And>x y . - x \<squnion> - (x \<squnion> (y \<squnion> (y \<squnion> (x \<squnion> - (x \<squnion> (y \<squnion> x)))))) = - x \<squnion> - (x \<squnion> y)"
using 3 4 10 266 1122 1167 by smt
have 1377: "\<And>x y . - x \<squnion> (y \<squnion> - (- x \<squnion> (y \<squnion> (- x \<squnion> - y)))) = y \<squnion> - (x \<squnion> y)"
using 78 1122 1239 by metis
have 1394: "\<And>x y . - (- x \<squnion> - y) \<squnion> - (y \<squnion> (y \<squnion> (- x \<squnion> - (x \<squnion> y)))) = x"
using 3 4 10 20 30 1122 1239 by smt
have 1427: "\<And>x y . - (- x \<squnion> - y) \<squnion> - (y \<squnion> - (x \<squnion> (x \<squnion> - (x \<squnion> y)))) = x \<squnion> (x \<squnion> - (x \<squnion> y))"
using 3 4 30 40 1240 by smt
have 1436: "\<And>x . - x \<squnion> - (x \<squnion> (x \<squnion> (- x \<squnion> - x))) = - x \<squnion> (- x \<squnion> - (x \<squnion> - x))"
using 3 4 30 1140 1239 by smt
have 1437: "\<And>x y . - (x \<squnion> y) \<squnion> - (x \<squnion> - y) = - x"
using 6 1122 by metis
have 1438: "\<And>x y . - (x \<squnion> y) \<squnion> - (y \<squnion> - x) = - y"
using 12 1122 by metis
have 1439: "\<And>x y . - (x \<squnion> y) \<squnion> - (- y \<squnion> x) = - x"
using 13 1122 by metis
have 1440: "\<And>x y . - (x \<squnion> - y) \<squnion> - (y \<squnion> x) = - x"
using 20 1122 by metis
have 1441: "\<And>x y . - (x \<squnion> y) \<squnion> - (- x \<squnion> y) = - y"
using 21 1122 by metis
have 1568: "\<And>x y . x \<squnion> (- y \<squnion> - x) = y \<squnion> (- y \<squnion> x)"
using 10 1122 1143 by metis
have 1598: "\<And>x . - x \<squnion> - (x \<squnion> (x \<squnion> (x \<squnion> - x))) = - x \<squnion> (- x \<squnion> - (x \<squnion> - x))"
using 4 1436 1568 by metis
have 1599: "\<And>x y . - x \<squnion> (y \<squnion> - (x \<squnion> (- x \<squnion> (- x \<squnion> y)))) = y \<squnion> - (x \<squnion> y)"
using 10 1377 1568 by smt
have 1617: "\<And>x . x \<squnion> (- x \<squnion> (- x \<squnion> - (x \<squnion> - x))) = x \<squnion> - x"
using 3 4 10 71 1122 1155 1568 1598 by metis
have 1632: "\<And>x y z . - (x \<squnion> - x) \<squnion> - (- y \<squnion> (- (z \<squnion> - z) \<squnion> - (y \<squnion> - (x \<squnion> - x)))) = y \<squnion> - (x \<squnion> - x)"
using 43 1157 by metis
have 1633: "\<And>x y z . - (x \<squnion> - x) \<squnion> - (- y \<squnion> (- (x \<squnion> - x) \<squnion> - (y \<squnion> - (z \<squnion> - z)))) = y \<squnion> - (x \<squnion> - x)"
using 43 1157 by metis
have 1636: "\<And>x y . x \<squnion> - (y \<squnion> (- y \<squnion> - (x \<squnion> x))) = x \<squnion> x"
using 43 1109 1122 by metis
have 1645: "\<And>x y . x \<squnion> - x = y \<squnion> (y \<squnion> - y)"
using 3 1110 1156 by metis
have 1648: "\<And>x y z . - (x \<squnion> (y \<squnion> (- y \<squnion> - x))) \<squnion> - (z \<squnion> - z) = - (y \<squnion> - y)"
using 3 1115 1156 by metis
have 1657: "\<And>x y z . x \<squnion> - x = y \<squnion> (z \<squnion> - z)"
using 1105 1645 by metis
have 1664: "\<And>x y z . x \<squnion> - x = y \<squnion> (z \<squnion> - y)"
using 1115 1645 by metis
have 1672: "\<And>x y z . x \<squnion> - x = y \<squnion> (- y \<squnion> z)"
using 3 4 1657 by metis
have 1697: "\<And>x y z . - x \<squnion> (y \<squnion> x) = z \<squnion> - z"
using 1122 1664 by metis
have 1733: "\<And>x y z . - (x \<squnion> y) \<squnion> - (- (z \<squnion> - z) \<squnion> - (- (- x \<squnion> - x) \<squnion> y)) = x \<squnion> - x"
using 4 47 1105 1122 by metis
have 1791: "\<And>x y z . x \<squnion> - (y \<squnion> (- y \<squnion> z)) = x \<squnion> - (x \<squnion> - x)"
using 4 71 1122 1672 by metis
have 1818: "\<And>x y z . x \<squnion> - (- y \<squnion> (z \<squnion> y)) = x \<squnion> - (x \<squnion> - x)"
using 4 71 1122 1697 by metis
have 1861: "\<And>x y z . - (x \<squnion> - x) \<squnion> - (y \<squnion> - (z \<squnion> - z)) = - y"
using 1437 1657 by metis
have 1867: "\<And>x y z . - (x \<squnion> - x) \<squnion> - (- y \<squnion> - (z \<squnion> y)) = y"
using 1122 1437 1697 by metis
have 1868: "\<And>x y . x \<squnion> - (y \<squnion> - y) = x"
using 1122 1155 1633 1861 by metis
have 1869: "\<And>x y z . - (x \<squnion> - x) \<squnion> - (- y \<squnion> (- (z \<squnion> - z) \<squnion> - y)) = y"
using 1632 1868 by metis
have 1870: "\<And>x y . - (x \<squnion> - x) \<squnion> - y = - y"
using 1861 1868 by metis
have 1872: "\<And>x y z . x \<squnion> - (- y \<squnion> (z \<squnion> y)) = x"
using 1818 1868 by metis
have 1875: "\<And>x y z . x \<squnion> - (y \<squnion> (- y \<squnion> z)) = x"
using 1791 1868 by metis
have 1883: "\<And>x y . - (x \<squnion> (y \<squnion> (- y \<squnion> - x))) = - (y \<squnion> - y)"
using 1648 1868 by metis
have 1885: "\<And>x . x \<squnion> (x \<squnion> - x) = x \<squnion> - x"
using 4 1568 1617 1868 by metis
have 1886: "\<And>x . - x \<squnion> - x = - x"
using 1598 1868 1885 by metis
have 1890: "\<And>x . - (x \<squnion> x) = - x"
using 1156 1868 by metis
have 1892: "\<And>x y . - (x \<squnion> - x) \<squnion> y = y"
using 1122 1869 1870 1886 by metis
have 1893: "\<And>x y . - (- x \<squnion> - (y \<squnion> x)) = x"
using 1867 1892 by metis
have 1902: "\<And>x y . x \<squnion> (y \<squnion> - (x \<squnion> y)) = x \<squnion> - x"
using 3 4 1122 1733 1886 1892 by metis
have 1908: "\<And>x . x \<squnion> x = x"
using 1636 1875 1890 by metis
have 1910: "\<And>x y . x \<squnion> - (y \<squnion> x) = - y \<squnion> x"
using 1599 1875 by metis
have 1921: "\<And>x y . x \<squnion> (- y \<squnion> - (y \<squnion> x)) = - y \<squnion> x"
using 1275 1875 1910 by metis
have 1951: "\<And>x y . - x \<squnion> - (y \<squnion> x) = - x"
using 1227 1872 1893 1908 by metis
have 1954: "\<And>x y z . x \<squnion> (y \<squnion> - (x \<squnion> z)) = y \<squnion> (- z \<squnion> x)"
using 745 1122 1910 1951 by metis
have 1956: "\<And>x y z . x \<squnion> (- (x \<squnion> y) \<squnion> z) = - y \<squnion> (x \<squnion> z)"
using 678 1122 1910 1951 by metis
have 1959: "\<And>x y . x \<squnion> - (x \<squnion> y) = - y \<squnion> x"
using 86 1122 1910 1951 by metis
have 1972: "\<And>x y . x \<squnion> (- x \<squnion> y) = x \<squnion> - x"
using 1902 1910 by metis
have 2000: "\<And>x y . - (- x \<squnion> - y) \<squnion> - (y \<squnion> (- x \<squnion> y)) = x \<squnion> - (y \<squnion> (- x \<squnion> y))"
using 4 1244 1910 1959 by metis
have 2054: "\<And>x y . x \<squnion> - (y \<squnion> (- x \<squnion> y)) = x"
using 1394 1921 2000 by metis
have 2057: "\<And>x y . - (x \<squnion> (y \<squnion> - y)) = - (y \<squnion> - y)"
using 1883 1972 by metis
have 2061: "\<And>x y . x \<squnion> (- y \<squnion> x) = x \<squnion> - y"
using 4 1122 1427 1910 1959 2054 by metis
have 2090: "\<And>x y z . x \<squnion> (- (y \<squnion> x) \<squnion> z) = x \<squnion> (- y \<squnion> z)"
using 1122 1169 1956 by metis
have 2100: "\<And>x y . - x \<squnion> - (x \<squnion> y) = - x"
using 4 1346 1868 1885 1910 1959 1972 2057 by metis
have 2144: "\<And>x y . x \<squnion> - (y \<squnion> - x) = x"
using 1122 1440 2000 2061 by metis
have 2199: "\<And>x y . x \<squnion> (x \<squnion> y) = x \<squnion> y"
using 3 1908 by metis
have 2208: "\<And>x y z . x \<squnion> (- (y \<squnion> - x) \<squnion> z) = x \<squnion> z"
using 3 2144 by metis
have 2349: "\<And>x y z . - (x \<squnion> y) \<squnion> - (x \<squnion> (y \<squnion> z)) = - (x \<squnion> y)"
using 3 2100 by metis
have 2432: "\<And>x y z . - (x \<squnion> (y \<squnion> z)) \<squnion> - (y \<squnion> (z \<squnion> - x)) = - (y \<squnion> z)"
using 3 1438 by metis
have 2530: "\<And>x y z . - (- (x \<squnion> y) \<squnion> z) = - (y \<squnion> (- x \<squnion> z)) \<squnion> - (- y \<squnion> z)"
using 4 1122 1439 2090 2208 by smt
have 3364: "\<And>x y z . - (- x \<squnion> y) \<squnion> (z \<squnion> - (x \<squnion> y)) = z \<squnion> - y"
using 3 4 1122 1441 1910 1954 2199 by metis
have 5763: "\<And>x y z . - (x \<squnion> y) \<squnion> - (- x \<squnion> (y \<squnion> z)) = - (x \<squnion> y) \<squnion> - (y \<squnion> z)"
using 4 2349 3364 by metis
have 6113: "\<And>x y z . - (x \<squnion> (y \<squnion> z)) \<squnion> - (z \<squnion> - x) = - (y \<squnion> z) \<squnion> - (z \<squnion> - x)"
using 4 2432 3364 5763 by metis
show "\<And>x y z. x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
proof -
fix x y z
have "- (y \<sqinter> z \<squnion> x) = - (- (- y \<squnion> z) \<squnion> - (- y \<squnion> - z) \<squnion> x) \<squnion> - (x \<squnion> - - z)"
using 1437 2530 6113 by (smt commutative inf_def)
thus "x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
using 12 1122 by (metis commutative inf_def)
qed
qed
show 14: "\<And>x. x \<sqinter> - x = bot"
proof -
fix x
have "(bot \<squnion> x) \<sqinter> (bot \<squnion> -x) = bot"
using huntington bot_def inf_def by auto
thus "x \<sqinter> -x = bot"
using 11 less_eq_def by force
qed
show 15: "\<And>x. x \<squnion> - x = top"
using 5 14 by (metis (no_types, lifting) huntington bot_def less_eq_def top_def)
show 16: "\<And>x y. x - y = x \<sqinter> - y"
using 15 by (metis commutative huntington inf_def minus_def)
show 7: "\<And>x y z. x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z"
by (simp add: 13 less_eq_def)
qed
end
context boolean_algebra
begin
sublocale ba_he: huntington_extended
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: sup_assoc)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: sup_commute)
show "\<And>x y. x = - (- x \<squnion> y) \<squnion> - (- x \<squnion> - y)"
by simp
show "top = (THE x. \<forall>y. x = y \<squnion> - y)"
by auto
show "bot = - (THE x. \<forall>y. x = y \<squnion> - y)"
by auto
show "\<And>x y. x \<sqinter> y = - (- x \<squnion> - y)"
by simp
show "\<And>x y. x - y = - (- x \<squnion> y)"
by (simp add: diff_eq)
show "\<And>x y. (x \<le> y) = (x \<squnion> y = y)"
by (simp add: le_iff_sup)
show "\<And>x y. (x < y) = (x \<squnion> y = y \<and> y \<squnion> x \<noteq> x)"
using sup.strict_order_iff sup_commute by auto
qed
end
subsection \<open>Stone Algebras\<close>
text \<open>
We relate Stone algebras to Boolean algebras.
\<close>
class stone_algebra_extended = stone_algebra + minus +
assumes stone_minus_def[simp]: "x - y = x \<sqinter> -y"
class regular_stone_algebra = stone_algebra_extended +
assumes double_complement[simp]: "--x = x"
begin
subclass boolean_algebra
proof
show "\<And>x. x \<sqinter> - x = bot"
by simp
show "\<And>x. x \<squnion> - x = top"
using regular_dense_top by fastforce
show "\<And>x y. x - y = x \<sqinter> - y"
by simp
qed
end
context boolean_algebra
begin
sublocale ba_rsa: regular_stone_algebra
proof
show "\<And>x y. x - y = x \<sqinter> - y"
by (simp add: diff_eq)
show "\<And>x. - - x = x"
by simp
qed
end
section \<open>Alternative Axiomatisations of Boolean Algebras\<close>
text \<open>
We consider four axiomatisations of Boolean algebras based only on join and complement.
The first three are from the literature and the fourth, a version using equational axioms, is new.
The motivation for Byrne's and the new axiomatisation is that the axioms are easier to understand than Huntington's third axiom.
We also include Meredith's axiomatisation.
\<close>
subsection \<open>Lee Byrne's Formulation A\<close>
text \<open>
The following axiomatisation is from \cite[Formulation A]{Byrne1946}; see also \cite{Frink1941}.
\<close>
text \<open>Theorem 3\<close>
class boolean_algebra_1 = sup + uminus +
assumes ba1_associative: "x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
assumes ba1_commutative: "x \<squnion> y = y \<squnion> x"
assumes ba1_complement: "x \<squnion> -y = z \<squnion> -z \<longleftrightarrow> x \<squnion> y = x"
begin
subclass huntington
proof
show 1: "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: ba1_associative)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: ba1_commutative)
show "\<And>x y. x = - (- x \<squnion> y) \<squnion> - (- x \<squnion> - y)"
proof -
have 2: "\<forall>x y. y \<squnion> (y \<squnion> x) = y \<squnion> x"
using 1 by (metis ba1_complement)
hence "\<forall>x. --x = x"
by (smt ba1_associative ba1_commutative ba1_complement)
hence "\<forall>x y. y \<squnion> -(y \<squnion> -x) = y \<squnion> x"
by (smt ba1_associative ba1_commutative ba1_complement)
thus "\<And>x y. x = -(-x \<squnion> y) \<squnion> -(-x \<squnion> - y)"
using 2 by (smt ba1_commutative ba1_complement)
qed
qed
end
context huntington
begin
sublocale h_ba1: boolean_algebra_1
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: associative)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: commutative)
show "\<And>x y z. (x \<squnion> - y = z \<squnion> - z) = (x \<squnion> y = x)"
proof
fix x y z
have 1: "\<And>x y z. -(-x \<squnion> y) \<squnion> (-(-x \<squnion> -y) \<squnion> z) = x \<squnion> z"
using associative huntington by force
have 2: "\<And>x y. -(x \<squnion> -y) \<squnion> -(-y \<squnion> -x) = y"
by (metis commutative huntington)
show "x \<squnion> - y = z \<squnion> - z \<Longrightarrow> x \<squnion> y = x"
by (metis 1 2 associative commutative top_unique)
show "x \<squnion> y = x \<Longrightarrow> x \<squnion> - y = z \<squnion> - z"
by (metis associative huntington commutative top_unique)
qed
qed
end
subsection \<open>Lee Byrne's Formulation B\<close>
text \<open>
The following axiomatisation is from \cite[Formulation B]{Byrne1946}.
\<close>
text \<open>Theorem 4\<close>
class boolean_algebra_2 = sup + uminus +
assumes ba2_associative_commutative: "(x \<squnion> y) \<squnion> z = (y \<squnion> z) \<squnion> x"
assumes ba2_complement: "x \<squnion> -y = z \<squnion> -z \<longleftrightarrow> x \<squnion> y = x"
begin
subclass boolean_algebra_1
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (smt ba2_associative_commutative ba2_complement)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (metis ba2_associative_commutative ba2_complement)
show "\<And>x y z. (x \<squnion> - y = z \<squnion> - z) = (x \<squnion> y = x)"
by (simp add: ba2_complement)
qed
end
context boolean_algebra_1
begin
sublocale ba1_ba2: boolean_algebra_2
proof
show "\<And>x y z. x \<squnion> y \<squnion> z = y \<squnion> z \<squnion> x"
using ba1_associative commutative by force
show "\<And>x y z. (x \<squnion> - y = z \<squnion> - z) = (x \<squnion> y = x)"
by (simp add: ba1_complement)
qed
end
subsection \<open>Meredith's Equational Axioms\<close>
text \<open>
The following axiomatisation is from \cite[page 221 (1) \{A,N\}]{MeredithPrior1968}.
\<close>
class boolean_algebra_mp = sup + uminus +
assumes ba_mp_1: "-(-x \<squnion> y) \<squnion> x = x"
assumes ba_mp_2: "-(-x \<squnion> y) \<squnion> (z \<squnion> y) = y \<squnion> (z \<squnion> x)"
begin
subclass huntington
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (metis ba_mp_1 ba_mp_2)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (metis ba_mp_1 ba_mp_2)
show "\<And>x y. x = - (- x \<squnion> y) \<squnion> - (- x \<squnion> - y)"
by (metis ba_mp_1 ba_mp_2)
qed
end
context huntington
begin
sublocale mp_h: boolean_algebra_mp
proof
show 1: "\<And>x y. - (- x \<squnion> y) \<squnion> x = x"
by (metis h_ba1.ba1_associative h_ba1.ba1_complement huntington)
show "\<And>x y z. - (- x \<squnion> y) \<squnion> (z \<squnion> y) = y \<squnion> (z \<squnion> x)"
proof -
fix x y z
have "y = -(-x \<squnion> -y) \<squnion> y"
using 1 h_ba1.ba1_commutative by auto
thus "-(-x \<squnion> y) \<squnion> (z \<squnion> y) = y \<squnion> (z \<squnion> x)"
by (metis h_ba1.ba1_associative h_ba1.ba1_commutative huntington)
qed
qed
end
subsection \<open>An Equational Axiomatisation based on Semilattices\<close>
text \<open>
The following version is an equational axiomatisation based on semilattices.
We add the double complement rule and that \<open>top\<close> is unique.
The final axiom \<open>ba3_export\<close> encodes the logical statement $P \vee Q = P \vee (\neg P \wedge Q)$.
Its dual appears in \cite{BalbesHorn1970}.
\<close>
text \<open>Theorem 5\<close>
class boolean_algebra_3 = sup + uminus +
assumes ba3_associative: "x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
assumes ba3_commutative: "x \<squnion> y = y \<squnion> x"
assumes ba3_idempotent[simp]: "x \<squnion> x = x"
assumes ba3_double_complement[simp]: "--x = x"
assumes ba3_top_unique: "x \<squnion> -x = y \<squnion> -y"
assumes ba3_export: "x \<squnion> -(x \<squnion> y) = x \<squnion> -y"
begin
subclass huntington
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: ba3_associative)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: ba3_commutative)
show "\<And>x y. x = - (- x \<squnion> y) \<squnion> - (- x \<squnion> - y)"
by (metis ba3_commutative ba3_double_complement ba3_export ba3_idempotent ba3_top_unique)
qed
end
context huntington
begin
sublocale h_ba3: boolean_algebra_3
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: h_ba1.ba1_associative)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: h_ba1.ba1_commutative)
show 3: "\<And>x. x \<squnion> x = x"
using h_ba1.ba1_complement by blast
show 4: "\<And>x. - - x = x"
by (metis h_ba1.ba1_commutative huntington top_unique)
show "\<And>x y. x \<squnion> - x = y \<squnion> - y"
by (simp add: top_unique)
show "\<And>x y. x \<squnion> - (x \<squnion> y) = x \<squnion> - y"
using 3 4 by (smt h_ba1.ba1_ba2.ba2_associative_commutative h_ba1.ba1_complement)
qed
end
section \<open>Subset Boolean Algebras\<close>
text \<open>
We apply Huntington's axioms to the range of a unary operation, which serves as complement on the range.
This gives a Boolean algebra structure on the range without imposing any further constraints on the set.
The obtained structure is used as a reference in the subsequent development and to inherit the results proved here.
This is taken from \cite{Guttmann2012c,GuttmannStruthWeber2011b} and follows the development of Boolean algebras in \cite{Maddux1996}.
\<close>
text \<open>Definition 6\<close>
class subset_boolean_algebra = sup + uminus +
assumes sub_associative: "-x \<squnion> (-y \<squnion> -z) = (-x \<squnion> -y) \<squnion> -z"
assumes sub_commutative: "-x \<squnion> -y = -y \<squnion> -x"
assumes sub_complement: "-x = -(--x \<squnion> -y) \<squnion> -(--x \<squnion> --y)"
assumes sub_sup_closed: "-x \<squnion> -y = --(-x \<squnion> -y)"
begin
text \<open>uniqueness of \<open>top\<close>, resulting in the lemma \<open>top_def\<close> to replace the assumption \<open>sub_top_def\<close>\<close>
lemma top_unique:
"-x \<squnion> --x = -y \<squnion> --y"
by (metis sub_associative sub_commutative sub_complement)
text \<open>consequences for join and complement\<close>
lemma double_negation[simp]:
"---x = -x"
by (metis sub_complement sub_sup_closed)
lemma complement_1:
"--x = -(-x \<squnion> -y) \<squnion> -(-x \<squnion> --y)"
by (metis double_negation sub_complement)
lemma sup_right_zero_var:
"-x \<squnion> (-y \<squnion> --y) = -z \<squnion> --z"
by (smt complement_1 sub_associative sub_sup_closed top_unique)
lemma sup_right_unit_idempotent:
"-x \<squnion> -x = -x \<squnion> -(-y \<squnion> --y)"
by (metis complement_1 double_negation sub_sup_closed sup_right_zero_var)
lemma sup_idempotent[simp]:
"-x \<squnion> -x = -x"
by (smt complement_1 double_negation sub_associative sup_right_unit_idempotent)
lemma complement_2:
"-x = -(-(-x \<squnion> -y) \<squnion> -(-x \<squnion> --y))"
using complement_1 by auto
lemma sup_eq_cases:
"-x \<squnion> -y = -x \<squnion> -z \<Longrightarrow> --x \<squnion> -y = --x \<squnion> -z \<Longrightarrow> -y = -z"
by (metis complement_2 sub_commutative)
lemma sup_eq_cases_2:
"-y \<squnion> -x = -z \<squnion> -x \<Longrightarrow> -y \<squnion> --x = -z \<squnion> --x \<Longrightarrow> -y = -z"
using sub_commutative sup_eq_cases by auto
end
text \<open>Definition 7\<close>
class subset_extended = sup + inf + minus + uminus + bot + top + ord +
assumes sub_top_def: "top = (THE x . \<forall>y . x = -y \<squnion> --y)" (* define without imposing uniqueness *)
assumes sub_bot_def: "bot = -(THE x . \<forall>y . x = -y \<squnion> --y)"
assumes sub_inf_def: "-x \<sqinter> -y = -(--x \<squnion> --y)"
assumes sub_minus_def: "-x - -y = -(--x \<squnion> -y)"
assumes sub_less_eq_def: "-x \<le> -y \<longleftrightarrow> -x \<squnion> -y = -y"
assumes sub_less_def: "-x < -y \<longleftrightarrow> -x \<squnion> -y = -y \<and> \<not> (-y \<squnion> -x = -x)"
class subset_boolean_algebra_extended = subset_boolean_algebra + subset_extended
begin
lemma top_def:
"top = -x \<squnion> --x"
using sub_top_def top_unique by blast
text \<open>consequences for meet\<close>
lemma inf_closed:
"-x \<sqinter> -y = --(-x \<sqinter> -y)"
by (simp add: sub_inf_def)
lemma inf_associative:
"-x \<sqinter> (-y \<sqinter> -z) = (-x \<sqinter> -y) \<sqinter> -z"
using sub_associative sub_inf_def sub_sup_closed by auto
lemma inf_commutative:
"-x \<sqinter> -y = -y \<sqinter> -x"
by (simp add: sub_commutative sub_inf_def)
lemma inf_idempotent[simp]:
"-x \<sqinter> -x = -x"
by (simp add: sub_inf_def)
lemma inf_absorb[simp]:
"(-x \<squnion> -y) \<sqinter> -x = -x"
by (metis complement_1 sup_idempotent sub_inf_def sub_associative sub_sup_closed)
lemma sup_absorb[simp]:
"-x \<squnion> (-x \<sqinter> -y) = -x"
by (metis sub_associative sub_complement sub_inf_def sup_idempotent)
lemma inf_demorgan:
"-(-x \<sqinter> -y) = --x \<squnion> --y"
using sub_inf_def sub_sup_closed by auto
lemma sub_sup_demorgan:
"-(-x \<squnion> -y) = --x \<sqinter> --y"
by (simp add: sub_inf_def)
lemma sup_cases:
"-x = (-x \<sqinter> -y) \<squnion> (-x \<sqinter> --y)"
by (metis inf_closed inf_demorgan sub_complement)
lemma inf_cases:
"-x = (-x \<squnion> -y) \<sqinter> (-x \<squnion> --y)"
by (metis complement_2 sub_sup_closed sub_sup_demorgan)
lemma inf_complement_intro:
"(-x \<squnion> -y) \<sqinter> --x = -y \<sqinter> --x"
proof -
have "(-x \<squnion> -y) \<sqinter> --x = (-x \<squnion> -y) \<sqinter> (--x \<squnion> -y) \<sqinter> --x"
by (metis inf_absorb inf_associative sub_sup_closed)
also have "... = -y \<sqinter> --x"
by (metis inf_cases sub_commutative)
finally show ?thesis
.
qed
lemma sup_complement_intro:
"-x \<squnion> -y = -x \<squnion> (--x \<sqinter> -y)"
by (metis inf_absorb inf_commutative inf_complement_intro sub_sup_closed sup_cases)
lemma inf_left_dist_sup:
"-x \<sqinter> (-y \<squnion> -z) = (-x \<sqinter> -y) \<squnion> (-x \<sqinter> -z)"
proof -
have "-x \<sqinter> (-y \<squnion> -z) = (-x \<sqinter> (-y \<squnion> -z) \<sqinter> -y) \<squnion> (-x \<sqinter> (-y \<squnion> -z) \<sqinter> --y)"
by (metis sub_inf_def sub_sup_closed sup_cases)
also have "... = (-x \<sqinter> -y) \<squnion> (-x \<sqinter> -z \<sqinter> --y)"
by (metis inf_absorb inf_associative inf_complement_intro sub_sup_closed)
also have "... = (-x \<sqinter> -y) \<squnion> ((-x \<sqinter> -y \<sqinter> -z) \<squnion> (-x \<sqinter> -z \<sqinter> --y))"
using sub_associative sub_inf_def sup_absorb by auto
also have "... = (-x \<sqinter> -y) \<squnion> ((-x \<sqinter> -z \<sqinter> -y) \<squnion> (-x \<sqinter> -z \<sqinter> --y))"
by (metis inf_associative inf_commutative)
also have "... = (-x \<sqinter> -y) \<squnion> (-x \<sqinter> -z)"
by (metis sub_inf_def sup_cases)
finally show ?thesis
.
qed
lemma sup_left_dist_inf:
"-x \<squnion> (-y \<sqinter> -z) = (-x \<squnion> -y) \<sqinter> (-x \<squnion> -z)"
proof -
have "-x \<squnion> (-y \<sqinter> -z) = -(--x \<sqinter> (--y \<squnion> --z))"
by (metis sub_inf_def sub_sup_closed sub_sup_demorgan)
also have "... = (-x \<squnion> -y) \<sqinter> (-x \<squnion> -z)"
by (metis inf_left_dist_sup sub_sup_closed sub_sup_demorgan)
finally show ?thesis
.
qed
lemma sup_right_dist_inf:
"(-y \<sqinter> -z) \<squnion> -x = (-y \<squnion> -x) \<sqinter> (-z \<squnion> -x)"
using sub_commutative sub_inf_def sup_left_dist_inf by auto
lemma inf_right_dist_sup:
"(-y \<squnion> -z) \<sqinter> -x = (-y \<sqinter> -x) \<squnion> (-z \<sqinter> -x)"
by (metis inf_commutative inf_left_dist_sup sub_sup_closed)
lemma case_duality:
"(--x \<sqinter> -y) \<squnion> (-x \<sqinter> -z) = (-x \<squnion> -y) \<sqinter> (--x \<squnion> -z)"
proof -
have 1: "-(--x \<sqinter> --y) \<sqinter> ----x = --x \<sqinter> -y"
using inf_commutative inf_complement_intro sub_sup_closed sub_sup_demorgan by auto
have 2: "-(----x \<squnion> -(--x \<squnion> -z)) = -----x \<sqinter> ---z"
by (metis (no_types) double_negation sup_complement_intro sub_sup_demorgan)
have 3: "-(--x \<sqinter> --y) \<sqinter> -x = -x"
using inf_commutative inf_left_dist_sup sub_sup_closed sub_sup_demorgan by auto
hence "-(--x \<sqinter> --y) = -x \<squnion> -y"
using sub_sup_closed sub_sup_demorgan by auto
thus ?thesis
by (metis double_negation 1 2 3 inf_associative inf_left_dist_sup sup_complement_intro)
qed
lemma case_duality_2:
"(-x \<sqinter> -y) \<squnion> (--x \<sqinter> -z) = (-x \<squnion> -z) \<sqinter> (--x \<squnion> -y)"
using case_duality sub_commutative sub_inf_def by auto
lemma complement_cases:
"((-v \<sqinter> -w) \<squnion> (--v \<sqinter> -x)) \<sqinter> -((-v \<sqinter> -y) \<squnion> (--v \<sqinter> -z)) = (-v \<sqinter> -w \<sqinter> --y) \<squnion> (--v \<sqinter> -x \<sqinter> --z)"
proof -
have 1: "(--v \<squnion> -w) = --(--v \<squnion> -w) \<and> (-v \<squnion> -x) = --(-v \<squnion> -x) \<and> (--v \<squnion> --y) = --(--v \<squnion> --y) \<and> (-v \<squnion> --z) = --(-v \<squnion> --z)"
using sub_inf_def sub_sup_closed by auto
have 2: "(-v \<squnion> (-x \<sqinter> --z)) = --(-v \<squnion> (-x \<sqinter> --z))"
using sub_inf_def sub_sup_closed by auto
have "((-v \<sqinter> -w) \<squnion> (--v \<sqinter> -x)) \<sqinter> -((-v \<sqinter> -y) \<squnion> (--v \<sqinter> -z)) = ((-v \<sqinter> -w) \<squnion> (--v \<sqinter> -x)) \<sqinter> (-(-v \<sqinter> -y) \<sqinter> -(--v \<sqinter> -z))"
using sub_inf_def by auto
also have "... = ((-v \<sqinter> -w) \<squnion> (--v \<sqinter> -x)) \<sqinter> ((--v \<squnion> --y) \<sqinter> (-v \<squnion> --z))"
using inf_demorgan by auto
also have "... = (--v \<squnion> -w) \<sqinter> (-v \<squnion> -x) \<sqinter> ((--v \<squnion> --y) \<sqinter> (-v \<squnion> --z))"
by (metis case_duality double_negation)
also have "... = (--v \<squnion> -w) \<sqinter> ((-v \<squnion> -x) \<sqinter> ((--v \<squnion> --y) \<sqinter> (-v \<squnion> --z)))"
by (metis 1 inf_associative sub_inf_def)
also have "... = (--v \<squnion> -w) \<sqinter> ((-v \<squnion> -x) \<sqinter> (--v \<squnion> --y) \<sqinter> (-v \<squnion> --z))"
by (metis 1 inf_associative)
also have "... = (--v \<squnion> -w) \<sqinter> ((--v \<squnion> --y) \<sqinter> (-v \<squnion> -x) \<sqinter> (-v \<squnion> --z))"
by (metis 1 inf_commutative)
also have "... = (--v \<squnion> -w) \<sqinter> ((--v \<squnion> --y) \<sqinter> ((-v \<squnion> -x) \<sqinter> (-v \<squnion> --z)))"
by (metis 1 inf_associative)
also have "... = (--v \<squnion> -w) \<sqinter> ((--v \<squnion> --y) \<sqinter> (-v \<squnion> (-x \<sqinter> --z)))"
by (simp add: sup_left_dist_inf)
also have "... = (--v \<squnion> -w) \<sqinter> (--v \<squnion> --y) \<sqinter> (-v \<squnion> (-x \<sqinter> --z))"
using 1 2 by (metis inf_associative)
also have "... = (--v \<squnion> (-w \<sqinter> --y)) \<sqinter> (-v \<squnion> (-x \<sqinter> --z))"
by (simp add: sup_left_dist_inf)
also have "... = (-v \<sqinter> (-w \<sqinter> --y)) \<squnion> (--v \<sqinter> (-x \<sqinter> --z))"
by (metis case_duality complement_1 complement_2 sub_inf_def)
also have "... = (-v \<sqinter> -w \<sqinter> --y) \<squnion> (--v \<sqinter> -x \<sqinter> --z)"
by (simp add: inf_associative)
finally show ?thesis
.
qed
lemma inf_cases_2: "--x = -(-x \<sqinter> -y) \<sqinter> -(-x \<sqinter> --y)"
using sub_inf_def sup_cases by auto
text \<open>consequences for \<open>top\<close> and \<open>bot\<close>\<close>
lemma sup_complement[simp]:
"-x \<squnion> --x = top"
using top_def by auto
lemma inf_complement[simp]:
"-x \<sqinter> --x = bot"
by (metis sub_bot_def sub_inf_def sub_top_def top_def)
lemma complement_bot[simp]:
"-bot = top"
using inf_complement inf_demorgan sup_complement by fastforce
lemma complement_top[simp]:
"-top = bot"
using sub_bot_def sub_top_def by blast
lemma sup_right_zero[simp]:
"-x \<squnion> top = top"
using sup_right_zero_var by auto
lemma sup_left_zero[simp]:
"top \<squnion> -x = top"
by (metis complement_bot sub_commutative sup_right_zero)
lemma inf_right_unit[simp]:
"-x \<sqinter> bot = bot"
by (metis complement_bot complement_top double_negation sub_sup_demorgan sup_right_zero)
lemma inf_left_unit[simp]:
"bot \<sqinter> -x = bot"
by (metis complement_top inf_commutative inf_right_unit)
lemma sup_right_unit[simp]:
"-x \<squnion> bot = -x"
using sup_right_unit_idempotent by auto
lemma sup_left_unit[simp]:
"bot \<squnion> -x = -x"
by (metis complement_top sub_commutative sup_right_unit)
lemma inf_right_zero[simp]:
"-x \<sqinter> top = -x"
by (metis inf_left_dist_sup sup_cases top_def)
lemma sub_inf_left_zero[simp]:
"top \<sqinter> -x = -x"
using inf_absorb top_def by fastforce
lemma bot_double_complement[simp]:
"--bot = bot"
by simp
lemma top_double_complement[simp]:
"--top = top"
by simp
text \<open>consequences for the order\<close>
lemma reflexive:
"-x \<le> -x"
by (simp add: sub_less_eq_def)
lemma transitive:
"-x \<le> -y \<Longrightarrow> -y \<le> -z \<Longrightarrow> -x \<le> -z"
by (metis sub_associative sub_less_eq_def)
lemma antisymmetric:
"-x \<le> -y \<Longrightarrow> -y \<le> -x \<Longrightarrow> -x = -y"
by (simp add: sub_commutative sub_less_eq_def)
lemma sub_bot_least:
"bot \<le> -x"
using sup_left_unit complement_top sub_less_eq_def by blast
lemma top_greatest:
"-x \<le> top"
using complement_bot sub_less_eq_def sup_right_zero by blast
lemma upper_bound_left:
"-x \<le> -x \<squnion> -y"
by (metis sub_associative sub_less_eq_def sub_sup_closed sup_idempotent)
lemma upper_bound_right:
"-y \<le> -x \<squnion> -y"
using sub_commutative upper_bound_left by fastforce
lemma sub_sup_left_isotone:
assumes "-x \<le> -y"
shows "-x \<squnion> -z \<le> -y \<squnion> -z"
proof -
have "-x \<squnion> -y = -y"
by (meson assms sub_less_eq_def)
thus ?thesis
by (metis (full_types) sub_associative sub_commutative sub_sup_closed upper_bound_left)
qed
lemma sub_sup_right_isotone:
"-x \<le> -y \<Longrightarrow> -z \<squnion> -x \<le> -z \<squnion> -y"
by (simp add: sub_commutative sub_sup_left_isotone)
lemma sup_isotone:
assumes "-p \<le> -q"
and "-r \<le> -s"
shows "-p \<squnion> -r \<le> -q \<squnion> -s"
proof -
have "\<And>x y. \<not> -x \<le> -y \<squnion> -r \<or> -x \<le> -y \<squnion> -s"
by (metis (full_types) assms(2) sub_sup_closed sub_sup_right_isotone transitive)
thus ?thesis
by (metis (no_types) assms(1) sub_sup_closed sub_sup_left_isotone)
qed
lemma sub_complement_antitone:
"-x \<le> -y \<Longrightarrow> --y \<le> --x"
by (metis inf_absorb inf_demorgan sub_less_eq_def)
lemma less_eq_inf:
"-x \<le> -y \<longleftrightarrow> -x \<sqinter> -y = -x"
by (metis inf_absorb inf_commutative sub_less_eq_def upper_bound_right sup_absorb)
lemma inf_complement_left_antitone:
"-x \<le> -y \<Longrightarrow> -(-y \<sqinter> -z) \<le> -(-x \<sqinter> -z)"
by (simp add: sub_complement_antitone inf_demorgan sub_sup_left_isotone)
lemma sub_inf_left_isotone:
"-x \<le> -y \<Longrightarrow> -x \<sqinter> -z \<le> -y \<sqinter> -z"
using sub_complement_antitone inf_closed inf_complement_left_antitone by fastforce
lemma sub_inf_right_isotone:
"-x \<le> -y \<Longrightarrow> -z \<sqinter> -x \<le> -z \<sqinter> -y"
by (simp add: inf_commutative sub_inf_left_isotone)
lemma inf_isotone:
assumes "-p \<le> -q"
and "-r \<le> -s"
shows "-p \<sqinter> -r \<le> -q \<sqinter> -s"
proof -
have "\<forall>w x y z. (-w \<le> -x \<sqinter> -y \<or> \<not> -w \<le> -x \<sqinter> -z) \<or> \<not> -z \<le> -y"
by (metis (no_types) inf_closed sub_inf_right_isotone transitive)
thus ?thesis
by (metis (no_types) assms inf_closed sub_inf_left_isotone)
qed
lemma least_upper_bound:
"-x \<le> -z \<and> -y \<le> -z \<longleftrightarrow> -x \<squnion> -y \<le> -z"
by (metis sub_sup_closed transitive upper_bound_right sup_idempotent sup_isotone upper_bound_left)
lemma lower_bound_left:
"-x \<sqinter> -y \<le> -x"
by (metis sub_inf_def upper_bound_right sup_absorb)
lemma lower_bound_right:
"-x \<sqinter> -y \<le> -y"
using inf_commutative lower_bound_left by fastforce
lemma greatest_lower_bound:
"-x \<le> -y \<and> -x \<le> -z \<longleftrightarrow> -x \<le> -y \<sqinter> -z"
by (metis inf_closed sub_inf_left_isotone less_eq_inf transitive lower_bound_left lower_bound_right)
lemma less_eq_sup_top:
"-x \<le> -y \<longleftrightarrow> --x \<squnion> -y = top"
by (metis complement_1 inf_commutative inf_complement_intro sub_inf_left_zero less_eq_inf sub_complement sup_complement_intro top_def)
lemma less_eq_inf_bot:
"-x \<le> -y \<longleftrightarrow> -x \<sqinter> --y = bot"
by (metis complement_bot complement_top double_negation inf_demorgan less_eq_sup_top sub_inf_def)
lemma shunting:
"-x \<sqinter> -y \<le> -z \<longleftrightarrow> -y \<le> --x \<squnion> -z"
proof (cases "--x \<squnion> (-z \<squnion> --y) = top")
case True
have "\<forall>v w. -v \<le> -w \<or> -w \<squnion> --v \<noteq> top"
using less_eq_sup_top sub_commutative by blast
thus ?thesis
by (metis True sub_associative sub_commutative sub_inf_def sub_sup_closed)
next
case False
hence "--x \<squnion> (-z \<squnion> --y) \<noteq> top \<and> \<not> -y \<le> -z \<squnion> --x"
by (metis (no_types) less_eq_sup_top sub_associative sub_commutative sub_sup_closed)
thus ?thesis
using less_eq_sup_top sub_associative sub_commutative sub_inf_def sub_sup_closed by auto
qed
lemma shunting_right:
"-x \<sqinter> -y \<le> -z \<longleftrightarrow> -x \<le> -z \<squnion> --y"
by (metis inf_commutative sub_commutative shunting)
lemma sup_less_eq_cases:
assumes "-z \<le> -x \<squnion> -y"
and "-z \<le> --x \<squnion> -y"
shows "-z \<le> -y"
proof -
have "-z \<le> (-x \<squnion> -y) \<sqinter> (--x \<squnion> -y)"
by (metis assms greatest_lower_bound sub_sup_closed)
also have "... = -y"
by (metis inf_cases sub_commutative)
finally show ?thesis
.
qed
lemma sup_less_eq_cases_2:
"-x \<squnion> -y \<le> -x \<squnion> -z \<Longrightarrow> --x \<squnion> -y \<le> --x \<squnion> -z \<Longrightarrow> -y \<le> -z"
by (metis least_upper_bound sup_less_eq_cases sub_sup_closed)
lemma sup_less_eq_cases_3:
"-y \<squnion> -x \<le> -z \<squnion> -x \<Longrightarrow> -y \<squnion> --x \<le> -z \<squnion> --x \<Longrightarrow> -y \<le> -z"
by (simp add: sup_less_eq_cases_2 sub_commutative)
lemma inf_less_eq_cases:
"-x \<sqinter> -y \<le> -z \<Longrightarrow> --x \<sqinter> -y \<le> -z \<Longrightarrow> -y \<le> -z"
by (simp add: shunting sup_less_eq_cases)
lemma inf_less_eq_cases_2:
"-x \<sqinter> -y \<le> -x \<sqinter> -z \<Longrightarrow> --x \<sqinter> -y \<le> --x \<sqinter> -z \<Longrightarrow> -y \<le> -z"
by (metis greatest_lower_bound inf_closed inf_less_eq_cases)
lemma inf_less_eq_cases_3:
"-y \<sqinter> -x \<le> -z \<sqinter> -x \<Longrightarrow> -y \<sqinter> --x \<le> -z \<sqinter> --x \<Longrightarrow> -y \<le> -z"
by (simp add: inf_commutative inf_less_eq_cases_2)
lemma inf_eq_cases:
"-x \<sqinter> -y = -x \<sqinter> -z \<Longrightarrow> --x \<sqinter> -y = --x \<sqinter> -z \<Longrightarrow> -y = -z"
by (metis inf_commutative sup_cases)
lemma inf_eq_cases_2:
"-y \<sqinter> -x = -z \<sqinter> -x \<Longrightarrow> -y \<sqinter> --x = -z \<sqinter> --x \<Longrightarrow> -y = -z"
using inf_commutative inf_eq_cases by auto
lemma wnf_lemma_1:
"((-x \<squnion> -y) \<sqinter> (--x \<squnion> -z)) \<squnion> -x = -x \<squnion> -y"
proof -
have "\<forall>u v w. (-u \<sqinter> (-v \<squnion> --w)) \<squnion> -w = -u \<squnion> -w"
by (metis inf_right_zero sub_associative sub_sup_closed sup_complement sup_idempotent sup_right_dist_inf)
thus ?thesis
by (metis (no_types) sub_associative sub_commutative sub_sup_closed sup_idempotent)
qed
lemma wnf_lemma_2:
"((-x \<squnion> -y) \<sqinter> (-z \<squnion> --y)) \<squnion> -y = -x \<squnion> -y"
using sub_commutative wnf_lemma_1 by fastforce
lemma wnf_lemma_3:
"((-x \<squnion> -z) \<sqinter> (--x \<squnion> -y)) \<squnion> --x = --x \<squnion> -y"
by (metis case_duality case_duality_2 double_negation sub_commutative wnf_lemma_2)
lemma wnf_lemma_4:
"((-z \<squnion> -y) \<sqinter> (-x \<squnion> --y)) \<squnion> --y = -x \<squnion> --y"
using sub_commutative wnf_lemma_3 by auto
end
class subset_boolean_algebra' = sup + uminus +
assumes sub_associative': "-x \<squnion> (-y \<squnion> -z) = (-x \<squnion> -y) \<squnion> -z"
assumes sub_commutative': "-x \<squnion> -y = -y \<squnion> -x"
assumes sub_complement': "-x = -(--x \<squnion> -y) \<squnion> -(--x \<squnion> --y)"
assumes sub_sup_closed': "\<exists>z . -x \<squnion> -y = -z"
begin
subclass subset_boolean_algebra
proof
show "\<And>x y z. - x \<squnion> (- y \<squnion> - z) = - x \<squnion> - y \<squnion> - z"
by (simp add: sub_associative')
show "\<And>x y. - x \<squnion> - y = - y \<squnion> - x"
by (simp add: sub_commutative')
show "\<And>x y. - x = - (- - x \<squnion> - y) \<squnion> - (- - x \<squnion> - - y)"
by (simp add: sub_complement')
show "\<And>x y. - x \<squnion> - y = - - (- x \<squnion> - y)"
proof -
fix x y
have "\<forall>x y. -y \<squnion> (-(--y \<squnion> -x) \<squnion> -(---x \<squnion> -y)) = -y \<squnion> --x"
by (metis (no_types) sub_associative' sub_commutative' sub_complement')
hence "\<forall>x. ---x = -x"
by (metis (no_types) sub_commutative' sub_complement')
thus "-x \<squnion> -y = --(-x \<squnion> -y)"
by (metis sub_sup_closed')
qed
qed
end
text \<open>
We introduce a type for the range of complement and show that it is an instance of \<open>boolean_algebra\<close>.
\<close>
typedef (overloaded) 'a boolean_subset = "{ x::'a::uminus . \<exists>y . x = -y }"
by auto
lemma simp_boolean_subset[simp]:
"\<exists>y . Rep_boolean_subset x = -y"
using Rep_boolean_subset by simp
setup_lifting type_definition_boolean_subset
text \<open>Theorem 8.1\<close>
instantiation boolean_subset :: (subset_boolean_algebra) huntington
begin
lift_definition sup_boolean_subset :: "'a boolean_subset \<Rightarrow> 'a boolean_subset \<Rightarrow> 'a boolean_subset" is sup
using sub_sup_closed by auto
lift_definition uminus_boolean_subset :: "'a boolean_subset \<Rightarrow> 'a boolean_subset" is uminus
by auto
instance
proof
show "\<And>x y z::'a boolean_subset. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
apply transfer
using sub_associative by blast
show "\<And>x y::'a boolean_subset. x \<squnion> y = y \<squnion> x"
apply transfer
using sub_commutative by blast
show "\<And>x y::'a boolean_subset. x = - (- x \<squnion> y) \<squnion> - (- x \<squnion> - y)"
apply transfer
using sub_complement by blast
qed
end
text \<open>Theorem 8.2\<close>
instantiation boolean_subset :: (subset_boolean_algebra_extended) huntington_extended
begin
lift_definition inf_boolean_subset :: "'a boolean_subset \<Rightarrow> 'a boolean_subset \<Rightarrow> 'a boolean_subset" is inf
using inf_closed by auto
lift_definition minus_boolean_subset :: "'a boolean_subset \<Rightarrow> 'a boolean_subset \<Rightarrow> 'a boolean_subset" is minus
using sub_minus_def by auto
lift_definition bot_boolean_subset :: "'a boolean_subset" is bot
by (metis complement_top)
lift_definition top_boolean_subset :: "'a boolean_subset" is top
by (metis complement_bot)
lift_definition less_eq_boolean_subset :: "'a boolean_subset \<Rightarrow> 'a boolean_subset \<Rightarrow> bool" is less_eq .
lift_definition less_boolean_subset :: "'a boolean_subset \<Rightarrow> 'a boolean_subset \<Rightarrow> bool" is less .
instance
proof
show 1: "top = (THE x. \<forall>y::'a boolean_subset. x = y \<squnion> - y)"
proof (rule the_equality[symmetric])
show "\<forall>y::'a boolean_subset. top = y \<squnion> - y"
apply transfer
by auto
show "\<And>x::'a boolean_subset. \<forall>y. x = y \<squnion> - y \<Longrightarrow> x = top"
apply transfer
by force
qed
have "(bot::'a boolean_subset) = - top"
apply transfer
by simp
thus "bot = - (THE x. \<forall>y::'a boolean_subset. x = y \<squnion> - y)"
using 1 by simp
show "\<And>x y::'a boolean_subset. x \<sqinter> y = - (- x \<squnion> - y)"
apply transfer
using sub_inf_def by blast
show "\<And>x y::'a boolean_subset. x - y = - (- x \<squnion> y)"
apply transfer
using sub_minus_def by blast
show "\<And>x y::'a boolean_subset. (x \<le> y) = (x \<squnion> y = y)"
apply transfer
using sub_less_eq_def by blast
show "\<And>x y::'a boolean_subset. (x < y) = (x \<squnion> y = y \<and> y \<squnion> x \<noteq> x)"
apply transfer
using sub_less_def by blast
qed
end
section \<open>Subset Boolean algebras with Additional Structure\<close>
text \<open>
We now discuss axioms that make the range of a unary operation a Boolean algebra, but add further properties that are common to the intended models.
In the intended models, the unary operation can be a complement, a pseudocomplement or the antidomain operation.
For simplicity, we mostly call the unary operation `complement'.
We first look at structures based only on join and complement, and then add axioms for the remaining operations of Boolean algebras.
In the intended models, the operation that is meet on the range of the complement can be a meet in the whole algebra or composition.
\<close>
subsection \<open>Axioms Derived from the New Axiomatisation\<close>
text \<open>
The axioms of the first algebra are based on \<open>boolean_algebra_3\<close>.
\<close>
text \<open>Definition 9\<close>
class subset_boolean_algebra_1 = sup + uminus +
assumes sba1_associative: "x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
assumes sba1_commutative: "x \<squnion> y = y \<squnion> x"
assumes sba1_idempotent[simp]: "x \<squnion> x = x"
assumes sba1_double_complement[simp]: "---x = -x"
assumes sba1_bot_unique: "-(x \<squnion> -x) = -(y \<squnion> -y)"
assumes sba1_export: "-x \<squnion> -(-x \<squnion> y) = -x \<squnion> -y"
begin
text \<open>Theorem 11.1\<close>
subclass subset_boolean_algebra
proof
show "\<And>x y z. - x \<squnion> (- y \<squnion> - z) = - x \<squnion> - y \<squnion> - z"
by (simp add: sba1_associative)
show "\<And>x y. - x \<squnion> - y = - y \<squnion> - x"
by (simp add: sba1_commutative)
show "\<And>x y. - x = - (- - x \<squnion> - y) \<squnion> - (- - x \<squnion> - - y)"
by (smt sba1_bot_unique sba1_commutative sba1_double_complement sba1_export sba1_idempotent)
thus "\<And>x y. - x \<squnion> - y = - - (- x \<squnion> - y)"
by (metis sba1_double_complement sba1_export)
qed
definition "sba1_bot \<equiv> THE x . \<forall>z . x = -(z \<squnion> -z)"
lemma sba1_bot:
"sba1_bot = -(z \<squnion> -z)"
using sba1_bot_def sba1_bot_unique by auto
end
text \<open>Boolean algebra operations based on join and complement\<close>
text \<open>Definition 10\<close>
class subset_extended_1 = sup + inf + minus + uminus + bot + top + ord +
assumes ba_bot: "bot = (THE x . \<forall>z . x = -(z \<squnion> -z))"
assumes ba_top: "top = -(THE x . \<forall>z . x = -(z \<squnion> -z))"
assumes ba_inf: "-x \<sqinter> -y = -(--x \<squnion> --y)"
assumes ba_minus: "-x - -y = -(--x \<squnion> -y)"
assumes ba_less_eq: "x \<le> y \<longleftrightarrow> x \<squnion> y = y"
assumes ba_less: "x < y \<longleftrightarrow> x \<squnion> y = y \<and> \<not> (y \<squnion> x = x)"
class subset_extended_2 = subset_extended_1 +
assumes ba_bot_unique: "-(x \<squnion> -x) = -(y \<squnion> -y)"
begin
lemma ba_bot_def:
"bot = -(z \<squnion> -z)"
using ba_bot ba_bot_unique by auto
lemma ba_top_def:
"top = --(z \<squnion> -z)"
using ba_bot_def ba_top by simp
end
text \<open>Subset forms Boolean Algebra, extended by Boolean algebra operations\<close>
class subset_boolean_algebra_1_extended = subset_boolean_algebra_1 + subset_extended_1
begin
subclass subset_extended_2
proof
show "\<And>x y. - (x \<squnion> - x) = - (y \<squnion> - y)"
by (simp add: sba1_bot_unique)
qed
subclass semilattice_sup
proof
show "\<And>x y. (x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (simp add: ba_less ba_less_eq)
show "\<And>x. x \<le> x"
by (simp add: ba_less_eq)
show "\<And>x y z. x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
by (metis sba1_associative ba_less_eq)
show "\<And>x y. x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
by (simp add: sba1_commutative ba_less_eq)
show "\<And>x y. x \<le> x \<squnion> y"
by (simp add: sba1_associative ba_less_eq)
thus "\<And>y x. y \<le> x \<squnion> y"
by (simp add: sba1_commutative)
show "\<And>y x z. y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> y \<squnion> z \<le> x"
by (metis sba1_associative ba_less_eq)
qed
text \<open>Theorem 11.2\<close>
subclass subset_boolean_algebra_extended
proof
show "top = (THE x. \<forall>y. x = - y \<squnion> - - y)"
by (smt ba_bot ba_bot_def ba_top sub_sup_closed the_equality)
thus "bot = - (THE x. \<forall>y. x = - y \<squnion> - - y)"
using ba_bot_def ba_top_def by force
show "\<And>x y. - x \<sqinter> - y = - (- - x \<squnion> - - y)"
by (simp add: ba_inf)
show "\<And>x y. - x - - y = - (- - x \<squnion> - y)"
by (simp add: ba_minus)
show "\<And>x y. (- x \<le> - y) = (- x \<squnion> - y = - y)"
using le_iff_sup by auto
show "\<And>x y. (- x < - y) = (- x \<squnion> - y = - y \<and> - y \<squnion> - x \<noteq> - x)"
by (simp add: ba_less)
qed
end
subsection \<open>Stronger Assumptions based on Join and Complement\<close>
text \<open>
We add further axioms covering properties common to the antidomain and (pseudo)complement instances.
\<close>
text \<open>Definition 12\<close>
class subset_boolean_algebra_2 = sup + uminus +
assumes sba2_associative: "x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
assumes sba2_commutative: "x \<squnion> y = y \<squnion> x"
assumes sba2_idempotent[simp]: "x \<squnion> x = x"
assumes sba2_bot_unit: "x \<squnion> -(y \<squnion> -y) = x"
assumes sba2_sub_sup_demorgan: "-(x \<squnion> y) = -(--x \<squnion> --y)"
assumes sba2_export: "-x \<squnion> -(-x \<squnion> y) = -x \<squnion> -y"
begin
text \<open>Theorem 13.1\<close>
subclass subset_boolean_algebra_1
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: sba2_associative)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: sba2_commutative)
show "\<And>x. x \<squnion> x = x"
by simp
show "\<And>x. - - - x = - x"
by (metis sba2_idempotent sba2_sub_sup_demorgan)
show "\<And>x y. - (x \<squnion> - x) = - (y \<squnion> - y)"
by (metis sba2_bot_unit sba2_commutative)
show "\<And>x y. - x \<squnion> - (- x \<squnion> y) = - x \<squnion> - y"
by (simp add: sba2_export)
qed
text \<open>Theorem 13.2\<close>
lemma double_complement_dist_sup:
"--(x \<squnion> y) = --x \<squnion> --y"
by (metis sba2_commutative sba2_export sba2_idempotent sba2_sub_sup_demorgan)
lemma maddux_3_3[simp]:
"-(x \<squnion> y) \<squnion> -(x \<squnion> -y) = -x"
by (metis double_complement_dist_sup sba1_double_complement sba2_commutative sub_complement)
lemma huntington_3_pp[simp]:
"-(-x \<squnion> -y) \<squnion> -(-x \<squnion> y) = --x"
using sba2_commutative maddux_3_3 by fastforce
end
class subset_boolean_algebra_2_extended = subset_boolean_algebra_2 + subset_extended_1
begin
subclass subset_boolean_algebra_1_extended ..
subclass bounded_semilattice_sup_bot
proof
show "\<And>x. bot \<le> x"
using sba2_bot_unit ba_bot_def sup_right_divisibility by auto
qed
text \<open>Theorem 13.3\<close>
lemma complement_antitone:
"x \<le> y \<Longrightarrow> -y \<le> -x"
by (metis le_iff_sup maddux_3_3 sba2_export sup_monoid.add_commute)
lemma double_complement_isotone:
"x \<le> y \<Longrightarrow> --x \<le> --y"
by (simp add: complement_antitone)
lemma sup_demorgan:
"-(x \<squnion> y) = -x \<sqinter> -y"
using sba2_sub_sup_demorgan ba_inf by auto
end
subsection \<open>Axioms for Meet\<close>
text \<open>
We add further axioms of \<open>inf\<close> covering properties common to the antidomain and pseudocomplement instances.
We omit the left distributivity rule and the right zero rule as they do not hold in some models.
In particular, the operation \<open>inf\<close> does not have to be commutative.
\<close>
text \<open>Definition 14\<close>
class subset_boolean_algebra_3_extended = subset_boolean_algebra_2_extended +
assumes sba3_inf_associative: "x \<sqinter> (y \<sqinter> z) = (x \<sqinter> y) \<sqinter> z"
assumes sba3_inf_right_dist_sup: "(x \<squnion> y) \<sqinter> z = (x \<sqinter> z) \<squnion> (y \<sqinter> z)"
assumes sba3_inf_complement_bot: "-x \<sqinter> x = bot"
assumes sba3_inf_left_unit[simp]: "top \<sqinter> x = x"
assumes sba3_complement_inf_double_complement: "-(x \<sqinter> --y) = -(x \<sqinter> y)"
begin
text \<open>Theorem 15\<close>
lemma inf_left_zero:
"bot \<sqinter> x = bot"
by (metis inf_right_unit sba3_inf_associative sba3_inf_complement_bot)
lemma inf_double_complement_export:
"--(--x \<sqinter> y) = --x \<sqinter> --y"
by (metis inf_closed sba3_complement_inf_double_complement)
lemma inf_left_isotone:
"x \<le> y \<Longrightarrow> x \<sqinter> z \<le> y \<sqinter> z"
using sba3_inf_right_dist_sup sup_right_divisibility by auto
lemma inf_complement_export:
"--(-x \<sqinter> y) = -x \<sqinter> --y"
by (metis inf_double_complement_export sba1_double_complement)
lemma double_complement_above:
"--x \<sqinter> x = x"
by (metis sup_monoid.add_0_right complement_bot inf_demorgan sba1_double_complement sba3_inf_complement_bot sba3_inf_right_dist_sup sba3_inf_left_unit)
lemma "x \<le> y \<Longrightarrow> z \<sqinter> x \<le> z \<sqinter> y" nitpick [expect=genuine] oops
lemma "x \<sqinter> top = x" nitpick [expect=genuine] oops
lemma "x \<sqinter> y = y \<sqinter> x" nitpick [expect=genuine] oops
end
subsection \<open>Stronger Assumptions for Meet\<close>
text \<open>
The following axioms also hold in both models, but follow from the axioms of \<open>subset_boolean_algebra_5_operations\<close>.
\<close>
text \<open>Definition 16\<close>
class subset_boolean_algebra_4_extended = subset_boolean_algebra_3_extended +
assumes sba4_inf_right_unit[simp]: "x \<sqinter> top = x"
assumes inf_right_isotone: "x \<le> y \<Longrightarrow> z \<sqinter> x \<le> z \<sqinter> y"
begin
lemma "x \<squnion> top = top" nitpick [expect=genuine] oops
lemma "x \<sqinter> bot = bot" nitpick [expect=genuine] oops
lemma "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" nitpick [expect=genuine] oops
lemma "(x \<sqinter> y = bot) = (x \<le> - y)" nitpick [expect=genuine] oops
end
section \<open>Boolean Algebras in Stone Algebras\<close>
text \<open>
We specialise \<open>inf\<close> to meet and complement to pseudocomplement.
This puts Stone algebras into the picture; for these it is well known that regular elements form a Boolean subalgebra \cite{Graetzer1971}.
\<close>
text \<open>Definition 17\<close>
class subset_boolean_algebra_5_extended = subset_boolean_algebra_3_extended +
assumes sba5_inf_commutative: "x \<sqinter> y = y \<sqinter> x"
assumes sba5_inf_absorb: "x \<sqinter> (x \<squnion> y) = x"
begin
subclass distrib_lattice_bot
proof
show "\<And>x y. x \<sqinter> y \<le> x"
by (metis sba5_inf_commutative sba3_inf_right_dist_sup sba5_inf_absorb sup_right_divisibility)
show "\<And>x y. x \<sqinter> y \<le> y"
by (metis inf_left_isotone sba5_inf_absorb sba5_inf_commutative sup_ge2)
show "\<And>x y z. x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z"
by (metis inf_left_isotone sba5_inf_absorb sup.orderE sup_monoid.add_commute)
show "\<And>x y z. x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z) "
by (metis sba3_inf_right_dist_sup sba5_inf_absorb sba5_inf_commutative sup_assoc)
qed
lemma inf_demorgan_2:
"-(x \<sqinter> y) = -x \<squnion> -y"
using sba3_complement_inf_double_complement sba5_inf_commutative sub_sup_closed sub_sup_demorgan by auto
lemma inf_export:
"x \<sqinter> -(x \<sqinter> y) = x \<sqinter> -y"
using inf_demorgan_2 sba3_inf_complement_bot sba3_inf_right_dist_sup sba5_inf_commutative by auto
lemma complement_inf[simp]:
"x \<sqinter> -x = bot"
using sba3_inf_complement_bot sba5_inf_commutative by auto
text \<open>Theorem 18.2\<close>
subclass stone_algebra
proof
show "\<And>x. x \<le> top"
by (simp add: inf.absorb_iff2)
show "\<And>x y. (x \<sqinter> y = bot) = (x \<le> - y)"
by (metis (full_types) complement_bot complement_inf inf.cobounded1 inf.order_iff inf_export sba3_complement_inf_double_complement sba3_inf_left_unit)
show "\<And>x. - x \<squnion> - - x = top"
by simp
qed
text \<open>Theorem 18.1\<close>
subclass subset_boolean_algebra_4_extended
proof
show "\<And>x. x \<sqinter> top = x"
by simp
show "\<And>x y z. x \<le> y \<Longrightarrow> z \<sqinter> x \<le> z \<sqinter> y"
using inf.sup_right_isotone by blast
qed
end
context stone_algebra_extended
begin
text \<open>Theorem 18.3\<close>
subclass subset_boolean_algebra_5_extended
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
using sup_assoc by auto
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: sup_commute)
show "\<And>x. x \<squnion> x = x"
by simp
show "\<And>x y. x \<squnion> - (y \<squnion> - y) = x"
by simp
show "\<And>x y. - (x \<squnion> y) = - (- - x \<squnion> - - y)"
by auto
show "\<And>x y. - x \<squnion> - (- x \<squnion> y) = - x \<squnion> - y"
by (metis maddux_3_21_pp p_dist_sup regular_closed_p)
show "bot = (THE x. \<forall>z. x = - (z \<squnion> - z))"
by simp
thus "top = - (THE x. \<forall>z. x = - (z \<squnion> - z))"
using p_bot by blast
show "\<And>x y. - x \<sqinter> - y = - (- - x \<squnion> - - y)"
by simp
show "\<And>x y. - x - - y = - (- - x \<squnion> - y)"
by auto
show "\<And>x y. (x \<le> y) = (x \<squnion> y = y)"
by (simp add: le_iff_sup)
thus "\<And>x y. (x < y) = (x \<squnion> y = y \<and> y \<squnion> x \<noteq> x)"
by (simp add: less_le_not_le)
show "\<And>x y z. x \<sqinter> (y \<sqinter> z) = x \<sqinter> y \<sqinter> z"
by (simp add: inf.sup_monoid.add_assoc)
show "\<And>x y z. (x \<squnion> y) \<sqinter> z = x \<sqinter> z \<squnion> y \<sqinter> z"
by (simp add: inf_sup_distrib2)
show "\<And>x. - x \<sqinter> x = bot"
by simp
show "\<And>x. top \<sqinter> x = x"
by simp
show "\<And>x y. - (x \<sqinter> - - y) = - (x \<sqinter> y)"
by simp
show "\<And>x y. x \<sqinter> y = y \<sqinter> x"
by (simp add: inf_commute)
show "\<And>x y. x \<sqinter> (x \<squnion> y) = x"
by simp
qed
end
section \<open>Domain Semirings\<close>
text \<open>
The following development of tests in IL-semirings, prepredomain semirings, predomain semirings and domain semirings is mostly based on \cite{MoellerDesharnais2019}; see also \cite{DesharnaisMoeller2014}.
See \cite{DesharnaisMoellerStruth2006b} for domain axioms in idempotent semirings.
See \cite{DesharnaisJipsenStruth2009,JacksonStokes2004} for domain axioms in semigroups and monoids.
Some variants have been implemented in \cite{GomesGuttmannHoefnerStruthWeber2016}.
\<close>
subsection \<open>Idempotent Left Semirings\<close>
text \<open>Definition 19\<close>
class il_semiring = sup + inf + bot + top + ord +
assumes il_associative: "x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
assumes il_commutative: "x \<squnion> y = y \<squnion> x"
assumes il_idempotent[simp]: "x \<squnion> x = x"
assumes il_bot_unit: "x \<squnion> bot = x"
assumes il_inf_associative: "x \<sqinter> (y \<sqinter> z) = (x \<sqinter> y) \<sqinter> z"
assumes il_inf_right_dist_sup: "(x \<squnion> y) \<sqinter> z = (x \<sqinter> z) \<squnion> (y \<sqinter> z)"
assumes il_inf_left_unit[simp]: "top \<sqinter> x = x"
assumes il_inf_right_unit[simp]: "x \<sqinter> top = x"
assumes il_sub_inf_left_zero[simp]: "bot \<sqinter> x = bot"
assumes il_sub_inf_right_isotone: "x \<le> y \<Longrightarrow> z \<sqinter> x \<le> z \<sqinter> y"
assumes il_less_eq: "x \<le> y \<longleftrightarrow> x \<squnion> y = y"
assumes il_less_def: "x < y \<longleftrightarrow> x \<le> y \<and> \<not>(y \<le> x)"
begin
lemma il_unit_bot: "bot \<squnion> x = x"
using il_bot_unit il_commutative by fastforce
subclass order
proof
show "\<And>x y. (x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (simp add: il_less_def)
show "\<And>x. x \<le> x"
by (simp add: il_less_eq)
show "\<And>x y z. x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
by (metis il_associative il_less_eq)
show "\<And>x y. x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
by (simp add: il_commutative il_less_eq)
qed
lemma il_sub_inf_right_isotone_var:
"(x \<sqinter> y) \<squnion> (x \<sqinter> z) \<le> x \<sqinter> (y \<squnion> z)"
by (smt il_associative il_commutative il_idempotent il_less_eq il_sub_inf_right_isotone)
lemma il_sub_inf_left_isotone:
"x \<le> y \<Longrightarrow> x \<sqinter> z \<le> y \<sqinter> z"
by (metis il_inf_right_dist_sup il_less_eq)
lemma il_sub_inf_left_isotone_var:
"(y \<sqinter> x) \<squnion> (z \<sqinter> x) \<le> (y \<squnion> z) \<sqinter> x"
by (simp add: il_inf_right_dist_sup)
lemma sup_left_isotone:
"x \<le> y \<Longrightarrow> x \<squnion> z \<le> y \<squnion> z"
by (smt il_associative il_commutative il_idempotent il_less_eq)
lemma sup_right_isotone:
"x \<le> y \<Longrightarrow> z \<squnion> x \<le> z \<squnion> y"
by (simp add: il_commutative sup_left_isotone)
lemma bot_least:
"bot \<le> x"
by (simp add: il_less_eq il_unit_bot)
lemma less_eq_bot:
"x \<le> bot \<longleftrightarrow> x = bot"
by (simp add: il_bot_unit il_less_eq)
abbreviation are_complementary :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
where "are_complementary x y \<equiv> x \<squnion> y = top \<and> x \<sqinter> y = bot \<and> y \<sqinter> x = bot"
abbreviation test :: "'a \<Rightarrow> bool"
where "test x \<equiv> \<exists>y . are_complementary x y"
definition tests :: "'a set"
where "tests = { x . test x }"
lemma bot_test:
"test bot"
by (simp add: il_unit_bot)
lemma top_test:
"test top"
by (simp add: il_bot_unit)
lemma test_sub_identity:
"test x \<Longrightarrow> x \<le> top"
using il_associative il_less_eq by auto
lemma neg_unique:
"are_complementary x y \<Longrightarrow> are_complementary x z \<Longrightarrow> y = z"
- by (metis antisym il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone_var)
+ by (metis order.antisym il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone_var)
definition neg :: "'a \<Rightarrow> 'a" ("!")
where "!x \<equiv> THE y . are_complementary x y"
lemma neg_char:
assumes "test x"
shows "are_complementary x (!x)"
proof (unfold neg_def)
from assms obtain y where 1: "are_complementary x y"
by auto
show "are_complementary x (THE y. are_complementary x y)"
proof (rule theI)
show "are_complementary x y"
using 1 by simp
show "\<And>z. are_complementary x z \<Longrightarrow> z = y"
using 1 neg_unique by blast
qed
qed
lemma are_complementary_symmetric:
"are_complementary x y \<longleftrightarrow> are_complementary y x"
using il_commutative by auto
lemma neg_test:
"test x \<Longrightarrow> test (!x)"
using are_complementary_symmetric neg_char by blast
lemma are_complementary_test:
"test x \<Longrightarrow> are_complementary x y \<Longrightarrow> test y"
using il_commutative by auto
lemma neg_involutive:
"test x \<Longrightarrow> !(!x) = x"
using are_complementary_symmetric neg_char neg_unique by blast
lemma test_inf_left_below:
"test x \<Longrightarrow> x \<sqinter> y \<le> y"
by (metis il_associative il_idempotent il_inf_left_unit il_inf_right_dist_sup il_less_eq)
lemma test_inf_right_below:
"test x \<Longrightarrow> y \<sqinter> x \<le> y"
by (metis il_inf_right_unit il_sub_inf_right_isotone test_sub_identity)
lemma neg_bot:
"!bot = top"
using il_unit_bot neg_char by fastforce
lemma neg_top:
"!top = bot"
using bot_test neg_bot neg_involutive by fastforce
lemma test_inf_idempotent:
"test x \<Longrightarrow> x \<sqinter> x = x"
by (metis il_bot_unit il_inf_left_unit il_inf_right_dist_sup)
lemma test_inf_semicommutative:
assumes "test x"
and "test y"
shows "x \<sqinter> y \<le> y \<sqinter> x"
proof -
have "x \<sqinter> y = (y \<sqinter> x \<sqinter> y) \<squnion> (!y \<sqinter> x \<sqinter> y)"
by (metis assms(2) il_inf_left_unit il_inf_right_dist_sup neg_char)
also have "... \<le> (y \<sqinter> x \<sqinter> y) \<squnion> (!y \<sqinter> y)"
proof -
obtain z where "are_complementary y z"
using assms(2) by blast
hence "y \<sqinter> (x \<sqinter> y) \<squnion> !y \<sqinter> (x \<sqinter> y) \<le> y \<sqinter> (x \<sqinter> y)"
by (metis assms(1) calculation il_sub_inf_left_isotone il_bot_unit il_idempotent il_inf_associative il_less_eq neg_char test_inf_right_below)
thus ?thesis
by (simp add: il_associative il_inf_associative il_less_eq)
qed
also have "... \<le> (y \<sqinter> x) \<squnion> (!y \<sqinter> y)"
by (metis assms(2) il_bot_unit il_inf_right_unit il_sub_inf_right_isotone neg_char test_sub_identity)
also have "... = y \<sqinter> x"
by (simp add: assms(2) il_bot_unit neg_char)
finally show ?thesis
.
qed
lemma test_inf_commutative:
"test x \<Longrightarrow> test y \<Longrightarrow> x \<sqinter> y = y \<sqinter> x"
- by (simp add: antisym test_inf_semicommutative)
+ by (simp add: order.antisym test_inf_semicommutative)
lemma test_inf_bot:
"test x \<Longrightarrow> x \<sqinter> bot = bot"
using il_inf_associative test_inf_idempotent by fastforce
lemma test_absorb_1:
"test x \<Longrightarrow> test y \<Longrightarrow> x \<squnion> (x \<sqinter> y) = x"
using il_commutative il_less_eq test_inf_right_below by auto
lemma test_absorb_2:
"test x \<Longrightarrow> test y \<Longrightarrow> x \<squnion> (y \<sqinter> x) = x"
by (metis test_absorb_1 test_inf_commutative)
lemma test_absorb_3:
"test x \<Longrightarrow> test y \<Longrightarrow> x \<sqinter> (x \<squnion> y) = x"
- apply (rule antisym)
+ apply (rule order.antisym)
apply (metis il_associative il_inf_right_unit il_less_eq il_sub_inf_right_isotone test_sub_identity)
by (metis il_sub_inf_right_isotone_var test_absorb_1 test_inf_idempotent)
lemma test_absorb_4:
"test x \<Longrightarrow> test y \<Longrightarrow> (x \<squnion> y) \<sqinter> x = x"
by (smt il_inf_right_dist_sup test_inf_idempotent il_commutative il_less_eq test_inf_left_below)
lemma test_import_1:
assumes "test x"
and "test y"
shows "x \<squnion> (!x \<sqinter> y) = x \<squnion> y"
proof -
have "x \<squnion> (!x \<sqinter> y) = x \<squnion> ((y \<squnion> !y) \<sqinter> x) \<squnion> (!x \<sqinter> y)"
by (simp add: assms(2) neg_char)
also have "... = x \<squnion> (!y \<sqinter> x) \<squnion> (x \<sqinter> y) \<squnion> (!x \<sqinter> y)"
by (smt assms il_associative il_commutative il_inf_right_dist_sup test_inf_commutative)
also have "... = x \<squnion> ((x \<squnion> !x) \<sqinter> y)"
by (smt calculation il_associative il_commutative il_idempotent il_inf_right_dist_sup)
also have "... = x \<squnion> y"
by (simp add: assms(1) neg_char)
finally show ?thesis
.
qed
lemma test_import_2:
assumes "test x"
and "test y"
shows "x \<squnion> (y \<sqinter> !x) = x \<squnion> y"
proof -
obtain z where 1: "are_complementary y z"
using assms(2) by moura
obtain w where 2: "are_complementary x w"
using assms(1) by auto
hence "x \<sqinter> !x = bot"
using neg_char by blast
hence "!x \<sqinter> y = y \<sqinter> !x"
using 1 2 by (metis il_commutative neg_char test_inf_commutative)
thus ?thesis
using 1 2 by (metis test_import_1)
qed
lemma test_import_3:
assumes "test x"
shows "(!x \<squnion> y) \<sqinter> x = y \<sqinter> x"
by (simp add: assms(1) il_inf_right_dist_sup il_unit_bot neg_char)
lemma test_import_4:
assumes "test x"
and "test y"
shows "(!x \<squnion> y) \<sqinter> x = x \<sqinter> y"
by (metis assms test_import_3 test_inf_commutative)
lemma test_inf:
"test x \<Longrightarrow> test y \<Longrightarrow> test z \<Longrightarrow> z \<le> x \<sqinter> y \<longleftrightarrow> z \<le> x \<and> z \<le> y"
apply (rule iffI)
using dual_order.trans test_inf_left_below test_inf_right_below apply blast
by (smt il_less_eq il_sub_inf_right_isotone test_absorb_4)
lemma test_shunting:
assumes "test x"
and "test y"
shows "x \<sqinter> y \<le> z \<longleftrightarrow> x \<le> !y \<squnion> z"
proof
assume 1: "x \<sqinter> y \<le> z"
have "x = (!y \<sqinter> x) \<squnion> (y \<sqinter> x)"
by (metis assms(2) il_commutative il_inf_left_unit il_inf_right_dist_sup neg_char)
also have "... \<le> !y \<squnion> (y \<sqinter> x)"
by (simp add: assms(1) sup_left_isotone test_inf_right_below)
also have "... \<le> !y \<squnion> z"
using 1 by (simp add: assms sup_right_isotone test_inf_commutative)
finally show "x \<le> !y \<squnion> z"
.
next
assume "x \<le> !y \<squnion> z"
hence "x \<sqinter> y \<le> (!y \<squnion> z) \<sqinter> y"
using il_sub_inf_left_isotone by blast
also have "... = z \<sqinter> y"
by (simp add: assms(2) test_import_3)
also have "... \<le> z"
by (simp add: assms(2) test_inf_right_below)
finally show "x \<sqinter> y \<le> z"
.
qed
lemma test_shunting_bot:
assumes "test x"
and "test y"
shows "x \<le> y \<longleftrightarrow> x \<sqinter> !y \<le> bot"
by (simp add: assms il_bot_unit neg_involutive neg_test test_shunting)
lemma test_shunting_bot_eq:
assumes "test x"
and "test y"
shows "x \<le> y \<longleftrightarrow> x \<sqinter> !y = bot"
by (simp add: assms test_shunting_bot less_eq_bot)
lemma neg_antitone:
assumes "test x"
and "test y"
and "x \<le> y"
shows "!y \<le> !x"
proof -
have 1: "x \<sqinter> !y = bot"
using assms test_shunting_bot_eq by blast
have 2: "x \<squnion> !x = top"
by (simp add: assms(1) neg_char)
have "are_complementary y (!y)"
by (simp add: assms(2) neg_char)
thus ?thesis
using 1 2 by (metis il_unit_bot il_commutative il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone test_sub_identity)
qed
lemma test_sup_neg_1:
assumes "test x"
and "test y"
shows "(x \<squnion> y) \<squnion> (!x \<sqinter> !y) = top"
proof -
have "x \<squnion> !x = top"
by (simp add: assms(1) neg_char)
hence "x \<squnion> (y \<squnion> !x) = top"
by (metis assms(2) il_associative il_commutative il_idempotent)
hence "x \<squnion> (y \<squnion> !x \<sqinter> !y) = top"
by (simp add: assms neg_test test_import_2)
thus ?thesis
by (simp add: il_associative)
qed
lemma test_sup_neg_2:
assumes "test x"
and "test y"
shows "(x \<squnion> y) \<sqinter> (!x \<sqinter> !y) = bot"
proof -
have 1: "are_complementary y (!y)"
by (simp add: assms(2) neg_char)
obtain z where 2: "are_complementary x z"
using assms(1) by auto
hence "!x = z"
using neg_char neg_unique by blast
thus ?thesis
using 1 2 by (metis are_complementary_symmetric il_inf_associative neg_involutive test_import_3 test_inf_bot test_inf_commutative)
qed
lemma de_morgan_1:
assumes "test x"
and "test y"
and "test (x \<sqinter> y)"
shows "!(x \<sqinter> y) = !x \<squnion> !y"
-proof (rule antisym)
+proof (rule order.antisym)
have 1: "test (!(x \<sqinter> y))"
by (simp add: assms neg_test)
have "x \<le> (x \<sqinter> y) \<squnion> !y"
by (metis (full_types) assms il_commutative neg_char test_shunting test_shunting_bot_eq)
hence "x \<sqinter> !(x \<sqinter> y) \<le> !y"
using 1 by (simp add: assms(1,3) neg_involutive test_shunting)
hence "!(x \<sqinter> y) \<sqinter> x \<le> !y"
using 1 by (metis assms(1) test_inf_commutative)
thus "!(x \<sqinter> y) \<le> !x \<squnion> !y"
using 1 assms(1) test_shunting by blast
have 2: "!x \<le> !(x \<sqinter> y)"
by (simp add: assms neg_antitone test_inf_right_below)
have "!y \<le> !(x \<sqinter> y)"
by (simp add: assms neg_antitone test_inf_left_below)
thus "!x \<squnion> !y \<le> !(x \<sqinter> y)"
using 2 by (metis il_associative il_less_eq)
qed
lemma de_morgan_2:
assumes "test x"
and "test y"
and "test (x \<squnion> y)"
shows "!(x \<squnion> y) = !x \<sqinter> !y"
-proof (rule antisym)
+proof (rule order.antisym)
have 1: "!(x \<squnion> y) \<le> !x"
by (metis assms il_inf_left_unit il_sub_inf_left_isotone neg_antitone test_absorb_3 test_sub_identity)
have "!(x \<squnion> y) \<le> !y"
by (metis assms il_commutative il_inf_left_unit il_sub_inf_left_isotone neg_antitone test_absorb_3 test_sub_identity)
thus "!(x \<squnion> y) \<le> !x \<sqinter> !y"
using 1 by (simp add: assms neg_test test_inf)
have "top \<le> x \<squnion> y \<squnion> !(x \<squnion> y)"
by (simp add: assms(3) neg_char)
hence "top \<sqinter> !x \<le> y \<squnion> !(x \<squnion> y)"
by (smt assms(1) assms(3) il_commutative il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone il_unit_bot neg_char test_sub_identity)
thus "!x \<sqinter> !y \<le> !(x \<squnion> y)"
by (simp add: assms(1) assms(2) neg_involutive neg_test test_shunting)
qed
lemma test_inf_closed_sup_complement:
assumes "test x"
and "test y"
and "\<forall>u v . test u \<and> test v \<longrightarrow> test (u \<sqinter> v)"
shows "!x \<sqinter> !y \<sqinter> (x \<squnion> y) = bot"
proof -
have 1: "!(!x \<sqinter> !y) = x \<squnion> y"
by (simp add: assms de_morgan_1 neg_involutive neg_test)
have "test (!(!x \<sqinter> !y))"
by (metis assms neg_test)
thus ?thesis
using 1 by (metis assms(1,2) de_morgan_2 neg_char)
qed
lemma test_sup_complement_sup_closed:
assumes "test x"
and "test y"
and "\<forall>u v . test u \<and> test v \<longrightarrow> !u \<sqinter> !v \<sqinter> (u \<squnion> v) = bot"
shows "test (x \<squnion> y)"
by (meson assms test_sup_neg_1 test_sup_neg_2)
lemma test_inf_closed_sup_closed:
assumes "test x"
and "test y"
and "\<forall>u v . test u \<and> test v \<longrightarrow> test (u \<sqinter> v)"
shows "test (x \<squnion> y)"
using assms test_inf_closed_sup_complement test_sup_complement_sup_closed by simp
end
subsection \<open>Prepredomain Semirings\<close>
class dom =
fixes d :: "'a \<Rightarrow> 'a"
class ppd_semiring = il_semiring + dom +
assumes d_closed: "test (d x)"
assumes d1: "x \<le> d x \<sqinter> x"
begin
lemma d_sub_identity:
"d x \<le> top"
using d_closed test_sub_identity by blast
lemma d1_eq:
"x = d x \<sqinter> x"
proof -
have "x = (d x \<squnion> top) \<sqinter> x"
using d_sub_identity il_less_eq by auto
thus ?thesis
using d1 il_commutative il_inf_right_dist_sup il_less_eq by force
qed
lemma d_increasing_sub_identity:
"x \<le> top \<Longrightarrow> x \<le> d x"
by (metis d1_eq il_inf_right_unit il_sub_inf_right_isotone)
lemma d_top:
"d top = top"
by (simp add: d_increasing_sub_identity d_sub_identity dual_order.antisym)
lemma d_bot_only:
"d x = bot \<Longrightarrow> x = bot"
by (metis d1_eq il_sub_inf_left_zero)
lemma d_strict: "d bot \<le> bot" nitpick [expect=genuine] oops
lemma d_isotone_var: "d x \<le> d (x \<squnion> y)" nitpick [expect=genuine] oops
lemma d_fully_strict: "d x = bot \<longleftrightarrow> x = bot" nitpick [expect=genuine] oops
lemma test_d_fixpoint: "test x \<Longrightarrow> d x = x" nitpick [expect=genuine] oops
end
subsection \<open>Predomain Semirings\<close>
class pd_semiring = ppd_semiring +
assumes d2: "test p \<Longrightarrow> d (p \<sqinter> x) \<le> p"
begin
lemma d_strict:
"d bot \<le> bot"
using bot_test d2 by fastforce
lemma d_strict_eq:
"d bot = bot"
using d_strict il_bot_unit il_less_eq by auto
lemma test_d_fixpoint:
"test x \<Longrightarrow> d x = x"
- by (metis antisym d1_eq d2 test_inf_idempotent test_inf_right_below)
+ by (metis order.antisym d1_eq d2 test_inf_idempotent test_inf_right_below)
lemma d_surjective:
"test x \<Longrightarrow> \<exists>y . d y = x"
using test_d_fixpoint by blast
lemma test_d_fixpoint_iff:
"test x \<longleftrightarrow> d x = x"
by (metis d_closed test_d_fixpoint)
lemma d_surjective_iff:
"test x \<longleftrightarrow> (\<exists>y . d y = x)"
using d_surjective d_closed by blast
lemma tests_d_range:
"tests = range d"
using tests_def image_def d_surjective_iff by auto
lemma llp:
assumes "test y"
shows "d x \<le> y \<longleftrightarrow> x \<le> y \<sqinter> x"
- by (metis assms d1_eq d2 eq_iff il_sub_inf_left_isotone test_inf_left_below)
+ by (metis assms d1_eq d2 order.eq_iff il_sub_inf_left_isotone test_inf_left_below)
lemma gla:
assumes "test y"
shows "y \<le> !(d x) \<longleftrightarrow> y \<sqinter> x \<le> bot"
proof -
obtain ad where 1: "\<forall>x. are_complementary (d x) (ad x)"
using d_closed by moura
hence 2: "\<forall>x y. d (d y \<sqinter> x) \<le> d y"
using d2 by blast
have 3: "\<forall>x. ad x \<sqinter> x = bot"
using 1 by (metis d1_eq il_inf_associative il_sub_inf_left_zero)
have 4: "\<forall>x y. d y \<sqinter> x \<squnion> ad y \<sqinter> x = top \<sqinter> x"
using 1 by (metis il_inf_right_dist_sup)
have 5: "\<forall>x y z. z \<sqinter> y \<le> x \<sqinter> y \<or> (z \<squnion> x) \<sqinter> y \<noteq> x \<sqinter> y"
by (simp add: il_inf_right_dist_sup il_less_eq)
have 6: "\<forall>x. !(d x) = ad x"
using 1 neg_char neg_unique by blast
have 7: "\<forall>x. top \<sqinter> x = x"
by auto
hence "\<forall>x. y \<sqinter> x \<squnion> !y \<sqinter> x = x"
by (metis assms il_inf_right_dist_sup neg_char)
thus ?thesis
using 1 2 3 4 5 6 7 by (metis assms d1_eq il_commutative il_less_eq test_d_fixpoint)
qed
lemma gla_var:
"test y \<Longrightarrow> y \<sqinter> d x \<le> bot \<longleftrightarrow> y \<sqinter> x \<le> bot"
using gla d_closed il_bot_unit test_shunting by auto
lemma llp_var:
assumes "test y"
shows "y \<le> !(d x) \<longleftrightarrow> x \<le> !y \<sqinter> x"
apply (rule iffI)
apply (metis (no_types, hide_lams) assms gla Least_equality il_inf_left_unit il_inf_right_dist_sup il_less_eq il_unit_bot order.refl neg_char)
by (metis assms gla gla_var llp il_commutative il_sub_inf_right_isotone neg_char)
lemma d_idempotent:
"d (d x) = d x"
using d_closed test_d_fixpoint_iff by auto
lemma d_neg:
"test x \<Longrightarrow> d (!x) = !x"
using il_commutative neg_char test_d_fixpoint_iff by fastforce
lemma d_fully_strict:
"d x = bot \<longleftrightarrow> x = bot"
using d_strict_eq d_bot_only by blast
lemma d_ad_comp:
"!(d x) \<sqinter> x = bot"
proof -
have "\<forall>x. !(d x) \<sqinter> d x = bot"
by (simp add: d_closed neg_char)
thus ?thesis
by (metis d1_eq il_inf_associative il_sub_inf_left_zero)
qed
lemma d_isotone:
assumes "x \<le> y"
shows "d x \<le> d y"
proof -
obtain ad where 1: "\<forall>x. are_complementary (d x) (ad x)"
using d_closed by moura
hence "ad y \<sqinter> x \<le> bot"
by (metis assms d1_eq il_inf_associative il_sub_inf_left_zero il_sub_inf_right_isotone)
thus ?thesis
using 1 by (metis d2 il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_less_eq)
qed
lemma d_isotone_var:
"d x \<le> d (x \<squnion> y)"
using d_isotone il_associative il_less_eq by auto
lemma d3_conv:
"d (x \<sqinter> y) \<le> d (x \<sqinter> d y)"
by (metis (mono_tags, hide_lams) d1_eq d2 d_closed il_inf_associative)
lemma d_test_inf_idempotent:
"d x \<sqinter> d x = d x"
by (metis d_idempotent d1_eq)
lemma d_test_inf_closed:
assumes "test x"
and "test y"
shows "d (x \<sqinter> y) = x \<sqinter> y"
-proof (rule antisym)
+proof (rule order.antisym)
have "d (x \<sqinter> y) = d (x \<sqinter> y) \<sqinter> d (x \<sqinter> y)"
by (simp add: d_test_inf_idempotent)
also have "... \<le> x \<sqinter> d (x \<sqinter> y)"
by (simp add: assms(1) d2 il_sub_inf_left_isotone)
also have "... \<le> x \<sqinter> y"
by (metis assms d_isotone il_sub_inf_right_isotone test_inf_left_below test_d_fixpoint)
finally show "d (x \<sqinter> y) \<le> x \<sqinter> y"
.
show "x \<sqinter> y \<le> d (x \<sqinter> y)"
using assms d_increasing_sub_identity dual_order.trans test_inf_left_below test_sub_identity by blast
qed
lemma test_inf_closed:
"test x \<Longrightarrow> test y \<Longrightarrow> test (x \<sqinter> y)"
using d_test_inf_closed test_d_fixpoint_iff by simp
lemma test_sup_closed:
"test x \<Longrightarrow> test y \<Longrightarrow> test (x \<squnion> y)"
using test_inf_closed test_inf_closed_sup_closed by simp
lemma d_export:
assumes "test x"
shows "d (x \<sqinter> y) = x \<sqinter> d y"
-proof (rule antisym)
+proof (rule order.antisym)
have 1: "d (x \<sqinter> y) \<le> x"
by (simp add: assms d2)
have "d (x \<sqinter> y) \<le> d y"
by (metis assms d_isotone_var il_inf_left_unit il_inf_right_dist_sup)
thus "d (x \<sqinter> y) \<le> x \<sqinter> d y"
using 1 by (metis assms d_idempotent llp dual_order.trans il_sub_inf_right_isotone)
have "y = (!x \<sqinter> y) \<squnion> (x \<sqinter> y)"
by (metis assms il_commutative il_inf_left_unit il_inf_right_dist_sup neg_char)
also have "... = (!x \<sqinter> y) \<squnion> (d (x \<sqinter> y) \<sqinter> x \<sqinter> y)"
by (metis d1_eq il_inf_associative)
also have "... = (!x \<sqinter> y) \<squnion> (d (x \<sqinter> y) \<sqinter> y)"
using 1 by (smt calculation d1_eq il_associative il_commutative il_inf_associative il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone_var)
also have "... = (!x \<squnion> d (x \<sqinter> y)) \<sqinter> y"
by (simp add: il_inf_right_dist_sup)
finally have "y \<le> (!x \<squnion> d (x \<sqinter> y)) \<sqinter> y"
by simp
hence "d y \<le> !x \<squnion> d (x \<sqinter> y)"
using assms llp test_sup_closed neg_test d_closed by simp
hence "d y \<sqinter> x \<le> d (x \<sqinter> y)"
by (simp add: assms d_closed test_shunting)
thus "x \<sqinter> d y \<le> d (x \<sqinter> y)"
by (metis assms d_closed test_inf_commutative)
qed
lemma test_inf_left_dist_sup:
assumes "test x"
and "test y"
and "test z"
shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
proof -
have "x \<sqinter> (y \<squnion> z) = (y \<squnion> z) \<sqinter> x"
using assms test_sup_closed test_inf_commutative by smt
also have "... = (y \<sqinter> x) \<squnion> (z \<sqinter> x)"
using il_inf_right_dist_sup by simp
also have "... = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
using assms test_sup_closed test_inf_commutative by smt
finally show ?thesis
.
qed
lemma "!x \<squnion> !y = !(!(!x \<squnion> !y))" nitpick [expect=genuine] oops
lemma "d x = !(!x)" nitpick [expect=genuine] oops
sublocale subset_boolean_algebra where uminus = "\<lambda> x . !(d x)"
proof
show "\<And>x y z. !(d x) \<squnion> (!(d y) \<squnion> !(d z)) = !(d x) \<squnion> !(d y) \<squnion> !(d z)"
using il_associative by blast
show "\<And>x y. !(d x) \<squnion> !(d y) = !(d y) \<squnion> !(d x)"
by (simp add: il_commutative)
show "\<And>x y. !(d x) \<squnion> !(d y) = !(d (!(d (!(d x) \<squnion> !(d y)))))"
proof -
fix x y
have "test (!(d x)) \<and> test (!(d y))"
by (simp add: d_closed neg_test)
hence "test (!(d x) \<squnion> !(d y))"
by (simp add: test_sup_closed)
thus "!(d x) \<squnion> !(d y) = !(d (!(d (!(d x) \<squnion> !(d y)))))"
by (simp add: d_neg neg_involutive test_d_fixpoint)
qed
show "\<And>x y. !(d x) = !(d (!(d (!(d x))) \<squnion> !(d y))) \<squnion> !(d (!(d (!(d x))) \<squnion> !(d (!(d y)))))"
proof -
fix x y
have "!(d (!(d (!(d x))) \<squnion> !(d y))) \<squnion> !(d (!(d (!(d x))) \<squnion> !(d (!(d y))))) = !(d x \<squnion> !(d y)) \<squnion> !(d x \<squnion> d y)"
using d_closed neg_test test_sup_closed neg_involutive test_d_fixpoint by auto
also have "... = (!(d x) \<sqinter> d y) \<squnion> (!(d x) \<sqinter> !(d y))"
using d_closed neg_test test_sup_closed neg_involutive de_morgan_2 by auto
also have "... = !(d x) \<sqinter> (d y \<squnion> !(d y))"
using d_closed neg_test test_inf_left_dist_sup by auto
also have "... = !(d x) \<sqinter> top"
by (simp add: neg_char d_closed)
finally show "!(d x) = !(d (!(d (!(d x))) \<squnion> !(d y))) \<squnion> !(d (!(d (!(d x))) \<squnion> !(d (!(d y)))))"
by simp
qed
qed
lemma d_dist_sup:
"d (x \<squnion> y) = d x \<squnion> d y"
-proof (rule antisym)
+proof (rule order.antisym)
have "x \<le> d x \<sqinter> x"
by (simp add: d1)
also have "... \<le> (d x \<squnion> d y) \<sqinter> (x \<squnion> y)"
using il_associative il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone by auto
finally have 1: "x \<le> (d x \<squnion> d y) \<sqinter> (x \<squnion> y)"
.
have "y \<le> d y \<sqinter> y"
by (simp add: d1)
also have "... \<le> (d y \<squnion> d x) \<sqinter> (y \<squnion> x)"
using il_associative il_idempotent il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone by simp
finally have "y \<le> (d x \<squnion> d y) \<sqinter> (x \<squnion> y)"
using il_commutative by auto
hence "x \<squnion> y \<le> (d x \<squnion> d y) \<sqinter> (x \<squnion> y)"
using 1 by (metis il_associative il_less_eq)
thus "d (x \<squnion> y) \<le> d x \<squnion> d y"
using llp test_sup_closed neg_test d_closed by simp
show "d x \<squnion> d y \<le> d (x \<squnion> y)"
using d_isotone_var il_associative il_commutative il_less_eq by fastforce
qed
end
class pd_semiring_extended = pd_semiring + uminus +
assumes uminus_def: "-x = !(d x)"
begin
subclass subset_boolean_algebra
by (metis subset_boolean_algebra_axioms uminus_def ext)
end
subsection \<open>Domain Semirings\<close>
class d_semiring = pd_semiring +
assumes d3: "d (x \<sqinter> d y) \<le> d (x \<sqinter> y)"
begin
lemma d3_eq: "d (x \<sqinter> d y) = d (x \<sqinter> y)"
- by (simp add: antisym d3 d3_conv)
+ by (simp add: order.antisym d3 d3_conv)
end
text \<open>
Axioms (d1), (d2) and (d3) are independent in IL-semirings.
\<close>
context il_semiring
begin
context
fixes d :: "'a \<Rightarrow> 'a"
assumes d_closed: "test (d x)"
begin
context
assumes d1: "x \<le> d x \<sqinter> x"
assumes d2: "test p \<Longrightarrow> d (p \<sqinter> x) \<le> p"
begin
lemma d3: "d (x \<sqinter> d y) \<le> d (x \<sqinter> y)" nitpick [expect=genuine] oops
end
context
assumes d1: "x \<le> d x \<sqinter> x"
assumes d3: "d (x \<sqinter> d y) \<le> d (x \<sqinter> y)"
begin
lemma d2: "test p \<Longrightarrow> d (p \<sqinter> x) \<le> p" nitpick [expect=genuine] oops
end
context
assumes d2: "test p \<Longrightarrow> d (p \<sqinter> x) \<le> p"
assumes d3: "d (x \<sqinter> d y) \<le> d (x \<sqinter> y)"
begin
lemma d1: "x \<le> d x \<sqinter> x" nitpick [expect=genuine] oops
end
end
end
class d_semiring_var = ppd_semiring +
assumes d3_var: "d (x \<sqinter> d y) \<le> d (x \<sqinter> y)"
assumes d_strict_eq_var: "d bot = bot"
begin
lemma d2_var:
assumes "test p"
shows "d (p \<sqinter> x) \<le> p"
proof -
have "!p \<sqinter> p \<sqinter> x = bot"
by (simp add: assms neg_char)
hence "d (!p \<sqinter> p \<sqinter> x) = bot"
by (simp add: d_strict_eq_var)
hence "d (!p \<sqinter> d (p \<sqinter> x)) = bot"
by (metis d3_var il_inf_associative less_eq_bot)
hence "!p \<sqinter> d (p \<sqinter> x) = bot"
using d_bot_only by blast
thus ?thesis
by (metis (no_types, hide_lams) assms d_sub_identity il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone neg_char)
qed
subclass d_semiring
proof
show "\<And>p x. test p \<Longrightarrow> d (p \<sqinter> x) \<le> p"
by (simp add: d2_var)
show "\<And>x y. d (x \<sqinter> d y) \<le> d (x \<sqinter> y)"
by (simp add: d3_var)
qed
end
section \<open>Antidomain Semirings\<close>
text \<open>
We now develop prepreantidomain semirings, preantidomain semirings and antidomain semirings.
See \cite{DesharnaisStruth2008b,DesharnaisStruth2008a,DesharnaisStruth2011} for related work on internal axioms for antidomain.
\<close>
subsection \<open>Prepreantidomain Semirings\<close>
text \<open>Definition 20\<close>
class ppa_semiring = il_semiring + uminus +
assumes a_inf_complement_bot: "-x \<sqinter> x = bot"
assumes a_stone[simp]: "-x \<squnion> --x = top"
begin
text \<open>Theorem 21\<close>
lemma l1:
"-top = bot"
by (metis a_inf_complement_bot il_inf_right_unit)
lemma l2:
"-bot = top"
by (metis l1 a_stone il_unit_bot)
lemma l3:
"-x \<le> -y \<Longrightarrow> -x \<sqinter> y = bot"
by (metis a_inf_complement_bot il_bot_unit il_inf_right_dist_sup il_less_eq)
lemma l5:
"--x \<le> --y \<Longrightarrow> -y \<le> -x"
by (metis (mono_tags, hide_lams) l3 a_stone bot_least il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone sup_right_isotone)
lemma l4:
"---x = -x"
- by (metis l5 a_inf_complement_bot a_stone antisym bot_least il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone il_unit_bot)
+ by (metis l5 a_inf_complement_bot a_stone order.antisym bot_least il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone il_unit_bot)
lemma l6:
"-x \<sqinter> --x = bot"
by (metis l3 l5 a_inf_complement_bot a_stone il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_less_eq il_sub_inf_right_isotone il_unit_bot)
lemma l7:
"-x \<sqinter> -y = -y \<sqinter> -x"
using l6 a_inf_complement_bot a_stone test_inf_commutative by blast
lemma l8:
"x \<le> --x \<sqinter> x"
by (metis a_inf_complement_bot a_stone il_idempotent il_inf_left_unit il_inf_right_dist_sup il_less_eq il_unit_bot)
sublocale ppa_ppd: ppd_semiring where d = "\<lambda>x . --x"
proof
show "\<And>x. test (- - x)"
using l4 l6 by force
show "\<And>x. x \<le> - - x \<sqinter> x"
by (simp add: l8)
qed
(*
The following statements have counterexamples, but they take a while to find.
lemma "- x = - (- - x \<squnion> - y) \<squnion> - (- - x \<squnion> - - y)" nitpick [card=8, expect=genuine] oops
lemma "- x \<squnion> - y = - - (- x \<squnion> - y)" nitpick [card=8, expect=genuine] oops
*)
end
subsection \<open>Preantidomain Semirings\<close>
text \<open>Definition 22\<close>
class pa_semiring = ppa_semiring +
assumes pad2: "--x \<le> -(-x \<sqinter> y)"
begin
text \<open>Theorem 23\<close>
lemma l10:
"-x \<sqinter> y = bot \<Longrightarrow> -x \<le> -y"
by (metis a_stone il_inf_left_unit il_inf_right_dist_sup il_unit_bot l4 pad2)
lemma l10_iff:
"-x \<sqinter> y = bot \<longleftrightarrow> -x \<le> -y"
using l10 l3 by blast
lemma l13:
"--(--x \<sqinter> y) \<le> --x"
by (metis l4 l5 pad2)
lemma l14:
"-(x \<sqinter> --y) \<le> -(x \<sqinter> y)"
by (metis il_inf_associative l4 pad2 ppa_ppd.d1_eq)
lemma l9:
"x \<le> y \<Longrightarrow> -y \<le> -x"
by (metis l10 a_inf_complement_bot il_commutative il_less_eq il_sub_inf_right_isotone il_unit_bot)
lemma l11:
"- x \<squnion> - y = - (- - x \<sqinter> - - y)"
proof -
have 1: "\<And>x y . x \<le> y \<longleftrightarrow> x \<squnion> y = y"
by (simp add: il_less_eq)
have 4: "\<And>x y . \<not>(x \<le> y) \<or> x \<squnion> y = y"
using 1 by metis
have 5: "\<And>x y z . (x \<sqinter> y) \<squnion> (x \<sqinter> z) \<le> x \<sqinter> (y \<squnion> z)"
by (simp add: il_sub_inf_right_isotone_var)
have 6: "\<And>x y . - - x \<le> - (- x \<sqinter> y)"
by (simp add: pad2)
have 7: "\<And>x y z . x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
by (simp add: il_associative)
have 8: "\<And>x y z . (x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
using 7 by metis
have 9: "\<And>x y . x \<squnion> y = y \<squnion> x"
by (simp add: il_commutative)
have 10: "\<And>x . x \<squnion> bot = x"
by (simp add: il_bot_unit)
have 11: "\<And>x . x \<squnion> x = x"
by simp
have 12: "\<And>x y z . x \<sqinter> (y \<sqinter> z) = (x \<sqinter> y) \<sqinter> z"
by (simp add: il_inf_associative)
have 13: "\<And>x y z . (x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
using 12 by metis
have 14: "\<And>x . top \<sqinter> x = x"
by simp
have 15: "\<And>x . x \<sqinter> top = x"
by simp
have 16: "\<And>x y z . (x \<squnion> y) \<sqinter> z = (x \<sqinter> z) \<squnion> (y \<sqinter> z)"
by (simp add: il_inf_right_dist_sup)
have 17: "\<And>x y z . (x \<sqinter> y) \<squnion> (z \<sqinter> y) = (x \<squnion> z) \<sqinter> y"
using 16 by metis
have 18: "\<And>x . bot \<sqinter> x = bot"
by simp
have 19: "\<And>x . - x \<squnion> - - x = top"
by simp
have 20: "\<And>x . - x \<sqinter> x = bot"
by (simp add: a_inf_complement_bot)
have 23: "\<And>x y z . ((x \<sqinter> y) \<squnion> (x \<sqinter> z)) \<squnion> (x \<sqinter> (y \<squnion> z)) = x \<sqinter> (y \<squnion> z)"
using 4 5 by metis
have 24: "\<And>x y z . (x \<sqinter> (y \<squnion> z)) \<squnion> ((x \<sqinter> y) \<squnion> (x \<sqinter> z)) = x \<sqinter> (y \<squnion> z)"
using 9 23 by metis
have 25: "\<And>x y . - - x \<squnion> - (- x \<sqinter> y) = - (- x \<sqinter> y)"
using 4 6 by metis
have 26: "\<And>x y z . x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
using 8 9 by metis
have 27: "\<And>x y z . (x \<sqinter> y) \<squnion> ((x \<sqinter> z) \<squnion> (x \<sqinter> (y \<squnion> z))) = x \<sqinter> (y \<squnion> z)"
using 9 24 26 by metis
have 30: "\<And>x . bot \<squnion> x = x"
using 9 10 by metis
have 31: "\<And>x y . x \<squnion> (x \<squnion> y) = x \<squnion> y"
using 8 11 by metis
have 34: "\<And>u x y z . ((x \<squnion> y) \<sqinter> z) \<squnion> u = (x \<sqinter> z) \<squnion> ((y \<sqinter> z) \<squnion> u)"
using 8 17 by metis
have 35: "\<And>u x y z . (x \<sqinter> (y \<sqinter> z)) \<squnion> (u \<sqinter> z) = ((x \<sqinter> y) \<squnion> u) \<sqinter> z"
using 13 17 by metis
have 36: "\<And>u x y z . (x \<sqinter> y) \<squnion> (z \<sqinter> (u \<sqinter> y)) = (x \<squnion> (z \<sqinter> u)) \<sqinter> y"
using 13 17 by metis
have 39: "\<And>x y . - x \<squnion> (- - x \<squnion> y) = top \<squnion> y"
using 8 19 by metis
have 41: "\<And>x y . - x \<sqinter> (x \<sqinter> y) = bot"
using 13 18 20 by metis
have 42: "- top = bot"
using 15 20 by metis
have 43: "\<And>x y . (- x \<squnion> y) \<sqinter> x = y \<sqinter> x"
using 17 20 30 by metis
have 44: "\<And>x y . (x \<squnion> - y) \<sqinter> y = x \<sqinter> y"
using 9 17 20 30 by metis
have 46: "\<And>x . - bot \<squnion> - - x = - bot"
using 9 20 25 by metis
have 50: "- bot = top"
using 19 30 42 by metis
have 51: "\<And>x . top \<squnion> - - x = top"
using 46 50 by metis
have 63: "\<And>x y . x \<squnion> ((x \<sqinter> - y) \<squnion> (x \<sqinter> - - y)) = x"
using 9 15 19 26 27 by metis
have 66: "\<And>x y . (- (x \<squnion> y) \<sqinter> x) \<squnion> (- (x \<squnion> y) \<sqinter> y) = bot"
using 9 20 27 30 by metis
have 67: "\<And>x y z . (x \<sqinter> - - y) \<squnion> (x \<sqinter> - (- y \<sqinter> z)) = x \<sqinter> - (- y \<sqinter> z)"
using 11 25 27 by metis
have 70: "\<And>x y . x \<squnion> (x \<sqinter> - - y) = x"
using 9 15 27 31 51 by metis
have 82: "\<And>x . top \<squnion> - x = top"
using 9 19 31 by metis
have 89: "\<And>x y . x \<squnion> (- y \<sqinter> x) = x"
using 14 17 82 by metis
have 102: "\<And>x y z . x \<squnion> (y \<squnion> (x \<sqinter> - - z)) = y \<squnion> x"
using 26 70 by metis
have 104: "\<And>x y . x \<squnion> (x \<sqinter> - y) = x"
using 9 63 102 by metis
have 112: "\<And>x y z . (- x \<sqinter> y) \<squnion> ((- - x \<sqinter> y) \<squnion> z) = y \<squnion> z"
using 14 19 34 by metis
have 117: "\<And>x y z . x \<squnion> ((x \<sqinter> - y) \<squnion> z) = x \<squnion> z"
using 8 104 by metis
have 120: "\<And>x y z . x \<squnion> (y \<squnion> (x \<sqinter> - z)) = y \<squnion> x"
using 26 104 by metis
have 124: "\<And>x . - - x \<sqinter> x = x"
using 14 19 43 by metis
have 128: "\<And>x y . - - x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
using 13 124 by metis
have 131: "\<And>x . - x \<squnion> - - - x = - x"
using 9 25 124 by metis
have 133: "\<And>x . - - - x = - x"
using 9 104 124 131 by metis
have 135: "\<And>x y . - x \<squnion> - (- - x \<sqinter> y) = - (- - x \<sqinter> y)"
using 25 133 by metis
have 137: "\<And>x y . (- x \<squnion> y) \<sqinter> - - x = y \<sqinter> - - x"
using 43 133 by metis
have 145: "\<And>x y z . ((- (x \<sqinter> y) \<sqinter> x) \<squnion> z) \<sqinter> y = z \<sqinter> y"
using 20 30 35 by metis
have 183: "\<And>x y z . (x \<squnion> (- - (y \<sqinter> z) \<sqinter> y)) \<sqinter> z = (x \<squnion> y) \<sqinter> z"
using 17 36 124 by metis
have 289: "\<And>x y . - x \<squnion> - (- x \<sqinter> y) = top"
using 25 39 82 by metis
have 316: "\<And>x y . - (- x \<sqinter> y) \<sqinter> x = x"
using 14 43 289 by metis
have 317: "\<And>x y z . - (- x \<sqinter> y) \<sqinter> (x \<sqinter> z) = x \<sqinter> z"
using 13 316 by metis
have 320: "\<And>x y . - x \<squnion> - - (- x \<sqinter> y) = - x"
using 9 25 316 by metis
have 321: "\<And>x y . - - (- x \<sqinter> y) \<sqinter> x = bot"
using 41 316 by metis
have 374: "\<And>x y . - x \<squnion> - (x \<sqinter> y) = - (x \<sqinter> y)"
using 25 128 133 by metis
have 388: "\<And>x y . - (x \<sqinter> y) \<sqinter> - x = - x"
using 128 316 by metis
have 389: "\<And>x y . - - (x \<sqinter> y) \<sqinter> - x = bot"
using 128 321 by metis
have 405: "\<And>x y z . - (x \<sqinter> y) \<sqinter> (- x \<sqinter> z) = - x \<sqinter> z"
using 13 388 by metis
have 406: "\<And>x y z . - (x \<sqinter> (y \<sqinter> z)) \<sqinter> - (x \<sqinter> y) = - (x \<sqinter> y)"
using 13 388 by metis
have 420: "\<And>x y . - x \<sqinter> - - (- x \<sqinter> y) = - - (- x \<sqinter> y)"
using 316 388 by metis
have 422: "\<And>x y z . - - (x \<sqinter> y) \<sqinter> (- x \<sqinter> z) = bot"
using 13 18 389 by metis
have 758: "\<And>x y z . x \<squnion> (x \<sqinter> (- y \<sqinter> - z)) = x"
using 13 104 117 by metis
have 1092: "\<And>x y . - (x \<squnion> y) \<sqinter> x = bot"
using 9 30 31 66 by metis
have 1130: "\<And>x y z . (- (x \<squnion> y) \<squnion> z) \<sqinter> x = z \<sqinter> x"
using 17 30 1092 by metis
have 1156: "\<And>x y . - - x \<sqinter> - (- x \<sqinter> y) = - - x"
using 67 104 124 133 by metis
have 2098: "\<And>x y . - - (x \<squnion> y) \<sqinter> x = x"
using 14 19 1130 by metis
have 2125: "\<And>x y . - - (x \<squnion> y) \<sqinter> y = y"
using 9 2098 by metis
have 2138: "\<And>x y . - x \<squnion> - - (x \<squnion> y) = top"
using 9 289 2098 by metis
have 2139: "\<And>x y . - x \<sqinter> - (x \<squnion> y) = - (x \<squnion> y)"
using 316 2098 by metis
have 2192: "\<And>x y . - - x \<sqinter> (- y \<sqinter> x) = - y \<sqinter> x"
using 89 2125 by metis
have 2202: "\<And>x y . - x \<squnion> - - (y \<squnion> x) = top"
using 9 289 2125 by metis
have 2344: "\<And>x y . - (- x \<sqinter> y) \<squnion> - - y = top"
using 89 2202 by metis
have 2547: "\<And>x y z . - x \<squnion> ((- - x \<sqinter> - y) \<squnion> z) = - x \<squnion> (- y \<squnion> z)"
using 112 117 by metis
have 3023: "\<And>x y . - x \<squnion> - (- y \<sqinter> - x) = top"
using 9 133 2344 by metis
have 3134: "\<And>x y . - (- x \<sqinter> - y) \<sqinter> y = y"
using 14 43 3023 by metis
have 3135: "\<And>x y . - x \<sqinter> (- y \<sqinter> - x) = - y \<sqinter> - x"
using 14 44 3023 by metis
have 3962: "\<And>x y . - - (x \<squnion> y) \<sqinter> - - x = - - x"
using 14 137 2138 by metis
have 5496: "\<And>x y z . - - (x \<sqinter> y) \<sqinter> - (x \<squnion> z) = bot"
using 422 2139 by metis
have 9414: "\<And>x y . - - (- x \<sqinter> y) \<sqinter> y = - x \<sqinter> y"
using 9 104 183 320 by metis
have 9520: "\<And>x y z . - - (- x \<sqinter> y) \<sqinter> - - (x \<sqinter> z) = bot"
using 374 5496 by metis
have 11070: "\<And>x y z . - (- - x \<sqinter> y) \<squnion> (- x \<sqinter> - z) = - (- - x \<sqinter> y)"
using 317 758 by metis
have 12371: "\<And>x y . - x \<sqinter> - (- - x \<sqinter> y) = - x"
using 133 1156 by metis
have 12377: "\<And>x y . - x \<sqinter> - (x \<sqinter> y) = - x"
using 128 133 1156 by metis
have 12384: "\<And>x y . - (x \<squnion> y) \<sqinter> - y = - (x \<squnion> y)"
using 133 1156 2125 by metis
have 12394: "\<And>x y . - - (- x \<sqinter> - y) = - x \<sqinter> - y"
using 1156 3134 9414 by metis
have 12640: "\<And>x y . - x \<sqinter> - (- y \<sqinter> x) = - x"
using 89 12384 by metis
have 24648: "\<And>x y . (- x \<sqinter> - y) \<squnion> - (- x \<sqinter> - y) = top"
using 19 12394 by metis
have 28270: "\<And>x y z . - - (x \<sqinter> y) \<squnion> - (- x \<sqinter> z) = - (- x \<sqinter> z)"
using 374 405 by metis
have 28339: "\<And>x y . - (- - (x \<sqinter> y) \<sqinter> x) = - (x \<sqinter> y)"
using 124 406 12371 by metis
have 28423: "\<And>x y . - (- x \<sqinter> - y) = - (- y \<sqinter> - x)"
using 13 3135 12394 28339 by metis
have 28487: "\<And>x y . - x \<sqinter> - y = - y \<sqinter> - x"
using 2098 3962 12394 28423 by metis
have 52423: "\<And>x y . - (- x \<sqinter> - (- x \<sqinter> y)) \<sqinter> y = y"
using 14 145 24648 28487 by metis
have 52522: "\<And>x y . - x \<sqinter> - (- x \<sqinter> y) = - x \<sqinter> - y"
using 13 12377 12394 12640 28487 52423 by metis
have 61103: "\<And>x y z . - (- - x \<sqinter> y) \<squnion> z = - x \<squnion> (- y \<squnion> z)"
using 112 2547 12371 52522 by metis
have 61158: "\<And>x y . - - (- x \<sqinter> y) = - x \<sqinter> - - y"
using 420 52522 by metis
have 61231: "\<And>x y z . - x \<sqinter> (- - y \<sqinter> - (x \<sqinter> z)) = - x \<sqinter> - - y"
using 13 15 50 133 9520 52522 61158 by metis
have 61313: "\<And>x y . - x \<squnion> - y = - (- - y \<sqinter> x)"
using 120 11070 61103 by metis
have 61393: "\<And>x y . - (- x \<sqinter> - - y) = - (- x \<sqinter> y)"
using 13 28270 61158 61231 61313 by metis
have 61422: "\<And>x y . - (- - x \<sqinter> y) = - (- - y \<sqinter> x)"
using 13 135 2192 61158 61313 by metis
show ?thesis
using 61313 61393 61422 by metis
qed
lemma l12:
"- x \<sqinter> - y = - (x \<squnion> y)"
proof -
have 1: "\<And>x y . x \<le> y \<longleftrightarrow> x \<squnion> y = y"
by (simp add: il_less_eq)
have 4: "\<And>x y . \<not>(x \<le> y) \<or> x \<squnion> y = y"
using 1 by metis
have 5: "\<And>x y z . (x \<sqinter> y) \<squnion> (x \<sqinter> z) \<le> x \<sqinter> (y \<squnion> z)"
by (simp add: il_sub_inf_right_isotone_var)
have 6: "\<And>x y . - - x \<le> - (- x \<sqinter> y)"
by (simp add: pad2)
have 7: "\<And>x y z . x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
by (simp add: il_associative)
have 8: "\<And>x y z . (x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
using 7 by metis
have 9: "\<And>x y . x \<squnion> y = y \<squnion> x"
by (simp add: il_commutative)
have 10: "\<And>x . x \<squnion> bot = x"
by (simp add: il_bot_unit)
have 11: "\<And>x . x \<squnion> x = x"
by simp
have 12: "\<And>x y z . x \<sqinter> (y \<sqinter> z) = (x \<sqinter> y) \<sqinter> z"
by (simp add: il_inf_associative)
have 13: "\<And>x y z . (x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
using 12 by metis
have 14: "\<And>x . top \<sqinter> x = x"
by simp
have 15: "\<And>x . x \<sqinter> top = x"
by simp
have 16: "\<And>x y z . (x \<squnion> y) \<sqinter> z = (x \<sqinter> z) \<squnion> (y \<sqinter> z)"
by (simp add: il_inf_right_dist_sup)
have 17: "\<And>x y z . (x \<sqinter> y) \<squnion> (z \<sqinter> y) = (x \<squnion> z) \<sqinter> y"
using 16 by metis
have 18: "\<And>x . bot \<sqinter> x = bot"
by simp
have 19: "\<And>x . - x \<squnion> - - x = top"
by simp
have 20: "\<And>x . - x \<sqinter> x = bot"
by (simp add: a_inf_complement_bot)
have 22: "\<And>x y z . ((x \<sqinter> y) \<squnion> (x \<sqinter> z)) \<squnion> (x \<sqinter> (y \<squnion> z)) = x \<sqinter> (y \<squnion> z)"
using 4 5 by metis
have 23: "\<And>x y z . (x \<sqinter> (y \<squnion> z)) \<squnion> ((x \<sqinter> y) \<squnion> (x \<sqinter> z)) = x \<sqinter> (y \<squnion> z)"
using 9 22 by metis
have 24: "\<And>x y . - - x \<squnion> - (- x \<sqinter> y) = - (- x \<sqinter> y)"
using 4 6 by metis
have 25: "\<And>x y z . x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
using 8 9 by metis
have 26: "\<And>x y z . (x \<sqinter> y) \<squnion> ((x \<sqinter> z) \<squnion> (x \<sqinter> (y \<squnion> z))) = x \<sqinter> (y \<squnion> z)"
using 9 23 25 by metis
have 29: "\<And>x . bot \<squnion> x = x"
using 9 10 by metis
have 30: "\<And>x y . x \<squnion> (x \<squnion> y) = x \<squnion> y"
using 8 11 by metis
have 32: "\<And>x y . x \<squnion> (y \<squnion> x) = y \<squnion> x"
using 8 9 11 by metis
have 33: "\<And>u x y z . ((x \<squnion> y) \<sqinter> z) \<squnion> u = (x \<sqinter> z) \<squnion> ((y \<sqinter> z) \<squnion> u)"
using 8 17 by metis
have 34: "\<And>u x y z . (x \<sqinter> (y \<sqinter> z)) \<squnion> (u \<sqinter> z) = ((x \<sqinter> y) \<squnion> u) \<sqinter> z"
using 13 17 by metis
have 35: "\<And>u x y z . (x \<sqinter> y) \<squnion> (z \<sqinter> (u \<sqinter> y)) = (x \<squnion> (z \<sqinter> u)) \<sqinter> y"
using 13 17 by metis
have 36: "\<And>x y . (top \<squnion> x) \<sqinter> y = y \<squnion> (x \<sqinter> y)"
using 14 17 by metis
have 37: "\<And>x y . (x \<squnion> top) \<sqinter> y = y \<squnion> (x \<sqinter> y)"
using 9 14 17 by metis
have 38: "\<And>x y . - x \<squnion> (- - x \<squnion> y) = top \<squnion> y"
using 8 19 by metis
have 40: "\<And>x y . - x \<sqinter> (x \<sqinter> y) = bot"
using 13 18 20 by metis
have 41: "- top = bot"
using 15 20 by metis
have 42: "\<And>x y . (- x \<squnion> y) \<sqinter> x = y \<sqinter> x"
using 17 20 29 by metis
have 43: "\<And>x y . (x \<squnion> - y) \<sqinter> y = x \<sqinter> y"
using 9 17 20 29 by metis
have 45: "\<And>x . - bot \<squnion> - - x = - bot"
using 9 20 24 by metis
have 46: "\<And>u x y z . (x \<sqinter> y) \<squnion> (z \<squnion> (u \<sqinter> y)) = z \<squnion> ((x \<squnion> u) \<sqinter> y)"
using 17 25 by metis
have 47: "\<And>x y . - x \<squnion> (y \<squnion> - - x) = y \<squnion> top"
using 19 25 by metis
have 49: "- bot = top"
using 19 29 41 by metis
have 50: "\<And>x . top \<squnion> - - x = top"
using 45 49 by metis
have 54: "\<And>u x y z . (x \<sqinter> y) \<squnion> ((x \<sqinter> z) \<squnion> ((x \<sqinter> (y \<squnion> z)) \<squnion> u)) = (x \<sqinter> (y \<squnion> z)) \<squnion> u"
using 8 26 by metis
have 58: "\<And>u x y z . (x \<sqinter> (y \<sqinter> z)) \<squnion> ((x \<sqinter> (y \<sqinter> u)) \<squnion> (x \<sqinter> (y \<sqinter> (z \<squnion> u)))) = x \<sqinter> (y \<sqinter> (z \<squnion> u))"
using 13 26 by metis
have 60: "\<And>x y . x \<squnion> ((x \<sqinter> y) \<squnion> (x \<sqinter> (y \<squnion> top))) = x \<sqinter> (y \<squnion> top)"
using 15 25 26 by metis
have 62: "\<And>x y . x \<squnion> ((x \<sqinter> - y) \<squnion> (x \<sqinter> - - y)) = x"
using 9 15 19 25 26 by metis
have 65: "\<And>x y . (- (x \<squnion> y) \<sqinter> x) \<squnion> (- (x \<squnion> y) \<sqinter> y) = bot"
using 9 20 26 29 by metis
have 66: "\<And>x y z . (x \<sqinter> - - y) \<squnion> (x \<sqinter> - (- y \<sqinter> z)) = x \<sqinter> - (- y \<sqinter> z)"
using 11 24 26 by metis
have 69: "\<And>x y . x \<squnion> (x \<sqinter> - - y) = x"
using 9 15 26 30 50 by metis
have 81: "\<And>x . top \<squnion> - x = top"
using 9 19 30 by metis
have 82: "\<And>x y z . (x \<sqinter> y) \<squnion> (x \<sqinter> (y \<squnion> z)) = x \<sqinter> (y \<squnion> z)"
using 11 26 30 by metis
have 83: "\<And>x y . x \<squnion> (x \<sqinter> (y \<squnion> top)) = x \<sqinter> (y \<squnion> top)"
using 60 82 by metis
have 88: "\<And>x y . x \<squnion> (- y \<sqinter> x) = x"
using 14 17 81 by metis
have 89: "\<And>x y . top \<squnion> (x \<squnion> - y) = x \<squnion> top"
using 25 81 by metis
have 91: "\<And>x y z . x \<squnion> (y \<squnion> (z \<squnion> x)) = y \<squnion> (z \<squnion> x)"
using 8 32 by metis
have 94: "\<And>x y z . x \<squnion> (y \<squnion> (- z \<sqinter> x)) = y \<squnion> x"
using 25 88 by metis
have 101: "\<And>x y z . x \<squnion> (y \<squnion> (x \<sqinter> - - z)) = y \<squnion> x"
using 25 69 by metis
have 102: "\<And>x . x \<squnion> (x \<sqinter> bot) = x"
using 41 49 69 by metis
have 103: "\<And>x y . x \<squnion> (x \<sqinter> - y) = x"
using 9 62 101 by metis
have 109: "\<And>x y . x \<squnion> (y \<squnion> (x \<sqinter> bot)) = y \<squnion> x"
using 25 102 by metis
have 111: "\<And>x y z . (- x \<sqinter> y) \<squnion> ((- - x \<sqinter> y) \<squnion> z) = y \<squnion> z"
using 14 19 33 by metis
have 116: "\<And>x y z . x \<squnion> ((x \<sqinter> - y) \<squnion> z) = x \<squnion> z"
using 8 103 by metis
have 119: "\<And>x y z . x \<squnion> (y \<squnion> (x \<sqinter> - z)) = y \<squnion> x"
using 25 103 by metis
have 123: "\<And>x . - - x \<sqinter> x = x"
using 14 19 42 by metis
have 127: "\<And>x y . - - x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
using 13 123 by metis
have 130: "\<And>x . - x \<squnion> - - - x = - x"
using 9 24 123 by metis
have 132: "\<And>x . - - - x = - x"
using 9 103 123 130 by metis
have 134: "\<And>x y . - x \<squnion> - (- - x \<sqinter> y) = - (- - x \<sqinter> y)"
using 24 132 by metis
have 136: "\<And>x y . (- x \<squnion> y) \<sqinter> - - x = y \<sqinter> - - x"
using 42 132 by metis
have 138: "\<And>x . - x \<sqinter> - x = - x"
using 123 132 by metis
have 144: "\<And>x y z . ((- (x \<sqinter> y) \<sqinter> x) \<squnion> z) \<sqinter> y = z \<sqinter> y"
using 20 29 34 by metis
have 157: "\<And>x y . (- x \<squnion> y) \<sqinter> - x = (top \<squnion> y) \<sqinter> - x"
using 17 36 138 by metis
have 182: "\<And>x y z . (x \<squnion> (- - (y \<sqinter> z) \<sqinter> y)) \<sqinter> z = (x \<squnion> y) \<sqinter> z"
using 17 35 123 by metis
have 288: "\<And>x y . - x \<squnion> - (- x \<sqinter> y) = top"
using 24 38 81 by metis
have 315: "\<And>x y . - (- x \<sqinter> y) \<sqinter> x = x"
using 14 42 288 by metis
have 316: "\<And>x y z . - (- x \<sqinter> y) \<sqinter> (x \<sqinter> z) = x \<sqinter> z"
using 13 315 by metis
have 319: "\<And>x y . - x \<squnion> - - (- x \<sqinter> y) = - x"
using 9 24 315 by metis
have 320: "\<And>x y . - - (- x \<sqinter> y) \<sqinter> x = bot"
using 40 315 by metis
have 373: "\<And>x y . - x \<squnion> - (x \<sqinter> y) = - (x \<sqinter> y)"
using 24 127 132 by metis
have 387: "\<And>x y . - (x \<sqinter> y) \<sqinter> - x = - x"
using 127 315 by metis
have 388: "\<And>x y . - - (x \<sqinter> y) \<sqinter> - x = bot"
using 127 320 by metis
have 404: "\<And>x y z . - (x \<sqinter> y) \<sqinter> (- x \<sqinter> z) = - x \<sqinter> z"
using 13 387 by metis
have 405: "\<And>x y z . - (x \<sqinter> (y \<sqinter> z)) \<sqinter> - (x \<sqinter> y) = - (x \<sqinter> y)"
using 13 387 by metis
have 419: "\<And>x y . - x \<sqinter> - - (- x \<sqinter> y) = - - (- x \<sqinter> y)"
using 315 387 by metis
have 420: "\<And>x y . - - x \<sqinter> - - (x \<sqinter> y) = - - (x \<sqinter> y)"
using 387 by metis
have 421: "\<And>x y z . - - (x \<sqinter> y) \<sqinter> (- x \<sqinter> z) = bot"
using 13 18 388 by metis
have 536: "\<And>x y . (x \<squnion> - - y) \<sqinter> y = (x \<squnion> top) \<sqinter> y"
using 42 47 by metis
have 662: "\<And>u x y z . (x \<sqinter> y) \<squnion> ((x \<sqinter> (z \<squnion> y)) \<squnion> u) = (x \<sqinter> (z \<squnion> y)) \<squnion> u"
using 9 32 54 by metis
have 705: "\<And>u x y z . (x \<sqinter> (y \<squnion> z)) \<squnion> ((x \<sqinter> (y \<squnion> (z \<sqinter> bot))) \<squnion> u) = (x \<sqinter> (y \<squnion> z)) \<squnion> u"
using 25 54 109 662 by metis
have 755: "\<And>x y z . (x \<sqinter> - y) \<squnion> (z \<squnion> x) = z \<squnion> x"
using 32 91 116 by metis
have 757: "\<And>x y z . x \<squnion> (x \<sqinter> (- y \<sqinter> - z)) = x"
using 13 103 116 by metis
have 930: "\<And>x y z . (- (x \<sqinter> (y \<squnion> z)) \<sqinter> (x \<sqinter> y)) \<squnion> (- (x \<sqinter> (y \<squnion> z)) \<sqinter> (x \<sqinter> z)) = bot"
using 9 20 29 58 by metis
have 1091: "\<And>x y . - (x \<squnion> y) \<sqinter> x = bot"
using 9 29 30 65 by metis
have 1092: "\<And>x y . - (x \<squnion> y) \<sqinter> y = bot"
using 29 30 65 1091 by metis
have 1113: "\<And>u x y z . - (x \<squnion> ((y \<squnion> z) \<sqinter> u)) \<sqinter> (x \<squnion> (z \<sqinter> u)) = bot"
using 29 46 65 1091 by metis
have 1117: "\<And>x y z . - (x \<squnion> y) \<sqinter> (x \<squnion> (- z \<sqinter> y)) = bot"
using 29 65 94 1092 by metis
have 1128: "\<And>x y z . - (x \<squnion> (y \<squnion> z)) \<sqinter> (x \<squnion> y) = bot"
using 8 1091 by metis
have 1129: "\<And>x y z . (- (x \<squnion> y) \<squnion> z) \<sqinter> x = z \<sqinter> x"
using 17 29 1091 by metis
have 1155: "\<And>x y . - - x \<sqinter> - (- x \<sqinter> y) = - - x"
using 66 103 123 132 by metis
have 1578: "\<And>x y z . - (x \<sqinter> (y \<squnion> z)) \<sqinter> (x \<sqinter> y) = bot"
using 82 1091 by metis
have 1594: "\<And>x y z . - (x \<sqinter> (y \<squnion> z)) \<sqinter> (x \<sqinter> z) = bot"
using 29 930 1578 by metis
have 2094: "\<And>x y z . - (x \<squnion> (y \<sqinter> (z \<squnion> top))) \<sqinter> (x \<squnion> y) = bot"
using 83 1128 by metis
have 2097: "\<And>x y . - - (x \<squnion> y) \<sqinter> x = x"
using 14 19 1129 by metis
have 2124: "\<And>x y . - - (x \<squnion> y) \<sqinter> y = y"
using 9 2097 by metis
have 2135: "\<And>x y . - - ((top \<squnion> x) \<sqinter> y) \<sqinter> y = y"
using 36 2097 by metis
have 2136: "\<And>x y . - - ((x \<squnion> top) \<sqinter> y) \<sqinter> y = y"
using 37 2097 by metis
have 2137: "\<And>x y . - x \<squnion> - - (x \<squnion> y) = top"
using 9 288 2097 by metis
have 2138: "\<And>x y . - x \<sqinter> - (x \<squnion> y) = - (x \<squnion> y)"
using 315 2097 by metis
have 2151: "\<And>x y . - x \<squnion> - (x \<squnion> y) = - x"
using 9 132 373 2097 by metis
have 2191: "\<And>x y . - - x \<sqinter> (- y \<sqinter> x) = - y \<sqinter> x"
using 88 2124 by metis
have 2201: "\<And>x y . - x \<squnion> - - (y \<squnion> x) = top"
using 9 288 2124 by metis
have 2202: "\<And>x y . - x \<sqinter> - (y \<squnion> x) = - (y \<squnion> x)"
using 315 2124 by metis
have 2320: "\<And>x y . - (x \<sqinter> (y \<squnion> top)) = - x"
using 83 373 2151 by metis
have 2343: "\<And>x y . - (- x \<sqinter> y) \<squnion> - - y = top"
using 88 2201 by metis
have 2546: "\<And>x y z . - x \<squnion> ((- - x \<sqinter> - y) \<squnion> z) = - x \<squnion> (- y \<squnion> z)"
using 111 116 by metis
have 2706: "\<And>x y z . - x \<squnion> (y \<squnion> - - ((top \<squnion> z) \<sqinter> - x)) = y \<squnion> - - ((top \<squnion> z) \<sqinter> - x)"
using 755 2135 by metis
have 2810: "\<And>x y . - x \<sqinter> - ((y \<squnion> top) \<sqinter> x) = - ((y \<squnion> top) \<sqinter> x)"
using 315 2136 by metis
have 3022: "\<And>x y . - x \<squnion> - (- y \<sqinter> - x) = top"
using 9 132 2343 by metis
have 3133: "\<And>x y . - (- x \<sqinter> - y) \<sqinter> y = y"
using 14 42 3022 by metis
have 3134: "\<And>x y . - x \<sqinter> (- y \<sqinter> - x) = - y \<sqinter> - x"
using 14 43 3022 by metis
have 3961: "\<And>x y . - - (x \<squnion> y) \<sqinter> - - x = - - x"
using 14 136 2137 by metis
have 4644: "\<And>x y z . - (x \<sqinter> - y) \<sqinter> (x \<sqinter> - (y \<squnion> z)) = bot"
using 1594 2151 by metis
have 5495: "\<And>x y z . - - (x \<sqinter> y) \<sqinter> - (x \<squnion> z) = bot"
using 421 2138 by metis
have 9413: "\<And>x y . - - (- x \<sqinter> y) \<sqinter> y = - x \<sqinter> y"
using 9 103 182 319 by metis
have 9519: "\<And>x y z . - - (- x \<sqinter> y) \<sqinter> - - (x \<sqinter> z) = bot"
using 373 5495 by metis
have 11069: "\<And>x y z . - (- - x \<sqinter> y) \<squnion> (- x \<sqinter> - z) = - (- - x \<sqinter> y)"
using 316 757 by metis
have 12370: "\<And>x y . - x \<sqinter> - (- - x \<sqinter> y) = - x"
using 132 1155 by metis
have 12376: "\<And>x y . - x \<sqinter> - (x \<sqinter> y) = - x"
using 127 132 1155 by metis
have 12383: "\<And>x y . - (x \<squnion> y) \<sqinter> - y = - (x \<squnion> y)"
using 132 1155 2124 by metis
have 12393: "\<And>x y . - - (- x \<sqinter> - y) = - x \<sqinter> - y"
using 1155 3133 9413 by metis
have 12407: "\<And>x y . - - x \<sqinter> - - (x \<squnion> y) = - - x"
using 1155 2138 by metis
have 12639: "\<And>x y . - x \<sqinter> - (- y \<sqinter> x) = - x"
using 88 12383 by metis
have 24647: "\<And>x y . (- x \<sqinter> - y) \<squnion> - (- x \<sqinter> - y) = top"
using 19 12393 by metis
have 28269: "\<And>x y z . - - (x \<sqinter> y) \<squnion> - (- x \<sqinter> z) = - (- x \<sqinter> z)"
using 373 404 by metis
have 28338: "\<And>x y . - (- - (x \<sqinter> y) \<sqinter> x) = - (x \<sqinter> y)"
using 123 405 12370 by metis
have 28422: "\<And>x y . - (- x \<sqinter> - y) = - (- y \<sqinter> - x)"
using 13 3134 12393 28338 by metis
have 28485: "\<And>x y . - x \<sqinter> - y = - y \<sqinter> - x"
using 2097 3961 12393 28422 by metis
have 30411: "\<And>x y . - x \<sqinter> (x \<squnion> (x \<sqinter> y)) = bot"
using 9 82 2094 2320 by metis
have 30469: "\<And>x . - x \<sqinter> (x \<squnion> - - x) = bot"
using 9 123 132 30411 by metis
have 37513: "\<And>x y . - (- x \<sqinter> - y) \<sqinter> - (y \<squnion> x) = bot"
using 2202 4644 by metis
have 52421: "\<And>x y . - (- x \<sqinter> - (- x \<sqinter> y)) \<sqinter> y = y"
using 14 144 24647 28485 by metis
have 52520: "\<And>x y . - x \<sqinter> - (- x \<sqinter> y) = - x \<sqinter> - y"
using 13 12376 12393 12639 28485 52421 by metis
have 52533: "\<And>x y z . - - (x \<squnion> (y \<sqinter> (z \<squnion> top))) \<sqinter> (x \<squnion> y) = x \<squnion> y"
using 15 49 2094 52421 by metis
have 61101: "\<And>x y z . - (- - x \<sqinter> y) \<squnion> z = - x \<squnion> (- y \<squnion> z)"
using 111 2546 12370 52520 by metis
have 61156: "\<And>x y . - - (- x \<sqinter> y) = - x \<sqinter> - - y"
using 419 52520 by metis
have 61162: "\<And>x y . - (x \<squnion> (x \<sqinter> y)) = - x"
using 15 49 2138 30411 52520 by metis
have 61163: "\<And>x . - (x \<squnion> - - x) = - x"
using 15 49 2138 30469 52520 by metis
have 61229: "\<And>x y z . - x \<sqinter> (- - y \<sqinter> - (x \<sqinter> z)) = - x \<sqinter> - - y"
using 13 15 49 132 9519 52520 61156 by metis
have 61311: "\<And>x y . - x \<squnion> - y = - (- - y \<sqinter> x)"
using 119 11069 61101 by metis
have 61391: "\<And>x y . - (- x \<sqinter> - - y) = - (- x \<sqinter> y)"
using 13 28269 61156 61229 61311 by metis
have 61420: "\<And>x y . - (- - x \<sqinter> y) = - (- - y \<sqinter> x)"
using 13 134 2191 61156 61311 by metis
have 61454: "\<And>x y . - (x \<squnion> - (- y \<sqinter> - x)) = - y \<sqinter> - x"
using 9 132 3133 61156 61162 by metis
have 61648: "\<And>x y . - x \<sqinter> (x \<squnion> (- y \<sqinter> - - x)) = bot"
using 1117 61163 by metis
have 62434: "\<And>x y . - (- - x \<sqinter> y) \<sqinter> x = - y \<sqinter> x"
using 43 61311 by metis
have 63947: "\<And>x y . - (- x \<sqinter> y) \<sqinter> - (- y \<squnion> x) = bot"
using 37513 61391 by metis
have 64227: "\<And>x y . - (x \<squnion> (- y \<sqinter> - - x)) = - x"
using 15 49 2138 52520 61648 by metis
have 64239: "\<And>x y . - (x \<squnion> (- - x \<squnion> y)) = - (x \<squnion> y)"
using 9 25 12407 64227 by metis
have 64241: "\<And>x y . - (x \<squnion> (- - x \<sqinter> - y)) = - x"
using 28485 64227 by metis
have 64260: "\<And>x y . - (x \<squnion> - - (x \<sqinter> y)) = - x"
using 420 64241 by metis
have 64271: "\<And>x y . - (- x \<squnion> (y \<squnion> - - (y \<sqinter> x))) = - (- x \<squnion> y)"
using 9 25 42 64260 by metis
have 64281: "\<And>x y . - (- x \<squnion> y) = - (y \<squnion> - - ((top \<squnion> y) \<sqinter> - x))"
using 9 25 157 2706 64260 by metis
have 64282: "\<And>x y . - (x \<squnion> - - ((x \<squnion> top) \<sqinter> y)) = - (x \<squnion> - - y)"
using 9 25 132 536 2810 28485 61311 64260 by metis
have 65110: "\<And>x y . - ((- x \<sqinter> y) \<squnion> (- y \<squnion> x)) = bot"
using 9 14 49 37513 63947 by metis
have 65231: "\<And>x y . - (x \<squnion> ((- x \<sqinter> y) \<squnion> - y)) = bot"
using 9 25 65110 by metis
have 65585: "\<And>x y . - (x \<squnion> - y) = - - y \<sqinter> - x"
using 61311 61454 64239 by metis
have 65615: "\<And>x y . - x \<sqinter> - ((x \<squnion> top) \<sqinter> y) = - y \<sqinter> - x"
using 132 28485 64282 65585 by metis
have 65616: "\<And>x y . - (- x \<squnion> y) = - y \<sqinter> - ((top \<squnion> y) \<sqinter> - x)"
using 132 28485 64281 65585 by metis
have 65791: "\<And>x y . - x \<sqinter> - ((top \<squnion> x) \<sqinter> - y) = - - y \<sqinter> - x"
using 89 132 12376 28485 64271 65585 65615 65616 by metis
have 65933: "\<And>x y . - (- x \<squnion> y) = - - x \<sqinter> - y"
using 65616 65791 by metis
have 66082: "\<And>x y z . - (x \<squnion> (y \<squnion> - z)) = - - z \<sqinter> - (x \<squnion> y)"
using 8 65585 by metis
have 66204: "\<And>x y . - - x \<sqinter> - (y \<squnion> (- y \<sqinter> x)) = bot"
using 65231 66082 by metis
have 66281: "\<And>x y z . - (x \<squnion> (- y \<squnion> z)) = - - y \<sqinter> - (x \<squnion> z)"
using 25 65933 by metis
have 67527: "\<And>x y . - - (x \<squnion> (- x \<sqinter> y)) \<sqinter> y = y"
using 14 49 62434 66204 by metis
have 67762: "\<And>x y . - (- - x \<sqinter> (y \<squnion> (- y \<sqinter> x))) = - x"
using 61420 67527 by metis
have 68018: "\<And>x y z . - (x \<squnion> y) \<sqinter> (x \<squnion> (y \<sqinter> (z \<squnion> top))) = bot"
using 8 83 1113 2320 by metis
have 71989: "\<And>x y z . - (x \<squnion> (y \<sqinter> (z \<squnion> top))) = - (x \<squnion> y)"
using 9 29 52533 67762 68018 by metis
have 71997: "\<And>x y z . - ((x \<sqinter> (y \<squnion> top)) \<squnion> z) = - (x \<squnion> z)"
using 17 2320 71989 by metis
have 72090: "\<And>x y z . - (x \<squnion> ((x \<sqinter> y) \<squnion> z)) = - (x \<squnion> z)"
using 10 14 705 71997 by metis
have 72139: "\<And>x y . - (x \<squnion> y) = - x \<sqinter> - y"
using 25 123 132 2138 65933 66281 72090 by metis
show ?thesis
using 72139 by metis
qed
lemma l15:
"--(x \<squnion> y) = --x \<squnion> --y"
by (simp add: l11 l12 l4)
lemma l13_var:
"- - (- x \<sqinter> y) = - x \<sqinter> - - y"
proof -
have 1: "\<And>x y . x \<le> y \<longleftrightarrow> x \<squnion> y = y"
by (simp add: il_less_eq)
have 4: "\<And>x y . \<not>(x \<le> y) \<or> x \<squnion> y = y"
using 1 by metis
have 5: "\<And>x y z . (x \<sqinter> y) \<squnion> (x \<sqinter> z) \<le> x \<sqinter> (y \<squnion> z)"
by (simp add: il_sub_inf_right_isotone_var)
have 6: "\<And>x y . - - x \<le> - (- x \<sqinter> y)"
by (simp add: pad2)
have 7: "\<And>x y z . x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
by (simp add: il_associative)
have 8: "\<And>x y z . (x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
using 7 by metis
have 9: "\<And>x y . x \<squnion> y = y \<squnion> x"
by (simp add: il_commutative)
have 10: "\<And>x . x \<squnion> bot = x"
by (simp add: il_bot_unit)
have 11: "\<And>x . x \<squnion> x = x"
by simp
have 12: "\<And>x y z . x \<sqinter> (y \<sqinter> z) = (x \<sqinter> y) \<sqinter> z"
by (simp add: il_inf_associative)
have 13: "\<And>x y z . (x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
using 12 by metis
have 14: "\<And>x . top \<sqinter> x = x"
by simp
have 15: "\<And>x . x \<sqinter> top = x"
by simp
have 16: "\<And>x y z . (x \<squnion> y) \<sqinter> z = (x \<sqinter> z) \<squnion> (y \<sqinter> z)"
by (simp add: il_inf_right_dist_sup)
have 17: "\<And>x y z . (x \<sqinter> y) \<squnion> (z \<sqinter> y) = (x \<squnion> z) \<sqinter> y"
using 16 by metis
have 19: "\<And>x . - x \<squnion> - - x = top"
by simp
have 20: "\<And>x . - x \<sqinter> x = bot"
by (simp add: a_inf_complement_bot)
have 22: "\<And>x y z . ((x \<sqinter> y) \<squnion> (x \<sqinter> z)) \<squnion> (x \<sqinter> (y \<squnion> z)) = x \<sqinter> (y \<squnion> z)"
using 4 5 by metis
have 23: "\<And>x y z . (x \<sqinter> (y \<squnion> z)) \<squnion> ((x \<sqinter> y) \<squnion> (x \<sqinter> z)) = x \<sqinter> (y \<squnion> z)"
using 9 22 by metis
have 24: "\<And>x y . - - x \<squnion> - (- x \<sqinter> y) = - (- x \<sqinter> y)"
using 4 6 by metis
have 25: "\<And>x y z . x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
using 8 9 by metis
have 26: "\<And>x y z . (x \<sqinter> y) \<squnion> ((x \<sqinter> z) \<squnion> (x \<sqinter> (y \<squnion> z))) = x \<sqinter> (y \<squnion> z)"
using 9 23 25 by metis
have 29: "\<And>x . bot \<squnion> x = x"
using 9 10 by metis
have 30: "\<And>x y . x \<squnion> (x \<squnion> y) = x \<squnion> y"
using 8 11 by metis
have 34: "\<And>u x y z . (x \<sqinter> (y \<sqinter> z)) \<squnion> (u \<sqinter> z) = ((x \<sqinter> y) \<squnion> u) \<sqinter> z"
using 13 17 by metis
have 35: "\<And>u x y z . (x \<sqinter> y) \<squnion> (z \<sqinter> (u \<sqinter> y)) = (x \<squnion> (z \<sqinter> u)) \<sqinter> y"
using 13 17 by metis
have 38: "\<And>x y . - x \<squnion> (- - x \<squnion> y) = top \<squnion> y"
using 8 19 by metis
have 41: "- top = bot"
using 15 20 by metis
have 42: "\<And>x y . (- x \<squnion> y) \<sqinter> x = y \<sqinter> x"
using 17 20 29 by metis
have 43: "\<And>x y . (x \<squnion> - y) \<sqinter> y = x \<sqinter> y"
using 9 17 20 29 by metis
have 45: "\<And>x . - bot \<squnion> - - x = - bot"
using 9 20 24 by metis
have 49: "- bot = top"
using 19 29 41 by metis
have 50: "\<And>x . top \<squnion> - - x = top"
using 45 49 by metis
have 62: "\<And>x y . x \<squnion> ((x \<sqinter> - y) \<squnion> (x \<sqinter> - - y)) = x"
using 9 15 19 25 26 by metis
have 65: "\<And>x y . (- (x \<squnion> y) \<sqinter> x) \<squnion> (- (x \<squnion> y) \<sqinter> y) = bot"
using 9 20 26 29 by metis
have 66: "\<And>x y z . (x \<sqinter> - - y) \<squnion> (x \<sqinter> - (- y \<sqinter> z)) = x \<sqinter> - (- y \<sqinter> z)"
using 11 24 26 by metis
have 69: "\<And>x y . x \<squnion> (x \<sqinter> - - y) = x"
using 9 15 26 30 50 by metis
have 81: "\<And>x . top \<squnion> - x = top"
using 9 19 30 by metis
have 88: "\<And>x y . x \<squnion> (- y \<sqinter> x) = x"
using 14 17 81 by metis
have 101: "\<And>x y z . x \<squnion> (y \<squnion> (x \<sqinter> - - z)) = y \<squnion> x"
using 25 69 by metis
have 103: "\<And>x y . x \<squnion> (x \<sqinter> - y) = x"
using 9 62 101 by metis
have 123: "\<And>x . - - x \<sqinter> x = x"
using 14 19 42 by metis
have 127: "\<And>x y . - - x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
using 13 123 by metis
have 130: "\<And>x . - x \<squnion> - - - x = - x"
using 9 24 123 by metis
have 132: "\<And>x . - - - x = - x"
using 9 103 123 130 by metis
have 136: "\<And>x y . (- x \<squnion> y) \<sqinter> - - x = y \<sqinter> - - x"
using 42 132 by metis
have 144: "\<And>x y z . ((- (x \<sqinter> y) \<sqinter> x) \<squnion> z) \<sqinter> y = z \<sqinter> y"
using 20 29 34 by metis
have 182: "\<And>x y z . (x \<squnion> (- - (y \<sqinter> z) \<sqinter> y)) \<sqinter> z = (x \<squnion> y) \<sqinter> z"
using 17 35 123 by metis
have 288: "\<And>x y . - x \<squnion> - (- x \<sqinter> y) = top"
using 24 38 81 by metis
have 315: "\<And>x y . - (- x \<sqinter> y) \<sqinter> x = x"
using 14 42 288 by metis
have 319: "\<And>x y . - x \<squnion> - - (- x \<sqinter> y) = - x"
using 9 24 315 by metis
have 387: "\<And>x y . - (x \<sqinter> y) \<sqinter> - x = - x"
using 127 315 by metis
have 405: "\<And>x y z . - (x \<sqinter> (y \<sqinter> z)) \<sqinter> - (x \<sqinter> y) = - (x \<sqinter> y)"
using 13 387 by metis
have 419: "\<And>x y . - x \<sqinter> - - (- x \<sqinter> y) = - - (- x \<sqinter> y)"
using 315 387 by metis
have 1091: "\<And>x y . - (x \<squnion> y) \<sqinter> x = bot"
using 9 29 30 65 by metis
have 1129: "\<And>x y z . (- (x \<squnion> y) \<squnion> z) \<sqinter> x = z \<sqinter> x"
using 17 29 1091 by metis
have 1155: "\<And>x y . - - x \<sqinter> - (- x \<sqinter> y) = - - x"
using 66 103 123 132 by metis
have 2097: "\<And>x y . - - (x \<squnion> y) \<sqinter> x = x"
using 14 19 1129 by metis
have 2124: "\<And>x y . - - (x \<squnion> y) \<sqinter> y = y"
using 9 2097 by metis
have 2137: "\<And>x y . - x \<squnion> - - (x \<squnion> y) = top"
using 9 288 2097 by metis
have 2201: "\<And>x y . - x \<squnion> - - (y \<squnion> x) = top"
using 9 288 2124 by metis
have 2343: "\<And>x y . - (- x \<sqinter> y) \<squnion> - - y = top"
using 88 2201 by metis
have 3022: "\<And>x y . - x \<squnion> - (- y \<sqinter> - x) = top"
using 9 132 2343 by metis
have 3133: "\<And>x y . - (- x \<sqinter> - y) \<sqinter> y = y"
using 14 42 3022 by metis
have 3134: "\<And>x y . - x \<sqinter> (- y \<sqinter> - x) = - y \<sqinter> - x"
using 14 43 3022 by metis
have 3961: "\<And>x y . - - (x \<squnion> y) \<sqinter> - - x = - - x"
using 14 136 2137 by metis
have 9413: "\<And>x y . - - (- x \<sqinter> y) \<sqinter> y = - x \<sqinter> y"
using 9 103 182 319 by metis
have 12370: "\<And>x y . - x \<sqinter> - (- - x \<sqinter> y) = - x"
using 132 1155 by metis
have 12376: "\<And>x y . - x \<sqinter> - (x \<sqinter> y) = - x"
using 127 132 1155 by metis
have 12383: "\<And>x y . - (x \<squnion> y) \<sqinter> - y = - (x \<squnion> y)"
using 132 1155 2124 by metis
have 12393: "\<And>x y . - - (- x \<sqinter> - y) = - x \<sqinter> - y"
using 1155 3133 9413 by metis
have 12639: "\<And>x y . - x \<sqinter> - (- y \<sqinter> x) = - x"
using 88 12383 by metis
have 24647: "\<And>x y . (- x \<sqinter> - y) \<squnion> - (- x \<sqinter> - y) = top"
using 19 12393 by metis
have 28338: "\<And>x y . - (- - (x \<sqinter> y) \<sqinter> x) = - (x \<sqinter> y)"
using 123 405 12370 by metis
have 28422: "\<And>x y . - (- x \<sqinter> - y) = - (- y \<sqinter> - x)"
using 13 3134 12393 28338 by metis
have 28485: "\<And>x y . - x \<sqinter> - y = - y \<sqinter> - x"
using 2097 3961 12393 28422 by metis
have 52421: "\<And>x y . - (- x \<sqinter> - (- x \<sqinter> y)) \<sqinter> y = y"
using 14 144 24647 28485 by metis
have 52520: "\<And>x y . - x \<sqinter> - (- x \<sqinter> y) = - x \<sqinter> - y"
using 13 12376 12393 12639 28485 52421 by metis
have 61156: "\<And>x y . - - (- x \<sqinter> y) = - x \<sqinter> - - y"
using 419 52520 by metis
show ?thesis
using 61156 by metis
qed
text \<open>Theorem 25.1\<close>
subclass subset_boolean_algebra_2
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: il_associative)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: il_commutative)
show "\<And>x. x \<squnion> x = x"
by simp
show "\<And>x y. x \<squnion> - (y \<squnion> - y) = x"
using il_bot_unit l12 l6 by auto
show "\<And>x y. - (x \<squnion> y) = - (- - x \<squnion> - - y)"
by (metis l15 l4)
show "\<And>x y. - x \<squnion> - (- x \<squnion> y) = - x \<squnion> - y"
by (smt l11 l15 il_inf_right_dist_sup il_unit_bot l6 l7)
qed
lemma aa_test:
"p = --p \<Longrightarrow> test p"
by (metis ppa_ppd.d_closed)
lemma test_aa_increasing:
"test p \<Longrightarrow> p \<le> --p"
by (simp add: ppa_ppd.d_increasing_sub_identity test_sub_identity)
lemma "test p \<Longrightarrow> - - (p \<sqinter> x) \<le> p" nitpick [expect=genuine] oops
lemma "test p \<Longrightarrow> --p \<le> p" nitpick [expect=genuine] oops
end
class pa_algebra = pa_semiring + minus +
assumes pa_minus_def: "-x - -y = -(--x \<squnion> -y)"
begin
subclass subset_boolean_algebra_2_extended
proof
show "bot = (THE x. \<forall>z. x = - (z \<squnion> - z))"
using l12 l6 by auto
thus "top = - (THE x. \<forall>z. x = - (z \<squnion> - z))"
using l2 by blast
show "\<And>x y. - x \<sqinter> - y = - (- - x \<squnion> - - y)"
by (metis l12 l4)
show "\<And>x y. - x - - y = - (- - x \<squnion> - y)"
by (simp add: pa_minus_def)
show "\<And>x y. (x \<le> y) = (x \<squnion> y = y)"
by (simp add: il_less_eq)
show "\<And>x y. (x < y) = (x \<squnion> y = y \<and> y \<squnion> x \<noteq> x)"
by (simp add: il_less_eq less_le_not_le)
qed
lemma "\<And>x y. - (x \<sqinter> - - y) = - (x \<sqinter> y)" nitpick [expect=genuine] oops
end
subsection \<open>Antidomain Semirings\<close>
text \<open>Definition 24\<close>
class a_semiring = ppa_semiring +
assumes ad3: "-(x \<sqinter> y) \<le> -(x \<sqinter> --y)"
begin
lemma l16:
"- - x \<le> - (- x \<sqinter> y)"
proof -
have 1: "\<And>x y . x \<le> y \<longleftrightarrow> x \<squnion> y = y"
by (simp add: il_less_eq)
have 3: "\<And>x y z . x \<squnion> (y \<squnion> z) = (x \<squnion> y) \<squnion> z"
by (simp add: il_associative)
have 4: "\<And>x y z . (x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
using 3 by metis
have 5: "\<And>x y . x \<squnion> y = y \<squnion> x"
by (simp add: il_commutative)
have 6: "\<And>x . x \<squnion> bot = x"
by (simp add: il_bot_unit)
have 7: "\<And>x . x \<squnion> x = x"
by simp
have 8: "\<And>x y . \<not>(x \<le> y) \<or> x \<squnion> y = y"
using 1 by metis
have 9: "\<And>x y . x \<le> y \<or> x \<squnion> y \<noteq> y"
using 1 by metis
have 10: "\<And>x y z . x \<sqinter> (y \<sqinter> z) = (x \<sqinter> y) \<sqinter> z"
by (simp add: il_inf_associative)
have 11: "\<And>x y z . (x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
using 10 by metis
have 12: "\<And>x . top \<sqinter> x = x"
by simp
have 13: "\<And>x . x \<sqinter> top = x"
by simp
have 14: "\<And>x y z . (x \<sqinter> y) \<squnion> (x \<sqinter> z) \<le> x \<sqinter> (y \<squnion> z)"
by (simp add: il_sub_inf_right_isotone_var)
have 15: "\<And>x y z . (x \<squnion> y) \<sqinter> z = (x \<sqinter> z) \<squnion> (y \<sqinter> z)"
by (simp add: il_inf_right_dist_sup)
have 16: "\<And>x y z . (x \<sqinter> y) \<squnion> (z \<sqinter> y) = (x \<squnion> z) \<sqinter> y"
using 15 by metis
have 17: "\<And>x . bot \<sqinter> x = bot"
by simp
have 18: "\<And>x . - x \<squnion> - - x = top"
by simp
have 19: "\<And>x . - x \<sqinter> x = bot"
by (simp add: a_inf_complement_bot)
have 20: "\<And>x y . - (x \<sqinter> y) \<le> - (x \<sqinter> - - y)"
by (simp add: ad3)
have 22: "\<And>x y z . x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
using 4 5 by metis
have 25: "\<And>x . bot \<squnion> x = x"
using 5 6 by metis
have 26: "\<And>x y . x \<squnion> (x \<squnion> y) = x \<squnion> y"
using 4 7 by metis
have 33: "\<And>x y z . (x \<sqinter> y) \<squnion> ((x \<sqinter> z) \<squnion> (x \<sqinter> (y \<squnion> z))) = x \<sqinter> (y \<squnion> z)"
using 5 8 14 22 by metis
have 47: "\<And>x y . - x \<squnion> (- - x \<squnion> y) = top \<squnion> y"
using 4 18 by metis
have 48: "\<And>x y . - - x \<squnion> (y \<squnion> - x) = y \<squnion> top"
using 4 5 18 by metis
have 51: "\<And>x y . - x \<sqinter> (x \<sqinter> y) = bot"
using 11 17 19 by metis
have 52: "- top = bot"
using 13 19 by metis
have 56: "\<And>x y . (- x \<squnion> y) \<sqinter> x = y \<sqinter> x"
using 16 19 25 by metis
have 57: "\<And>x y . (x \<squnion> - y) \<sqinter> y = x \<sqinter> y"
using 5 16 19 25 by metis
have 58: "\<And>x y . - (x \<sqinter> y) \<squnion> - (x \<sqinter> - - y) = - (x \<sqinter> - - y)"
using 8 20 by metis
have 60: "\<And>x . - x \<le> - - - x"
using 12 20 by metis
have 69: "- bot = top"
using 18 25 52 by metis
have 74: "\<And>x y . x \<le> x \<squnion> y"
using 9 26 by metis
have 78: "\<And>x . top \<squnion> - x = top"
using 5 18 26 by metis
have 80: "\<And>x y . x \<le> y \<squnion> x"
using 5 74 by metis
have 86: "\<And>x y z . x \<squnion> y \<le> x \<squnion> (z \<squnion> y)"
using 22 80 by metis
have 95: "\<And>x . - x \<squnion> - - - x = - - - x"
using 8 60 by metis
have 143: "\<And>x y . x \<squnion> (x \<sqinter> - y) = x"
using 5 13 26 33 78 by metis
have 370: "\<And>x y z . x \<squnion> (y \<sqinter> - z) \<le> x \<squnion> y"
using 86 143 by metis
have 907: "\<And>x . - x \<sqinter> - x = - x"
using 12 18 57 by metis
have 928: "\<And>x y . - x \<sqinter> (- x \<sqinter> y) = - x \<sqinter> y"
using 11 907 by metis
have 966: "\<And>x y . - (- x \<sqinter> - - (x \<sqinter> y)) = top"
using 51 58 69 78 by metis
have 1535: "\<And>x . - x \<squnion> - - - - x = top"
using 47 78 95 by metis
have 1630: "\<And>x y z . (x \<squnion> y) \<sqinter> - z \<le> (x \<sqinter> - z) \<squnion> y"
using 16 370 by metis
have 2422: "\<And>x . - x \<sqinter> - - - x = - - - x"
using 12 57 1535 by metis
have 6567: "\<And>x y . - x \<sqinter> - - (x \<sqinter> y) = bot"
using 12 19 966 by metis
have 18123: "\<And>x . - - - x = - x"
using 95 143 2422 by metis
have 26264: "\<And>x y . - x \<le> (- y \<sqinter> - x) \<squnion> - - y"
using 12 18 1630 by metis
have 26279: "\<And>x y . - - (x \<sqinter> y) \<le> - - x"
using 25 6567 26264 by metis
have 26307: "\<And>x y . - - (- x \<sqinter> y) \<le> - x"
using 928 18123 26279 by metis
have 26339: "\<And>x y . - x \<squnion> - - (- x \<sqinter> y) = - x"
using 5 8 26307 by metis
have 26564: "\<And>x y . - x \<squnion> - (- x \<sqinter> y) = top"
using 5 48 78 18123 26339 by metis
have 26682: "\<And>x y . - (- x \<sqinter> y) \<sqinter> x = x"
using 12 56 26564 by metis
have 26864: "\<And>x y . - - x \<le> - (- x \<sqinter> y)"
using 18123 26279 26682 by metis
show ?thesis
using 26864 by metis
qed
text \<open>Theorem 25.2\<close>
subclass pa_semiring
proof
show "\<And>x y. - - x \<le> - (- x \<sqinter> y)"
by (rule l16)
qed
lemma l17:
"-(x \<sqinter> y) = -(x \<sqinter> --y)"
- by (simp add: ad3 antisym l14)
+ by (simp add: ad3 order.antisym l14)
lemma a_complement_inf_double_complement:
"-(x \<sqinter> --y) = -(x \<sqinter> y)"
using l17 by auto
sublocale a_d: d_semiring_var where d = "\<lambda>x . --x"
proof
show "\<And>x y. - - (x \<sqinter> - - y) \<le> - - (x \<sqinter> y)"
using l17 by auto
show "- - bot = bot"
by (simp add: l1 l2)
qed
lemma "test p \<Longrightarrow> - - (p \<sqinter> x) \<le> p"
by (fact a_d.d2)
end
class a_algebra = a_semiring + minus +
assumes a_minus_def: "-x - -y = -(--x \<squnion> -y)"
begin
subclass pa_algebra
proof
show "\<And>x y. - x - - y = - (- - x \<squnion> - y)"
by (simp add: a_minus_def)
qed
text \<open>Theorem 25.4\<close>
subclass subset_boolean_algebra_4_extended
proof
show "\<And>x y z. x \<sqinter> (y \<sqinter> z) = x \<sqinter> y \<sqinter> z"
by (simp add: il_inf_associative)
show "\<And>x y z. (x \<squnion> y) \<sqinter> z = x \<sqinter> z \<squnion> y \<sqinter> z"
by (simp add: il_inf_right_dist_sup)
show "\<And>x. - x \<sqinter> x = bot"
by (simp add: a_inf_complement_bot)
show "\<And>x. top \<sqinter> x = x"
by simp
show "\<And>x y. - (x \<sqinter> - - y) = - (x \<sqinter> y)"
using l17 by auto
show "\<And>x. x \<sqinter> top = x"
by simp
show "\<And>x y z. x \<le> y \<Longrightarrow> z \<sqinter> x \<le> z \<sqinter> y"
by (simp add: il_sub_inf_right_isotone)
qed
end
context subset_boolean_algebra_4_extended
begin
subclass il_semiring
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: sup_assoc)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: sup_commute)
show "\<And>x. x \<squnion> x = x"
by simp
show "\<And>x. x \<squnion> bot = x"
by simp
show "\<And>x y z. x \<sqinter> (y \<sqinter> z) = x \<sqinter> y \<sqinter> z"
by (simp add: sba3_inf_associative)
show "\<And>x y z. (x \<squnion> y) \<sqinter> z = x \<sqinter> z \<squnion> y \<sqinter> z"
by (simp add: sba3_inf_right_dist_sup)
show "\<And>x. top \<sqinter> x = x"
by simp
show "\<And>x. x \<sqinter> top = x"
by simp
show "\<And>x. bot \<sqinter> x = bot"
by (simp add: inf_left_zero)
show "\<And>x y z. x \<le> y \<Longrightarrow> z \<sqinter> x \<le> z \<sqinter> y"
by (simp add: inf_right_isotone)
show "\<And>x y. (x \<le> y) = (x \<squnion> y = y)"
by (simp add: le_iff_sup)
show "\<And>x y. (x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (simp add: less_le_not_le)
qed
subclass a_semiring
proof
show "\<And>x. - x \<sqinter> x = bot"
by (simp add: sba3_inf_complement_bot)
show "\<And>x. - x \<squnion> - - x = top"
by simp
show "\<And>x y. - (x \<sqinter> y) \<le> - (x \<sqinter> - - y)"
by (simp add: sba3_complement_inf_double_complement)
qed
sublocale sba4_a: a_algebra
proof
show "\<And>x y. - x - - y = - (- - x \<squnion> - y)"
by (simp add: sub_minus_def)
qed
end
context stone_algebra
begin
text \<open>Theorem 25.3\<close>
subclass il_semiring
proof
show "\<And>x y z. x \<squnion> (y \<squnion> z) = x \<squnion> y \<squnion> z"
by (simp add: sup_assoc)
show "\<And>x y. x \<squnion> y = y \<squnion> x"
by (simp add: sup_commute)
show "\<And>x. x \<squnion> x = x"
by simp
show "\<And>x. x \<squnion> bot = x"
by simp
show "\<And>x y z. x \<sqinter> (y \<sqinter> z) = x \<sqinter> y \<sqinter> z"
by (simp add: inf.sup_monoid.add_assoc)
show "\<And>x y z. (x \<squnion> y) \<sqinter> z = x \<sqinter> z \<squnion> y \<sqinter> z"
by (simp add: inf_sup_distrib2)
show "\<And>x. top \<sqinter> x = x"
by simp
show "\<And>x. x \<sqinter> top = x"
by simp
show "\<And>x. bot \<sqinter> x = bot"
by simp
show "\<And>x y z. x \<le> y \<Longrightarrow> z \<sqinter> x \<le> z \<sqinter> y"
using inf.sup_right_isotone by blast
show "\<And>x y. (x \<le> y) = (x \<squnion> y = y)"
by (simp add: le_iff_sup)
show "\<And>x y. (x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (simp add: less_le_not_le)
qed
subclass a_semiring
proof
show "\<And>x. - x \<sqinter> x = bot"
by simp
show "\<And>x. - x \<squnion> - - x = top"
by simp
show "\<And>x y. - (x \<sqinter> y) \<le> - (x \<sqinter> - - y)"
by simp
qed
end
end
diff --git a/thys/SuperCalc/multisets_continued.thy b/thys/SuperCalc/multisets_continued.thy
--- a/thys/SuperCalc/multisets_continued.thy
+++ b/thys/SuperCalc/multisets_continued.thy
@@ -1,441 +1,441 @@
theory multisets_continued
(* N. Peltier - http://membres-lig.imag.fr/peltier/ *)
imports Main "HOL-Library.Multiset"
begin
subsection \<open>Multisets\<close>
text \<open>We use the Multiset theory provided in Isabelle. We prove some additional
(mostly trivial) lemmata.\<close>
lemma mset_set_inclusion:
assumes "finite E2"
assumes "E1 \<subset> E2"
shows "mset_set E1 \<subset># (mset_set E2)"
proof (rule ccontr)
let ?s1 = "mset_set E1"
let ?s2 = "mset_set E2"
assume "\<not> ?s1 \<subset># ?s2"
from assms(1) and assms(2) have "finite E1" using finite_subset less_imp_le by auto
from \<open>\<not> ?s1 \<subset># ?s2\<close> obtain x where "(count ?s1 x > count ?s2 x)" using subseteq_mset_def [of ?s1 ?s2]
by (metis assms(1) assms(2) finite_set_mset_mset_set finite_subset less_imp_le
less_le not_le_imp_less subset_mset.le_less)
from this have "count ?s1 x > 0" by linarith
from this and \<open>finite E1\<close> have "count ?s1 x = 1" and "x \<in> E1" using subseteq_mset_def by auto
from this and assms(2) have "x \<in> E2" by auto
from this and \<open>finite E2\<close> have "count ?s2 x = 1" by auto
from this and \<open>count ?s1 x = 1\<close> and \<open>(count ?s1 x > count ?s2 x)\<close> show False by auto
qed
lemma mset_ordering_addition:
assumes "A = B + C"
shows "B \<subseteq># A"
using assms by simp
lemma equal_image_mset:
assumes "\<forall>x \<in> E. (f x) = (g x)"
shows "{# (f x). x \<in># (mset_set E) #} = {# (g x). x \<in># (mset_set E) #}"
by (meson assms count_eq_zero_iff count_mset_set(3) image_mset_cong)
lemma multiset_order_inclusion:
assumes "E \<subset># F"
assumes "trans r"
shows "(E,F) \<in> (mult r)"
proof -
let ?G = "F-E"
from assms(1) have "F = E +?G"
by (simp add: subset_mset.add_diff_inverse subset_mset_def)
from this assms(1) have "?G \<noteq> {#}"
by fastforce
have "E = E + {#}" by auto
from this \<open>F = E +?G\<close> \<open>?G \<noteq> {#}\<close> assms(2) show ?thesis using one_step_implies_mult [of ?G "{#}" r E] by auto
qed
lemma multiset_order_inclusion_eq:
assumes "E \<subseteq># F"
assumes "trans r"
shows "E = F \<or> (E,F) \<in> (mult r)"
proof (cases)
assume "E = F"
then show ?thesis by auto
next
assume "E \<noteq> F"
from this and assms(1) have "E \<subset># F" by auto
from this assms(2) and multiset_order_inclusion show ?thesis by auto
qed
lemma image_mset_ordering:
assumes "M1 = {# (f1 u). u \<in># L #}"
assumes "M2 = {# (f2 u). u \<in># L #}"
assumes "\<forall>u. (u \<in># L \<longrightarrow> (((f1 u), (f2 u)) \<in> r \<or> (f1 u) = (f2 u)))"
assumes "\<exists>u. (u \<in># L \<and> ((f1 u), (f2 u)) \<in> r)"
assumes "irrefl r"
shows "( (M1,M2) \<in> (mult r) )"
proof -
let ?L' = "{# u \<in># L. (f1 u) = (f2 u) #}"
let ?L'' = "{# u \<in># L. (f1 u) \<noteq> (f2 u) #}"
have "L = ?L' + ?L''" by (simp)
from assms(3) have "\<forall>u. (u \<in># ?L'' \<longrightarrow> ((f1 u),(f2 u)) \<in> r)" by auto
let ?M1' = "{# (f1 u). u \<in># ?L' #}"
let ?M2' = "{# (f2 u). u \<in># ?L' #}"
have "?M1' = ?M2'"
by (metis (mono_tags, lifting) mem_Collect_eq multiset.map_cong0 set_mset_filter)
let ?M1'' = "{# (f1 u). u \<in># ?L'' #}"
let ?M2'' = "{# (f2 u). u \<in># ?L'' #}"
from \<open>L = ?L' + ?L''\<close> have "M1 = ?M1' + ?M1''" by (metis assms(1) image_mset_union)
from \<open>L = ?L' + ?L''\<close> have "M2 = ?M2' + ?M2''" by (metis assms(2) image_mset_union)
have dom: "(\<forall>k \<in> set_mset ?M1''. \<exists>j \<in> set_mset ?M2''. (k, j) \<in> r)"
proof
fix k assume "k \<in> set_mset ?M1''"
from this obtain u where "k = (f1 u)" and "u \<in># ?L''" by auto
from \<open>u \<in># ?L''\<close> have "(f2 u) \<in># ?M2''" by simp
from \<open>\<forall>u. (u \<in># ?L'' \<longrightarrow> ((f1 u),(f2 u)) \<in> r)\<close> and \<open>u \<in># ?L''\<close>
have "((f1 u),(f2 u)) \<in> r" by auto
from this and \<open>k = (f1 u)\<close> and \<open>(f2 u) \<in> set_mset ?M2''\<close>
show "\<exists>j \<in> set_mset ?M2''. (k, j) \<in> r" by auto
qed
have "?L'' \<noteq> {#}"
proof -
from assms(4) obtain u where "u \<in># L" and "( (f1 u),(f2 u) ) \<in> r" by auto
from assms(5) \<open>( (f1 u),(f2 u) ) \<in> r\<close> have "( (f1 u) \<noteq> (f2 u) )"
unfolding irrefl_def by fastforce
from \<open>u \<in># L\<close> \<open>( (f1 u) \<noteq> (f2 u) )\<close> have "u \<in># ?L''" by auto
from this show ?thesis by force
qed
from this have "?M2'' \<noteq> {#}" by auto
from this and dom and \<open>M1 = ?M1' + ?M1''\<close> \<open>M2 = ?M2' + ?M2''\<close> \<open>?M1'=?M2'\<close>
show "(M1,M2) \<in> (mult r)" by (simp add: one_step_implies_mult)
qed
lemma image_mset_ordering_eq:
assumes "M1 = {# (f1 u). u \<in># L #}"
assumes "M2 = {# (f2 u). u \<in># L #}"
assumes "\<forall>u. (u \<in># L \<longrightarrow> (((f1 u), (f2 u)) \<in> r \<or> (f1 u) = (f2 u)))"
shows "(M1 = M2) \<or> ( (M1,M2) \<in> (mult r) )"
proof (cases)
assume "M1 = M2" then show ?thesis by auto
next assume "M1 \<noteq> M2"
let ?L' = "{# u \<in># L. (f1 u) = (f2 u) #}"
let ?L'' = "{# u \<in># L. (f1 u) \<noteq> (f2 u) #}"
have "L = ?L' + ?L''" by (simp)
from assms(3) have "\<forall>u. (u \<in># ?L'' \<longrightarrow> ((f1 u),(f2 u)) \<in> r)" by auto
let ?M1' = "{# (f1 u). u \<in># ?L' #}"
let ?M2' = "{# (f2 u). u \<in># ?L' #}"
have "?M1' = ?M2'"
by (metis (mono_tags, lifting) mem_Collect_eq multiset.map_cong0 set_mset_filter)
let ?M1'' = "{# (f1 u). u \<in># ?L'' #}"
let ?M2'' = "{# (f2 u). u \<in># ?L'' #}"
from \<open>L = ?L' + ?L''\<close> have "M1 = ?M1' + ?M1''" by (metis assms(1) image_mset_union)
from \<open>L = ?L' + ?L''\<close> have "M2 = ?M2' + ?M2''" by (metis assms(2) image_mset_union)
have dom: "(\<forall>k \<in> set_mset ?M1''. \<exists>j \<in> set_mset ?M2''. (k, j) \<in> r)"
proof
fix k assume "k \<in> set_mset ?M1''"
from this obtain u where "k = (f1 u)" and "u \<in># ?L''" by auto
from \<open>u \<in># ?L''\<close> have "(f2 u) \<in># ?M2''" by simp
from \<open>\<forall>u. (u \<in># ?L'' \<longrightarrow> ((f1 u),(f2 u)) \<in> r)\<close> and \<open>u \<in># ?L''\<close>
have "((f1 u),(f2 u)) \<in> r" by auto
from this and \<open>k = (f1 u)\<close> and \<open>(f2 u) \<in> set_mset ?M2''\<close>
show "\<exists>j \<in> set_mset ?M2''. (k, j) \<in> r" by auto
qed
from \<open>M1 \<noteq> M2\<close> have "?M2'' \<noteq> {#}"
using \<open>M1 = image_mset f1 {# u \<in># L. f1 u = f2 u#} + image_mset f1 {# u \<in># L. f1 u \<noteq> f2 u#}\<close> \<open>M2 = image_mset f2 {# u \<in># L. f1 u = f2 u#} + image_mset f2 {# u \<in># L. f1 u \<noteq> f2 u#}\<close> \<open>image_mset f1 {# u \<in># L. f1 u = f2 u#} = image_mset f2 {# u \<in># L. f1 u = f2 u#}\<close> by auto
from this and dom and \<open>M1 = ?M1' + ?M1''\<close> \<open>M2 = ?M2' + ?M2''\<close> \<open>?M1'=?M2'\<close>
have "(M1,M2) \<in> (mult r)" by (simp add: one_step_implies_mult)
from this show ?thesis by auto
qed
lemma mult1_def_lemma :
assumes "M = M0 + {#a#} \<and> N = M0 + K \<and> (\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r)"
shows "(N,M) \<in> (mult1 r)"
proof -
from assms(1) show ?thesis using mult1_def [of r] by auto
qed
lemma mset_ordering_add1:
assumes "(E1,E2) \<in> (mult r)"
shows "(E1,E2 + {# a #}) \<in> (mult r)"
proof -
have i: "(E2,E2 + {# a #}) \<in> (mult1 r)" using mult1_def_lemma [of "E2 + {# a #}" E2 a E2 "{#}" r]
by auto
have i: "(E2,E2 + {# a #}) \<in> (mult1 r)" using mult1_def_lemma [of "E2 + {# a #}" E2 a E2 "{#}" r]
by auto
from assms(1) have "(E1,E2) \<in> (mult1 r)\<^sup>+" using mult_def by auto
from this and i have "(E1,E2 + {# a #}) \<in> (mult1 r)\<^sup>+" by auto
then show ?thesis using mult_def by auto
qed
lemma mset_ordering_singleton:
assumes "\<forall>x. (x \<in># E1 \<longrightarrow> (x,a) \<in> r)"
shows "(E1, {# a #}) \<in> (mult r)"
proof -
let ?K = "E1"
let ?M0 = "{#}"
have i: "E1 = ?M0 + ?K" by auto
have ii: "{# a #} = ?M0 + {# a #}" by auto
from assms(1) have iii: "\<forall>x. (x \<in># ?K \<longrightarrow> (x,a) \<in> r)" by auto
from i and ii and iii show ?thesis using mult1_def_lemma [of "{# a #}" ?M0 a E1 ?K r] mult_def by auto
qed
lemma monotonic_fun_mult1:
assumes "\<And> t s. ((t,s) \<in> r \<Longrightarrow> ((f t), (f s)) \<in> r)"
assumes "(E1,E2) \<in> (mult1 r)"
shows "({# (f x). x \<in># E1 #},{# (f x). x \<in># E2 #}) \<in> (mult1 r)"
proof -
let ?E1 = "{# (f x). x \<in># E1 #}"
let ?E2 = "{# (f x). x \<in># E2 #}"
from assms(2) obtain M0 a K where "E2 = M0 + {#a#}" and "E1 = M0 + K" and "(\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r)"
unfolding mult1_def [of r] by auto
let ?K = "{# (f x). x \<in># K #}"
let ?M0 = "{# (f x). x \<in># M0 #}"
from \<open>E2 = M0 + {#a#}\<close> have "?E2 = ?M0 + {# (f a) #}" by simp
from \<open>E1 = M0 + K\<close> have "?E1 = ?M0 + ?K" by simp
have "(\<forall>b. b \<in># ?K \<longrightarrow> (b, (f a)) \<in> r)"
proof ((rule allI),(rule impI))
fix b assume "b \<in># ?K"
from \<open>b \<in># ?K\<close> obtain b' where "b = (f b')" and "b' \<in># K"
by (auto simp: insert_DiffM2 msed_map_invR union_single_eq_member)
from \<open>b' \<in># K\<close> and \<open>(\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r)\<close> have "(b',a) \<in> r" by auto
from assms(1) and this and \<open>b = (f b')\<close> show "(b, (f a)) \<in> r" by auto
qed
from \<open>?E1 = ?M0 + ?K\<close> and \<open>?E2 = ?M0 + {# (f a) #}\<close> and \<open>(\<forall>b. b \<in># ?K \<longrightarrow> (b, (f a)) \<in> r)\<close>
show "(?E1,?E2) \<in> (mult1 r)" by (metis mult1_def_lemma)
qed
lemma monotonic_fun_mult:
assumes "\<And> t s. ((t,s) \<in> r \<Longrightarrow> ((f t), (f s)) \<in> r)"
assumes "(E1,E2) \<in> (mult r)"
shows "({# (f x). x \<in># E1 #},{# (f x). x \<in># E2 #}) \<in> (mult r)"
proof -
let ?E1 = "{# (f x). x \<in># E1 #}"
let ?E2 = "{# (f x). x \<in># E2 #}"
let ?P = "\<lambda>x. (?E1,{# (f y). y \<in># x #}) \<in> (mult r)"
show ?thesis
proof (rule trancl_induct [of E1 E2 "(mult1 r)" ?P])
from assms(1) show "(E1, E2) \<in> (mult1 r)\<^sup>+" using assms(2) mult_def by blast
next
fix x assume "(E1, x) \<in> mult1 r"
have "(image_mset f E1, image_mset f x) \<in> mult1 r"
by (simp add: \<open>(E1, x) \<in> mult1 r\<close> assms(1) monotonic_fun_mult1)
from this show "(image_mset f E1, image_mset f x) \<in> mult r" by (simp add: mult_def)
next
fix x z assume "(E1, x) \<in> (mult1 r)\<^sup>+"
"(x, z) \<in> mult1 r" and "(image_mset f E1, image_mset f x) \<in> mult r"
from \<open>(x, z) \<in> mult1 r\<close> have "(image_mset f x, image_mset f z) \<in> mult1 r"
by (simp add: assms(1) monotonic_fun_mult1)
from this and \<open>(image_mset f E1, image_mset f x) \<in> mult r\<close>
show "(image_mset f E1, image_mset f z) \<in> mult r"
using mult_def trancl.trancl_into_trancl by fastforce
qed
qed
lemma mset_set_insert_eq:
assumes "finite E"
shows "mset_set (E \<union> { x }) \<subseteq># mset_set E + {# x #}"
proof (rule ccontr)
assume "\<not> ?thesis"
from this obtain y where "(count (mset_set (E \<union> { x })) y)
> (count (mset_set E + {# x #}) y)"
by (meson leI subseteq_mset_def)
from assms(1) have "finite (E \<union> { x })" by auto
have "(count (mset_set E + {# x #}) y) = (count (mset_set E) y) + (count {# x #} y)" by auto
have "x \<noteq> y"
proof
assume "x = y"
then have "y \<in> E \<union> { x }" by auto
from \<open>finite (E \<union> { x })\<close> this have "(count (mset_set (E \<union> { x })) y) = 1"
using count_mset_set(1) by auto
from this and \<open>(count (mset_set (E \<union> { x })) y) > (count (mset_set E + {# x #}) y)\<close> have
"(count (mset_set E + {# x #}) y) = 0" by auto
from \<open>(count (mset_set E + {# x #}) y) = 0\<close> have "count {# x #} y = 0" by auto
from \<open>x = y\<close> have "count {# x #} y = 1" using count_mset_set by auto
from this and \<open>count {# x #} y = 0\<close> show False by auto
qed
have "y \<notin> E"
proof
assume "y \<in> E"
then have "y \<in> E \<union> { x }" by auto
from \<open>finite (E \<union> { x })\<close> this have "(count (mset_set (E \<union> { x })) y) = 1"
using count_mset_set(1) by auto
from this and \<open>(count (mset_set (E \<union> { x })) y) > (count (mset_set E + {# x #}) y)\<close> have
"(count (mset_set E + {# x #}) y) = 0" by auto
from \<open>(count (mset_set E + {# x #}) y) = 0\<close> have "count (mset_set E) y = 0" by (simp split: if_splits)
from \<open>y \<in> E\<close> \<open>finite E\<close> have "count (mset_set E) y = 1" using count_mset_set(1) by auto
from this and \<open>count (mset_set E) y = 0\<close> show False by auto
qed
from this and \<open>x \<noteq> y\<close> have "y \<notin> E \<union> { x }" by auto
from this have "(count (mset_set (E \<union> { x })) y) = 0" by auto
from this and \<open>(count (mset_set (E \<union> { x })) y)
> (count (mset_set E + {# x #}) y)\<close> show False by auto
qed
lemma mset_set_insert:
assumes "x \<notin> E"
assumes "finite E"
shows "mset_set (E \<union> { x }) = mset_set E + {# x #}"
proof (rule ccontr)
assume "\<not> ?thesis"
from this obtain y where "(count (mset_set (E \<union> { x })) y)
\<noteq> (count (mset_set E + {# x #}) y)" by (meson multiset_eqI)
have "(count (mset_set E + {# x #}) y) = (count (mset_set E) y) + (count {# x #} y)" by auto
from assms(2) have "finite (E \<union> { x })" by auto
have "x \<noteq> y"
proof
assume "x = y"
then have "y \<in> E \<union> { x }" by auto
from \<open>finite (E \<union> { x })\<close> this have "(count (mset_set (E \<union> { x })) y) = 1"
using count_mset_set(1) by auto
from \<open>x = y\<close> have "count {# x #} y = 1" using count_mset_set by auto
from \<open>x = y\<close> \<open>x \<notin> E\<close> have "(count (mset_set E) y) = 0" using count_mset_set by auto
from \<open>count {# x #} y = 1\<close> \<open>(count (mset_set E) y) = 0\<close>
\<open>(count (mset_set E + {# x #}) y) = (count (mset_set E) y) + (count {# x #} y)\<close>
have "(count (mset_set E + {# x #}) y) = 1" by auto
from this and \<open>(count (mset_set (E \<union> { x })) y) = 1\<close> and \<open>(count (mset_set (E \<union> { x })) y)
\<noteq> (count (mset_set E + {# x #}) y)\<close> show False by auto
qed
from \<open>x \<noteq> y\<close> have "count {# x #} y = 0" using count_mset_set by auto
have "y \<notin> E"
proof
assume "y \<in> E"
then have "y \<in> E \<union> { x }" by auto
from \<open>finite (E \<union> { x })\<close> this have "(count (mset_set (E \<union> { x })) y) = 1"
using count_mset_set(1) by auto
from assms(2) \<open>y \<in> E\<close> have "(count (mset_set E) y) = 1" using count_mset_set by auto
from \<open>count {# x #} y = 0\<close> \<open>(count (mset_set E) y) = 1\<close>
\<open>(count (mset_set E + {# x #}) y) = (count (mset_set E) y) + (count {# x #} y)\<close>
have "(count (mset_set E + {# x #}) y) = 1" by auto
from this and \<open>(count (mset_set (E \<union> { x })) y) = 1\<close> and \<open>(count (mset_set (E \<union> { x })) y)
\<noteq> (count (mset_set E + {# x #}) y)\<close> show False by auto
qed
from this and \<open>x \<noteq> y\<close> have "y \<notin> E \<union> { x }" by auto
from this have "(count (mset_set (E \<union> { x })) y) = 0" by auto
from \<open>y \<notin> E\<close> have "(count (mset_set E) y) = 0" using count_mset_set by auto
from \<open>count {# x #} y = 0\<close> \<open>(count (mset_set E) y) = 0\<close>
\<open>(count (mset_set E + {# x #}) y) = (count (mset_set E) y) + (count {# x #} y)\<close>
have "(count (mset_set E + {# x #}) y) = 0" by auto
from this and \<open>(count (mset_set (E \<union> { x })) y) = 0\<close> and \<open>(count (mset_set (E \<union> { x })) y)
\<noteq> (count (mset_set E + {# x #}) y)\<close> show False by auto
qed
lemma mset_image_comp:
shows "{# (f x). x \<in># {# (g x). x \<in># E #} #} = {# (f (g x)). x \<in># E #}"
by (simp add: image_mset.compositionality comp_def)
lemma mset_set_mset_image:
shows "\<And> E. card E = N \<Longrightarrow> finite E \<Longrightarrow> mset_set (g ` E) \<subseteq># {# (g x). x \<in># mset_set (E) #}"
proof (induction N)
case 0
assume "card E = 0"
assume "finite E"
from this and \<open>card E = 0\<close> have "E = {}" by auto
then show "mset_set (g ` E) \<subseteq># {# (g x). x \<in># mset_set (E) #}" by auto
next
case (Suc N)
assume "card E = (Suc N)"
assume "finite E"
from this and \<open>card E = (Suc N)\<close> have "E \<noteq> {}" by auto
from this obtain x where "x \<in> E" by auto
let ?E = "E - { x }"
from \<open>finite E\<close> \<open>card E = (Suc N)\<close> and \<open>x \<in> E\<close> have "card ?E = N" by auto
from \<open>finite E\<close> have "finite ?E" by auto
from this and "Suc.IH" [of ?E] and \<open>card ?E = N\<close>
have ind: "mset_set (g ` ?E) \<subseteq># {# (g x). x \<in># mset_set (?E) #}" by force
from \<open>x \<in> E\<close> have "E = ?E \<union> { x }" by auto
have "x \<notin> ?E" by auto
from \<open>finite ?E\<close> \<open>E = ?E \<union> { x }\<close> and \<open>x \<notin> ?E\<close> have "mset_set (?E \<union> { x }) = mset_set ?E + {# x #}"
using mset_set_insert [of x ?E] by auto
from this have
"{# (g x). x \<in># mset_set (?E \<union> { x }) #} = {# (g x). x \<in># mset_set ?E #} + {# (g x) #}"
by auto
have "(g ` (?E \<union> { x }) = (g ` ?E) \<union> { g x })" by auto
from this have i: "mset_set (g ` (?E \<union> { x })) = mset_set ( (g ` ?E) \<union> { g x } )" by auto
from \<open>finite ?E\<close> have "finite (g ` ?E)" by auto
from this have "mset_set ( (g ` ?E) \<union> { g x } ) \<subseteq># mset_set (g ` ?E) + {# (g x) #}"
using mset_set_insert_eq [of "(g ` ?E)" "(g x)" ] by meson
from this i have ii: "mset_set (g ` (?E \<union> { x })) \<subseteq># mset_set (g ` ?E) + {# (g x) #}" by auto
from ind have "mset_set (g ` ?E) + {# (g x) #} \<subseteq># {# (g x). x \<in># mset_set (?E) #} + {# (g x) #}"
using Multiset.subset_mset.add_right_mono by metis
from this and ii have "mset_set (g ` (?E \<union> { x })) \<subseteq># {# (g x). x \<in># mset_set (?E) #} + {# (g x) #}"
- using Multiset.subset_mset.order.trans [of "mset_set (g ` (?E \<union> { x }))" ] by metis
+ using subset_mset.trans [of "mset_set (g ` (?E \<union> { x }))" ] by metis
from this and \<open>E = ?E \<union> { x }\<close> \<open>{# (g x). x \<in># mset_set (?E \<union> { x }) #} = {# (g x). x \<in># mset_set ?E #} + {# (g x) #}\<close>
show "mset_set (g ` E) \<subseteq># {# (g x). x \<in># mset_set E #}"
by metis
qed
lemma split_mset_set:
assumes "C = C1 \<union> C2"
assumes "C1 \<inter> C2 = {}"
assumes "finite C1"
assumes "finite C2"
shows "(mset_set C) = (mset_set C1) + (mset_set C2)"
proof (rule ccontr)
assume "(mset_set C) \<noteq> (mset_set C1) + (mset_set C2)"
then obtain x where "count (mset_set C) x \<noteq> count ((mset_set C1) + (mset_set C2)) x"
by (meson multiset_eqI)
from assms(3) assms(4) assms(1) have "finite C" by auto
have "count ((mset_set C1) + (mset_set C2)) x = (count (mset_set C1) x) + (count (mset_set C2) x)"
by auto
from this and \<open>count (mset_set C) x \<noteq> count ((mset_set C1) + (mset_set C2)) x\<close> have
"count (mset_set C) x \<noteq> (count (mset_set C1) x) + (count (mset_set C2) x)" by auto
have "x \<in> C1 \<or> x \<in> C2"
proof (rule ccontr)
assume "\<not> (x \<in> C1 \<or> x \<in> C2)"
then have "x \<notin> C1" and "x\<notin> C2" by auto
from assms(1) \<open>x \<notin> C1\<close> and \<open>x\<notin> C2\<close> have "x \<notin> C" by auto
from \<open>x \<notin> C1\<close> have "(count (mset_set C1) x) = 0" by auto
from \<open>x \<notin> C2\<close> have "(count (mset_set C2) x) = 0" by auto
from \<open>x \<notin> C\<close> have "(count (mset_set C) x) = 0" by auto
from \<open>(count (mset_set C1) x) = 0\<close> \<open>(count (mset_set C2) x) = 0\<close>
\<open>(count (mset_set C) x) = 0\<close>
\<open>count (mset_set C) x \<noteq> (count (mset_set C1) x) + (count (mset_set C2) x)\<close>
show False by auto
qed
have "(x \<notin> C1 \<or> x \<in> C2)"
proof (rule ccontr)
assume "\<not> (x \<notin> C1 \<or> x \<in> C2)"
then have "x \<in> C1" and " x\<notin> C2" by auto
from assms(1) \<open>x \<in> C1\<close> have "x \<in> C" by auto
from assms(3) \<open>x \<in> C1\<close> have "(count (mset_set C1) x) = 1" by auto
from \<open>x \<notin> C2\<close> have "(count (mset_set C2) x) = 0" by auto
from assms(3) assms(4) assms(1) have "finite C" by auto
from \<open>finite C\<close> \<open>x \<in> C\<close> have "(count (mset_set C) x) = 1" by auto
from \<open>(count (mset_set C1) x) = 1\<close> \<open>(count (mset_set C2) x) = 0\<close>
\<open>(count (mset_set C) x) = 1\<close>
\<open>count (mset_set C) x \<noteq> (count (mset_set C1) x) + (count (mset_set C2) x)\<close>
show False by auto
qed
have "(x \<notin> C2 \<or> x \<in> C1)"
proof (rule ccontr)
assume "\<not> (x \<notin> C2 \<or> x \<in> C1)"
then have "x \<in> C2" and " x\<notin> C1" by auto
from assms(1) \<open>x \<in> C2\<close> have "x \<in> C" by auto
from assms(4) \<open>x \<in> C2\<close> have "(count (mset_set C2) x) = 1" by auto
from \<open>x \<notin> C1\<close> have "(count (mset_set C1) x) = 0" by auto
from \<open>finite C\<close> \<open>x \<in> C\<close> have "(count (mset_set C) x) = 1" by auto
from \<open>(count (mset_set C2) x) = 1\<close> \<open>(count (mset_set C1) x) = 0\<close>
\<open>(count (mset_set C) x) = 1\<close>
\<open>count (mset_set C) x \<noteq> (count (mset_set C1) x) + (count (mset_set C2) x)\<close>
show False by auto
qed
from \<open>x \<in> C1 \<or> x \<in> C2\<close> \<open>(x \<notin> C1 \<or> x \<in> C2)\<close> \<open>(x \<notin> C2 \<or> x \<in> C1)\<close>
have "x \<in> C1 \<and> x \<in> C2" by blast
from this and assms(2) show False by auto
qed
lemma image_mset_thm:
assumes "E = {# (f x). x \<in># E' #}"
assumes "x \<in># E"
shows "\<exists> y. ((y \<in># E') \<and> x = (f y))"
using assms by auto
lemma split_image_mset:
assumes "M = M1 + M2"
shows "{# (f x). x \<in># M #} = {# (f x). x \<in># M1 #} + {# (f x). x \<in># M2 #}"
by (simp add: assms)
end
diff --git a/thys/Timed_Automata/Floyd_Warshall.thy b/thys/Timed_Automata/Floyd_Warshall.thy
--- a/thys/Timed_Automata/Floyd_Warshall.thy
+++ b/thys/Timed_Automata/Floyd_Warshall.thy
@@ -1,1736 +1,1736 @@
theory Floyd_Warshall
imports Main
begin
chapter \<open>Floyd-Warshall Algorithm for the All-Pairs Shortest Paths Problem\<close>
subsubsection \<open>Auxiliary\<close>
lemma distinct_list_single_elem_decomp: "{xs. set xs \<subseteq> {0} \<and> distinct xs} = {[], [0]}"
proof (standard, goal_cases)
case 1
{ fix xs :: "'a list" assume xs: "xs \<in> {xs. set xs \<subseteq> {0} \<and> distinct xs}"
have "xs \<in> {[], [0]}"
proof (cases xs)
case (Cons y ys)
hence y: "y = 0" using xs by auto
with Cons xs have "ys = []" by (cases ys, auto)
thus ?thesis using y Cons by simp
qed simp
}
thus ?case by blast
qed simp
section \<open>Cycles in Lists\<close>
abbreviation "cnt x xs \<equiv> length (filter (\<lambda>y. x = y) xs)"
fun remove_cycles :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
"remove_cycles [] _ acc = rev acc" |
"remove_cycles (x#xs) y acc =
(if x = y then remove_cycles xs y [x] else remove_cycles xs y (x#acc))"
lemma cnt_rev: "cnt x (rev xs) = cnt x xs" by (metis length_rev rev_filter)
value "as @ [x] @ bs @ [x] @ cs @ [x] @ ds"
lemma remove_cycles_removes: "cnt x (remove_cycles xs x ys) \<le> max 1 (cnt x ys)"
proof (induction xs arbitrary: ys)
case Nil thus ?case
by (simp, cases "x \<in> set ys", (auto simp: cnt_rev[of x ys]))
next
case (Cons y xs)
thus ?case
proof (cases "x = y")
case True
thus ?thesis using Cons[of "[y]"] True by auto
next
case False
thus ?thesis using Cons[of "y # ys"] by auto
qed
qed
lemma remove_cycles_id: "x \<notin> set xs \<Longrightarrow> remove_cycles xs x ys = rev ys @ xs"
by (induction xs arbitrary: ys) auto
lemma remove_cycles_cnt_id:
"x \<noteq> y \<Longrightarrow> cnt y (remove_cycles xs x ys) \<le> cnt y ys + cnt y xs"
proof (induction xs arbitrary: ys x)
case Nil thus ?case by (simp add: cnt_rev)
next
case (Cons z xs)
thus ?case
proof (cases "x = z")
case True thus ?thesis using Cons.IH[of z "[z]"] Cons.prems by auto
next
case False
thus ?thesis using Cons.IH[of x "z # ys"] Cons.prems False by auto
qed
qed
lemma remove_cycles_ends_cycle: "remove_cycles xs x ys \<noteq> rev ys @ xs \<Longrightarrow> x \<in> set xs"
using remove_cycles_id by fastforce
lemma remove_cycles_begins_with: "x \<in> set xs \<Longrightarrow> \<exists> zs. remove_cycles xs x ys = x # zs \<and> x \<notin> set zs"
proof (induction xs arbitrary: ys)
case Nil thus ?case by auto
next
case (Cons y xs)
thus ?case
proof (cases "x = y")
case True thus ?thesis
proof (cases "x \<in> set xs", goal_cases)
case 1 with Cons show ?case by auto
next
case 2 with remove_cycles_id[of x xs "[y]"] show ?case by auto
qed
next
case False
with Cons show ?thesis by auto
qed
qed
lemma remove_cycles_self:
"x \<in> set xs \<Longrightarrow> remove_cycles (remove_cycles xs x ys) x zs = remove_cycles xs x ys"
proof -
assume x:"x \<in> set xs"
then obtain ws where ws: "remove_cycles xs x ys = x # ws" "x \<notin> set ws"
using remove_cycles_begins_with[OF x, of ys] by blast
from remove_cycles_id[OF this(2)] have "remove_cycles ws x [x] = x # ws" by auto
with ws(1) show "remove_cycles (remove_cycles xs x ys) x zs = remove_cycles xs x ys" by simp
qed
lemma remove_cycles_one: "remove_cycles (as @ x # xs) x ys = remove_cycles (x#xs) x ys"
by (induction as arbitrary: ys) auto
lemma remove_cycles_cycles:
"x \<in> set xs \<Longrightarrow> \<exists> xxs as. as @ concat (map (\<lambda> xs. x # xs) xxs) @ remove_cycles xs x ys = xs \<and> x \<notin> set as"
proof (induction xs arbitrary: ys)
case Nil thus ?case by auto
next
case (Cons y xs)
thus ?case
proof (cases "x = y")
case True thus ?thesis
proof (cases "x \<in> set xs", goal_cases)
case 1
then obtain as xxs where "as @ concat (map (\<lambda>xs. y#xs) xxs) @ remove_cycles xs y [y] = xs"
using Cons.IH[of "[y]"] by auto
hence "[] @ concat (map (\<lambda>xs. x#xs) (as#xxs)) @ remove_cycles (y#xs) x ys = y # xs"
by (simp add: \<open>x = y\<close>)
thus ?thesis by fastforce
next
case 2
hence "remove_cycles (y # xs) x ys = y # xs" using remove_cycles_id[of x xs "[y]"] by auto
hence "[] @ concat (map (\<lambda>xs. x # xs) []) @ remove_cycles (y#xs) x ys = y # xs" by auto
thus ?thesis by fastforce
qed
next
case False
then obtain as xxs where as:
"as @ concat (map (\<lambda>xs. x # xs) xxs) @ remove_cycles xs x (y#ys) = xs" "x \<notin> set as"
using Cons.IH[of "y # ys"] Cons.prems by auto
hence "(y # as) @ concat (map (\<lambda>xs. x # xs) xxs) @ remove_cycles (y#xs) x ys = y # xs"
using \<open>x \<noteq> y\<close> by auto
thus ?thesis using as(2) \<open>x \<noteq> y\<close> by fastforce
qed
qed
fun start_remove :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
"start_remove [] _ acc = rev acc" |
"start_remove (x#xs) y acc =
(if x = y then rev acc @ remove_cycles xs y [y] else start_remove xs y (x # acc))"
lemma start_remove_decomp:
"x \<in> set xs \<Longrightarrow> \<exists> as bs. xs = as @ x # bs \<and> start_remove xs x ys = rev ys @ as @ remove_cycles bs x [x]"
proof (induction xs arbitrary: ys)
case Nil thus ?case by auto
next
case (Cons y xs)
thus ?case
proof (auto, goal_cases)
case 1
from 1(1)[of "y # ys"]
obtain as bs where
"xs = as @ x # bs" "start_remove xs x (y # ys) = rev (y # ys) @ as @ remove_cycles bs x [x]"
by blast
hence "y # xs = (y # as) @ x # bs"
"start_remove xs x (y # ys) = rev ys @ (y # as) @ remove_cycles bs x [x]" by simp+
thus ?case by blast
qed
qed
lemma start_remove_removes: "cnt x (start_remove xs x ys) \<le> Suc (cnt x ys)"
proof (induction xs arbitrary: ys)
case Nil thus ?case using cnt_rev[of x ys] by auto
next
case (Cons y xs)
thus ?case
proof (cases "x = y")
case True
thus ?thesis using remove_cycles_removes[of y xs "[y]"] cnt_rev[of y ys] by auto
next
case False
thus ?thesis using Cons[of "y # ys"] by auto
qed
qed
lemma start_remove_id[simp]: "x \<notin> set xs \<Longrightarrow> start_remove xs x ys = rev ys @ xs"
by (induction xs arbitrary: ys) auto
lemma start_remove_cnt_id:
"x \<noteq> y \<Longrightarrow> cnt y (start_remove xs x ys) \<le> cnt y ys + cnt y xs"
proof (induction xs arbitrary: ys)
case Nil thus ?case by (simp add: cnt_rev)
next
case (Cons z xs)
thus ?case
proof (cases "x = z", goal_cases)
case 1 thus ?case using remove_cycles_cnt_id[of x y xs "[x]"] by (simp add: cnt_rev)
next
case 2 from this(1)[of "(z # ys)"] this(2,3) show ?case by auto
qed
qed
fun remove_all_cycles :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
"remove_all_cycles [] xs = xs" |
"remove_all_cycles (x # xs) ys = remove_all_cycles xs (start_remove ys x [])"
lemma cnt_remove_all_mono:"cnt y (remove_all_cycles xs ys) \<le> max 1 (cnt y ys)"
proof (induction xs arbitrary: ys)
case Nil thus ?case by auto
next
case (Cons x xs)
thus ?case
proof (cases "x = y")
case True thus ?thesis using start_remove_removes[of y ys "[]"] Cons[of "start_remove ys y []"]
by auto
next
case False
hence "cnt y (start_remove ys x []) \<le> cnt y ys"
using start_remove_cnt_id[of x y ys "[]"] by auto
thus ?thesis using Cons[of "start_remove ys x []"] by auto
qed
qed
lemma cnt_remove_all_cycles: "x \<in> set xs \<Longrightarrow> cnt x (remove_all_cycles xs ys) \<le> 1"
proof (induction xs arbitrary: ys)
case Nil thus ?case by auto
next
case (Cons y xs)
thus ?case
using start_remove_removes[of x ys "[]"] cnt_remove_all_mono[of y xs "start_remove ys y []"]
by auto
qed
lemma cnt_mono:
"cnt a (b # xs) \<le> cnt a (b # c # xs)"
by (induction xs) auto
lemma cnt_distinct_intro: "\<forall> x \<in> set xs. cnt x xs \<le> 1 \<Longrightarrow> distinct xs"
proof (induction xs)
case Nil thus ?case by auto
next
case (Cons x xs)
from this(2) have "\<forall> x \<in> set xs. cnt x xs \<le> 1"
by (metis filter.simps(2) impossible_Cons linorder_class.linear list.set_intros(2)
preorder_class.order_trans)
with Cons.IH have "distinct xs" by auto
moreover have "x \<notin> set xs" using Cons.prems
proof (induction xs)
case Nil then show ?case by auto
next
case (Cons a xs)
from this(2) have "\<forall>xa\<in>set (x # xs). cnt xa (x # a # xs) \<le> 1"
by auto
then have *: "\<forall>xa\<in>set (x # xs). cnt xa (x # xs) \<le> 1"
proof (safe, goal_cases)
case (1 b)
then have "cnt b (x # a # xs) \<le> 1" by auto
with cnt_mono[of b x xs a] show ?case by fastforce
qed
with Cons(1) have "x \<notin> set xs" by auto
moreover have "x \<noteq> a"
by (metis (full_types) Cons.prems One_nat_def * empty_iff filter.simps(2) impossible_Cons
le_0_eq le_Suc_eq length_0_conv list.set(1) list.set_intros(1))
ultimately show ?case by auto
qed
ultimately show ?case by auto
qed
lemma remove_cycles_subs:
"set (remove_cycles xs x ys) \<subseteq> set xs \<union> set ys"
by (induction xs arbitrary: ys; auto; fastforce)
lemma start_remove_subs:
"set (start_remove xs x ys) \<subseteq> set xs \<union> set ys"
using remove_cycles_subs by (induction xs arbitrary: ys; auto; fastforce)
lemma remove_all_cycles_subs:
"set (remove_all_cycles xs ys) \<subseteq> set ys"
using start_remove_subs by (induction xs arbitrary: ys, auto) (fastforce+)
lemma remove_all_cycles_distinct: "set ys \<subseteq> set xs \<Longrightarrow> distinct (remove_all_cycles xs ys)"
proof -
assume "set ys \<subseteq> set xs"
hence "\<forall> x \<in> set ys. cnt x (remove_all_cycles xs ys) \<le> 1" using cnt_remove_all_cycles by fastforce
hence "\<forall> x \<in> set (remove_all_cycles xs ys). cnt x (remove_all_cycles xs ys) \<le> 1"
using remove_all_cycles_subs by fastforce
thus "distinct (remove_all_cycles xs ys)" using cnt_distinct_intro by auto
qed
lemma distinct_remove_cycles_inv: "distinct (xs @ ys) \<Longrightarrow> distinct (remove_cycles xs x ys)"
proof (induction xs arbitrary: ys)
case Nil thus ?case by auto
next
case (Cons y xs)
thus ?case by auto
qed
definition "remove_all x xs = (if x \<in> set xs then tl (remove_cycles xs x []) else xs)"
definition "remove_all_rev x xs = (if x \<in> set xs then rev (tl (remove_cycles (rev xs) x [])) else xs)"
lemma remove_all_distinct:
"distinct xs \<Longrightarrow> distinct (x # remove_all x xs)"
proof (cases "x \<in> set xs", goal_cases)
case 1
from remove_cycles_begins_with[OF 1(2), of "[]"] obtain zs
where "remove_cycles xs x [] = x # zs" "x \<notin> set zs" by auto
thus ?thesis using 1(1) distinct_remove_cycles_inv[of "xs" "[]" x] by (simp add: remove_all_def)
next
case 2 thus ?thesis by (simp add: remove_all_def)
qed
lemma remove_all_removes:
"x \<notin> set (remove_all x xs)"
by (metis list.sel(3) remove_all_def remove_cycles_begins_with)
lemma remove_all_subs:
"set (remove_all x xs) \<subseteq> set xs"
using remove_cycles_subs remove_all_def
by (metis (no_types, lifting) append_Nil2 list.sel(2) list.set_sel(2) set_append subsetCE subsetI)
lemma remove_all_rev_distinct: "distinct xs \<Longrightarrow> distinct (x # remove_all_rev x xs)"
proof (cases "x \<in> set xs", goal_cases)
case 1
then have "x \<in> set (rev xs)" by auto
from remove_cycles_begins_with[OF this, of "[]"] obtain zs
where "remove_cycles (rev xs) x [] = x # zs" "x \<notin> set zs" by auto
thus ?thesis using 1(1) distinct_remove_cycles_inv[of "rev xs" "[]" x] by (simp add: remove_all_rev_def)
next
case 2 thus ?thesis by (simp add: remove_all_rev_def)
qed
lemma remove_all_rev_removes: "x \<notin> set (remove_all_rev x xs)"
by (metis remove_all_def remove_all_removes remove_all_rev_def set_rev)
lemma remove_all_rev_subs: "set (remove_all_rev x xs) \<subseteq> set xs"
by (metis remove_all_def remove_all_subs set_rev remove_all_rev_def)
abbreviation "rem_cycles i j xs \<equiv> remove_all i (remove_all_rev j (remove_all_cycles xs xs))"
lemma rem_cycles_distinct': "i \<noteq> j \<Longrightarrow> distinct (i # j # rem_cycles i j xs)"
proof -
assume "i \<noteq> j"
have "distinct (remove_all_cycles xs xs)" by (simp add: remove_all_cycles_distinct)
from remove_all_rev_distinct[OF this] have
"distinct (remove_all_rev j (remove_all_cycles xs xs))"
by simp
from remove_all_distinct[OF this] have "distinct (i # rem_cycles i j xs)" by simp
moreover have
"j \<notin> set (rem_cycles i j xs)"
using remove_all_subs remove_all_rev_removes remove_all_removes by fastforce
ultimately show ?thesis by (simp add: \<open>i \<noteq> j\<close>)
qed
lemma rem_cycles_removes_last: "j \<notin> set (rem_cycles i j xs)"
by (meson remove_all_rev_removes remove_all_subs rev_subsetD)
lemma rem_cycles_distinct: "distinct (rem_cycles i j xs)"
by (meson distinct.simps(2) order_refl remove_all_cycles_distinct
remove_all_distinct remove_all_rev_distinct)
lemma rem_cycles_subs: "set (rem_cycles i j xs) \<subseteq> set xs"
by (meson order_trans remove_all_cycles_subs remove_all_subs remove_all_rev_subs)
section \<open>Definition of the Algorithm\<close>
text \<open>
We formalize the Floyd-Warshall algorithm on a linearly ordered abelian semigroup.
However, we would not need an \<open>abelian\<close> monoid if we had the right type class.
\<close>
class linordered_ab_monoid_add = linordered_ab_semigroup_add +
fixes neutral :: 'a ("\<one>")
assumes neutl[simp]: "\<one> + x = x"
assumes neutr[simp]: "x + \<one> = x"
begin
lemmas assoc = add.assoc
type_synonym 'c mat = "nat \<Rightarrow> nat \<Rightarrow> 'c"
definition (in -) upd :: "'c mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'c \<Rightarrow> 'c mat"
where
"upd m x y v = m (x := (m x) (y := v))"
definition fw_upd :: "'a mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a mat" where
"fw_upd m k i j \<equiv> upd m i j (min (m i j) (m i k + m k j))"
lemma fw_upd_mono:
"fw_upd m k i j i' j' \<le> m i' j'"
by (cases "i = i'", cases "j = j'") (auto simp: fw_upd_def upd_def)
fun fw :: "'a mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a mat" where
"fw m n 0 0 0 = fw_upd m 0 0 0" |
"fw m n (Suc k) 0 0 = fw_upd (fw m n k n n) (Suc k) 0 0" |
"fw m n k (Suc i) 0 = fw_upd (fw m n k i n) k (Suc i) 0" |
"fw m n k i (Suc j) = fw_upd (fw m n k i j) k i (Suc j)"
lemma fw_invariant_aux_1:
"j'' \<le> j \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> k \<le> n \<Longrightarrow> fw m n k i j i' j' \<le> fw m n k i j'' i' j'"
proof (induction j)
case 0 thus ?case by simp
next
case (Suc j) thus ?case
proof (cases "j'' = Suc j")
case True thus ?thesis by simp
next
case False
have "fw_upd (fw m n k i j) k i (Suc j) i' j' \<le> fw m n k i j i' j'" by (simp add: fw_upd_mono)
thus ?thesis using Suc False by simp
qed
qed
lemma fw_invariant_aux_2:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> k \<le> n \<Longrightarrow> i'' \<le> i \<Longrightarrow> j'' \<le> j
\<Longrightarrow> fw m n k i j i' j' \<le> fw m n k i'' j'' i' j'"
proof (induction i)
case 0 thus ?case using fw_invariant_aux_1 by auto
next
case (Suc i) thus ?case
proof (cases "i'' = Suc i")
case True thus ?thesis using Suc fw_invariant_aux_1 by simp
next
case False
have "fw m n k (Suc i) j i' j' \<le> fw m n k (Suc i) 0 i' j'"
using fw_invariant_aux_1[of 0 j "Suc i" n k] Suc(2-) by simp
also have "\<dots> \<le> fw m n k i n i' j'" by (simp add: fw_upd_mono)
also have "\<dots> \<le> fw m n k i j i' j'" using fw_invariant_aux_1[of j n i n k] False Suc by simp
also have "\<dots> \<le> fw m n k i'' j'' i' j'" using Suc False by simp
finally show ?thesis by simp
qed
qed
lemma fw_invariant:
"k' \<le> k \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> k \<le> n \<Longrightarrow> j'' \<le> j \<Longrightarrow> i'' \<le> i
\<Longrightarrow> fw m n k i j i' j' \<le> fw m n k' i'' j'' i' j'"
proof (induction k)
case 0 thus ?case using fw_invariant_aux_2 by auto
next
case (Suc k) thus ?case
proof (cases "k' = Suc k")
case True thus ?thesis using Suc fw_invariant_aux_2 by simp
next
case False
have "fw m n (Suc k) i j i' j' \<le> fw m n (Suc k) 0 0 i' j'"
using fw_invariant_aux_2[of i n j "Suc k" 0 0] Suc(2-) by simp
also have "\<dots> \<le> fw m n k n n i' j'" by (simp add: fw_upd_mono)
also have "\<dots> \<le> fw m n k i j i' j'" using fw_invariant_aux_2[of n n n k] False Suc by simp
also have "\<dots> \<le> fw m n k' i'' j'' i' j'" using Suc False by simp
finally show ?thesis by simp
qed
qed
lemma single_row_inv:
"j' < j \<Longrightarrow> j \<le> n \<Longrightarrow> i' \<le> n \<Longrightarrow> fw m n k i' j i' j' = fw m n k i' j' i' j'"
proof (induction j)
case 0 thus ?case by simp
next
case (Suc j) thus ?case by (cases "j' = j") (simp add: fw_upd_def upd_def)+
qed
lemma single_iteration_inv':
"i' < i \<Longrightarrow> j' \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> i \<le> n \<Longrightarrow> fw m n k i j i' j' = fw m n k i' j' i' j'"
proof (induction i arbitrary: j)
case 0 thus ?case by simp
next
case (Suc i) thus ?case
proof (induction j)
case 0 thus ?case
proof (cases "i = i'", goal_cases)
case 2 thus ?case by (simp add: fw_upd_def upd_def)
next
case 1 thus ?case using single_row_inv[of j' n n i' m k]
by (cases "j' = n") (fastforce simp add: fw_upd_def upd_def)+
qed
next
case (Suc j) thus ?case by (simp add: fw_upd_def upd_def)
qed
qed
lemma single_iteration_inv:
"i' \<le> i \<Longrightarrow> j' \<le> j \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n\<Longrightarrow> fw m n k i j i' j' = fw m n k i' j' i' j'"
proof (induction i arbitrary: j)
case 0 thus ?case
proof (induction j)
case 0 thus ?case by simp
next
case (Suc j) thus ?case using 0 by (cases "j' = Suc j") (simp add: fw_upd_def upd_def)+
qed
next
case (Suc i) thus ?case
proof (induction j)
case 0 thus ?case by (cases "i' = Suc i") (simp add: fw_upd_def upd_def)+
next
case (Suc j) thus ?case
proof (cases "i' = Suc i", goal_cases)
case 1 thus ?case
proof (cases "j' = Suc j", goal_cases)
case 1 thus ?case by simp
next
case 2 thus ?case by (simp add: fw_upd_def upd_def)
qed
next
case 2 thus ?case
proof (cases "j' = Suc j", goal_cases)
case 1 thus ?case using single_iteration_inv'[of i' "Suc i" j' n "Suc j" m k] by simp
next
case 2 thus ?case by (simp add: fw_upd_def upd_def)
qed
qed
qed
qed
lemma fw_innermost_id:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> j' \<le> n \<Longrightarrow> i' < i \<Longrightarrow> fw m n 0 i' j' i j = m i j"
proof (induction i' arbitrary: j')
case 0 thus ?case
proof (induction j')
case 0 thus ?case by (simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp: fw_upd_def upd_def)
qed
next
case (Suc i') thus ?case
proof (induction j')
case 0 thus ?case by (auto simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp add: fw_upd_def upd_def)
qed
qed
lemma fw_middle_id:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> j' < j \<Longrightarrow> i' \<le> i \<Longrightarrow> fw m n 0 i' j' i j = m i j"
proof (induction i' arbitrary: j')
case 0 thus ?case
proof (induction j')
case 0 thus ?case by (simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp: fw_upd_def upd_def)
qed
next
case (Suc i') thus ?case
proof (induction j')
case 0 thus ?case using fw_innermost_id by (auto simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp add: fw_upd_def upd_def)
qed
qed
lemma fw_outermost_mono:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> fw m n 0 i j i j \<le> m i j"
proof (cases j)
case 0
assume "i \<le> n"
thus ?thesis
proof (cases i)
case 0 thus ?thesis using \<open>j = 0\<close> by (simp add: fw_upd_def upd_def)
next
case (Suc i')
hence "fw m n 0 i' n (Suc i') 0 = m (Suc i') 0" using fw_innermost_id[of "Suc i'" n 0 n i' m]
using \<open>i \<le> n\<close> by simp
thus ?thesis using \<open>j = 0\<close> Suc by (simp add: fw_upd_def upd_def)
qed
next
case (Suc j')
assume "i \<le> n" "j \<le> n"
hence "fw m n 0 i j' i (Suc j') = m i (Suc j')"
using fw_middle_id[of i n "Suc j'" j' i m] Suc by simp
thus ?thesis using Suc by (simp add: fw_upd_def upd_def)
qed
lemma Suc_innermost_id1:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> j' \<le> n \<Longrightarrow> i' < i \<Longrightarrow> fw m n (Suc k) i' j' i j = fw m n k i j i j"
proof (induction i' arbitrary: j')
case 0 thus ?case
proof (induction j')
case 0
hence "fw m n k n n i j = fw m n k i j i j" using single_iteration_inv[of i n j n n m k] by simp
thus ?case using 0 by (simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp: fw_upd_def upd_def)
qed
next
case (Suc i') thus ?case
proof (induction j')
case 0 thus ?case by (auto simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp add: fw_upd_def upd_def)
qed
qed
lemma Suc_innermost_id2:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> j' < j \<Longrightarrow> i' \<le> i \<Longrightarrow> fw m n (Suc k) i' j' i j = fw m n k i j i j"
proof (induction i' arbitrary: j')
case 0
hence "fw m n k n n i j = fw m n k i j i j" using single_iteration_inv[of i n j n n m k] by simp
with 0 show ?case
proof (induction j')
case 0
thus ?case by (auto simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp: fw_upd_def upd_def)
qed
next
case (Suc i') thus ?case
proof (induction j')
case 0 thus ?case using Suc_innermost_id1 by (auto simp add: fw_upd_def upd_def)
next
case (Suc j') thus ?case by (auto simp add: fw_upd_def upd_def)
qed
qed
lemma Suc_innermost_id1':
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> j' \<le> n \<Longrightarrow> i' < i \<Longrightarrow> fw m n (Suc k) i' j' i j = fw m n k n n i j"
proof goal_cases
case 1
hence "fw m n (Suc k) i' j' i j = fw m n k i j i j" using Suc_innermost_id1 by simp
thus ?thesis using 1 single_iteration_inv[of i n] by simp
qed
lemma Suc_innermost_id2':
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> j' < j \<Longrightarrow> i' \<le> i \<Longrightarrow> fw m n (Suc k) i' j' i j = fw m n k n n i j"
proof goal_cases
case 1
hence "fw m n (Suc k) i' j' i j = fw m n k i j i j" using Suc_innermost_id2 by simp
thus ?thesis using 1 single_iteration_inv[of i n] by simp
qed
lemma Suc_innermost_mono:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> fw m n (Suc k) i j i j \<le> fw m n k i j i j"
proof (cases j)
case 0
assume "i \<le> n"
thus ?thesis
proof (cases i)
case 0 thus ?thesis using \<open>j = 0\<close> single_iteration_inv[of 0 n 0 n n m k]
by (simp add: fw_upd_def upd_def)
next
case (Suc i')
thus ?thesis using Suc_innermost_id1 \<open>i \<le> n\<close> \<open>j = 0\<close>
by (auto simp: fw_upd_def upd_def local.min.coboundedI1)
qed
next
case (Suc j')
assume "i \<le> n" "j \<le> n"
thus ?thesis using Suc Suc_innermost_id2 by (auto simp: fw_upd_def upd_def local.min.coboundedI1)
qed
lemma fw_mono':
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> fw m n k i j i j \<le> m i j"
proof (induction k)
case 0 thus ?case using fw_outermost_mono by simp
next
case (Suc k) thus ?case using Suc_innermost_mono[OF Suc.prems, of m k] by simp
qed
lemma fw_mono:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> i' \<le> n \<Longrightarrow> j' \<le> n \<Longrightarrow> fw m n k i j i' j' \<le> m i' j'"
proof (cases k)
case 0
assume 0: "i \<le> n" "j \<le> n" "i' \<le> n" "j' \<le> n" "k = 0"
thus ?thesis
proof (cases "i' \<le> i")
case False thus ?thesis using 0 fw_innermost_id by simp
next
case True thus ?thesis
proof (cases "j' \<le> j")
case True
have "fw m n 0 i j i' j' \<le> fw m n 0 i' j' i' j'" using fw_invariant True \<open>i' \<le> i\<close> 0 by simp
also have "fw m n 0 i' j' i' j' \<le> m i' j'" using 0 fw_outermost_mono by blast
finally show ?thesis by (simp add: \<open>k = 0\<close>)
next
case False thus ?thesis
proof (cases "i = i'", goal_cases)
case 1 then show ?thesis using fw_middle_id[of i' n j' j i' m] 0 by simp
next
case 2
then show ?case
using single_iteration_inv'[of i' i j' n j m 0] \<open>i' \<le> i\<close> fw_middle_id[of i' n j' j i' m]
fw_outermost_mono[of i' n j' m] 0
by simp
qed
qed
qed
next
case (Suc k)
assume prems: "i \<le> n" "j \<le> n" "i' \<le> n" "j' \<le> n"
thus ?thesis
proof (cases "i' \<le> i \<and> j' \<le> j")
case True
hence "fw m n (Suc k) i j i' j' = fw m n (Suc k) i' j' i' j'"
using prems single_iteration_inv by blast
thus ?thesis using Suc prems fw_mono' by auto
next
case False thus ?thesis
proof auto
assume "\<not> i' \<le> i"
thus ?thesis using Suc prems fw_mono' Suc_innermost_id1 by auto
next
assume "\<not> j' \<le> j"
hence "j < j'" by simp
show ?thesis
proof (cases "i \<le> i'")
case True
thus ?thesis using Suc prems Suc_innermost_id2 \<open>j < j'\<close> fw_mono' by auto
next
case False
thus ?thesis using single_iteration_inv' Suc prems fw_mono' by auto
qed
qed
qed
qed
lemma add_mono_neutr:
assumes "\<one> \<le> b"
shows "a \<le> a + b"
using neutr add_mono assms by force
lemma add_mono_neutl:
assumes "\<one> \<le> b"
shows "a \<le> b + a"
using neutr add_mono assms by force
lemma fw_step_0:
"m 0 0 \<ge> \<one> \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> fw m n 0 i j i j = min (m i j) (m i 0 + m 0 j)"
proof (induction i)
case 0 thus ?case
proof (cases j)
case 0 thus ?thesis by (simp add: fw_upd_def upd_def)
next
case (Suc j)
hence "fw m n 0 0 j 0 (Suc j) = m 0 (Suc j)" using 0 fw_middle_id[of 0 n "Suc j" j 0 m] by fast
moreover have "fw m n 0 0 j 0 0 = m 0 0" using single_iteration_inv[of 0 0 0 j n m 0] Suc 0
by (auto simp add: fw_upd_def upd_def min_def intro: add_mono_neutl)
ultimately show ?thesis using Suc by (simp add: fw_upd_def upd_def)
qed
next
case (Suc i)
note A = this
show ?case
proof (cases j)
case 0
have "fw m n 0 i n (Suc i) 0 = m (Suc i) 0" using fw_innermost_id[of "Suc i" n 0 n i m] Suc by simp
moreover have "fw m n 0 i n 0 0 = m 0 0" using Suc single_iteration_inv[of 0 i 0 n n m 0]
by (auto simp add: fw_upd_def upd_def min_def intro: add_mono_neutl)
ultimately show ?thesis using 0 by (simp add: fw_upd_def upd_def)
next
case (Suc j)
have *: "fw m n 0 0 j 0 0 = m 0 0" using single_iteration_inv[ of 0 0 0 j n m 0] A Suc
by (auto simp add: fw_upd_def upd_def min_def intro: add_mono_neutl)
have **: "fw m n 0 i n 0 0 = m 0 0" using single_iteration_inv[of 0 i 0 n n m 0] A
by (auto simp add: fw_upd_def upd_def min_def intro: add_mono_neutl)
have "m 0 (Suc j) = fw_upd m 0 0 (Suc j) 0 (Suc j)" using \<open>m 0 0 >= \<one>\<close>
by (auto simp add: fw_upd_def upd_def min_def intro: add_mono_neutl)
also have "\<dots> = fw m n 0 0 (Suc j) 0 (Suc j)" using fw_middle_id[of 0 n "Suc j" j 0 m] Suc A(4)
by (simp add: fw_upd_def upd_def *)
finally have ***: "fw m n 0 (Suc i) j 0 (Suc j) = m 0 (Suc j)"
using single_iteration_inv'[of 0 "Suc i" "Suc j" n j m 0] A Suc by simp
have "m (Suc i) 0 = fw_upd m 0 (Suc i) 0 (Suc i) 0" using \<open>m 0 0 >= \<one>\<close>
by (auto simp add: fw_upd_def upd_def min_def intro: add_mono_neutr)
also have "\<dots> = fw m n 0 (Suc i) 0 (Suc i) 0"
using fw_innermost_id[of "Suc i" n 0 n i m] \<open>Suc i \<le> n\<close> ** by (simp add: fw_upd_def upd_def)
finally have "fw m n 0 (Suc i) j (Suc i) 0 = m (Suc i) 0"
using single_iteration_inv A Suc by auto
moreover have "fw m n 0 (Suc i) j (Suc i) (Suc j) = m (Suc i) (Suc j)"
using fw_middle_id A Suc by simp
ultimately show ?thesis using Suc *** by (simp add: fw_upd_def upd_def)
qed
qed
lemma fw_step_Suc:
"\<forall> k'\<le>n. fw m n k n n k' k' \<ge> \<one> \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> Suc k \<le> n
\<Longrightarrow> fw m n (Suc k) i j i j = min (fw m n k n n i j) (fw m n k n n i (Suc k) + fw m n k n n (Suc k) j)"
proof (induction i)
case 0 thus ?case
proof (cases j)
case 0 thus ?thesis by (simp add: fw_upd_def upd_def)
next
case (Suc j)
then have
"fw m n k n n 0 (Suc j) = fw m n (Suc k) 0 j 0 (Suc j)"
using 0(2-) Suc_innermost_id2' by simp
moreover have "fw m n (Suc k) 0 j 0 (Suc k) = fw m n k n n 0 (Suc k)"
proof (cases "j < Suc k")
case True thus ?thesis using 0 Suc_innermost_id2' by simp
next
case False
hence
"fw m n (Suc k) 0 k 0 (Suc k) = fw m n k n n 0 (Suc k)"
using 0(2-) Suc Suc_innermost_id2' by simp
moreover have "fw m n (Suc k) 0 k (Suc k) (Suc k) = fw m n k n n (Suc k) (Suc k)"
using 0(2-) Suc Suc_innermost_id2' by simp
moreover have "fw m n (Suc k) 0 j 0 (Suc k) = fw m n (Suc k) 0 (Suc k) 0 (Suc k)"
using False single_iteration_inv 0(2-) Suc by force
ultimately show ?thesis using 0(1)
by (auto simp add: fw_upd_def upd_def \<open>Suc k \<le> n\<close> min_def intro: add_mono_neutr)
qed
moreover have "fw m n k n n (Suc k) (Suc j) = fw m n (Suc k) 0 j (Suc k) (Suc j)"
using 0(2-) Suc Suc_innermost_id2' by simp
ultimately show ?thesis using Suc by (simp add: fw_upd_def upd_def)
qed
next
case (Suc i)
note A = this
show ?case
proof (cases j)
case 0
hence
"fw m n (Suc k) i n (Suc i) 0 = fw m n k n n (Suc i) 0"
using Suc_innermost_id1' \<open>Suc i \<le> n\<close> by simp
moreover have "fw m n (Suc k) i n (Suc i) (Suc k) = fw m n k n n (Suc i) (Suc k)"
using Suc_innermost_id1' A(3,5) by simp
moreover have "fw m n (Suc k) i n (Suc k) 0 = fw m n k n n (Suc k) 0"
proof (cases "i < Suc k")
case True thus ?thesis using Suc_innermost_id1' A(3,5) by simp
next
case False
have "fw m n (Suc k) k n (Suc k) (Suc k) = fw m n k n n (Suc k) (Suc k)"
using Suc_innermost_id1' \<open>Suc i \<le> n\<close> False by simp
moreover have "fw m n (Suc k) k n (Suc k) 0 = fw m n k n n (Suc k) 0"
using Suc_innermost_id1' \<open>Suc i \<le> n\<close> False by simp
moreover have "fw m n (Suc k) i n (Suc k) 0 = fw m n (Suc k) (Suc k) 0 (Suc k) 0"
using single_iteration_inv \<open>Suc i \<le> n\<close> False by simp
ultimately show ?thesis using Suc(2)
by (auto simp add: fw_upd_def upd_def \<open>Suc k \<le> n\<close> min_def intro: add_mono_neutl)
qed
ultimately show ?thesis using 0 by (simp add: fw_upd_def upd_def)
next
case (Suc j)
hence "fw m n (Suc k) (Suc i) j (Suc i) (Suc j) = fw m n k n n (Suc i) (Suc j)"
using Suc_innermost_id2' A(3,4) by simp
moreover have "fw m n (Suc k) (Suc i) j (Suc i) (Suc k) = fw m n k n n (Suc i) (Suc k)"
proof (cases "j < Suc k")
case True thus ?thesis using Suc A(3-) Suc_innermost_id2' by simp
next
case False
have *:"fw m n (Suc k) (Suc i) k (Suc i) (Suc k) = fw m n k n n (Suc i) (Suc k)"
using Suc_innermost_id2' A(3,5) by simp
have **:"fw m n (Suc k) (Suc i) k (Suc k) (Suc k) = fw m n k n n (Suc k) (Suc k)"
proof (cases "Suc i \<le> Suc k")
case True thus ?thesis using Suc_innermost_id2' A(5) by simp
next
case False
hence "fw m n (Suc k) (Suc i) k (Suc k) (Suc k) = fw m n (Suc k) (Suc k) (Suc k) (Suc k) (Suc k)"
using single_iteration_inv'[of "Suc k" "Suc i" "Suc k" n k m "Suc k"] A(3) by simp
moreover have "fw m n (Suc k) (Suc k) k (Suc k) (Suc k) = fw m n k n n (Suc k) (Suc k)"
using Suc_innermost_id2' A(5) by simp
ultimately show ?thesis using A(2)
by (auto simp add: fw_upd_def upd_def \<open>Suc k \<le> n\<close> min_def intro: add_mono_neutl)
qed
have "fw m n (Suc k) (Suc i) j (Suc i) (Suc k) = fw m n (Suc k) (Suc i) (Suc k) (Suc i) (Suc k)"
using False single_iteration_inv[of "Suc i" "Suc i" "Suc k" j n m "Suc k"] A(3-) Suc by simp
also have "\<dots> = fw m n k n n (Suc i) (Suc k)" using * ** A(2)
by (auto simp add: fw_upd_def upd_def \<open>Suc k \<le> n\<close> min_def intro: add_mono_neutr)
finally show ?thesis by simp
qed
moreover have "fw m n (Suc k) (Suc i) j (Suc k) (Suc j) = fw m n k n n (Suc k) (Suc j)"
proof (cases "Suc i \<le> Suc k")
case True thus ?thesis using Suc_innermost_id2' Suc A(3-5) by simp
next
case False
have "fw m n (Suc k) (Suc k) j (Suc k) (Suc k) = fw m n k n n (Suc k) (Suc k)"
proof (cases "j < Suc k")
case True thus ?thesis using Suc_innermost_id2' A(5) by simp
next
case False
hence "fw m n (Suc k) (Suc k) j (Suc k) (Suc k) = fw m n (Suc k) (Suc k) (Suc k) (Suc k) (Suc k)"
using single_iteration_inv A(3,4) Suc by simp
moreover have "fw m n (Suc k) (Suc k) k (Suc k) (Suc k) = fw m n k n n (Suc k) (Suc k)"
using Suc_innermost_id2' A(5) by simp
ultimately show ?thesis using A(2)
by (auto simp add: fw_upd_def upd_def \<open>Suc k \<le> n\<close> min_def intro: add_mono_neutl)
qed
moreover have "fw m n (Suc k) (Suc k) j (Suc k) (Suc j) = fw m n k n n (Suc k) (Suc j)"
using Suc_innermost_id2' Suc A(3-5) by simp
ultimately have "fw m n (Suc k) (Suc k) (Suc j) (Suc k) (Suc j) = fw m n k n n (Suc k) (Suc j)"
using A(2) by (auto simp add: fw_upd_def upd_def \<open>Suc k \<le> n\<close> min_def intro: add_mono_neutl)
moreover have "fw m n (Suc k) (Suc i) j (Suc k) (Suc j) = fw m n (Suc k) (Suc k) (Suc j) (Suc k) (Suc j)"
using single_iteration_inv'[of "Suc k" "Suc i" "Suc j" n j m "Suc k"] Suc A(3-) False by simp
moreover have "fw m n (Suc k) (Suc k) k (Suc k) (Suc k) = fw m n k n n (Suc k) (Suc k)"
using Suc_innermost_id2' A(5) by simp
ultimately show ?thesis using A(2) by (simp add: fw_upd_def upd_def)
qed
ultimately show ?thesis using Suc by (simp add: fw_upd_def upd_def)
qed
qed
subsection \<open>Length of Paths\<close>
fun len :: "'a mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> 'a" where
"len m u v [] = m u v" |
"len m u v (w#ws) = m u w + len m w v ws"
lemma len_decomp: "xs = ys @ y # zs \<Longrightarrow> len m x z xs = len m x y ys + len m y z zs"
by (induction ys arbitrary: x xs) (simp add: assoc)+
lemma len_comp: "len m a c (xs @ b # ys) = len m a b xs + len m b c ys"
by (induction xs arbitrary: a) (auto simp: assoc)
subsection \<open>Shortening Negative Cycles\<close>
lemma remove_cycles_neg_cycles_aux:
fixes i xs ys
defines "xs' \<equiv> i # ys"
assumes "i \<notin> set ys"
assumes "i \<in> set xs"
assumes "xs = as @ concat (map ((#) i) xss) @ xs'"
assumes "len m i j ys > len m i j xs"
shows "\<exists> ys. set ys \<subseteq> set xs \<and> len m i i ys < \<one>" using assms
proof (induction xss arbitrary: xs as)
case Nil
with Nil show ?case
proof (cases "len m i i as \<ge> \<one>", goal_cases)
case 1
from this(4,6) len_decomp[of xs as i ys m i j]
have "len m i j xs = len m i i as + len m i j ys" by simp
with 1(11)
have "len m i j ys \<le> len m i j xs" using add_mono by fastforce
thus ?thesis using Nil(5) by auto
next
case 2 thus ?case by auto
qed
next
case (Cons zs xss)
let ?xs = "zs @ concat (map ((#) i) xss) @ xs'"
from Cons show ?case
proof (cases "len m i i as \<ge> \<one>", goal_cases)
case 1
from this(5,7) len_decomp add_mono
have "len m i j ?xs \<le> len m i j xs" by fastforce
hence 4:"len m i j ?xs < len m i j ys" using 1(6) by simp
have 2:"i \<in> set ?xs" using Cons(2) by auto
have "set ?xs \<subseteq> set xs" using Cons(5) by auto
moreover from Cons(1)[OF 1(2,3) 2 _ 4] have "\<exists>ys. set ys \<subseteq> set ?xs \<and> len m i i ys < \<one>" by auto
ultimately show ?case by blast
next
case 2
from this(5,7) show ?case by auto
qed
qed
lemma add_lt_neutral: "a + b < b \<Longrightarrow> a < \<one>"
proof (rule ccontr)
assume "a + b < b" "\<not> a < \<one>"
hence "a \<ge> \<one>" by auto
from add_mono[OF this, of b b] \<open>a + b < b\<close> show False by auto
qed
lemma remove_cycles_neg_cycles_aux':
fixes j xs ys
assumes "j \<notin> set ys"
assumes "j \<in> set xs"
assumes "xs = ys @ j # concat (map (\<lambda> xs. xs @ [j]) xss) @ as"
assumes "len m i j ys > len m i j xs"
shows "\<exists> ys. set ys \<subseteq> set xs \<and> len m j j ys < \<one>" using assms
proof (induction xss arbitrary: xs as)
case Nil
show ?case
proof (cases "len m j j as \<ge> \<one>")
case True
from Nil(3) len_decomp[of xs ys j as m i j]
have "len m i j xs = len m i j ys + len m j j as" by simp
with True
have "len m i j ys \<le> len m i j xs" using add_mono by fastforce
with Nil show ?thesis by auto
next
case False with Nil show ?thesis by auto
qed
next
case (Cons zs xss)
let ?xs = "ys @ j # concat (map (\<lambda>xs. xs @ [j]) xss) @ as"
let ?t = "concat (map (\<lambda>xs. xs @ [j]) xss) @ as"
show ?case
proof (cases "len m i j ?xs \<le> len m i j xs")
case True
hence 4:"len m i j ?xs < len m i j ys" using Cons(5) by simp
have 2:"j \<in> set ?xs" using Cons(2) by auto
have "set ?xs \<subseteq> set xs" using Cons(4) by auto
moreover from Cons(1)[OF Cons(2) 2 _ 4] have "\<exists>ys. set ys \<subseteq> set ?xs \<and> len m j j ys < \<one>" by blast
ultimately show ?thesis by blast
next
case False
hence "len m i j xs < len m i j ?xs" by auto
from this len_decomp Cons(4) add_mono
have "len m j j (concat (map (\<lambda>xs. xs @ [j]) (zs # xss)) @ as) < len m j j ?t"
using False local.leI by fastforce
hence "len m j j (zs @ j # ?t) < len m j j ?t" by simp
with len_decomp[of "zs @ j # ?t" zs j ?t m j j]
have "len m j j zs + len m j j ?t < len m j j ?t" by auto
hence "len m j j zs < \<one>" using add_lt_neutral by auto
thus ?thesis using Cons.prems(3) by auto
qed
qed
lemma add_le_impl: "a + b < a + c \<Longrightarrow> b < c"
proof (rule ccontr)
assume "a + b < a + c" "\<not> b < c"
hence "b \<ge> c" by auto
from add_mono[OF _ this, of a a] \<open>a + b < a + c\<close> show False by auto
qed
lemma start_remove_neg_cycles:
"len m i j (start_remove xs k []) > len m i j xs \<Longrightarrow> \<exists> ys. set ys \<subseteq> set xs \<and> len m k k ys < \<one>"
proof-
let ?xs = "start_remove xs k []"
assume len_lt:"len m i j ?xs > len m i j xs"
hence "k \<in> set xs" using start_remove_id by fastforce
from start_remove_decomp[OF this, of "[]"] obtain as bs where as_bs:
"xs = as @ k # bs" "?xs = as @ remove_cycles bs k [k]"
by fastforce
let ?xs' = "remove_cycles bs k [k]"
have "k \<in> set bs" using as_bs len_lt remove_cycles_id by fastforce
then obtain ys where ys: "?xs = as @ k # ys" "?xs' = k # ys" "k \<notin> set ys"
using as_bs(2) remove_cycles_begins_with[OF \<open>k \<in> set bs\<close>] by auto
have len_lt': "len m k j bs < len m k j ys"
using len_decomp[OF as_bs(1), of m i j] len_decomp[OF ys(1), of m i j] len_lt add_le_impl by metis
from remove_cycles_cycles[OF \<open>k \<in> set bs\<close>] obtain xss as'
where "as' @ concat (map ((#) k) xss) @ ?xs' = bs" by fastforce
hence "as' @ concat (map ((#) k) xss) @ k # ys = bs" using ys(2) by simp
from remove_cycles_neg_cycles_aux[OF \<open>k \<notin> set ys\<close> \<open>k \<in> set bs\<close> this[symmetric] len_lt']
show ?thesis using as_bs(1) by auto
qed
lemma remove_all_cycles_neg_cycles:
"len m i j (remove_all_cycles ys xs) > len m i j xs
\<Longrightarrow> \<exists> ys k. set ys \<subseteq> set xs \<and> k \<in> set xs \<and> len m k k ys < \<one>"
proof (induction ys arbitrary: xs)
case Nil thus ?case by auto
next
case (Cons y ys)
let ?xs = "start_remove xs y []"
show ?case
proof (cases "len m i j xs < len m i j ?xs")
case True
with start_remove_id have "y \<in> set xs" by fastforce
with start_remove_neg_cycles[OF True] show ?thesis by blast
next
case False
with Cons(2) have "len m i j ?xs < len m i j (remove_all_cycles (y # ys) xs)" by auto
hence "len m i j ?xs < len m i j (remove_all_cycles ys ?xs)" by auto
from Cons(1)[OF this] show ?thesis using start_remove_subs[of xs y "[]"] by auto
qed
qed
lemma (in -) concat_map_cons_rev:
"rev (concat (map ((#) j) xss)) = concat (map (\<lambda> xs. xs @ [j]) (rev (map rev xss)))"
by (induction xss) auto
lemma negative_cycle_dest: "len m i j (rem_cycles i j xs) > len m i j xs
\<Longrightarrow> \<exists> i' ys. len m i' i' ys < \<one> \<and> set ys \<subseteq> set xs \<and> i' \<in> set (i # j # xs)"
proof -
let ?xsij = "rem_cycles i j xs"
let ?xsj = "remove_all_rev j (remove_all_cycles xs xs)"
let ?xs = "remove_all_cycles xs xs"
assume len_lt: "len m i j ?xsij > len m i j xs"
show ?thesis
proof (cases "len m i j ?xsij \<le> len m i j ?xsj")
case True
hence len_lt: "len m i j ?xsj > len m i j xs" using len_lt by simp
show ?thesis
proof (cases "len m i j ?xsj \<le> len m i j ?xs")
case True
hence "len m i j ?xs > len m i j xs" using len_lt by simp
with remove_all_cycles_neg_cycles[OF this] show ?thesis by auto
next
case False
then have len_lt': "len m i j ?xsj > len m i j ?xs" by simp
show ?thesis
proof (cases "j \<in> set ?xs")
case False
thus ?thesis using len_lt' by (simp add: remove_all_rev_def)
next
case True
from remove_all_rev_removes[of j] have 1: "j \<notin> set ?xsj" by simp
from True have "j \<in> set (rev ?xs)" by auto
from remove_cycles_cycles[OF this] obtain xss as where as:
"as @ concat (map ((#) j) xss) @ remove_cycles (rev ?xs) j [] = rev ?xs" "j \<notin> set as"
by blast
from True have "?xsj = rev (tl (remove_cycles (rev ?xs) j []))" by (simp add: remove_all_rev_def)
with remove_cycles_begins_with[OF \<open>j \<in> set (rev ?xs)\<close>, of "[]"]
have "remove_cycles (rev ?xs) j [] = j # rev ?xsj" "j \<notin> set ?xsj"
by auto
with as(1) have xss: "as @ concat (map ((#) j) xss) @ j # rev ?xsj = rev ?xs" by simp
hence "rev (as @ concat (map ((#) j) xss) @ j # rev ?xsj) = ?xs" by simp
hence "?xsj @ j # rev (concat (map ((#) j) xss)) @ rev as = ?xs" by simp
hence "?xsj @ j # concat (map (\<lambda> xs. xs @ [j]) (rev (map rev xss))) @ rev as = ?xs"
by (simp add: concat_map_cons_rev)
from remove_cycles_neg_cycles_aux'[OF 1 True this[symmetric] len_lt']
show ?thesis using remove_all_cycles_subs by fastforce
qed
qed
next
case False
hence len_lt': "len m i j ?xsij > len m i j ?xsj" by simp
show ?thesis
proof (cases "i \<in> set ?xsj")
case False
thus ?thesis using len_lt' by (simp add: remove_all_def)
next
case True
from remove_all_removes[of i] have 1: "i \<notin> set ?xsij" by (simp add: remove_all_def)
from remove_cycles_cycles[OF True] obtain xss as where as:
"as @ concat (map ((#) i) xss) @ remove_cycles ?xsj i [] = ?xsj" "i \<notin> set as" by blast
from True have "?xsij = tl (remove_cycles ?xsj i [])" by (simp add: remove_all_def)
with remove_cycles_begins_with[OF True, of "[]"]
have "remove_cycles ?xsj i [] = i # ?xsij" "i \<notin> set ?xsij"
by auto
with as(1) have xss: "as @ concat (map ((#) i) xss) @ i # ?xsij = ?xsj" by simp
from remove_cycles_neg_cycles_aux[OF 1 True this[symmetric] len_lt']
show ?thesis using remove_all_rev_subs remove_all_cycles_subs by fastforce
qed
qed
qed
section \<open>Definition of Shortest Paths\<close>
definition D :: "'a mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where
"D m i j k \<equiv> Min {len m i j xs | xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
lemma (in -) distinct_length_le:"finite s \<Longrightarrow> set xs \<subseteq> s \<Longrightarrow> distinct xs \<Longrightarrow> length xs \<le> card s"
by (metis card_mono distinct_card)
lemma (in -) finite_distinct: "finite s \<Longrightarrow> finite {xs . set xs \<subseteq> s \<and> distinct xs}"
proof -
assume "finite s"
hence "{xs . set xs \<subseteq> s \<and> distinct xs} \<subseteq> {xs. set xs \<subseteq> s \<and> length xs \<le> card s}"
using distinct_length_le by auto
moreover have "finite {xs. set xs \<subseteq> s \<and> length xs \<le> card s}"
using finite_lists_length_le[OF \<open>finite s\<close>] by auto
ultimately show ?thesis by (rule finite_subset)
qed
lemma D_base_finite:
"finite {len m i j xs | xs. set xs \<subseteq> {0..k} \<and> distinct xs}"
using finite_distinct finite_image_set by blast
lemma D_base_finite':
"finite {len m i j xs | xs. set xs \<subseteq> {0..k} \<and> distinct (i # j # xs)}"
proof -
have "{len m i j xs | xs. set xs \<subseteq> {0..k} \<and> distinct (i # j # xs)}
\<subseteq> {len m i j xs | xs. set xs \<subseteq> {0..k} \<and> distinct xs}" by auto
with D_base_finite[of m i j k] show ?thesis by (rule rev_finite_subset)
qed
lemma D_base_finite'':
"finite {len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
using D_base_finite[of m i j k] by - (rule finite_subset, auto)
definition cycle_free :: "'a mat \<Rightarrow> nat \<Rightarrow> bool" where
"cycle_free m n \<equiv> \<forall> i xs. i \<le> n \<and> set xs \<subseteq> {0..n} \<longrightarrow>
(\<forall> j. j \<le> n \<longrightarrow> len m i j (rem_cycles i j xs) \<le> len m i j xs) \<and> len m i i xs \<ge> \<one>"
lemma D_eqI:
fixes m n i j k
defines "A \<equiv> {len m i j xs | xs. set xs \<subseteq> {0..k}}"
defines "A_distinct \<equiv> {len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
assumes "cycle_free m n" "i \<le> n" "j \<le> n" "k \<le> n" "(\<And>y. y \<in> A_distinct \<Longrightarrow> x \<le> y)" "x \<in> A"
shows "D m i j k = x" using assms
proof -
let ?S = "{len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
show ?thesis unfolding D_def
proof (rule Min_eqI)
have "?S \<subseteq> {len m i j xs |xs. set xs \<subseteq> {0..k} \<and> distinct xs}" by auto
thus "finite {len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
using D_base_finite[of m i j k] by (rule finite_subset)
next
fix y assume "y \<in> ?S"
hence "y \<in> A_distinct" using assms(2,7) by fastforce
thus "x \<le> y" using assms by meson
next
from assms obtain xs where xs: "x = len m i j xs" "set xs \<subseteq> {0..k}" by auto
let ?ys = "rem_cycles i j xs"
let ?y = "len m i j ?ys"
from assms(3-6) xs have *:"?y \<le> x" by (fastforce simp add: cycle_free_def)
have distinct: "i \<notin> set ?ys" "j \<notin> set ?ys" "distinct ?ys"
using rem_cycles_distinct remove_all_removes rem_cycles_removes_last by fast+
with xs(2) have "?y \<in> A_distinct" unfolding A_distinct_def using rem_cycles_subs by fastforce
hence "x \<le> ?y" using assms by meson
moreover have "?y \<le> x" using assms(3-6) xs by (fastforce simp add: cycle_free_def)
ultimately have "x = ?y" by simp
thus "x \<in> ?S" using distinct xs(2) rem_cycles_subs[of i j xs] by fastforce
qed
qed
lemma D_base_not_empty:
"{len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs} \<noteq> {}"
proof -
have "len m i j [] \<in> {len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
by fastforce
thus ?thesis by auto
qed
lemma Min_elem_dest: "finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> x = Min A \<Longrightarrow> x \<in> A" by simp
lemma D_dest: "x = D m i j k \<Longrightarrow>
x \<in> {len m i j xs |xs. set xs \<subseteq> {0..Suc k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
using Min_elem_dest[OF D_base_finite'' D_base_not_empty] by (fastforce simp add: D_def)
lemma D_dest': "x = D m i j k \<Longrightarrow> x \<in> {len m i j xs |xs. set xs \<subseteq> {0..Suc k}}"
using Min_elem_dest[OF D_base_finite'' D_base_not_empty] by (fastforce simp add: D_def)
lemma D_dest'': "x = D m i j k \<Longrightarrow> x \<in> {len m i j xs |xs. set xs \<subseteq> {0..k}}"
using Min_elem_dest[OF D_base_finite'' D_base_not_empty] by (fastforce simp add: D_def)
lemma cycle_free_loop_dest: "i \<le> n \<Longrightarrow> set xs \<subseteq> {0..n} \<Longrightarrow> cycle_free m n \<Longrightarrow> len m i i xs \<ge> \<one>"
unfolding cycle_free_def by auto
lemma cycle_free_dest:
"cycle_free m n \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> set xs \<subseteq> {0..n}
\<Longrightarrow> len m i j (rem_cycles i j xs) \<le> len m i j xs"
by (auto simp add: cycle_free_def)
definition cycle_free_up_to :: "'a mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
"cycle_free_up_to m k n \<equiv> \<forall> i xs. i \<le> n \<and> set xs \<subseteq> {0..k} \<longrightarrow>
(\<forall> j. j \<le> n \<longrightarrow> len m i j (rem_cycles i j xs) \<le> len m i j xs) \<and> len m i i xs \<ge> \<one>"
lemma cycle_free_up_to_loop_dest:
"i \<le> n \<Longrightarrow> set xs \<subseteq> {0..k} \<Longrightarrow> cycle_free_up_to m k n \<Longrightarrow> len m i i xs \<ge> \<one>"
unfolding cycle_free_up_to_def by auto
lemma cycle_free_up_to_diag:
assumes "cycle_free_up_to m k n" "i \<le> n"
shows "m i i \<ge> \<one>"
using cycle_free_up_to_loop_dest[OF assms(2) _ assms(1), of "[]"] by auto
lemma D_eqI2:
fixes m n i j k
defines "A \<equiv> {len m i j xs | xs. set xs \<subseteq> {0..k}}"
defines "A_distinct \<equiv> {len m i j xs | xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
assumes "cycle_free_up_to m k n" "i \<le> n" "j \<le> n" "k \<le> n"
"(\<And>y. y \<in> A_distinct \<Longrightarrow> x \<le> y)" "x \<in> A"
shows "D m i j k = x" using assms
proof -
show ?thesis
proof (simp add: D_def A_distinct_def[symmetric], rule Min_eqI)
show "finite A_distinct" using D_base_finite''[of m i j k] unfolding A_distinct_def by auto
next
fix y assume "y \<in> A_distinct"
thus "x \<le> y" using assms by meson
next
from assms obtain xs where xs: "x = len m i j xs" "set xs \<subseteq> {0..k}" by auto
let ?ys = "rem_cycles i j xs"
let ?y = "len m i j ?ys"
from assms(3-6) xs have *:"?y \<le> x" by (fastforce simp add: cycle_free_up_to_def)
have distinct: "i \<notin> set ?ys" "j \<notin> set ?ys" "distinct ?ys"
using rem_cycles_distinct remove_all_removes rem_cycles_removes_last by fast+
with xs(2) have "?y \<in> A_distinct" unfolding A_distinct_def using rem_cycles_subs by fastforce
hence "x \<le> ?y" using assms by meson
moreover have "?y \<le> x" using assms(3-6) xs by (fastforce simp add: cycle_free_up_to_def)
ultimately have "x = ?y" by simp
then show "x \<in> A_distinct" using distinct xs(2) rem_cycles_subs[of i j xs]
unfolding A_distinct_def by fastforce
qed
qed
section \<open>Result Under The Absence of Negative Cycles\<close>
text \<open>
This proves that the algorithm correctly computes shortest paths under the absence of negative
cycles by a standard argument.
\<close>
theorem fw_shortest_path_up_to:
"cycle_free_up_to m k n \<Longrightarrow> i' \<le> i \<Longrightarrow> j' \<le> j \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> k \<le> n
\<Longrightarrow> D m i' j' k = fw m n k i j i' j'"
proof (induction k arbitrary: i j i' j')
case 0
from cycle_free_up_to_diag[OF 0(1)] have diag: "\<forall> k \<le> n. m k k \<ge> \<one>" by auto
then have m_diag: "m 0 0 \<ge> \<one>" by simp
let ?S = "{len m i' j' xs |xs. set xs \<subseteq> {0} \<and> i' \<notin> set xs \<and> j' \<notin> set xs \<and> distinct xs}"
show ?case unfolding D_def
proof (simp, rule Min_eqI)
have "?S \<subseteq> {len m i' j' xs |xs. set xs \<subseteq> {0..0} \<and> distinct xs}" by auto
thus "finite ?S" using D_base_finite[of m i' j' 0] by (rule finite_subset)
next
fix l assume "l \<in> ?S"
then obtain xs where l: "l = len m i' j' xs" and xs: "xs = [] \<or> xs = [0]"
using distinct_list_single_elem_decomp by auto
{ assume "xs = []"
have "fw m n 0 i j i' j' \<le> fw m n 0 0 0 i' j'" using fw_invariant 0 by blast
also have "\<dots> \<le> m i' j'" by (cases "i' = 0 \<and> j' = 0") (simp add: fw_upd_def upd_def)+
finally have "fw m n 0 i j i' j' \<le> l" using \<open>xs = []\<close> l by simp
}
moreover
{ assume "xs = [0]"
have "fw m n 0 i j i' j' \<le> fw m n 0 i' j' i' j'" using fw_invariant 0 by blast
also have "\<dots> \<le> m i' 0 + m 0 j'"
proof (cases j')
assume "j' = 0"
show ?thesis
proof (cases i')
assume "i' = 0"
thus ?thesis using \<open>j' = 0\<close> by (simp add: fw_upd_def upd_def)
next
fix i'' assume i'': "i' = Suc i''"
have "fw_upd (fw m n 0 i'' n) 0 (Suc i'') 0 (Suc i'') 0 \<le> fw m n 0 i'' n (Suc i'') 0"
by (simp add: fw_upd_mono)
also have "\<dots> \<le> m (Suc i'') 0" using fw_mono 0 i'' by simp
finally show ?thesis using \<open>j' = 0\<close> m_diag i'' neutr add_mono by fastforce
qed
next
fix j'' assume j'': "j' = Suc j''"
have "fw_upd (fw m n 0 i' j'') 0 i' (Suc j'') i' (Suc j'')
\<le> fw m n 0 i' j'' i' 0 + fw m n 0 i' j'' 0 (Suc j'') "
by (simp add: fw_upd_def upd_def)
also have "\<dots> \<le> m i' 0 + m 0 (Suc j'')"
using fw_mono[of i' n j'' i' 0 m 0] fw_mono[of i' n j'' 0 "Suc j''" m 0 ] j'' 0
by (simp add: add_mono)
finally show ?thesis using j'' by simp
qed
finally have "fw m n 0 i j i' j' \<le> l" using \<open>xs = [0]\<close> l by simp
}
ultimately show "fw m n 0 i j i' j' \<le> l" using xs by auto
next
have A: "fw m n 0 i j i' j' = fw m n 0 i' j' i' j'" using single_iteration_inv 0 by blast
have "fw m n 0 i' j' i' j' = min (m i' j') (m i' 0 + m 0 j')"
using 0 by (simp add: fw_step_0[of m, OF m_diag])
hence
"fw m n 0 i' j' i' j' = m i' j'
\<or> (fw m n 0 i' j' i' j' = m i' 0 + m 0 j'\<and> m i' 0 + m 0 j' \<le> m i' j')"
by (auto simp add: ord.min_def)
thus "fw m n 0 i j i' j' \<in> ?S"
proof (standard, goal_cases)
case 1
hence "fw m n 0 i j i' j' = len m i' j' []" using A by auto
thus ?case by fastforce
next
case 2
hence *:"fw m n 0 i j i' j' = len m i' j' [0]" using A by auto
thus ?case
proof (cases "i' = 0 \<or> j' = 0")
case False thus ?thesis using * by fastforce
next
case True
{ assume "i' = 0"
from diag have "m 0 0 + m 0 j' \<ge> m 0 j'" by (auto intro: add_mono_neutl)
with \<open>i' = 0\<close> have "fw m n 0 i j i' j' = len m 0 j' []" using 0 A 2 by auto
} moreover
{ assume "j' = 0"
from diag have "m i' 0 + m 0 0 \<ge> m i' 0" by (auto intro: add_mono_neutr)
with \<open>j' = 0\<close> have "fw m n 0 i j i' j' = len m i' 0 []" using 0 A 2 by auto
}
ultimately have "fw m n 0 i j i' j' = len m i' j' []" using True by auto
then show ?thesis by fastforce
qed
qed
qed
next
case (Suc k)
from cycle_free_up_to_diag[OF Suc.prems(1)] have diag: "\<forall> k \<le> n. m k k \<ge> \<one>" by auto
from Suc.prems have cycle_free_to_k:
"cycle_free_up_to m k n" by (fastforce simp add: cycle_free_up_to_def)
{ fix k' assume "k' \<le> n"
with Suc cycle_free_to_k have "D m k' k' k = fw m n k n n k' k'" by auto
from D_dest''[OF this[symmetric]] obtain xs where
"set xs \<subseteq> {0..k}" "fw m n k n n k' k'= len m k' k' xs"
by auto
with Suc(2) \<open>Suc k \<le> n\<close> \<open>k' \<le> n\<close> have "fw m n k n n k' k' \<ge> \<one>"
unfolding cycle_free_up_to_def by force
}
hence K: "\<forall>k'\<le>n. fw m n k n n k' k' \<ge> \<one>" by simp
let ?S = "\<lambda> k i j. {len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
show ?case
proof (rule D_eqI2)
show "cycle_free_up_to m (Suc k) n" using Suc.prems(1) .
next
show "i' \<le> n" using Suc.prems by simp
next
show "j' \<le> n" using Suc.prems by simp
next
show "Suc k \<le> n" using Suc.prems by simp
next
fix l assume "l \<in> {len m i' j' xs | xs. set xs \<subseteq> {0..Suc k} \<and> i' \<notin> set xs \<and> j' \<notin> set xs \<and> distinct xs}"
then obtain xs where xs:
"l = len m i' j' xs" "set xs \<subseteq> {0..Suc k}" "i' \<notin> set xs" "j' \<notin> set xs" "distinct xs"
by auto
have IH: "D m i' j' k = fw m n k i j i' j'" using cycle_free_to_k Suc by auto
have fin:
"finite {len m i' j' xs |xs. set xs \<subseteq> {0..k} \<and> i' \<notin> set xs \<and> j' \<notin> set xs \<and> distinct xs}"
using D_base_finite'' by simp
show "fw m n (Suc k) i j i' j' \<le> l"
proof (cases "Suc k \<in> set xs")
case False
hence "set xs \<subseteq> {0..k}" using xs(2) using atLeastAtMostSuc_conv by auto
hence
"l \<in> {len m i' j' xs | xs. set xs \<subseteq> {0..k} \<and> i' \<notin> set xs \<and> j' \<notin> set xs \<and> distinct xs}"
using xs by auto
with Min_le[OF fin this] have "fw m n k i j i' j' \<le> l" using IH by (simp add: D_def)
thus ?thesis using fw_invariant[of k "Suc k" i n j j i m i' j'] Suc.prems by simp
next
case True
then obtain ys zs where ys_zs_id: "xs = ys @ Suc k # zs" by (meson split_list)
with xs(5) have ys_zs: "distinct ys" "distinct zs" "Suc k \<notin> set ys" "Suc k \<notin> set zs"
"set ys \<inter> set zs = {}" by auto
have "i' \<noteq> Suc k" "j' \<noteq> Suc k" using xs(3,4) True by auto
have "set ys \<subseteq> {0..k}" using ys_zs(3) xs(2) ys_zs_id using atLeastAtMostSuc_conv by auto
hence "len m i' (Suc k) ys \<in> ?S k i' (Suc k)" using ys_zs_id ys_zs xs(3) by fastforce
with Min_le[OF _ this] have "Min (?S k i' (Suc k)) \<le> len m i' (Suc k) ys"
using D_base_finite'[of m i' "Suc k" k] \<open>i' \<noteq> Suc k\<close> by fastforce
moreover have "fw m n k n n i' (Suc k) = D m i' (Suc k) k"
using Suc.IH[OF cycle_free_to_k, of i' n] Suc.prems by auto
ultimately have *:"fw m n k n n i' (Suc k) \<le> len m i' (Suc k) ys" using \<open>i' \<noteq> Suc k\<close>
by (auto simp: D_def)
have "set zs \<subseteq> {0..k}" using ys_zs(4) xs(2) ys_zs_id using atLeastAtMostSuc_conv by auto
hence "len m (Suc k) j' zs \<in> ?S k (Suc k) j'" using ys_zs_id ys_zs xs(3,4,5) by fastforce
with Min_le[OF _ this] have "Min (?S k (Suc k) j') \<le> len m (Suc k) j' zs"
using D_base_finite'[of m "Suc k" j' k] \<open>j' \<noteq> Suc k\<close> by fastforce
moreover have "fw m n k n n (Suc k) j' = D m (Suc k) j' k"
using Suc.IH[OF cycle_free_to_k, of "Suc k" n j' n] Suc.prems by auto
ultimately have **:"fw m n k n n (Suc k) j' \<le> len m (Suc k) j' zs" using \<open>j' \<noteq> Suc k\<close>
by (auto simp: D_def)
have len_eq: "l = len m i' (Suc k) ys + len m (Suc k) j' zs"
by (simp add: xs(1) len_decomp[OF ys_zs_id, symmetric] ys_zs_id)
have "fw m n (Suc k) i' j' i' j' \<le> fw m n k n n i' (Suc k) + fw m n k n n (Suc k) j'"
using fw_step_Suc[of n m k i' j', OF K] Suc.prems(2-) by simp
hence "fw m n (Suc k) i' j' i' j' \<le> l"
using fw_step_Suc[of n m k i j] Suc.prems(3-) * ** len_eq add_mono by fastforce
thus ?thesis using fw_invariant[of "Suc k" "Suc k" i n j j' i' m i' j'] Suc.prems(2-) by simp
qed
next
have "fw m n (Suc k) i j i' j' = fw m n (Suc k) i' j' i' j'"
using single_iteration_inv[OF Suc.prems(2-5)] .
also have "\<dots> = min (fw m n k n n i' j') (fw m n k n n i' (Suc k) + fw m n k n n (Suc k) j')"
using fw_step_Suc[OF K] Suc.prems(2-) by simp
finally show "fw m n (Suc k) i j i' j' \<in> {len m i' j' xs | xs. set xs \<subseteq> {0..Suc k}}"
proof (cases "fw m n (Suc k) i j i' j' = fw m n k n n i' j'", goal_cases)
case True
have "fw m n (Suc k) i j i' j' = D m i' j' k"
using Suc.IH[OF cycle_free_to_k, of i' n j' n] Suc.prems(2-) True by simp
from D_dest'[OF this] show ?thesis by blast
next
case 2
hence A:"fw m n (Suc k) i j i' j' = fw m n k n n i' (Suc k) + fw m n k n n (Suc k) j'"
by (metis ord.min_def)
have "fw m n k n n i' j' = D m i' j' k"
using Suc.IH[OF cycle_free_to_k, of i' n j' n] Suc.prems by simp
from D_dest[OF this] have B:"fw m n k n n i' j' \<in> ?S (Suc k) i' j'"
by blast
have "fw m n k n n i' (Suc k) = D m i' (Suc k) k"
using Suc.IH[OF cycle_free_to_k, of i' n "Suc k" n] Suc.prems by simp
from D_dest'[OF this] obtain xs where xs:
"fw m n k n n i' (Suc k) = len m i' (Suc k) xs" "set xs \<subseteq> {0..Suc k}" by blast
have "fw m n k n n (Suc k) j' = D m (Suc k) j' k"
using Suc.IH[OF cycle_free_to_k, of "Suc k" n j' n] Suc.prems by simp
from D_dest'[OF this] obtain ys where ys:
"fw m n k n n (Suc k) j' = len m (Suc k) j' ys" "set ys \<subseteq> {0..Suc k}" by blast
from A xs(1) ys(1) len_comp
have "fw m n (Suc k) i j i' j' = len m i' j' (xs @ Suc k # ys)" by simp
moreover have "set (xs @ Suc k # ys) \<subseteq> {0..Suc k}" using xs(2) ys(2) by auto
ultimately show ?thesis by blast
qed
qed
qed
lemma cycle_free_cycle_free_up_to:
"cycle_free m n \<Longrightarrow> k \<le> n \<Longrightarrow> cycle_free_up_to m k n"
unfolding cycle_free_def cycle_free_up_to_def by force
lemma cycle_free_diag:
"cycle_free m n \<Longrightarrow> i \<le> n \<Longrightarrow> \<one> \<le> m i i"
using cycle_free_up_to_diag[OF cycle_free_cycle_free_up_to] by blast
corollary fw_shortest_path:
"cycle_free m n \<Longrightarrow> i' \<le> i \<Longrightarrow> j' \<le> j \<Longrightarrow> i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> k \<le> n
\<Longrightarrow> D m i' j' k = fw m n k i j i' j'"
using fw_shortest_path_up_to[OF cycle_free_cycle_free_up_to] by auto
corollary fw_shortest:
assumes "cycle_free m n" "i \<le> n" "j \<le> n" "k \<le> n"
shows "fw m n n n n i j \<le> fw m n n n n i k + fw m n n n n k j"
proof (rule ccontr, goal_cases)
case 1
let ?S = "\<lambda> i j. {len m i j xs |xs. set xs \<subseteq> {0..n}}"
let ?FW = "fw m n n n n"
from assms fw_shortest_path
have FW: "?FW i j = D m i j n" "?FW i k = D m i k n" "?FW k j = D m k j n" by auto
with D_dest'' FW have "?FW i k \<in> ?S i k" "?FW k j \<in> ?S k j" by auto
then obtain xs ys where xs_ys:
"?FW i k = len m i k xs" "set xs \<subseteq> {0..n}" "?FW k j = len m k j ys" "set ys \<subseteq> {0..n}" by auto
let ?zs = "rem_cycles i j (xs @ k # ys)"
have *:"?FW i j = Min {len m i j xs |xs. set xs \<subseteq> {0..n} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
using FW(1) unfolding D_def .
have "set (xs @ k # ys) \<subseteq> {0..n}" using assms xs_ys by fastforce
from cycle_free_dest [OF \<open>cycle_free m n\<close> \<open>i \<le> n\<close> \<open>j \<le> n\<close> this]
have **:"len m i j ?zs \<le> len m i j (xs @ k # ys)" by auto
moreover have "i \<notin> set ?zs" "j \<notin> set ?zs" "distinct ?zs"
using rem_cycles_distinct remove_all_removes rem_cycles_removes_last by fast+
moreover have "set ?zs \<subseteq> {0..n}" using rem_cycles_subs[of i j"xs @ k # ys"] xs_ys assms by fastforce
ultimately have
"len m i j ?zs \<in> {len m i j xs |xs. set xs \<subseteq> {0..n} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}"
by blast
with * have "?FW i j \<le> len m i j ?zs" using D_base_finite'' by auto
with ** xs_ys len_comp 1 show ?case by auto
qed
section \<open>Result Under the Presence of Negative Cycles\<close>
lemma not_cylce_free_dest: "\<not> cycle_free m n \<Longrightarrow> \<exists> k \<le> n. \<not> cycle_free_up_to m k n"
by (auto simp add: cycle_free_def cycle_free_up_to_def)
lemma D_not_diag_le:
"(x :: 'a) \<in> {len m i j xs |xs. set xs \<subseteq> {0..k} \<and> i \<notin> set xs \<and> j \<notin> set xs \<and> distinct xs}
\<Longrightarrow> D m i j k \<le> x" using Min_le[OF D_base_finite''] by (auto simp add: D_def)
lemma D_not_diag_le': "set xs \<subseteq> {0..k} \<Longrightarrow> i \<notin> set xs \<Longrightarrow> j \<notin> set xs \<Longrightarrow> distinct xs
\<Longrightarrow> D m i j k \<le> len m i j xs" using Min_le[OF D_base_finite''] by (fastforce simp add: D_def)
lemma (in -) nat_upto_subs_top_removal':
"S \<subseteq> {0..Suc n} \<Longrightarrow> Suc n \<notin> S \<Longrightarrow> S \<subseteq> {0..n}"
apply (induction n)
apply safe
apply (rename_tac x)
apply (case_tac "x = Suc 0"; fastforce)
apply (rename_tac n x)
apply (case_tac "x = Suc (Suc n)"; fastforce)
done
lemma (in -) nat_upto_subs_top_removal:
"S \<subseteq> {0..n::nat} \<Longrightarrow> n \<notin> S \<Longrightarrow> S \<subseteq> {0..n - 1}"
using nat_upto_subs_top_removal' by (cases n; simp)
lemma fw_Suc:
"i \<le> n \<Longrightarrow> j \<le> n \<Longrightarrow> i' \<le> n \<Longrightarrow> j' \<le> n \<Longrightarrow> fw m n (Suc k) i' j' i j \<le> fw m n k n n i j"
-by (metis Suc_innermost_id1' Suc_innermost_id2 Suc_innermost_mono linorder_class.not_le local.eq_iff
+by (metis Suc_innermost_id1' Suc_innermost_id2 Suc_innermost_mono linorder_class.not_le order.eq_iff
preorder_class.order_refl single_iteration_inv single_iteration_inv')
lemma negative_len_shortest:
"length xs = n \<Longrightarrow> len m i i xs < \<one>
\<Longrightarrow> \<exists> j ys. distinct (j # ys) \<and> len m j j ys < \<one> \<and> j \<in> set (i # xs) \<and> set ys \<subseteq> set xs"
proof (induction n arbitrary: xs i rule: less_induct)
case (less n)
show ?case
proof (cases xs)
case Nil
thus ?thesis using less.prems by auto
next
case (Cons y ys)
then have "length xs \<ge> 1" by auto
show ?thesis
proof (cases "i \<in> set xs")
assume i: "i \<in> set xs"
then obtain as bs where xs: "xs = as @ i # bs" by (meson split_list)
show ?thesis
proof (cases "len m i i as < \<one>")
case True
from xs less.prems have "length as < n" by auto
from less.IH[OF this _ True] xs show ?thesis by auto
next
case False
from len_decomp[OF xs] have "len m i i xs = len m i i as + len m i i bs" by auto
with False less.prems have *: "len m i i bs < \<one>"
by (metis add_lt_neutral local.dual_order.strict_trans local.neqE)
from xs less.prems have "length bs < n" by auto
from less.IH[OF this _ *] xs show ?thesis by auto
qed
next
assume i: "i \<notin> set xs"
show ?thesis
proof (cases "distinct xs")
case True
with i less.prems show ?thesis by auto
next
case False
from not_distinct_decomp[OF this] obtain a as bs cs where xs:
"xs = as @ a # bs @ a # cs"
by auto
show ?thesis
proof (cases "len m a a bs < \<one>")
case True
from xs less.prems have "length bs < n" by auto
from less.IH[OF this _ True] xs show ?thesis by auto
next
case False
from len_decomp[OF xs, of m i i] len_decomp[of "bs @ a # cs" bs a cs m a i]
have *:"len m i i xs = len m i a as + (len m a a bs + len m a i cs)" by auto
from False have "len m a a bs \<ge> \<one>" by auto
with add_mono have "len m a a bs + len m a i cs \<ge> len m a i cs" by fastforce
with * have "len m i i xs \<ge> len m i a as + len m a i cs" by (simp add: add_mono)
with less.prems(2) have "len m i a as + len m a i cs < \<one>" by auto
with len_comp have "len m i i (as @ a # cs) < \<one>" by auto
from less.IH[OF _ _ this, of "length (as @ a # cs)"] xs less.prems
show ?thesis by auto
qed
qed
qed
qed
qed
theorem FW_neg_cycle_detect:
"\<not> cycle_free m n \<Longrightarrow> \<exists> i \<le> n. fw m n n n n i i < \<one>"
proof -
assume A: "\<not> cycle_free m n"
let ?K = "{k. k \<le> n \<and> \<not> cycle_free_up_to m k n}"
let ?k = "Min ?K"
have not_empty_K: "?K \<noteq> {}" using not_cylce_free_dest[OF A(1)] by auto
have "finite ?K" by auto
with not_empty_K have *:
"\<forall> k' < ?k. cycle_free_up_to m k' n"
by (auto, metis le_trans less_or_eq_imp_le preorder_class.less_irrefl)
from linorder_class.Min_in[OF \<open>finite ?K\<close> \<open>?K \<noteq> {}\<close>] have
"\<not> cycle_free_up_to m ?k n" "?k \<le> n"
by auto
then have "\<exists> xs j. j \<le> n \<and> len m j j xs < \<one> \<and> set xs \<subseteq> {0..?k}" unfolding cycle_free_up_to_def
proof (auto, goal_cases)
case (2 i xs) then have "len m i i xs < \<one>" by auto
with 2 show ?case by auto
next
case (1 i xs j)
then have "len m i j (rem_cycles i j xs) > len m i j xs" by auto
from negative_cycle_dest[OF this]
obtain i' ys where ys: "i' \<in> set (i # j # xs)" "len m i' i' ys < \<one>" "set ys \<subseteq> set xs" by blast
from ys(1) 1(2-4) show ?case
proof (auto, goal_cases)
case 1
with ys(2,3) show ?case by auto
next
case 2
with ys(2,3) show ?case by auto
next
case 3
with \<open>?k \<le> n\<close> have "i' \<le> n" unfolding cycle_free_up_to_def by auto
with 3 ys(2,3) show ?case by auto
qed
qed
then obtain a as where a_as: "a \<le> n \<and> len m a a as < \<one> \<and> set as \<subseteq> {0..?k}" by auto
with negative_len_shortest[of as "length as" m a] obtain j xs where j_xs:
"distinct (j # xs) \<and> len m j j xs < \<one> \<and> j \<in> set (a # as) \<and> set xs \<subseteq> set as" by auto
with a_as \<open>?k \<le> n\<close> have cyc: "j \<le> n" "set xs \<subseteq> {0..?k}" "len m j j xs < \<one>" "distinct (j # xs)"
by auto
{ assume "?k > 0"
then have "?k - 1 < ?k" by simp
with * have **:"cycle_free_up_to m (?k - 1) n" by blast
have "?k \<in> set xs"
proof (rule ccontr, goal_cases)
case 1
with \<open>set xs \<subseteq> {0..?k}\<close> nat_upto_subs_top_removal have "set xs \<subseteq> {0..?k-1}" by auto
from cycle_free_up_to_loop_dest[OF \<open>j \<le> n\<close> this \<open>cycle_free_up_to m (?k - 1) n\<close>] cyc(3)
show ?case by auto
qed
with cyc(4) have "j \<noteq> ?k" by auto
from \<open>?k \<in> set xs\<close> obtain ys zs where "xs = ys @ ?k # zs" by (meson split_list)
with \<open>distinct (j # xs)\<close>
have xs: "xs = ys @ ?k # zs" "distinct ys" "distinct zs" "?k \<notin> set ys" "?k \<notin> set zs"
"j \<notin> set ys" "j \<notin> set zs" by auto
from xs(1,4) \<open>set xs \<subseteq> {0..?k}\<close> nat_upto_subs_top_removal have ys: "set ys \<subseteq> {0..?k-1}" by auto
from xs(1,5) \<open>set xs \<subseteq> {0..?k}\<close> nat_upto_subs_top_removal have zs: "set zs \<subseteq> {0..?k-1}" by auto
have "D m j ?k (?k - 1) = fw m n (?k - 1) n n j ?k"
using \<open>?k \<le> n\<close> \<open>j \<le> n\<close> fw_shortest_path_up_to[OF **, of j n ?k n] by auto
moreover have "D m ?k j (?k - 1) = fw m n (?k - 1) n n ?k j"
using \<open>?k \<le> n\<close> \<open>j \<le> n\<close> fw_shortest_path_up_to[OF **, of ?k n j n] by auto
ultimately have "fw m n (?k - 1) n n j ?k + fw m n (?k - 1) n n ?k j \<le> len m j ?k ys + len m ?k j zs"
using D_not_diag_le'[OF zs(1) xs(5,7,3), of m]
D_not_diag_le'[OF ys(1) xs(6,4,2), of m] by (auto simp: add_mono)
then have neg: "fw m n (?k - 1) n n j ?k + fw m n (?k - 1) n n ?k j < \<one>"
using xs(1) \<open>len m j j xs < \<one>\<close> len_comp by auto
have "fw m n ?k j j j j \<le> fw m n (?k - 1) n n j ?k + fw m n (?k - 1) n n ?k j"
proof (cases "j = 0")
case True
with\<open>?k > 0\<close> fw.simps(2)[of m n "?k - 1"]
have "fw m n ?k j j = fw_upd (fw m n (?k - 1) n n) ?k j j" by auto
then have "fw m n ?k j j j j \<le> fw m n (?k - 1) n n j ?k + fw m n (?k - 1) n n ?k j"
by (simp add: fw_upd_def upd_def)
then show ?thesis by auto
next
case False
with fw.simps(4)[of m n ?k j "j - 1"]
have "fw m n ?k j j = fw_upd (fw m n ?k j (j -1)) ?k j j" by simp
then have *: "fw m n ?k j j j j \<le> fw m n ?k j (j -1) j ?k + fw m n ?k j (j -1) ?k j"
by (simp add: fw_upd_def upd_def)
have "j - 1 < n" using \<open>j \<le> n\<close> False by auto
then have "fw m n ?k j (j -1) j ?k \<le> fw m n (?k - 1) n n j ?k"
using fw_Suc[of j n ?k j "j - 1" m "?k - 1"] \<open>j \<le> n\<close> \<open>?k \<le> n\<close> \<open>?k > 0\<close> by auto
moreover have "fw m n ?k j (j -1) ?k j \<le> fw m n (?k - 1) n n ?k j"
using fw_Suc[of ?k n j j "j - 1" m "?k - 1"] \<open>j \<le> n\<close> \<open>?k \<le> n\<close> \<open>?k > 0\<close> by auto
ultimately have "fw m n ?k j j j j \<le> fw m n (?k - 1) n n j ?k + fw m n (?k - 1) n n ?k j"
using * add_mono by fastforce
then show ?thesis by auto
qed
with neg have "fw m n ?k j j j j < \<one>" by auto
moreover have "fw m n n n n j j \<le> fw m n ?k j j j j" using fw_invariant \<open>j\<le>n\<close> \<open>?k \<le> n\<close> by auto
ultimately have "fw m n n n n j j < \<one>" using neg by auto
with \<open>j\<le>n\<close> have ?thesis by auto
}
moreover
{ assume "?k = 0"
with cyc(2,4) have "xs = [] \<or> xs = [0]"
apply safe
apply (case_tac xs)
apply fastforce
apply (rename_tac ys)
apply (case_tac ys)
apply auto
done
then have ?thesis
proof
assume "xs = []"
with cyc have "m j j < \<one>" by auto
with fw_mono[of n n n j j m n] \<open>j \<le> n\<close> have "fw m n n n n j j < \<one>" by auto
with \<open>j \<le> n\<close> show ?thesis by auto
next
assume xs: "xs = [0]"
with cyc have "m j 0 + m 0 j < \<one>" by auto
then have "fw m n 0 j j j j < \<one>"
proof (cases "j = 0", goal_cases)
case 1
have "m j j < \<one>"
proof (rule ccontr)
assume "\<not> m j j < \<one>"
with 1 have "m 0 0 \<ge> \<one>" by simp
with add_mono have "m 0 0 + m 0 0 \<ge> \<one>" by fastforce
with 1 show False by simp
qed
with fw_mono[of j n j j j m 0] \<open>j \<le> n\<close> show ?thesis by auto
next
case 2
with fw.simps(4)[of m n 0 j "j - 1"]
have "fw m n 0 j j = fw_upd (fw m n 0 j (j - 1)) 0 j j" by simp
then have "fw m n 0 j j j j \<le> fw m n 0 j (j - 1) j 0 + fw m n 0 j (j - 1) 0 j"
by (simp add: fw_upd_def upd_def)
also have "\<dots> \<le> m j 0 + m 0 j" using \<open>j \<le> n\<close> add_mono fw_mono by auto
finally show ?thesis using 2 by auto
qed
then have "fw m n 0 n n j j < \<one>" by (metis cyc(1) less_or_eq_imp_le single_iteration_inv)
with fw_invariant[of 0 n n n n n n m j j] \<open>j \<le> n\<close> have "fw m n n n n j j < \<one>" by auto
with \<open>j \<le> n\<close> show ?thesis by blast
qed
}
ultimately show ?thesis by auto
qed
end (* End of local class context *)
end (* End of theory *)
diff --git a/thys/UTP/toolkit/List_Extra.thy b/thys/UTP/toolkit/List_Extra.thy
--- a/thys/UTP/toolkit/List_Extra.thy
+++ b/thys/UTP/toolkit/List_Extra.thy
@@ -1,864 +1,864 @@
(******************************************************************************)
(* Project: Isabelle/UTP Toolkit *)
(* File: List_Extra.thy *)
(* Authors: Simon Foster and Frank Zeyda *)
(* Emails: simon.foster@york.ac.uk and frank.zeyda@york.ac.uk *)
(******************************************************************************)
section \<open> Lists: extra functions and properties \<close>
theory List_Extra
imports
"HOL-Library.Sublist"
"HOL-Library.Monad_Syntax"
"HOL-Library.Prefix_Order"
"Optics.Lens_Instances"
begin
subsection \<open> Useful Abbreviations \<close>
abbreviation "list_sum xs \<equiv> foldr (+) xs 0"
subsection \<open> Folds \<close>
context abel_semigroup
begin
lemma foldr_snoc: "foldr (\<^bold>*) (xs @ [x]) k = (foldr (\<^bold>*) xs k) \<^bold>* x"
by (induct xs, simp_all add: commute left_commute)
end
subsection \<open> List Lookup \<close>
text \<open>
The following variant of the standard \<open>nth\<close> function returns \<open>\<bottom>\<close> if the index is
out of range.
\<close>
primrec
nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>\<^sub>l" [90, 0] 91)
where
"[]\<langle>i\<rangle>\<^sub>l = None"
| "(x # xs)\<langle>i\<rangle>\<^sub>l = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>\<^sub>l)"
lemma nth_el_appendl[simp]: "i < length xs \<Longrightarrow> (xs @ ys)\<langle>i\<rangle>\<^sub>l = xs\<langle>i\<rangle>\<^sub>l"
apply (induct xs arbitrary: i)
apply simp
apply (case_tac i)
apply simp_all
done
lemma nth_el_appendr[simp]: "length xs \<le> i \<Longrightarrow> (xs @ ys)\<langle>i\<rangle>\<^sub>l = ys\<langle>i - length xs\<rangle>\<^sub>l"
apply (induct xs arbitrary: i)
apply simp
apply (case_tac i)
apply simp_all
done
subsection \<open> Extra List Theorems \<close>
subsubsection \<open> Map \<close>
lemma map_nth_Cons_atLeastLessThan:
"map (nth (x # xs)) [Suc m..<n] = map (nth xs) [m..<n - 1]"
proof -
have "nth xs = nth (x # xs) \<circ> Suc"
by auto
hence "map (nth xs) [m..<n - 1] = map (nth (x # xs) \<circ> Suc) [m..<n - 1]"
by simp
also have "... = map (nth (x # xs)) (map Suc [m..<n - 1])"
by simp
also have "... = map (nth (x # xs)) [Suc m..<n]"
by (metis Suc_diff_1 le_0_eq length_upt list.simps(8) list.size(3) map_Suc_upt not_less upt_0)
finally show ?thesis ..
qed
subsubsection \<open> Sorted Lists \<close>
lemma sorted_last [simp]: "\<lbrakk> x \<in> set xs; sorted xs \<rbrakk> \<Longrightarrow> x \<le> last xs"
by (induct xs, auto)
lemma sorted_prefix:
assumes "xs \<le> ys" "sorted ys"
shows "sorted xs"
proof -
obtain zs where "ys = xs @ zs"
using Prefix_Order.prefixE assms(1) by auto
thus ?thesis
using assms(2) sorted_append by blast
qed
lemma sorted_map: "\<lbrakk> sorted xs; mono f \<rbrakk> \<Longrightarrow> sorted (map f xs)"
by (simp add: monoD sorted_iff_nth_mono)
lemma sorted_distinct [intro]: "\<lbrakk> sorted (xs); distinct(xs) \<rbrakk> \<Longrightarrow> (\<forall> i<length xs - 1. xs!i < xs!(i + 1))"
apply (induct xs)
apply (auto)
apply (metis (no_types, hide_lams) Suc_leI Suc_less_eq Suc_pred gr0_conv_Suc not_le not_less_iff_gr_or_eq nth_Cons_Suc nth_mem nth_non_equal_first_eq)
done
text \<open> Is the given list a permutation of the given set? \<close>
definition is_sorted_list_of_set :: "('a::ord) set \<Rightarrow> 'a list \<Rightarrow> bool" where
"is_sorted_list_of_set A xs = ((\<forall> i<length(xs) - 1. xs!i < xs!(i + 1)) \<and> set(xs) = A)"
lemma sorted_is_sorted_list_of_set:
assumes "is_sorted_list_of_set A xs"
shows "sorted(xs)" and "distinct(xs)"
using assms proof (induct xs arbitrary: A)
show "sorted []"
by auto
next
show "distinct []"
by auto
next
fix A :: "'a set"
case (Cons x xs) note hyps = this
assume isl: "is_sorted_list_of_set A (x # xs)"
hence srt: "(\<forall>i<length xs - Suc 0. xs ! i < xs ! Suc i)"
using less_diff_conv by (auto simp add: is_sorted_list_of_set_def)
with hyps(1) have srtd: "sorted xs"
by (simp add: is_sorted_list_of_set_def)
with isl show "sorted (x # xs)"
apply (auto simp add: is_sorted_list_of_set_def)
apply (metis (mono_tags, lifting) all_nth_imp_all_set less_le_trans linorder_not_less not_less_iff_gr_or_eq nth_Cons_0 sorted_iff_nth_mono zero_order(3))
done
from srt hyps(2) have "distinct xs"
by (simp add: is_sorted_list_of_set_def)
with isl show "distinct (x # xs)"
proof -
have "(\<forall>n. \<not> n < length (x # xs) - 1 \<or> (x # xs) ! n < (x # xs) ! (n + 1)) \<and> set (x # xs) = A"
by (meson \<open>is_sorted_list_of_set A (x # xs)\<close> is_sorted_list_of_set_def)
then show ?thesis
by (metis \<open>distinct xs\<close> add.commute add_diff_cancel_left' distinct.simps(2) leD length_Cons length_greater_0_conv length_pos_if_in_set less_le nth_Cons_0 nth_Cons_Suc plus_1_eq_Suc set_ConsD sorted.elims(2) srtd)
qed
qed
lemma is_sorted_list_of_set_alt_def:
"is_sorted_list_of_set A xs \<longleftrightarrow> sorted (xs) \<and> distinct (xs) \<and> set(xs) = A"
apply (auto intro: sorted_is_sorted_list_of_set)
apply (auto simp add: is_sorted_list_of_set_def)
apply (metis Nat.add_0_right One_nat_def add_Suc_right sorted_distinct)
done
definition sorted_list_of_set_alt :: "('a::ord) set \<Rightarrow> 'a list" where
"sorted_list_of_set_alt A =
(if (A = {}) then [] else (THE xs. is_sorted_list_of_set A xs))"
lemma is_sorted_list_of_set:
"finite A \<Longrightarrow> is_sorted_list_of_set A (sorted_list_of_set A)"
by (simp add: is_sorted_list_of_set_alt_def)
lemma sorted_list_of_set_other_def:
"finite A \<Longrightarrow> sorted_list_of_set(A) = (THE xs. sorted(xs) \<and> distinct(xs) \<and> set xs = A)"
apply (rule sym)
apply (rule the_equality)
apply (auto)
apply (simp add: sorted_distinct_set_unique)
done
lemma sorted_list_of_set_alt [simp]:
"finite A \<Longrightarrow> sorted_list_of_set_alt(A) = sorted_list_of_set(A)"
apply (rule sym)
apply (auto simp add: sorted_list_of_set_alt_def is_sorted_list_of_set_alt_def sorted_list_of_set_other_def)
done
text \<open> Sorting lists according to a relation \<close>
definition is_sorted_list_of_set_by :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" where
"is_sorted_list_of_set_by R A xs = ((\<forall> i<length(xs) - 1. (xs!i, xs!(i + 1)) \<in> R) \<and> set(xs) = A)"
definition sorted_list_of_set_by :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a list" where
"sorted_list_of_set_by R A = (THE xs. is_sorted_list_of_set_by R A xs)"
definition fin_set_lexord :: "'a rel \<Rightarrow> 'a set rel" where
"fin_set_lexord R = {(A, B). finite A \<and> finite B \<and>
(\<exists> xs ys. is_sorted_list_of_set_by R A xs \<and> is_sorted_list_of_set_by R B ys
\<and> (xs, ys) \<in> lexord R)}"
lemma is_sorted_list_of_set_by_mono:
"\<lbrakk> R \<subseteq> S; is_sorted_list_of_set_by R A xs \<rbrakk> \<Longrightarrow> is_sorted_list_of_set_by S A xs"
by (auto simp add: is_sorted_list_of_set_by_def)
lemma lexord_mono':
"\<lbrakk> (\<And> x y. f x y \<longrightarrow> g x y); (xs, ys) \<in> lexord {(x, y). f x y} \<rbrakk> \<Longrightarrow> (xs, ys) \<in> lexord {(x, y). g x y}"
by (metis case_prodD case_prodI lexord_take_index_conv mem_Collect_eq)
lemma fin_set_lexord_mono [mono]:
"(\<And> x y. f x y \<longrightarrow> g x y) \<Longrightarrow> (xs, ys) \<in> fin_set_lexord {(x, y). f x y} \<longrightarrow> (xs, ys) \<in> fin_set_lexord {(x, y). g x y}"
proof
assume
fin: "(xs, ys) \<in> fin_set_lexord {(x, y). f x y}" and
hyp: "(\<And> x y. f x y \<longrightarrow> g x y)"
from fin have "finite xs" "finite ys"
using fin_set_lexord_def by fastforce+
with fin hyp show "(xs, ys) \<in> fin_set_lexord {(x, y). g x y}"
apply (auto simp add: fin_set_lexord_def)
apply (rename_tac xs' ys')
apply (rule_tac x="xs'" in exI)
apply (auto)
apply (metis case_prodD case_prodI is_sorted_list_of_set_by_def mem_Collect_eq)
apply (metis case_prodD case_prodI is_sorted_list_of_set_by_def lexord_mono' mem_Collect_eq)
done
qed
definition distincts :: "'a set \<Rightarrow> 'a list set" where
"distincts A = {xs \<in> lists A. distinct(xs)}"
lemma tl_element:
"\<lbrakk> x \<in> set xs; x \<noteq> hd(xs) \<rbrakk> \<Longrightarrow> x \<in> set(tl(xs))"
by (metis in_set_insert insert_Nil list.collapse list.distinct(2) set_ConsD)
subsubsection \<open> List Update \<close>
lemma listsum_update:
fixes xs :: "'a::ring list"
assumes "i < length xs"
shows "list_sum (xs[i := v]) = list_sum xs - xs ! i + v"
using assms proof (induct xs arbitrary: i)
case Nil
then show ?case by (simp)
next
case (Cons a xs)
then show ?case
proof (cases i)
case 0
thus ?thesis
by (simp add: add.commute)
next
case (Suc i')
with Cons show ?thesis
by (auto)
qed
qed
subsubsection \<open> Drop While and Take While \<close>
lemma dropWhile_sorted_le_above:
"\<lbrakk> sorted xs; x \<in> set (dropWhile (\<lambda> x. x \<le> n) xs) \<rbrakk> \<Longrightarrow> x > n"
apply (induct xs)
apply (auto)
apply (rename_tac a xs)
apply (case_tac "a \<le> n")
apply (auto)
done
lemma set_dropWhile_le:
"sorted xs \<Longrightarrow> set (dropWhile (\<lambda> x. x \<le> n) xs) = {x\<in>set xs. x > n}"
apply (induct xs)
apply (simp)
apply (rename_tac x xs)
apply (subgoal_tac "sorted xs")
apply (simp)
apply (safe)
apply (auto)
done
lemma set_takeWhile_less_sorted:
"\<lbrakk> sorted I; x \<in> set I; x < n \<rbrakk> \<Longrightarrow> x \<in> set (takeWhile (\<lambda>x. x < n) I)"
proof (induct I arbitrary: x)
case Nil thus ?case
by (simp)
next
case (Cons a I) thus ?case
by auto
qed
lemma nth_le_takeWhile_ord: "\<lbrakk> sorted xs; i \<ge> length (takeWhile (\<lambda> x. x \<le> n) xs); i < length xs \<rbrakk> \<Longrightarrow> n \<le> xs ! i"
apply (induct xs arbitrary: i, auto)
apply (rename_tac x xs i)
apply (case_tac "x \<le> n")
apply (auto)
apply (metis One_nat_def Suc_eq_plus1 le_less_linear le_less_trans less_imp_le list.size(4) nth_mem set_ConsD)
done
lemma length_takeWhile_less:
"\<lbrakk> a \<in> set xs; \<not> P a \<rbrakk> \<Longrightarrow> length (takeWhile P xs) < length xs"
by (metis in_set_conv_nth length_takeWhile_le nat_neq_iff not_less set_takeWhileD takeWhile_nth)
lemma nth_length_takeWhile_less:
"\<lbrakk> sorted xs; distinct xs; (\<exists> a \<in> set xs. a \<ge> n) \<rbrakk> \<Longrightarrow> xs ! length (takeWhile (\<lambda>x. x < n) xs) \<ge> n"
by (induct xs, auto)
subsubsection \<open> Last and But Last \<close>
lemma length_gt_zero_butlast_concat:
assumes "length ys > 0"
shows "butlast (xs @ ys) = xs @ (butlast ys)"
using assms by (metis butlast_append length_greater_0_conv)
lemma length_eq_zero_butlast_concat:
assumes "length ys = 0"
shows "butlast (xs @ ys) = butlast xs"
using assms by (metis append_Nil2 length_0_conv)
lemma butlast_single_element:
shows "butlast [e] = []"
by (metis butlast.simps(2))
lemma last_single_element:
shows "last [e] = e"
by (metis last.simps)
lemma length_zero_last_concat:
assumes "length t = 0"
shows "last (s @ t) = last s"
by (metis append_Nil2 assms length_0_conv)
lemma length_gt_zero_last_concat:
assumes "length t > 0"
shows "last (s @ t) = last t"
by (metis assms last_append length_greater_0_conv)
subsubsection \<open> Prefixes and Strict Prefixes \<close>
lemma prefix_length_eq:
"\<lbrakk> length xs = length ys; prefix xs ys \<rbrakk> \<Longrightarrow> xs = ys"
by (metis not_equal_is_parallel parallel_def)
lemma prefix_Cons_elim [elim]:
assumes "prefix (x # xs) ys"
obtains ys' where "ys = x # ys'" "prefix xs ys'"
using assms
by (metis append_Cons prefix_def)
lemma prefix_map_inj:
"\<lbrakk> inj_on f (set xs \<union> set ys); prefix (map f xs) (map f ys) \<rbrakk> \<Longrightarrow>
prefix xs ys"
apply (induct xs arbitrary:ys)
apply (simp_all)
apply (erule prefix_Cons_elim)
apply (auto)
apply (metis image_insert insertI1 insert_Diff_if singletonE)
done
lemma prefix_map_inj_eq [simp]:
"inj_on f (set xs \<union> set ys) \<Longrightarrow>
prefix (map f xs) (map f ys) \<longleftrightarrow> prefix xs ys"
using map_mono_prefix prefix_map_inj by blast
lemma strict_prefix_Cons_elim [elim]:
assumes "strict_prefix (x # xs) ys"
obtains ys' where "ys = x # ys'" "strict_prefix xs ys'"
using assms
by (metis Sublist.strict_prefixE' Sublist.strict_prefixI' append_Cons)
lemma strict_prefix_map_inj:
"\<lbrakk> inj_on f (set xs \<union> set ys); strict_prefix (map f xs) (map f ys) \<rbrakk> \<Longrightarrow>
strict_prefix xs ys"
apply (induct xs arbitrary:ys)
- apply (auto)
- using prefix_bot.bot.not_eq_extremum apply fastforce
+ apply auto
+ apply (simp flip: prefix_bot.bot_less)
apply (erule strict_prefix_Cons_elim)
apply (auto)
- apply (metis (hide_lams, full_types) image_insert insertI1 insert_Diff_if singletonE)
+ apply (metis image_eqI insert_Diff_single insert_iff)
done
lemma strict_prefix_map_inj_eq [simp]:
"inj_on f (set xs \<union> set ys) \<Longrightarrow>
strict_prefix (map f xs) (map f ys) \<longleftrightarrow> strict_prefix xs ys"
by (simp add: inj_on_map_eq_map strict_prefix_def)
lemma prefix_drop:
"\<lbrakk> drop (length xs) ys = zs; prefix xs ys \<rbrakk>
\<Longrightarrow> ys = xs @ zs"
by (metis append_eq_conv_conj prefix_def)
lemma list_append_prefixD [dest]: "x @ y \<le> z \<Longrightarrow> x \<le> z"
using append_prefixD less_eq_list_def by blast
lemma prefix_not_empty:
assumes "strict_prefix xs ys" and "xs \<noteq> []"
shows "ys \<noteq> []"
using Sublist.strict_prefix_simps(1) assms(1) by blast
lemma prefix_not_empty_length_gt_zero:
assumes "strict_prefix xs ys" and "xs \<noteq> []"
shows "length ys > 0"
using assms prefix_not_empty by auto
lemma butlast_prefix_suffix_not_empty:
assumes "strict_prefix (butlast xs) ys"
shows "ys \<noteq> []"
using assms prefix_not_empty_length_gt_zero by fastforce
lemma prefix_and_concat_prefix_is_concat_prefix:
assumes "prefix s t" "prefix (e @ t) u"
shows "prefix (e @ s) u"
using Sublist.same_prefix_prefix assms(1) assms(2) prefix_order.dual_order.trans by blast
lemma prefix_eq_exists:
"prefix s t \<longleftrightarrow> (\<exists>xs . s @ xs = t)"
using prefix_def by auto
lemma strict_prefix_eq_exists:
"strict_prefix s t \<longleftrightarrow> (\<exists>xs . s @ xs = t \<and> (length xs) > 0)"
using prefix_def strict_prefix_def by auto
lemma butlast_strict_prefix_eq_butlast:
assumes "length s = length t" and "strict_prefix (butlast s) t"
shows "strict_prefix (butlast s) t \<longleftrightarrow> (butlast s) = (butlast t)"
by (metis append_butlast_last_id append_eq_append_conv assms(1) assms(2) length_0_conv length_butlast strict_prefix_eq_exists)
lemma butlast_eq_if_eq_length_and_prefix:
assumes "length s > 0" "length z > 0"
"length s = length z" "strict_prefix (butlast s) t" "strict_prefix (butlast z) t"
shows "(butlast s) = (butlast z)"
using assms by (auto simp add:strict_prefix_eq_exists)
lemma butlast_prefix_imp_length_not_gt:
assumes "length s > 0" "strict_prefix (butlast s) t"
shows "\<not> (length t < length s)"
using assms prefix_length_less by fastforce
lemma length_not_gt_iff_eq_length:
assumes "length s > 0" and "strict_prefix (butlast s) t"
shows "(\<not> (length s < length t)) = (length s = length t)"
proof -
have "(\<not> (length s < length t)) = ((length t < length s) \<or> (length s = length t))"
by (metis not_less_iff_gr_or_eq)
also have "... = (length s = length t)"
using assms
by (simp add:butlast_prefix_imp_length_not_gt)
finally show ?thesis .
qed
text \<open> Greatest common prefix \<close>
fun gcp :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"gcp [] ys = []" |
"gcp (x # xs) (y # ys) = (if (x = y) then x # gcp xs ys else [])" |
"gcp _ _ = []"
lemma gcp_right [simp]: "gcp xs [] = []"
by (induct xs, auto)
lemma gcp_append [simp]: "gcp (xs @ ys) (xs @ zs) = xs @ gcp ys zs"
by (induct xs, auto)
lemma gcp_lb1: "prefix (gcp xs ys) xs"
apply (induct xs arbitrary: ys, auto)
apply (case_tac ys, auto)
done
lemma gcp_lb2: "prefix (gcp xs ys) ys"
apply (induct ys arbitrary: xs, auto)
apply (case_tac xs, auto)
done
interpretation prefix_semilattice: semilattice_inf gcp prefix strict_prefix
proof
fix xs ys :: "'a list"
show "prefix (gcp xs ys) xs"
by (induct xs arbitrary: ys, auto, case_tac ys, auto)
show "prefix (gcp xs ys) ys"
by (induct ys arbitrary: xs, auto, case_tac xs, auto)
next
fix xs ys zs :: "'a list"
assume "prefix xs ys" "prefix xs zs"
thus "prefix xs (gcp ys zs)"
by (simp add: prefix_def, auto)
qed
subsubsection \<open> Lexicographic Order \<close>
lemma lexord_append:
assumes "(xs\<^sub>1 @ ys\<^sub>1, xs\<^sub>2 @ ys\<^sub>2) \<in> lexord R" "length(xs\<^sub>1) = length(xs\<^sub>2)"
shows "(xs\<^sub>1, xs\<^sub>2) \<in> lexord R \<or> (xs\<^sub>1 = xs\<^sub>2 \<and> (ys\<^sub>1, ys\<^sub>2) \<in> lexord R)"
using assms
proof (induct xs\<^sub>2 arbitrary: xs\<^sub>1)
case (Cons x\<^sub>2 xs\<^sub>2') note hyps = this
from hyps(3) obtain x\<^sub>1 xs\<^sub>1' where xs\<^sub>1: "xs\<^sub>1 = x\<^sub>1 # xs\<^sub>1'" "length(xs\<^sub>1') = length(xs\<^sub>2')"
by (auto, metis Suc_length_conv)
with hyps(2) have xcases: "(x\<^sub>1, x\<^sub>2) \<in> R \<or> (xs\<^sub>1' @ ys\<^sub>1, xs\<^sub>2' @ ys\<^sub>2) \<in> lexord R"
by (auto)
show ?case
proof (cases "(x\<^sub>1, x\<^sub>2) \<in> R")
case True with xs\<^sub>1 show ?thesis
by (auto)
next
case False
with xcases have "(xs\<^sub>1' @ ys\<^sub>1, xs\<^sub>2' @ ys\<^sub>2) \<in> lexord R"
by (auto)
with hyps(1) xs\<^sub>1 have dichot: "(xs\<^sub>1', xs\<^sub>2') \<in> lexord R \<or> (xs\<^sub>1' = xs\<^sub>2' \<and> (ys\<^sub>1, ys\<^sub>2) \<in> lexord R)"
by (auto)
have "x\<^sub>1 = x\<^sub>2"
using False hyps(2) xs\<^sub>1(1) by auto
with dichot xs\<^sub>1 show ?thesis
by (simp)
qed
next
case Nil thus ?case
by auto
qed
lemma strict_prefix_lexord_rel:
"strict_prefix xs ys \<Longrightarrow> (xs, ys) \<in> lexord R"
by (metis Sublist.strict_prefixE' lexord_append_rightI)
lemma strict_prefix_lexord_left:
assumes "trans R" "(xs, ys) \<in> lexord R" "strict_prefix xs' xs"
shows "(xs', ys) \<in> lexord R"
by (metis assms lexord_trans strict_prefix_lexord_rel)
lemma prefix_lexord_right:
assumes "trans R" "(xs, ys) \<in> lexord R" "strict_prefix ys ys'"
shows "(xs, ys') \<in> lexord R"
by (metis assms lexord_trans strict_prefix_lexord_rel)
lemma lexord_eq_length:
assumes "(xs, ys) \<in> lexord R" "length xs = length ys"
shows "\<exists> i. (xs!i, ys!i) \<in> R \<and> i < length xs \<and> (\<forall> j<i. xs!j = ys!j)"
using assms proof (induct xs arbitrary: ys)
case (Cons x xs) note hyps = this
then obtain y ys' where ys: "ys = y # ys'" "length ys' = length xs"
by (metis Suc_length_conv)
show ?case
proof (cases "(x, y) \<in> R")
case True with ys show ?thesis
by (rule_tac x="0" in exI, simp)
next
case False
with ys hyps(2) have xy: "x = y" "(xs, ys') \<in> lexord R"
by auto
with hyps(1,3) ys obtain i where "(xs!i, ys'!i) \<in> R" "i < length xs" "(\<forall> j<i. xs!j = ys'!j)"
by force
with xy ys show ?thesis
apply (rule_tac x="Suc i" in exI)
apply (auto simp add: less_Suc_eq_0_disj)
done
qed
next
case Nil thus ?case by (auto)
qed
lemma lexord_intro_elems:
assumes "length xs > i" "length ys > i" "(xs!i, ys!i) \<in> R" "\<forall> j<i. xs!j = ys!j"
shows "(xs, ys) \<in> lexord R"
using assms proof (induct i arbitrary: xs ys)
case 0 thus ?case
by (auto, metis lexord_cons_cons list.exhaust nth_Cons_0)
next
case (Suc i) note hyps = this
then obtain x' y' xs' ys' where "xs = x' # xs'" "ys = y' # ys'"
by (metis Suc_length_conv Suc_lessE)
moreover with hyps(5) have "\<forall>j<i. xs' ! j = ys' ! j"
by (auto)
ultimately show ?case using hyps
by (auto)
qed
subsection \<open> Distributed Concatenation \<close>
definition uncurry :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b \<Rightarrow> 'c)" where
[simp]: "uncurry f = (\<lambda>(x, y). f x y)"
definition dist_concat ::
"'a list set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" (infixr "\<^sup>\<frown>" 100) where
"dist_concat ls1 ls2 = (uncurry (@) ` (ls1 \<times> ls2))"
lemma dist_concat_left_empty [simp]:
"{} \<^sup>\<frown> ys = {}"
by (simp add: dist_concat_def)
lemma dist_concat_right_empty [simp]:
"xs \<^sup>\<frown> {} = {}"
by (simp add: dist_concat_def)
lemma dist_concat_insert [simp]:
"insert l ls1 \<^sup>\<frown> ls2 = ((@) l ` ( ls2)) \<union> (ls1 \<^sup>\<frown> ls2)"
by (auto simp add: dist_concat_def)
subsection \<open> List Domain and Range \<close>
abbreviation seq_dom :: "'a list \<Rightarrow> nat set" ("dom\<^sub>l") where
"seq_dom xs \<equiv> {0..<length xs}"
abbreviation seq_ran :: "'a list \<Rightarrow> 'a set" ("ran\<^sub>l") where
"seq_ran xs \<equiv> set xs"
subsection \<open> Extracting List Elements \<close>
definition seq_extract :: "nat set \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "\<upharpoonleft>\<^sub>l" 80) where
"seq_extract A xs = nths xs A"
lemma seq_extract_Nil [simp]: "A \<upharpoonleft>\<^sub>l [] = []"
by (simp add: seq_extract_def)
lemma seq_extract_Cons:
"A \<upharpoonleft>\<^sub>l (x # xs) = (if 0 \<in> A then [x] else []) @ {j. Suc j \<in> A} \<upharpoonleft>\<^sub>l xs"
by (simp add: seq_extract_def nths_Cons)
lemma seq_extract_empty [simp]: "{} \<upharpoonleft>\<^sub>l xs = []"
by (simp add: seq_extract_def)
lemma seq_extract_ident [simp]: "{0..<length xs} \<upharpoonleft>\<^sub>l xs = xs"
unfolding list_eq_iff_nth_eq
by (auto simp add: seq_extract_def length_nths atLeast0LessThan)
lemma seq_extract_split:
assumes "i \<le> length xs"
shows "{0..<i} \<upharpoonleft>\<^sub>l xs @ {i..<length xs} \<upharpoonleft>\<^sub>l xs = xs"
using assms
proof (induct xs arbitrary: i)
case Nil thus ?case by (simp add: seq_extract_def)
next
case (Cons x xs) note hyp = this
have "{j. Suc j < i} = {0..<i - 1}"
by (auto)
moreover have "{j. i \<le> Suc j \<and> j < length xs} = {i - 1..<length xs}"
by (auto)
ultimately show ?case
using hyp by (force simp add: seq_extract_def nths_Cons)
qed
lemma seq_extract_append:
"A \<upharpoonleft>\<^sub>l (xs @ ys) = (A \<upharpoonleft>\<^sub>l xs) @ ({j. j + length xs \<in> A} \<upharpoonleft>\<^sub>l ys)"
by (simp add: seq_extract_def nths_append)
lemma seq_extract_range: "A \<upharpoonleft>\<^sub>l xs = (A \<inter> dom\<^sub>l(xs)) \<upharpoonleft>\<^sub>l xs"
apply (auto simp add: seq_extract_def nths_def)
apply (metis (no_types, lifting) atLeastLessThan_iff filter_cong in_set_zip nth_mem set_upt)
done
lemma seq_extract_out_of_range:
"A \<inter> dom\<^sub>l(xs) = {} \<Longrightarrow> A \<upharpoonleft>\<^sub>l xs = []"
by (metis seq_extract_def seq_extract_range nths_empty)
lemma seq_extract_length [simp]:
"length (A \<upharpoonleft>\<^sub>l xs) = card (A \<inter> dom\<^sub>l(xs))"
proof -
have "{i. i < length(xs) \<and> i \<in> A} = (A \<inter> {0..<length(xs)})"
by (auto)
thus ?thesis
by (simp add: seq_extract_def length_nths)
qed
lemma seq_extract_Cons_atLeastLessThan:
assumes "m < n"
shows "{m..<n} \<upharpoonleft>\<^sub>l (x # xs) = (if (m = 0) then x # ({0..<n-1} \<upharpoonleft>\<^sub>l xs) else {m-1..<n-1} \<upharpoonleft>\<^sub>l xs)"
proof -
have "{j. Suc j < n} = {0..<n - Suc 0}"
by (auto)
moreover have "{j. m \<le> Suc j \<and> Suc j < n} = {m - Suc 0..<n - Suc 0}"
by (auto)
ultimately show ?thesis using assms
by (auto simp add: seq_extract_Cons)
qed
lemma seq_extract_singleton:
assumes "i < length xs"
shows "{i} \<upharpoonleft>\<^sub>l xs = [xs ! i]"
using assms
apply (induct xs arbitrary: i)
apply (auto simp add: seq_extract_Cons)
apply (rename_tac xs i)
apply (subgoal_tac "{j. Suc j = i} = {i - 1}")
apply (auto)
done
lemma seq_extract_as_map:
assumes "m < n" "n \<le> length xs"
shows "{m..<n} \<upharpoonleft>\<^sub>l xs = map (nth xs) [m..<n]"
using assms proof (induct xs arbitrary: m n)
case Nil thus ?case by simp
next
case (Cons x xs)
have "[m..<n] = m # [m+1..<n]"
using Cons.prems(1) upt_eq_Cons_conv by blast
moreover have "map (nth (x # xs)) [Suc m..<n] = map (nth xs) [m..<n-1]"
by (simp add: map_nth_Cons_atLeastLessThan)
ultimately show ?case
using Cons upt_rec
by (auto simp add: seq_extract_Cons_atLeastLessThan)
qed
lemma seq_append_as_extract:
"xs = ys @ zs \<longleftrightarrow> (\<exists> i\<le>length(xs). ys = {0..<i} \<upharpoonleft>\<^sub>l xs \<and> zs = {i..<length(xs)} \<upharpoonleft>\<^sub>l xs)"
proof
assume xs: "xs = ys @ zs"
moreover have "ys = {0..<length ys} \<upharpoonleft>\<^sub>l (ys @ zs)"
by (simp add: seq_extract_append)
moreover have "zs = {length ys..<length ys + length zs} \<upharpoonleft>\<^sub>l (ys @ zs)"
proof -
have "{length ys..<length ys + length zs} \<inter> {0..<length ys} = {}"
by auto
moreover have s1: "{j. j < length zs} = {0..<length zs}"
by auto
ultimately show ?thesis
by (simp add: seq_extract_append seq_extract_out_of_range)
qed
ultimately show "(\<exists> i\<le>length(xs). ys = {0..<i} \<upharpoonleft>\<^sub>l xs \<and> zs = {i..<length(xs)} \<upharpoonleft>\<^sub>l xs)"
by (rule_tac x="length ys" in exI, auto)
next
assume "\<exists>i\<le>length xs. ys = {0..<i} \<upharpoonleft>\<^sub>l xs \<and> zs = {i..<length xs} \<upharpoonleft>\<^sub>l xs"
thus "xs = ys @ zs"
by (auto simp add: seq_extract_split)
qed
subsection \<open> Filtering a list according to a set \<close>
definition seq_filter :: "'a list \<Rightarrow> 'a set \<Rightarrow> 'a list" (infix "\<restriction>\<^sub>l" 80) where
"seq_filter xs A = filter (\<lambda> x. x \<in> A) xs"
lemma seq_filter_Cons_in [simp]:
"x \<in> cs \<Longrightarrow> (x # xs) \<restriction>\<^sub>l cs = x # (xs \<restriction>\<^sub>l cs)"
by (simp add: seq_filter_def)
lemma seq_filter_Cons_out [simp]:
"x \<notin> cs \<Longrightarrow> (x # xs) \<restriction>\<^sub>l cs = (xs \<restriction>\<^sub>l cs)"
by (simp add: seq_filter_def)
lemma seq_filter_Nil [simp]: "[] \<restriction>\<^sub>l A = []"
by (simp add: seq_filter_def)
lemma seq_filter_empty [simp]: "xs \<restriction>\<^sub>l {} = []"
by (simp add: seq_filter_def)
lemma seq_filter_append: "(xs @ ys) \<restriction>\<^sub>l A = (xs \<restriction>\<^sub>l A) @ (ys \<restriction>\<^sub>l A)"
by (simp add: seq_filter_def)
lemma seq_filter_UNIV [simp]: "xs \<restriction>\<^sub>l UNIV = xs"
by (simp add: seq_filter_def)
lemma seq_filter_twice [simp]: "(xs \<restriction>\<^sub>l A) \<restriction>\<^sub>l B = xs \<restriction>\<^sub>l (A \<inter> B)"
by (simp add: seq_filter_def)
subsection \<open> Minus on lists \<close>
instantiation list :: (type) minus
begin
text \<open> We define list minus so that if the second list is not a prefix of the first, then an arbitrary
list longer than the combined length is produced. Thus we can always determined from the output
whether the minus is defined or not. \<close>
definition "xs - ys = (if (prefix ys xs) then drop (length ys) xs else [])"
instance ..
end
lemma minus_cancel [simp]: "xs - xs = []"
by (simp add: minus_list_def)
lemma append_minus [simp]: "(xs @ ys) - xs = ys"
by (simp add: minus_list_def)
lemma minus_right_nil [simp]: "xs - [] = xs"
by (simp add: minus_list_def)
lemma list_concat_minus_list_concat: "(s @ t) - (s @ z) = t - z"
by (simp add: minus_list_def)
lemma length_minus_list: "y \<le> x \<Longrightarrow> length(x - y) = length(x) - length(y)"
by (simp add: less_eq_list_def minus_list_def)
lemma map_list_minus:
"xs \<le> ys \<Longrightarrow> map f (ys - xs) = map f ys - map f xs"
by (simp add: drop_map less_eq_list_def map_mono_prefix minus_list_def)
lemma list_minus_first_tl [simp]:
"[x] \<le> xs \<Longrightarrow> (xs - [x]) = tl xs"
by (metis Prefix_Order.prefixE append.left_neutral append_minus list.sel(3) not_Cons_self2 tl_append2)
text \<open> Extra lemmas about @{term prefix} and @{term strict_prefix} \<close>
lemma prefix_concat_minus:
assumes "prefix xs ys"
shows "xs @ (ys - xs) = ys"
using assms by (metis minus_list_def prefix_drop)
lemma prefix_minus_concat:
assumes "prefix s t"
shows "(t - s) @ z = (t @ z) - s"
using assms by (simp add: Sublist.prefix_length_le minus_list_def)
lemma strict_prefix_minus_not_empty:
assumes "strict_prefix xs ys"
shows "ys - xs \<noteq> []"
using assms by (metis append_Nil2 prefix_concat_minus strict_prefix_def)
lemma strict_prefix_diff_minus:
assumes "prefix xs ys" and "xs \<noteq> ys"
shows "(ys - xs) \<noteq> []"
using assms by (simp add: strict_prefix_minus_not_empty)
lemma length_tl_list_minus_butlast_gt_zero:
assumes "length s < length t" and "strict_prefix (butlast s) t" and "length s > 0"
shows "length (tl (t - (butlast s))) > 0"
using assms
by (metis Nitpick.size_list_simp(2) butlast_snoc hd_Cons_tl length_butlast length_greater_0_conv length_tl less_trans nat_neq_iff strict_prefix_minus_not_empty prefix_order.dual_order.strict_implies_order prefix_concat_minus)
lemma list_minus_butlast_eq_butlast_list:
assumes "length t = length s" and "strict_prefix (butlast s) t"
shows "t - (butlast s) = [last t]"
using assms
by (metis append_butlast_last_id append_eq_append_conv butlast.simps(1) length_butlast less_numeral_extra(3) list.size(3) prefix_order.dual_order.strict_implies_order prefix_concat_minus prefix_length_less)
lemma butlast_strict_prefix_length_lt_imp_last_tl_minus_butlast_eq_last:
assumes "length s > 0" "strict_prefix (butlast s) t" "length s < length t"
shows "last (tl (t - (butlast s))) = (last t)"
using assms by (metis last_append last_tl length_tl_list_minus_butlast_gt_zero less_numeral_extra(3) list.size(3) append_minus strict_prefix_eq_exists)
lemma tl_list_minus_butlast_not_empty:
assumes "strict_prefix (butlast s) t" and "length s > 0" and "length t > length s"
shows "tl (t - (butlast s)) \<noteq> []"
using assms length_tl_list_minus_butlast_gt_zero by fastforce
lemma tl_list_minus_butlast_empty:
assumes "strict_prefix (butlast s) t" and "length s > 0" and "length t = length s"
shows "tl (t - (butlast s)) = []"
using assms by (simp add: list_minus_butlast_eq_butlast_list)
lemma tl_list_minus_butlast_eq_empty:
assumes "strict_prefix (butlast s) t" and "length s = length t"
shows "tl (t - (butlast s)) = []"
using assms by (metis list.sel(3) list_minus_butlast_eq_butlast_list)
lemma length_list_minus:
assumes "strict_prefix s t"
shows "length(t - s) = length(t) - length(s)"
using assms by (simp add: minus_list_def prefix_order.dual_order.strict_implies_order)
subsection \<open> Laws on @{term take}, @{term drop}, and @{term nths} \<close>
lemma take_prefix: "m \<le> n \<Longrightarrow> take m xs \<le> take n xs"
by (metis Prefix_Order.prefixI append_take_drop_id min_absorb2 take_append take_take)
lemma nths_atLeastAtMost_0_take: "nths xs {0..m} = take (Suc m) xs"
by (metis atLeast0AtMost lessThan_Suc_atMost nths_upt_eq_take)
lemma nths_atLeastLessThan_0_take: "nths xs {0..<m} = take m xs"
by (simp add: atLeast0LessThan)
lemma nths_atLeastAtMost_prefix: "m \<le> n \<Longrightarrow> nths xs {0..m} \<le> nths xs {0..n}"
by (simp add: nths_atLeastAtMost_0_take take_prefix)
lemma sorted_nths_atLeastAtMost_0: "\<lbrakk> m \<le> n; sorted (nths xs {0..n}) \<rbrakk> \<Longrightarrow> sorted (nths xs {0..m})"
using nths_atLeastAtMost_prefix sorted_prefix by blast
lemma sorted_nths_atLeastLessThan_0: "\<lbrakk> m \<le> n; sorted (nths xs {0..<n}) \<rbrakk> \<Longrightarrow> sorted (nths xs {0..<m})"
by (metis atLeast0LessThan nths_upt_eq_take sorted_prefix take_prefix)
lemma list_augment_as_update:
"k < length xs \<Longrightarrow> list_augment xs k x = list_update xs k x"
by (metis list_augment_def list_augment_idem list_update_overwrite)
lemma nths_list_update_out: "k \<notin> A \<Longrightarrow> nths (list_update xs k x) A = nths xs A"
apply (induct xs arbitrary: k x A)
apply (auto)
apply (rename_tac a xs k x A)
apply (case_tac k)
apply (auto simp add: nths_Cons)
done
lemma nths_list_augment_out: "\<lbrakk> k < length xs; k \<notin> A \<rbrakk> \<Longrightarrow> nths (list_augment xs k x) A = nths xs A"
by (simp add: list_augment_as_update nths_list_update_out)
end
\ No newline at end of file

File Metadata

Mime Type
application/octet-stream
Expires
Mon, May 6, 6:12 AM (1 d, 23 h)
Storage Engine
chunks
Storage Format
Chunks
Storage Handle
kJujfVMAIPW7
Default Alt Text
(4 MB)

Event Timeline