section ‹The Simplex Algorithm›
theory Simplex
imports
Linear_Poly_Maps
QDelta
Rel_Chain
Simplex_Algebra
"HOL-Library.RBT_Mapping"
"HOL-Library.Permutation"
"HOL-Library.Code_Target_Numeral"
begin
text‹Linear constraints are of the form ‹p ⨝ c› or ‹p⇩1 ⨝ p⇩2›, where ‹p›, ‹p⇩1›, and ‹p⇩2›, are
linear polynomials, ‹c› is a rational constant and ‹⨝ ∈
{<, >, ≤, ≥, =}›. Their abstract syntax is given by the ‹constraint› type, and semantics is given by the relation ‹⊨⇩c›, defined straightforwardly by primitive recursion over the
‹constraint› type. A set of constraints is satisfied,
denoted by ‹⊨⇩c⇩s›, if all constraints are. There is also an indexed
version ‹⊨⇩i⇩c⇩s› which takes an explicit set of indices and then only
demands that these constraints are satisfied.›
datatype constraint = LT linear_poly rat
| GT linear_poly rat
| LEQ linear_poly rat
| GEQ linear_poly rat
| EQ linear_poly rat
| LTPP linear_poly linear_poly
| GTPP linear_poly linear_poly
| LEQPP linear_poly linear_poly
| GEQPP linear_poly linear_poly
| EQPP linear_poly linear_poly
text ‹Indexed constraints are just pairs of indices and constraints. Indices will be used
to identify constraints, e.g., to easily specify an unsatisfiable core by a list of indices.›
type_synonym 'i i_constraint = "'i × constraint"
abbreviation (input) restrict_to :: "'i set ⇒ ('i × 'a) set ⇒ 'a set" where
"restrict_to I xs ≡ snd ` (xs ∩ (I × UNIV))"
text ‹The operation @{const restrict_to} is used to select constraints for a given index set.›
abbreviation (input) flat :: "('i × 'a) set ⇒ 'a set" where
"flat xs ≡ snd ` xs"
text ‹The operation @{const flat} is used to drop indices from a set of indexed constraints.›
abbreviation (input) flat_list :: "('i × 'a) list ⇒ 'a list" where
"flat_list xs ≡ map snd xs"
primrec
satisfies_constraint :: "'a :: lrv valuation ⇒ constraint ⇒ bool" (infixl "⊨⇩c" 100) where
"v ⊨⇩c (LT l r) ⟷ (l⦃v⦄) < r *R 1"
| "v ⊨⇩c GT l r ⟷ (l⦃v⦄) > r *R 1"
| "v ⊨⇩c LEQ l r ⟷ (l⦃v⦄) ≤ r *R 1"
| "v ⊨⇩c GEQ l r ⟷ (l⦃v⦄) ≥ r *R 1"
| "v ⊨⇩c EQ l r ⟷ (l⦃v⦄) = r *R 1"
| "v ⊨⇩c LTPP l1 l2 ⟷ (l1⦃v⦄) < (l2⦃v⦄)"
| "v ⊨⇩c GTPP l1 l2 ⟷ (l1⦃v⦄) > (l2⦃v⦄)"
| "v ⊨⇩c LEQPP l1 l2 ⟷ (l1⦃v⦄) ≤ (l2⦃v⦄)"
| "v ⊨⇩c GEQPP l1 l2 ⟷ (l1⦃v⦄) ≥ (l2⦃v⦄)"
| "v ⊨⇩c EQPP l1 l2 ⟷ (l1⦃v⦄) = (l2⦃v⦄)"
abbreviation satisfies_constraints :: "rat valuation ⇒ constraint set ⇒ bool" (infixl "⊨⇩c⇩s" 100) where
"v ⊨⇩c⇩s cs ≡ ∀ c ∈ cs. v ⊨⇩c c"
lemma unsat_mono: assumes "¬ (∃ v. v ⊨⇩c⇩s cs)"
and "cs ⊆ ds"
shows "¬ (∃ v. v ⊨⇩c⇩s ds)"
using assms by auto
fun i_satisfies_cs (infixl "⊨⇩i⇩c⇩s" 100) where
"(I,v) ⊨⇩i⇩c⇩s cs ⟷ v ⊨⇩c⇩s restrict_to I cs"
definition distinct_indices :: "('i × 'c) list ⇒ bool" where
"distinct_indices as = (distinct (map fst as))"
lemma distinct_indicesD: "distinct_indices as ⟹ (i,x) ∈ set as ⟹ (i,y) ∈ set as ⟹ x = y"
unfolding distinct_indices_def by (rule eq_key_imp_eq_value)
text ‹For the unsat-core predicate we only demand minimality in case that the indices are distinct.
Otherwise, minimality does in general not hold. For instance, consider the input
constraints $c_1: x < 0$, $c_2: x > 2$ and $c_2: x < 1$ where the index $c_2$ occurs twice.
If the simplex-method first encounters constraint $c_1$, then it will detect that there is a conflict
between $c_1$ and the first $c_2$-constraint. Consequently, the index-set $\{c_1,c_2\}$ will be returned,
but this set is not minimal since $\{c_2\}$ is already unsatisfiable.›
definition minimal_unsat_core :: "'i set ⇒ 'i i_constraint list ⇒ bool" where
"minimal_unsat_core I ics = ((I ⊆ fst ` set ics) ∧ (¬ (∃ v. (I,v) ⊨⇩i⇩c⇩s set ics))
∧ (distinct_indices ics ⟶ (∀ J. J ⊂ I ⟶ (∃ v. (J,v) ⊨⇩i⇩c⇩s set ics))))"
subsection ‹Procedure Specification›
abbreviation (input) Unsat where "Unsat ≡ Inl"
abbreviation (input) Sat where "Sat ≡ Inr"
text‹The specification for the satisfiability check procedure is given by:›
locale Solve =
fixes solve :: "'i i_constraint list ⇒ 'i list + rat valuation"
assumes simplex_sat: "solve cs = Sat v ⟹ v ⊨⇩c⇩s flat (set cs)"
assumes simplex_unsat: "solve cs = Unsat I ⟹ minimal_unsat_core (set I) cs"
abbreviation (input) look where "look ≡ Mapping.lookup"
abbreviation (input) upd where "upd ≡ Mapping.update"
lemma look_upd: "look (upd k v m) = (look m)(k ↦ v)"
by (transfer, auto)
lemmas look_upd_simps[simp] = look_upd Mapping.lookup_empty
definition map2fun:: "(var, 'a :: zero) mapping ⇒ var ⇒ 'a" where
"map2fun v ≡ λx. case look v x of None ⇒ 0 | Some y ⇒ y"
syntax
"_map2fun" :: "(var, 'a) mapping ⇒ var ⇒ 'a" ("⟨_⟩")
translations
"⟨v⟩" == "CONST map2fun v"
lemma map2fun_def':
"⟨v⟩ x ≡ case Mapping.lookup v x of None ⇒ 0 | Some y ⇒ y"
by (auto simp add: map2fun_def)
text‹Note that the above specification requires returning a
valuation (defined as a HOL function), which is not efficiently
executable. In order to enable more efficient data structures for
representing valuations, a refinement of this specification is needed
and the function ‹solve› is replaced by the function ‹solve_exec› returning optional ‹(var, rat) mapping› instead
of ‹var ⇒ rat› function. This way, efficient data structures
for representing mappings can be easily plugged-in during code
generation \cite{florian-refinement}. A conversion from the ‹mapping› datatype to HOL function is denoted by ‹⟨_⟩› and
given by: @{thm map2fun_def'[no_vars]}.›
locale SolveExec =
fixes solve_exec :: "'i i_constraint list ⇒ 'i list + (var, rat) mapping"
assumes simplex_sat0: "solve_exec cs = Sat v ⟹ ⟨v⟩ ⊨⇩c⇩s flat (set cs)"
assumes simplex_unsat0: "solve_exec cs = Unsat I ⟹ minimal_unsat_core (set I) cs"
begin
definition solve where
"solve cs ≡ case solve_exec cs of Sat v ⇒ Sat ⟨v⟩ | Unsat c ⇒ Unsat c"
end
sublocale SolveExec < Solve solve
by (unfold_locales, insert simplex_sat0 simplex_unsat0,
auto simp: solve_def split: sum.splits)
subsection ‹Handling Strict Inequalities›
text‹The first step of the procedure is removing all equalities and
strict inequalities. Equalities can be easily rewritten to non-strict
inequalities. Removing strict inequalities can be done by replacing
the list of constraints by a new one, formulated over an extension
‹ℚ'› of the space of rationals ‹ℚ›. ‹ℚ'› must
have a structure of a linearly ordered vector space over ‹ℚ›
(represented by the type class ‹lrv›) and must guarantee that
if some non-strict constraints are satisfied in ‹ℚ'›, then
there is a satisfying valuation for the original constraints in ‹ℚ›. Our final implementation uses the ‹ℚ⇩δ› space, defined in
\cite{simplex-rad} (basic idea is to replace ‹p < c› by ‹p ≤ c - δ› and ‹p > c› by ‹p ≥ c + δ› for a symbolic
parameter ‹δ›). So, all constraints are reduced to the form
‹p ⨝ b›, where ‹p› is a linear polynomial (still over
‹ℚ›), ‹b› is constant from ‹ℚ'› and ‹⨝
∈ {≤, ≥}›. The non-strict constraints are represented by the type
‹'a ns_constraint›, and their semantics is denoted by ‹⊨⇩n⇩s› and ‹⊨⇩n⇩s⇩s›. The indexed variant is ‹⊨⇩i⇩n⇩s⇩s›.›
datatype 'a ns_constraint = LEQ_ns linear_poly 'a | GEQ_ns linear_poly 'a
type_synonym ('i,'a) i_ns_constraint = "'i × 'a ns_constraint"
primrec satisfiable_ns_constraint :: "'a::lrv valuation ⇒ 'a ns_constraint ⇒ bool" (infixl "⊨⇩n⇩s" 100) where
"v ⊨⇩n⇩s LEQ_ns l r ⟷ l⦃v⦄ ≤ r"
| "v ⊨⇩n⇩s GEQ_ns l r ⟷ l⦃v⦄ ≥ r"
abbreviation satisfies_ns_constraints :: "'a::lrv valuation ⇒ 'a ns_constraint set ⇒ bool" (infixl "⊨⇩n⇩s⇩s " 100) where
"v ⊨⇩n⇩s⇩s cs ≡ ∀ c ∈ cs. v ⊨⇩n⇩s c"
fun i_satisfies_ns_constraints :: "'i set × 'a::lrv valuation ⇒ ('i,'a) i_ns_constraint set ⇒ bool" (infixl "⊨⇩i⇩n⇩s⇩s " 100) where
"(I,v) ⊨⇩i⇩n⇩s⇩s cs ⟷ v ⊨⇩n⇩s⇩s restrict_to I cs"
lemma i_satisfies_ns_constraints_mono:
"(I,v) ⊨⇩i⇩n⇩s⇩s cs ⟹ J ⊆ I ⟹ (J,v) ⊨⇩i⇩n⇩s⇩s cs"
by auto
primrec poly :: "'a ns_constraint ⇒ linear_poly" where
"poly (LEQ_ns p a) = p"
| "poly (GEQ_ns p a) = p"
primrec ns_constraint_const :: "'a ns_constraint ⇒ 'a" where
"ns_constraint_const (LEQ_ns p a) = a"
| "ns_constraint_const (GEQ_ns p a) = a"
definition distinct_indices_ns :: "('i,'a :: lrv) i_ns_constraint set ⇒ bool" where
"distinct_indices_ns ns = ((∀ n1 n2 i. (i,n1) ∈ ns ⟶ (i,n2) ∈ ns ⟶
poly n1 = poly n2 ∧ ns_constraint_const n1 = ns_constraint_const n2))"
definition minimal_unsat_core_ns :: "'i set ⇒ ('i,'a :: lrv) i_ns_constraint set ⇒ bool" where
"minimal_unsat_core_ns I cs = ((I ⊆ fst ` cs) ∧ (¬ (∃ v. (I,v) ⊨⇩i⇩n⇩s⇩s cs))
∧ (distinct_indices_ns cs ⟶ (∀ J ⊂ I. ∃ v. (J,v) ⊨⇩i⇩n⇩s⇩s cs)))"
text‹Specification of reduction of constraints to non-strict form is given by:›
locale To_ns =
fixes to_ns :: "'i i_constraint list ⇒ ('i,'a::lrv) i_ns_constraint list"
fixes from_ns :: "(var, 'a) mapping ⇒ 'a ns_constraint list ⇒ (var, rat) mapping"
assumes to_ns_unsat: "minimal_unsat_core_ns I (set (to_ns cs)) ⟹ minimal_unsat_core I cs"
assumes i_to_ns_sat: "(I,⟨v'⟩) ⊨⇩i⇩n⇩s⇩s set (to_ns cs) ⟹ (I,⟨from_ns v' (flat_list (to_ns cs))⟩) ⊨⇩i⇩c⇩s set cs"
assumes to_ns_indices: "fst ` set (to_ns cs) = fst ` set cs"
assumes distinct_cond: "distinct_indices cs ⟹ distinct_indices_ns (set (to_ns cs))"
begin
lemma to_ns_sat: "⟨v'⟩ ⊨⇩n⇩s⇩s flat (set (to_ns cs)) ⟹ ⟨from_ns v' (flat_list (to_ns cs))⟩ ⊨⇩c⇩s flat (set cs)"
using i_to_ns_sat[of UNIV v' cs] by auto
end
locale Solve_exec_ns =
fixes solve_exec_ns :: "('i,'a::lrv) i_ns_constraint list ⇒ 'i list + (var, 'a) mapping"
assumes simplex_ns_sat: "solve_exec_ns cs = Sat v ⟹ ⟨v⟩ ⊨⇩n⇩s⇩s flat (set cs)"
assumes simplex_ns_unsat: "solve_exec_ns cs = Unsat I ⟹ minimal_unsat_core_ns (set I) (set cs)"
text‹After the transformation, the procedure is reduced to solving
only the non-strict constraints, implemented in the ‹solve_exec_ns› function having an analogous specification to the
‹solve› function. If ‹to_ns›, ‹from_ns› and
‹solve_exec_ns› are available, the ‹solve_exec›
function can be easily defined and it can be easily shown that this
definition satisfies its specification (also analogous to ‹solve›).
›
locale SolveExec' = To_ns to_ns from_ns + Solve_exec_ns solve_exec_ns for
to_ns:: "'i i_constraint list ⇒ ('i,'a::lrv) i_ns_constraint list" and
from_ns :: "(var, 'a) mapping ⇒ 'a ns_constraint list ⇒ (var, rat) mapping" and
solve_exec_ns :: "('i,'a) i_ns_constraint list ⇒ 'i list + (var, 'a) mapping"
begin
definition solve_exec where
"solve_exec cs ≡ let cs' = to_ns cs in case solve_exec_ns cs'
of Sat v ⇒ Sat (from_ns v (flat_list cs'))
| Unsat is ⇒ Unsat is"
end
sublocale SolveExec' < SolveExec solve_exec
by (unfold_locales, insert simplex_ns_sat simplex_ns_unsat to_ns_sat to_ns_unsat,
(force simp: solve_exec_def Let_def split: sum.splits)+)
subsection ‹Preprocessing›
text‹The next step in the procedure rewrites a list of non-strict
constraints into an equisatisfiable form consisting of a list of
linear equations (called the \emph{tableau}) and of a list of
\emph{atoms} of the form ‹x⇩i ⨝ b⇩i› where ‹x⇩i› is a
variable and ‹b⇩i› is a constant (from the extension field). The
transformation is straightforward and introduces auxiliary variables
for linear polynomials occurring in the initial formula. For example,
‹[x⇩1 + x⇩2 ≤ b⇩1, x⇩1 + x⇩2 ≥ b⇩2, x⇩2 ≥ b⇩3]› can be transformed to
the tableau ‹[x⇩3 = x⇩1 + x⇩2]› and atoms ‹[x⇩3 ≤ b⇩1, x⇩3 ≥
b⇩2, x⇩2 ≥ b⇩3]›.›
type_synonym eq = "var × linear_poly"
primrec lhs :: "eq ⇒ var" where "lhs (l, r) = l"
primrec rhs :: "eq ⇒ linear_poly" where "rhs (l, r) = r"
abbreviation rvars_eq :: "eq ⇒ var set" where
"rvars_eq eq ≡ vars (rhs eq)"
definition satisfies_eq :: "'a::rational_vector valuation ⇒ eq ⇒ bool" (infixl "⊨⇩e" 100) where
"v ⊨⇩e eq ≡ v (lhs eq) = (rhs eq)⦃v⦄"
lemma satisfies_eq_iff: "v ⊨⇩e (x, p) ≡ v x = p⦃v⦄"
by (simp add: satisfies_eq_def)
type_synonym tableau ="eq list"
definition satisfies_tableau ::"'a::rational_vector valuation ⇒ tableau ⇒ bool" (infixl "⊨⇩t" 100) where
"v ⊨⇩t t ≡ ∀ e ∈ set t. v ⊨⇩e e"
definition lvars :: "tableau ⇒ var set" where
"lvars t = set (map lhs t)"
definition rvars :: "tableau ⇒ var set" where
"rvars t = ⋃ (set (map rvars_eq t))"
abbreviation tvars where "tvars t ≡ lvars t ∪ rvars t"
text ‹The condition that the rhss are non-zero is required to obtain minimal unsatisfiable cores.
To observe the problem with 0 as rhs, consider the tableau $x = 0$ in combination
with atom $(A: x \leq 0)$ where then $(B: x \geq 1)$ is asserted.
In this case, the unsat core would be computed as $\{A,B\}$, although already $\{B\}$ is unsatisfiable.›
definition normalized_tableau :: "tableau ⇒ bool" ("△") where
"normalized_tableau t ≡ distinct (map lhs t) ∧ lvars t ∩ rvars t = {} ∧ 0 ∉ rhs ` set t"
text‹Equations are of the form ‹x = p›, where ‹x› is
a variable and ‹p› is a polynomial, and are represented by the
type ‹eq = var × linear_poly›. Semantics of equations is given
by @{thm satisfies_eq_iff[no_vars]}. Tableau is represented as a list
of equations, by the type ‹tableau = eq list›. Semantics for a
tableau is given by @{thm satisfies_tableau_def[no_vars]}. Functions
‹lvars› and ‹rvars› return sets of variables appearing on
the left hand side (lhs) and the right hand side (rhs) of a
tableau. Lhs variables are called \emph{basic} while rhs variables are
called \emph{non-basic} variables. A tableau ‹t› is
\emph{normalized} (denoted by @{term "△ t"}) iff no variable occurs on
the lhs of two equations in a tableau and if sets of lhs and rhs
variables are distinct.›
lemma normalized_tableau_unique_eq_for_lvar:
assumes "△ t"
shows "∀ x ∈ lvars t. ∃! p. (x, p) ∈ set t"
proof (safe)
fix x
assume "x ∈ lvars t"
then show "∃p. (x, p) ∈ set t"
unfolding lvars_def
by auto
next
fix x p1 p2
assume *: "(x, p1) ∈ set t" "(x, p2) ∈ set t"
then show "p1 = p2"
using ‹△ t›
unfolding normalized_tableau_def
by (force simp add: distinct_map inj_on_def)
qed
lemma recalc_tableau_lvars:
assumes "△ t"
shows "∀ v. ∃ v'. (∀ x ∈ rvars t. v x = v' x) ∧ v' ⊨⇩t t"
proof
fix v
let ?v' = "λ x. if x ∈ lvars t then (THE p. (x, p) ∈ set t) ⦃ v ⦄ else v x"
show "∃ v'. (∀ x ∈ rvars t. v x = v' x) ∧ v' ⊨⇩t t"
proof (rule_tac x="?v'" in exI, rule conjI)
show "∀x∈rvars t. v x = ?v' x"
using ‹△ t›
unfolding normalized_tableau_def
by auto
show "?v' ⊨⇩t t"
unfolding satisfies_tableau_def satisfies_eq_def
proof
fix e
assume "e ∈ set t"
obtain l r where e: "e = (l,r)" by force
show "?v' (lhs e) = rhs e ⦃ ?v' ⦄"
proof-
have "(lhs e, rhs e) ∈ set t"
using ‹e ∈ set t› e by auto
have "∃!p. (lhs e, p) ∈ set t"
using ‹△ t› normalized_tableau_unique_eq_for_lvar[of t]
using ‹e ∈ set t›
unfolding lvars_def
by auto
let ?p = "THE p. (lhs e, p) ∈ set t"
have "(lhs e, ?p) ∈ set t"
apply (rule theI')
using ‹∃!p. (lhs e, p) ∈ set t›
by auto
then have "?p = rhs e"
using ‹(lhs e, rhs e) ∈ set t›
using ‹∃!p. (lhs e, p) ∈ set t›
by auto
moreover
have "?v' (lhs e) = ?p ⦃ v ⦄"
using ‹e ∈ set t›
unfolding lvars_def
by simp
moreover
have "rhs e ⦃ ?v' ⦄ = rhs e ⦃ v ⦄"
apply (rule valuate_depend)
using ‹△ t› ‹e ∈ set t›
unfolding normalized_tableau_def
by (auto simp add: lvars_def rvars_def)
ultimately
show ?thesis
by auto
qed
qed
qed
qed
lemma tableau_perm:
assumes "lvars t1 = lvars t2" "rvars t1 = rvars t2"
"△ t1" "△ t2" "⋀ v::'a::lrv valuation. v ⊨⇩t t1 ⟷ v ⊨⇩t t2"
shows "t1 <~~> t2"
proof-
{
fix t1 t2
assume "lvars t1 = lvars t2" "rvars t1 = rvars t2"
"△ t1" "⋀ v::'a::lrv valuation. v ⊨⇩t t1 ⟷ v ⊨⇩t t2"
have "set t1 ⊆ set t2"
proof (safe)
fix a b
assume "(a, b) ∈ set t1"
then have "a ∈ lvars t1"
unfolding lvars_def
by force
then have "a ∈ lvars t2"
using ‹lvars t1 = lvars t2›
by simp
then obtain b' where "(a, b') ∈ set t2"
unfolding lvars_def
by force
have "∀v::'a valuation. ∃v'. (∀x∈vars (b - b'). v' x = v x) ∧ (b - b') ⦃ v' ⦄ = 0"
proof
fix v::"'a valuation"
obtain v' where "v' ⊨⇩t t1" "∀ x ∈ rvars t1. v x = v' x"
using recalc_tableau_lvars[of t1] ‹△ t1›
by auto
have "v' ⊨⇩t t2"
using ‹v' ⊨⇩t t1› ‹⋀ v. v ⊨⇩t t1 ⟷ v ⊨⇩t t2›
by simp
have "b ⦃v'⦄ = b' ⦃v'⦄"
using ‹(a, b) ∈ set t1› ‹v' ⊨⇩t t1›
using ‹(a, b') ∈ set t2› ‹v' ⊨⇩t t2›
unfolding satisfies_tableau_def satisfies_eq_def
by force
then have "(b - b') ⦃v'⦄ = 0"
using valuate_minus[of b b' v']
by auto
moreover
have "vars b ⊆ rvars t1" "vars b' ⊆ rvars t1"
using ‹(a, b) ∈ set t1› ‹(a, b') ∈ set t2› ‹rvars t1 = rvars t2›
unfolding rvars_def
by force+
then have "vars (b - b') ⊆ rvars t1"
using vars_minus[of b b']
by blast
then have "∀x∈vars (b - b'). v' x = v x"
using ‹∀ x ∈ rvars t1. v x = v' x›
by auto
ultimately
show "∃v'. (∀x∈vars (b - b'). v' x = v x) ∧ (b - b') ⦃ v' ⦄ = 0"
by auto
qed
then have "b = b'"
using all_val[of "b - b'"]
by simp
then show "(a, b) ∈ set t2"
using ‹(a, b') ∈ set t2›
by simp
qed
}
note * = this
have "set t1 = set t2"
using *[of t1 t2] *[of t2 t1]
using assms
by auto
moreover
have "distinct t1" "distinct t2"
using ‹△ t1› ‹△ t2›
unfolding normalized_tableau_def
by (auto simp add: distinct_map)
ultimately
show ?thesis
by (auto simp add: set_eq_iff_mset_eq_distinct mset_eq_perm)
qed
text‹Elementary atoms are represented by the type ‹'a atom›
and semantics for atoms and sets of atoms is denoted by ‹⊨⇩a› and
‹⊨⇩a⇩s› and given by:
›
datatype 'a atom = Leq var 'a | Geq var 'a
primrec atom_var::"'a atom ⇒ var" where
"atom_var (Leq var a) = var"
| "atom_var (Geq var a) = var"
primrec atom_const::"'a atom ⇒ 'a" where
"atom_const (Leq var a) = a"
| "atom_const (Geq var a) = a"
primrec satisfies_atom :: "'a::linorder valuation ⇒ 'a atom ⇒ bool" (infixl "⊨⇩a" 100) where
"v ⊨⇩a Leq x c ⟷ v x ≤ c" | "v ⊨⇩a Geq x c ⟷ v x ≥ c"
definition satisfies_atom_set :: "'a::linorder valuation ⇒ 'a atom set ⇒ bool" (infixl "⊨⇩a⇩s" 100) where
"v ⊨⇩a⇩s as ≡ ∀ a ∈ as. v ⊨⇩a a"
definition satisfies_atom' :: "'a::linorder valuation ⇒ 'a atom ⇒ bool" (infixl "⊨⇩a⇩e" 100) where
"v ⊨⇩a⇩e a ⟷ v (atom_var a) = atom_const a"
lemma satisfies_atom'_stronger: "v ⊨⇩a⇩e a ⟹ v ⊨⇩a a" by (cases a, auto simp: satisfies_atom'_def)
abbreviation satisfies_atom_set' :: "'a::linorder valuation ⇒ 'a atom set ⇒ bool" (infixl "⊨⇩a⇩e⇩s" 100) where
"v ⊨⇩a⇩e⇩s as ≡ ∀ a ∈ as. v ⊨⇩a⇩e a"
lemma satisfies_atom_set'_stronger: "v ⊨⇩a⇩e⇩s as ⟹ v ⊨⇩a⇩s as"
using satisfies_atom'_stronger unfolding satisfies_atom_set_def by auto
text ‹There is also the indexed variant of an atom›
type_synonym ('i,'a) i_atom = "'i × 'a atom"
fun i_satisfies_atom_set :: "'i set × 'a::linorder valuation ⇒ ('i,'a) i_atom set ⇒ bool" (infixl "⊨⇩i⇩a⇩s" 100) where
"(I,v) ⊨⇩i⇩a⇩s as ⟷ v ⊨⇩a⇩s restrict_to I as"
fun i_satisfies_atom_set' :: "'i set × 'a::linorder valuation ⇒ ('i,'a) i_atom set ⇒ bool" (infixl "⊨⇩i⇩a⇩e⇩s" 100) where
"(I,v) ⊨⇩i⇩a⇩e⇩s as ⟷ v ⊨⇩a⇩e⇩s restrict_to I as"
lemma i_satisfies_atom_set'_stronger: "Iv ⊨⇩i⇩a⇩e⇩s as ⟹ Iv ⊨⇩i⇩a⇩s as"
using satisfies_atom_set'_stronger[of _ "snd Iv"] by (cases Iv, auto)
lemma satisifies_atom_restrict_to_Cons: "v ⊨⇩a⇩s restrict_to I (set as) ⟹ (i ∈ I ⟹ v ⊨⇩a a)
⟹ v ⊨⇩a⇩s restrict_to I (set ((i,a) # as))"
unfolding satisfies_atom_set_def by auto
lemma satisfies_tableau_Cons: "v ⊨⇩t t ⟹ v ⊨⇩e e ⟹ v ⊨⇩t (e # t)"
unfolding satisfies_tableau_def by auto
definition distinct_indices_atoms :: "('i,'a) i_atom set ⇒ bool" where
"distinct_indices_atoms as = (∀ i a b. (i,a) ∈ as ⟶ (i,b) ∈ as ⟶ atom_var a = atom_var b ∧ atom_const a = atom_const b)"
text‹
The specification of the preprocessing function is given by:›
locale Preprocess = fixes preprocess::"('i,'a::lrv) i_ns_constraint list ⇒ tableau× ('i,'a) i_atom list
× ((var,'a) mapping ⇒ (var,'a) mapping) × 'i list"
assumes
preprocess_tableau_normalized: "preprocess cs = (t,as,trans_v,U) ⟹ △ t" and
i_preprocess_sat: "⋀ v. preprocess cs = (t,as,trans_v,U) ⟹ I ∩ set U = {} ⟹ (I,⟨v⟩) ⊨⇩i⇩a⇩s set as ⟹ ⟨v⟩ ⊨⇩t t ⟹ (I,⟨trans_v v⟩) ⊨⇩i⇩n⇩s⇩s set cs" and
preprocess_unsat: "preprocess cs = (t, as,trans_v,U) ⟹ (I,v) ⊨⇩i⇩n⇩s⇩s set cs ⟹ ∃ v'. (I,v') ⊨⇩i⇩a⇩s set as ∧ v' ⊨⇩t t" and
preprocess_distinct: "preprocess cs = (t, as,trans_v, U) ⟹ distinct_indices_ns (set cs) ⟹ distinct_indices_atoms (set as)" and
preprocess_unsat_indices: "preprocess cs = (t, as,trans_v, U) ⟹ i ∈ set U ⟹ ¬ (∃ v. ({i},v) ⊨⇩i⇩n⇩s⇩s set cs)" and
preprocess_index: "preprocess cs = (t,as,trans_v, U) ⟹ fst ` set as ∪ set U ⊆ fst ` set cs"
begin
lemma preprocess_sat: "preprocess cs = (t,as,trans_v,U) ⟹ U = [] ⟹ ⟨v⟩ ⊨⇩a⇩s flat (set as) ⟹ ⟨v⟩ ⊨⇩t t ⟹ ⟨trans_v v⟩ ⊨⇩n⇩s⇩s flat (set cs)"
using i_preprocess_sat[of cs t as trans_v U UNIV v] by auto
end
definition minimal_unsat_core_tabl_atoms :: "'i set ⇒ tableau ⇒ ('i,'a::lrv) i_atom set ⇒ bool" where
"minimal_unsat_core_tabl_atoms I t as = ( I ⊆ fst ` as ∧ (¬ (∃ v. v ⊨⇩t t ∧ (I,v) ⊨⇩i⇩a⇩s as)) ∧
(distinct_indices_atoms as ⟶ (∀ J ⊂ I. ∃ v. v ⊨⇩t t ∧ (J,v) ⊨⇩i⇩a⇩e⇩s as)))"
lemma minimal_unsat_core_tabl_atomsD: assumes "minimal_unsat_core_tabl_atoms I t as"
shows "I ⊆ fst ` as"
"¬ (∃ v. v ⊨⇩t t ∧ (I,v) ⊨⇩i⇩a⇩s as)"
"distinct_indices_atoms as ⟹ J ⊂ I ⟹ ∃ v. v ⊨⇩t t ∧ (J,v) ⊨⇩i⇩a⇩e⇩s as"
using assms unfolding minimal_unsat_core_tabl_atoms_def by auto
locale AssertAll =
fixes assert_all :: "tableau ⇒ ('i,'a::lrv) i_atom list ⇒ 'i list + (var, 'a)mapping"
assumes assert_all_sat: "△ t ⟹ assert_all t as = Sat v ⟹ ⟨v⟩ ⊨⇩t t ∧ ⟨v⟩ ⊨⇩a⇩s flat (set as)"
assumes assert_all_unsat: "△ t ⟹ assert_all t as = Unsat I ⟹ minimal_unsat_core_tabl_atoms (set I) t (set as)"
text‹Once the preprocessing is done and tableau and atoms are
obtained, their satisfiability is checked by the
‹assert_all› function. Its precondition is that the starting
tableau is normalized, and its specification is analogue to the one for the
‹solve› function. If ‹preprocess› and ‹assert_all›
are available, the ‹solve_exec_ns› can be defined, and it
can easily be shown that this definition satisfies the specification.›
locale Solve_exec_ns' = Preprocess preprocess + AssertAll assert_all for
preprocess:: "('i,'a::lrv) i_ns_constraint list ⇒ tableau × ('i,'a) i_atom list × ((var,'a)mapping ⇒ (var,'a)mapping) × 'i list" and
assert_all :: "tableau ⇒ ('i,'a::lrv) i_atom list ⇒ 'i list + (var, 'a) mapping"
begin
definition solve_exec_ns where
"solve_exec_ns s ≡
case preprocess s of (t,as,trans_v,ui) ⇒
(case ui of i # _ ⇒ Inl [i] | _ ⇒
(case assert_all t as of Inl I ⇒ Inl I | Inr v ⇒ Inr (trans_v v))) "
end
context Preprocess
begin
lemma preprocess_unsat_index: assumes prep: "preprocess cs = (t,as,trans_v,ui)"
and i: "i ∈ set ui"
shows "minimal_unsat_core_ns {i} (set cs)"
proof -
from preprocess_index[OF prep] have "set ui ⊆ fst ` set cs" by auto
with i have i': "i ∈ fst ` set cs" by auto
from preprocess_unsat_indices[OF prep i]
show ?thesis unfolding minimal_unsat_core_ns_def using i' by auto
qed
lemma preprocess_minimal_unsat_core: assumes prep: "preprocess cs = (t,as,trans_v,ui)"
and unsat: "minimal_unsat_core_tabl_atoms I t (set as)"
and inter: "I ∩ set ui = {}"
shows "minimal_unsat_core_ns I (set cs)"
proof -
from preprocess_tableau_normalized[OF prep]
have t: "△ t" .
from preprocess_index[OF prep] have index: "fst ` set as ∪ set ui ⊆ fst ` set cs" by auto
from minimal_unsat_core_tabl_atomsD(1,2)[OF unsat] preprocess_unsat[OF prep, of I]
have 1: "I ⊆ fst ` set as" "¬ (∃ v. (I, v) ⊨⇩i⇩n⇩s⇩s set cs)" by force+
show "minimal_unsat_core_ns I (set cs)" unfolding minimal_unsat_core_ns_def
proof (intro conjI impI allI 1(2))
show "I ⊆ fst ` set cs" using 1 index by auto
fix J
assume "distinct_indices_ns (set cs)" "J ⊂ I"
with preprocess_distinct[OF prep]
have "distinct_indices_atoms (set as)" "J ⊂ I" by auto
from minimal_unsat_core_tabl_atomsD(3)[OF unsat this]
obtain v where model: "v ⊨⇩t t" "(J, v) ⊨⇩i⇩a⇩e⇩s set as" by auto
from i_satisfies_atom_set'_stronger[OF model(2)]
have model': "(J, v) ⊨⇩i⇩a⇩s set as" .
define w where "w = Mapping.Mapping (λ x. Some (v x))"
have "v = ⟨w⟩" unfolding w_def map2fun_def
by (intro ext, transfer, auto)
with model model' have "⟨w⟩ ⊨⇩t t" "(J, ⟨w⟩) ⊨⇩i⇩a⇩s set as" by auto
from i_preprocess_sat[OF prep _ this(2,1)] ‹J ⊂ I› inter
have "(J, ⟨trans_v w⟩) ⊨⇩i⇩n⇩s⇩s set cs" by auto
then show "∃ w. (J, w) ⊨⇩i⇩n⇩s⇩s set cs" by auto
qed
qed
end
sublocale Solve_exec_ns' < Solve_exec_ns solve_exec_ns
proof
fix cs
obtain t as trans_v ui where prep: "preprocess cs = (t,as,trans_v,ui)" by (cases "preprocess cs")
from preprocess_tableau_normalized[OF prep]
have t: "△ t" .
from preprocess_index[OF prep] have index: "fst ` set as ∪ set ui ⊆ fst ` set cs" by auto
note solve = solve_exec_ns_def[of cs, unfolded prep split]
{
fix v
assume "solve_exec_ns cs = Sat v"
from this[unfolded solve] t assert_all_sat[OF t] preprocess_sat[OF prep]
show " ⟨v⟩ ⊨⇩n⇩s⇩s flat (set cs)" by (auto split: sum.splits list.splits)
}
{
fix I
assume res: "solve_exec_ns cs = Unsat I"
show "minimal_unsat_core_ns (set I) (set cs)"
proof (cases ui)
case (Cons i uis)
hence I: "I = [i]" using res[unfolded solve] by auto
from preprocess_unsat_index[OF prep, of i] I Cons index show ?thesis by auto
next
case Nil
from res[unfolded solve Nil] have assert: "assert_all t as = Unsat I"
by (auto split: sum.splits)
from assert_all_unsat[OF t assert]
have "minimal_unsat_core_tabl_atoms (set I) t (set as)" .
from preprocess_minimal_unsat_core[OF prep this] Nil
show "minimal_unsat_core_ns (set I) (set cs)" by simp
qed
}
qed
subsection‹Incrementally Asserting Atoms›
text‹The function @{term assert_all} can be implemented by
iteratively asserting one by one atom from the given list of atoms.
›
type_synonym 'a bounds = "var ⇀ 'a"
text‹Asserted atoms will be stored in a form of \emph{bounds} for a
given variable. Bounds are of the form ‹l⇩i ≤ x⇩i ≤ u⇩i›, where
‹l⇩i› and ‹u⇩i› and are either scalars or $\pm
\infty$. Each time a new atom is asserted, a bound for the
corresponding variable is updated (checking for conflict with the
previous bounds). Since bounds for a variable can be either finite or
$\pm \infty$, they are represented by (partial) maps from variables to
values (‹'a bounds = var ⇀ 'a›). Upper and lower bounds are
represented separately. Infinite bounds map to @{term None} and this
is reflected in the semantics:
\begin{small}
‹c ≥⇩u⇩b b ⟷ case b of None ⇒ False | Some b' ⇒ c ≥ b'›
‹c ≤⇩u⇩b b ⟷ case b of None ⇒ True | Some b' ⇒ c ≤ b'›
\end{small}
\noindent Strict comparisons, and comparisons with lower bounds are performed similarly.
›
abbreviation (input) le where
"le lt x y ≡ lt x y ∨ x = y"
definition geub ("⊵⇩u⇩b") where
"⊵⇩u⇩b lt c b ≡ case b of None ⇒ False | Some b' ⇒ le lt b' c"
definition gtub ("⊳⇩u⇩b") where
"⊳⇩u⇩b lt c b ≡ case b of None ⇒ False | Some b' ⇒ lt b' c"
definition leub ("⊴⇩u⇩b") where
"⊴⇩u⇩b lt c b ≡ case b of None ⇒ True | Some b' ⇒ le lt c b'"
definition ltub ("⊲⇩u⇩b") where
"⊲⇩u⇩b lt c b ≡ case b of None ⇒ True | Some b' ⇒ lt c b'"
definition lelb ("⊴⇩l⇩b") where
"⊴⇩l⇩b lt c b ≡ case b of None ⇒ False | Some b' ⇒ le lt c b'"
definition ltlb ("⊲⇩l⇩b") where
"⊲⇩l⇩b lt c b ≡ case b of None ⇒ False | Some b' ⇒ lt c b'"
definition gelb ("⊵⇩l⇩b") where
"⊵⇩l⇩b lt c b ≡ case b of None ⇒ True | Some b' ⇒ le lt b' c"
definition gtlb ("⊳⇩l⇩b") where
"⊳⇩l⇩b lt c b ≡ case b of None ⇒ True | Some b' ⇒ lt b' c"
definition ge_ubound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl "≥⇩u⇩b" 100) where
"c ≥⇩u⇩b b = ⊵⇩u⇩b (<) c b"
definition gt_ubound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl ">⇩u⇩b" 100) where
"c >⇩u⇩b b = ⊳⇩u⇩b (<) c b"
definition le_ubound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl "≤⇩u⇩b" 100) where
"c ≤⇩u⇩b b = ⊴⇩u⇩b (<) c b"
definition lt_ubound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl "<⇩u⇩b" 100) where
"c <⇩u⇩b b = ⊲⇩u⇩b (<) c b"
definition le_lbound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl "≤⇩l⇩b" 100) where
"c ≤⇩l⇩b b = ⊴⇩l⇩b (<) c b"
definition lt_lbound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl "<⇩l⇩b" 100) where
"c <⇩l⇩b b = ⊲⇩l⇩b (<) c b"
definition ge_lbound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl "≥⇩l⇩b" 100) where
"c ≥⇩l⇩b b = ⊵⇩l⇩b (<) c b"
definition gt_lbound :: "'a::linorder ⇒ 'a option ⇒ bool" (infixl ">⇩l⇩b" 100) where
"c >⇩l⇩b b = ⊳⇩l⇩b (<) c b"
lemmas bound_compare'_defs =
geub_def gtub_def leub_def ltub_def
gelb_def gtlb_def lelb_def ltlb_def
lemmas bound_compare''_defs =
ge_ubound_def gt_ubound_def le_ubound_def lt_ubound_def
le_lbound_def lt_lbound_def ge_lbound_def gt_lbound_def
lemmas bound_compare_defs = bound_compare'_defs bound_compare''_defs
lemma opposite_dir [simp]:
"⊴⇩l⇩b (>) a b = ⊵⇩u⇩b (<) a b"
"⊴⇩u⇩b (>) a b = ⊵⇩l⇩b (<) a b"
"⊵⇩l⇩b (>) a b = ⊴⇩u⇩b (<) a b"
"⊵⇩u⇩b (>) a b = ⊴⇩l⇩b (<) a b"
"⊲⇩l⇩b (>) a b = ⊳⇩u⇩b (<) a b"
"⊲⇩u⇩b (>) a b = ⊳⇩l⇩b (<) a b"
"⊳⇩l⇩b (>) a b = ⊲⇩u⇩b (<) a b"
"⊳⇩u⇩b (>) a b = ⊲⇩l⇩b (<) a b"
by (case_tac[!] b) (auto simp add: bound_compare'_defs)
lemma [simp]: "¬ c ≥⇩u⇩b None " "¬ c ≤⇩l⇩b None"
by (auto simp add: bound_compare_defs)
lemma neg_bounds_compare:
"(¬ (c ≥⇩u⇩b b)) ⟹ c <⇩u⇩b b" "(¬ (c ≤⇩u⇩b b)) ⟹ c >⇩u⇩b b"
"(¬ (c >⇩u⇩b b)) ⟹ c ≤⇩u⇩b b" "(¬ (c <⇩u⇩b b)) ⟹ c ≥⇩u⇩b b"
"(¬ (c ≤⇩l⇩b b)) ⟹ c >⇩l⇩b b" "(¬ (c ≥⇩l⇩b b)) ⟹ c <⇩l⇩b b"
"(¬ (c <⇩l⇩b b)) ⟹ c ≥⇩l⇩b b" "(¬ (c >⇩l⇩b b)) ⟹ c ≤⇩l⇩b b"
by (case_tac[!] b) (auto simp add: bound_compare_defs)
lemma bounds_compare_contradictory [simp]:
"⟦c ≥⇩u⇩b b; c <⇩u⇩b b⟧ ⟹ False" "⟦c ≤⇩u⇩b b; c >⇩u⇩b b⟧ ⟹ False"
"⟦c >⇩u⇩b b; c ≤⇩u⇩b b⟧ ⟹ False" "⟦c <⇩u⇩b b; c ≥⇩u⇩b b⟧ ⟹ False"
"⟦c ≤⇩l⇩b b; c >⇩l⇩b b⟧ ⟹ False" "⟦c ≥⇩l⇩b b; c <⇩l⇩b b⟧ ⟹ False"
"⟦c <⇩l⇩b b; c ≥⇩l⇩b b⟧ ⟹ False" "⟦c >⇩l⇩b b; c ≤⇩l⇩b b⟧ ⟹ False"
by (case_tac[!] b) (auto simp add: bound_compare_defs)
lemma compare_strict_nonstrict:
"x <⇩u⇩b b ⟹ x ≤⇩u⇩b b"
"x >⇩u⇩b b ⟹ x ≥⇩u⇩b b"
"x <⇩l⇩b b ⟹ x ≤⇩l⇩b b"
"x >⇩l⇩b b ⟹ x ≥⇩l⇩b b"
by (case_tac[!] b) (auto simp add: bound_compare_defs)
lemma [simp]:
"⟦ x ≤ c; c <⇩u⇩b b ⟧ ⟹ x <⇩u⇩b b"
"⟦ x < c; c ≤⇩u⇩b b ⟧ ⟹ x <⇩u⇩b b"
"⟦ x ≤ c; c ≤⇩u⇩b b ⟧ ⟹ x ≤⇩u⇩b b"
"⟦ x ≥ c; c >⇩l⇩b b ⟧ ⟹ x >⇩l⇩b b"
"⟦ x > c; c ≥⇩l⇩b b ⟧ ⟹ x >⇩l⇩b b"
"⟦ x ≥ c; c ≥⇩l⇩b b ⟧ ⟹ x ≥⇩l⇩b b"
by (case_tac[!] b) (auto simp add: bound_compare_defs)
lemma bounds_lg [simp]:
"⟦ c >⇩u⇩b b; x ≤⇩u⇩b b⟧ ⟹ x < c"
"⟦ c ≥⇩u⇩b b; x <⇩u⇩b b⟧ ⟹ x < c"
"⟦ c ≥⇩u⇩b b; x ≤⇩u⇩b b⟧ ⟹ x ≤ c"
"⟦ c <⇩l⇩b b; x ≥⇩l⇩b b⟧ ⟹ x > c"
"⟦ c ≤⇩l⇩b b; x >⇩l⇩b b⟧ ⟹ x > c"
"⟦ c ≤⇩l⇩b b; x ≥⇩l⇩b b⟧ ⟹ x ≥ c"
by (case_tac[!] b) (auto simp add: bound_compare_defs)
lemma bounds_compare_Some [simp]:
"x ≤⇩u⇩b Some c ⟷ x ≤ c" "x ≥⇩u⇩b Some c ⟷ x ≥ c"
"x <⇩u⇩b Some c ⟷ x < c" "x >⇩u⇩b Some c ⟷ x > c"
"x ≥⇩l⇩b Some c ⟷ x ≥ c" "x ≤⇩l⇩b Some c ⟷ x ≤ c"
"x >⇩l⇩b Some c ⟷ x > c" "x <⇩l⇩b Some c ⟷ x < c"
by (auto simp add: bound_compare_defs)
fun in_bounds where
"in_bounds x v (lb, ub) = (v x ≥⇩l⇩b lb x ∧ v x ≤⇩u⇩b ub x)"
fun satisfies_bounds :: "'a::linorder valuation ⇒ 'a bounds × 'a bounds ⇒ bool" (infixl "⊨⇩b" 100) where
"v ⊨⇩b b ⟷ (∀ x. in_bounds x v b)"
declare satisfies_bounds.simps [simp del]
lemma satisfies_bounds_iff:
"v ⊨⇩b (lb, ub) ⟷ (∀ x. v x ≥⇩l⇩b lb x ∧ v x ≤⇩u⇩b ub x)"
by (auto simp add: satisfies_bounds.simps)
lemma not_in_bounds:
"¬ (in_bounds x v (lb, ub)) = (v x <⇩l⇩b lb x ∨ v x >⇩u⇩b ub x)"
using bounds_compare_contradictory(7)
using bounds_compare_contradictory(2)
using neg_bounds_compare(7)[of "v x" "lb x"]
using neg_bounds_compare(2)[of "v x" "ub x"]
by auto
fun atoms_equiv_bounds :: "'a::linorder atom set ⇒ 'a bounds × 'a bounds ⇒ bool" (infixl "≐" 100) where
"as ≐ (lb, ub) ⟷ (∀ v. v ⊨⇩a⇩s as ⟷ v ⊨⇩b (lb, ub))"
declare atoms_equiv_bounds.simps [simp del]
lemma atoms_equiv_bounds_simps:
"as ≐ (lb, ub) ≡ ∀ v. v ⊨⇩a⇩s as ⟷ v ⊨⇩b (lb, ub)"
by (simp add: atoms_equiv_bounds.simps)
text‹A valuation satisfies bounds iff the value of each variable
respects both its lower and upper bound, i.e, @{thm
satisfies_bounds_iff[no_vars]}. Asserted atoms are precisely encoded
by the current bounds in a state (denoted by ‹≐›) if every
valuation satisfies them iff it satisfies the bounds, i.e.,
@{thm atoms_equiv_bounds_simps[no_vars, iff]}.›
text‹The procedure also keeps track of a valuation that is a
candidate solution. Whenever a new atom is asserted, it is checked
whether the valuation is still satisfying. If not, the procedure tries
to fix that by changing it and changing the tableau if necessary (but
so that it remains equivalent to the initial tableau).›
text‹Therefore, the state of the procedure stores the tableau
(denoted by ‹𝒯›), lower and upper bounds (denoted by ‹ℬ⇩l› and ‹ℬ⇩u›, and ordered pair of lower and upper bounds
denoted by ‹ℬ›), candidate solution (denoted by ‹𝒱›)
and a flag (denoted by ‹𝒰›) indicating if unsatisfiability has
been detected so far:›
text‹Since we also need to now about the indices of atoms, actually,
the bounds are also indexed, and in addition to the flag for unsatisfiability,
we also store an optional unsat core.›
type_synonym 'i bound_index = "var ⇒ 'i"
type_synonym ('i,'a) bounds_index = "(var, ('i × 'a))mapping"
datatype ('i,'a) state = State
(𝒯: "tableau")
(ℬ⇩i⇩l: "('i,'a) bounds_index")
(ℬ⇩i⇩u: "('i,'a) bounds_index")
(𝒱: "(var, 'a) mapping")
(𝒰: bool)
(𝒰⇩c: "'i list option")
definition indexl :: "('i,'a) state ⇒ 'i bound_index" ("ℐ⇩l") where
"ℐ⇩l s = (fst o the) o look (ℬ⇩i⇩l s)"
definition boundsl :: "('i,'a) state ⇒ 'a bounds" ("ℬ⇩l") where
"ℬ⇩l s = map_option snd o look (ℬ⇩i⇩l s)"
definition indexu :: "('i,'a) state ⇒ 'i bound_index" ("ℐ⇩u") where
"ℐ⇩u s = (fst o the) o look (ℬ⇩i⇩u s)"
definition boundsu :: "('i,'a) state ⇒ 'a bounds" ("ℬ⇩u") where
"ℬ⇩u s = map_option snd o look (ℬ⇩i⇩u s)"
abbreviation BoundsIndicesMap ("ℬ⇩i") where "ℬ⇩i s ≡ (ℬ⇩i⇩l s, ℬ⇩i⇩u s)"
abbreviation Bounds :: "('i,'a) state ⇒ 'a bounds × 'a bounds" ("ℬ") where "ℬ s ≡ (ℬ⇩l s, ℬ⇩u s)"
abbreviation Indices :: "('i,'a) state ⇒ 'i bound_index × 'i bound_index" ("ℐ") where "ℐ s ≡ (ℐ⇩l s, ℐ⇩u s)"
abbreviation BoundsIndices :: "('i,'a) state ⇒ ('a bounds × 'a bounds) × ('i bound_index × 'i bound_index)" ("ℬℐ")
where "ℬℐ s ≡ (ℬ s, ℐ s)"
fun satisfies_bounds_index :: "'i set × 'a::lrv valuation ⇒ ('a bounds × 'a bounds) ×
('i bound_index × 'i bound_index) ⇒ bool" (infixl "⊨⇩i⇩b" 100) where
"(I,v) ⊨⇩i⇩b ((BL,BU),(IL,IU)) ⟷ (
(∀ x c. BL x = Some c ⟶ IL x ∈ I ⟶ v x ≥ c)
∧ (∀ x c. BU x = Some c ⟶ IU x ∈ I ⟶ v x ≤ c))"
declare satisfies_bounds_index.simps[simp del]
fun satisfies_bounds_index' :: "'i set × 'a::lrv valuation ⇒ ('a bounds × 'a bounds) ×
('i bound_index × 'i bound_index) ⇒ bool" (infixl "⊨⇩i⇩b⇩e" 100) where
"(I,v) ⊨⇩i⇩b⇩e ((BL,BU),(IL,IU)) ⟷ (
(∀ x c. BL x = Some c ⟶ IL x ∈ I ⟶ v x = c)
∧ (∀ x c. BU x = Some c ⟶ IU x ∈ I ⟶ v x = c))"
declare satisfies_bounds_index'.simps[simp del]
fun atoms_imply_bounds_index :: "('i,'a::lrv) i_atom set ⇒ ('a bounds × 'a bounds) × ('i bound_index × 'i bound_index)
⇒ bool" (infixl "⊨⇩i" 100) where
"as ⊨⇩i bi ⟷ (∀ I v. (I,v) ⊨⇩i⇩a⇩s as ⟶ (I,v) ⊨⇩i⇩b bi)"
declare atoms_imply_bounds_index.simps[simp del]
lemma i_satisfies_atom_set_mono: "as ⊆ as' ⟹ v ⊨⇩i⇩a⇩s as' ⟹ v ⊨⇩i⇩a⇩s as"
by (cases v, auto simp: satisfies_atom_set_def)
lemma atoms_imply_bounds_index_mono: "as ⊆ as' ⟹ as ⊨⇩i bi ⟹ as' ⊨⇩i bi"
unfolding atoms_imply_bounds_index.simps using i_satisfies_atom_set_mono by blast
definition satisfies_state :: "'a::lrv valuation ⇒ ('i,'a) state ⇒ bool" (infixl "⊨⇩s" 100) where
"v ⊨⇩s s ≡ v ⊨⇩b ℬ s ∧ v ⊨⇩t 𝒯 s"
definition curr_val_satisfies_state :: "('i,'a::lrv) state ⇒ bool" ("⊨") where
"⊨ s ≡ ⟨𝒱 s⟩ ⊨⇩s s"
fun satisfies_state_index :: "'i set × 'a::lrv valuation ⇒ ('i,'a) state ⇒ bool" (infixl "⊨⇩i⇩s" 100) where
"(I,v) ⊨⇩i⇩s s ⟷ (v ⊨⇩t 𝒯 s ∧ (I,v) ⊨⇩i⇩b ℬℐ s)"
declare satisfies_state_index.simps[simp del]
fun satisfies_state_index' :: "'i set × 'a::lrv valuation ⇒ ('i,'a) state ⇒ bool" (infixl "⊨⇩i⇩s⇩e" 100) where
"(I,v) ⊨⇩i⇩s⇩e s ⟷ (v ⊨⇩t 𝒯 s ∧ (I,v) ⊨⇩i⇩b⇩e ℬℐ s)"
declare satisfies_state_index'.simps[simp del]
definition indices_state :: "('i,'a)state ⇒ 'i set" where
"indices_state s = { i. ∃ x b. look (ℬ⇩i⇩l s) x = Some (i,b) ∨ look (ℬ⇩i⇩u s) x = Some (i,b)}"
text ‹distinctness requires that for each index $i$, there is at most one variable $x$ and bound
$b$ such that $x \leq b$ or $x \geq b$ or both are enforced.›
definition distinct_indices_state :: "('i,'a)state ⇒ bool" where
"distinct_indices_state s = (∀ i x b x' b'.
((look (ℬ⇩i⇩l s) x = Some (i,b) ∨ look (ℬ⇩i⇩u s) x = Some (i,b)) ⟶
(look (ℬ⇩i⇩l s) x' = Some (i,b') ∨ look (ℬ⇩i⇩u s) x' = Some (i,b')) ⟶
(x = x' ∧ b = b')))"
lemma distinct_indices_stateD: assumes "distinct_indices_state s"
shows "look (ℬ⇩i⇩l s) x = Some (i,b) ∨ look (ℬ⇩i⇩u s) x = Some (i,b) ⟹ look (ℬ⇩i⇩l s) x' = Some (i,b') ∨ look (ℬ⇩i⇩u s) x' = Some (i,b')
⟹ x = x' ∧ b = b'"
using assms unfolding distinct_indices_state_def by blast+
definition unsat_state_core :: "('i,'a::lrv) state ⇒ bool" where
"unsat_state_core s = (set (the (𝒰⇩c s)) ⊆ indices_state s ∧ (¬ (∃ v. (set (the (𝒰⇩c s)),v) ⊨⇩i⇩s s)))"
definition subsets_sat_core :: "('i,'a::lrv) state ⇒ bool" where
"subsets_sat_core s = ((∀ I. I ⊂ set (the (𝒰⇩c s)) ⟶ (∃ v. (I,v) ⊨⇩i⇩s⇩e s)))"
definition minimal_unsat_state_core :: "('i,'a::lrv) state ⇒ bool" where
"minimal_unsat_state_core s = (unsat_state_core s ∧ (distinct_indices_state s ⟶ subsets_sat_core s))"
lemma minimal_unsat_core_tabl_atoms_mono: assumes sub: "as ⊆ bs"
and unsat: "minimal_unsat_core_tabl_atoms I t as"
shows "minimal_unsat_core_tabl_atoms I t bs"
unfolding minimal_unsat_core_tabl_atoms_def
proof (intro conjI impI allI)
note min = unsat[unfolded minimal_unsat_core_tabl_atoms_def]
from min have I: "I ⊆ fst ` as" by blast
with sub show "I ⊆ fst ` bs" by blast
from min have "(∄v. v ⊨⇩t t ∧ (I, v) ⊨⇩i⇩a⇩s as)" by blast
with i_satisfies_atom_set_mono[OF sub]
show "(∄v. v ⊨⇩t t ∧ (I, v) ⊨⇩i⇩a⇩s bs)" by blast
fix J
assume J: "J ⊂ I" and dist_bs: "distinct_indices_atoms bs"
hence dist: "distinct_indices_atoms as"
using sub unfolding distinct_indices_atoms_def by blast
from min dist J obtain v where v: "v ⊨⇩t t" "(J, v) ⊨⇩i⇩a⇩e⇩s as" by blast
have "(J, v) ⊨⇩i⇩a⇩e⇩s bs"
unfolding i_satisfies_atom_set'.simps
proof (intro ballI)
fix a
assume "a ∈ snd ` (bs ∩ J × UNIV)"
then obtain i where ia: "(i,a) ∈ bs" and i: "i ∈ J"
by force
with J have "i ∈ I" by auto
with I obtain b where ib: "(i,b) ∈ as" by force
with sub have ib': "(i,b) ∈ bs" by auto
from dist_bs[unfolded distinct_indices_atoms_def, rule_format, OF ia ib']
have id: "atom_var a = atom_var b" "atom_const a = atom_const b" by auto
from v(2)[unfolded i_satisfies_atom_set'.simps] i ib
have "v ⊨⇩a⇩e b" by force
thus "v ⊨⇩a⇩e a" using id unfolding satisfies_atom'_def by auto
qed
with v show "∃v. v ⊨⇩t t ∧ (J, v) ⊨⇩i⇩a⇩e⇩s bs" by blast
qed
lemma state_satisfies_index: assumes "v ⊨⇩s s"
shows "(I,v) ⊨⇩i⇩s s"
unfolding satisfies_state_index.simps satisfies_bounds_index.simps
proof (intro conjI impI allI)
fix x c
from assms[unfolded satisfies_state_def satisfies_bounds.simps, simplified]
have "v ⊨⇩t 𝒯 s" and bnd: "v x ≥⇩l⇩b ℬ⇩l s x" "v x ≤⇩u⇩b ℬ⇩u s x" by auto
show "v ⊨⇩t 𝒯 s" by fact
show "ℬ⇩l s x = Some c ⟹ ℐ⇩l s x ∈ I ⟹ c ≤ v x"
using bnd(1) by auto
show "ℬ⇩u s x = Some c ⟹ ℐ⇩u s x ∈ I ⟹ v x ≤ c"
using bnd(2) by auto
qed
lemma unsat_state_core_unsat: "unsat_state_core s ⟹ ¬ (∃ v. v ⊨⇩s s)"
unfolding unsat_state_core_def using state_satisfies_index by blast
definition tableau_valuated ("∇") where
"∇ s ≡ ∀ x ∈ tvars (𝒯 s). Mapping.lookup (𝒱 s) x ≠ None"
definition index_valid where
"index_valid as (s :: ('i,'a) state) = (∀ x b i.
(look (ℬ⇩i⇩l s) x = Some (i,b) ⟶ ((i, Geq x b) ∈ as))
∧ (look (ℬ⇩i⇩u s) x = Some (i,b) ⟶ ((i, Leq x b) ∈ as)))"
lemma index_valid_indices_state: "index_valid as s ⟹ indices_state s ⊆ fst ` as"
unfolding index_valid_def indices_state_def by force
lemma index_valid_mono: "as ⊆ bs ⟹ index_valid as s ⟹ index_valid bs s"
unfolding index_valid_def by blast
lemma index_valid_distinct_indices: assumes "index_valid as s"
and "distinct_indices_atoms as"
shows "distinct_indices_state s"
unfolding distinct_indices_state_def
proof (intro allI impI, goal_cases)
case (1 i x b y c)
note valid = assms(1)[unfolded index_valid_def, rule_format]
from 1(1) valid[of x i b] have "(i, Geq x b) ∈ as ∨ (i, Leq x b) ∈ as" by auto
then obtain a where a: "(i,a) ∈ as" "atom_var a = x" "atom_const a = b" by auto
from 1(2) valid[of y i c] have y: "(i, Geq y c) ∈ as ∨ (i, Leq y c) ∈ as" by auto
then obtain a' where a': "(i,a') ∈ as" "atom_var a' = y" "atom_const a' = c" by auto
from assms(2)[unfolded distinct_indices_atoms_def, rule_format, OF a(1) a'(1)]
show ?case using a a' by auto
qed
text‹To be a solution of the initial problem, a valuation should
satisfy the initial tableau and list of atoms. Since tableau is
changed only by equivalency preserving transformations and asserted
atoms are encoded in the bounds, a valuation is a solution if it
satisfies both the tableau and the bounds in the final state (when all
atoms have been asserted). So, a valuation ‹v› satisfies a state
‹s› (denoted by ‹⊨⇩s›) if it satisfies the tableau and
the bounds, i.e., @{thm satisfies_state_def [no_vars]}. Since ‹𝒱› should be a candidate solution, it should satisfy the state
(unless the ‹𝒰› flag is raised). This is denoted by ‹⊨ s›
and defined by @{thm curr_val_satisfies_state_def[no_vars]}. ‹∇
s› will denote that all variables of ‹𝒯 s› are explicitly
valuated in ‹𝒱 s›.›
definition updateℬℐ where
[simp]: "updateℬℐ field_update i x c s = field_update (upd x (i,c)) s"
fun ℬ⇩i⇩u_update where
"ℬ⇩i⇩u_update up (State T BIL BIU V U UC) = State T BIL (up BIU) V U UC"
fun ℬ⇩i⇩l_update where
"ℬ⇩i⇩l_update up (State T BIL BIU V U UC) = State T (up BIL) BIU V U UC"
fun 𝒱_update where
"𝒱_update V (State T BIL BIU V_old U UC) = State T BIL BIU V U UC"
fun 𝒯_update where
"𝒯_update T (State T_old BIL BIU V U UC) = State T BIL BIU V U UC"
lemma update_simps[simp]:
"ℬ⇩i⇩u (ℬ⇩i⇩u_update up s) = up (ℬ⇩i⇩u s)"
"ℬ⇩i⇩l (ℬ⇩i⇩u_update up s) = ℬ⇩i⇩l s"
"𝒯 (ℬ⇩i⇩u_update up s) = 𝒯 s"
"𝒱 (ℬ⇩i⇩u_update up s) = 𝒱 s"
"𝒰 (ℬ⇩i⇩u_update up s) = 𝒰 s"
"𝒰⇩c (ℬ⇩i⇩u_update up s) = 𝒰⇩c s"
"ℬ⇩i⇩l (ℬ⇩i⇩l_update up s) = up (ℬ⇩i⇩l s)"
"ℬ⇩i⇩u (ℬ⇩i⇩l_update up s) = ℬ⇩i⇩u s"
"𝒯 (ℬ⇩i⇩l_update up s) = 𝒯 s"
"𝒱 (ℬ⇩i⇩l_update up s) = 𝒱 s"
"𝒰 (ℬ⇩i⇩l_update up s) = 𝒰 s"
"𝒰⇩c (ℬ⇩i⇩l_update up s) = 𝒰⇩c s"
"𝒱 (𝒱_update V s) = V"
"ℬ⇩i⇩l (𝒱_update V s) = ℬ⇩i⇩l s"
"ℬ⇩i⇩u (𝒱_update V s) = ℬ⇩i⇩u s"
"𝒯 (𝒱_update V s) = 𝒯 s"
"𝒰 (𝒱_update V s) = 𝒰 s"
"𝒰⇩c (𝒱_update V s) = 𝒰⇩c s"
"𝒯 (𝒯_update T s) = T"
"ℬ⇩i⇩l (𝒯_update T s) = ℬ⇩i⇩l s"
"ℬ⇩i⇩u (𝒯_update T s) = ℬ⇩i⇩u s"
"𝒱 (𝒯_update T s) = 𝒱 s"
"𝒰 (𝒯_update T s) = 𝒰 s"
"𝒰⇩c (𝒯_update T s) = 𝒰⇩c s"
by (atomize(full), cases s, auto)
declare
ℬ⇩i⇩u_update.simps[simp del]
ℬ⇩i⇩l_update.simps[simp del]
fun set_unsat :: "'i list ⇒ ('i,'a) state ⇒ ('i,'a) state" where
"set_unsat I (State T BIL BIU V U UC) = State T BIL BIU V True (Some (remdups I))"
lemma set_unsat_simps[simp]: "ℬ⇩i⇩l (set_unsat I s) = ℬ⇩i⇩l s"
"ℬ⇩i⇩u (set_unsat I s) = ℬ⇩i⇩u s"
"𝒯 (set_unsat I s) = 𝒯 s"
"𝒱 (set_unsat I s) = 𝒱 s"
"𝒰 (set_unsat I s) = True"
"𝒰⇩c (set_unsat I s) = Some (remdups I)"
by (atomize(full), cases s, auto)
datatype ('i,'a) Direction = Direction
(lt: "'a::linorder ⇒ 'a ⇒ bool")
(LBI: "('i,'a) state ⇒ ('i,'a) bounds_index")
(UBI: "('i,'a) state ⇒ ('i,'a) bounds_index")
(LB: "('i,'a) state ⇒ 'a bounds")
(UB: "('i,'a) state ⇒ 'a bounds")
(LI: "('i,'a) state ⇒ 'i bound_index")
(UI: "('i,'a) state ⇒ 'i bound_index")
(UBI_upd: "(('i,'a) bounds_index ⇒ ('i,'a) bounds_index) ⇒ ('i,'a) state ⇒ ('i,'a) state")
(LE: "var ⇒ 'a ⇒ 'a atom")
(GE: "var ⇒ 'a ⇒ 'a atom")
(le_rat: "rat ⇒ rat ⇒ bool")
definition Positive where
[simp]: "Positive ≡ Direction (<) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℐ⇩l ℐ⇩u ℬ⇩i⇩u_update Leq Geq (≤)"
definition Negative where
[simp]: "Negative ≡ Direction (>) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l ℐ⇩u ℐ⇩l ℬ⇩i⇩l_update Geq Leq (≥)"
text‹Assuming that the ‹𝒰› flag and the current valuation
‹𝒱› in the final state determine the solution of a problem, the
‹assert_all› function can be reduced to the ‹assert_all_state›
function that operates on the states:
@{text[display] "assert_all t as ≡ let s = assert_all_state t as in
if (𝒰 s) then (False, None) else (True, Some (𝒱 s))" }
›
text‹Specification for the ‹assert_all_state› can be directly
obtained from the specification of ‹assert_all›, and it describes
the connection between the valuation in the final state and the
initial tableau and atoms. However, we will make an additional
refinement step and give stronger assumptions about the ‹assert_all_state› function that describes the connection between
the initial tableau and atoms with the tableau and bounds in the final
state.›
locale AssertAllState = fixes assert_all_state::"tableau ⇒ ('i,'a::lrv) i_atom list ⇒ ('i,'a) state"
assumes
assert_all_state_tableau_equiv: "△ t ⟹ assert_all_state t as = s' ⟹ (v::'a valuation) ⊨⇩t t ⟷ v ⊨⇩t 𝒯 s'" and
assert_all_state_sat: "△ t ⟹ assert_all_state t as = s' ⟹ ¬ 𝒰 s' ⟹ ⊨ s'" and
assert_all_state_sat_atoms_equiv_bounds: "△ t ⟹ assert_all_state t as = s' ⟹ ¬ 𝒰 s' ⟹ flat (set as) ≐ ℬ s'" and
assert_all_state_unsat: "△ t ⟹ assert_all_state t as = s' ⟹ 𝒰 s' ⟹ minimal_unsat_state_core s'" and
assert_all_state_unsat_atoms_equiv_bounds: "△ t ⟹ assert_all_state t as = s' ⟹ 𝒰 s' ⟹ set as ⊨⇩i ℬℐ s'" and
assert_all_state_indices: "△ t ⟹ assert_all_state t as = s ⟹ indices_state s ⊆ fst ` set as" and
assert_all_index_valid: "△ t ⟹ assert_all_state t as = s ⟹ index_valid (set as) s"
begin
definition assert_all where
"assert_all t as ≡ let s = assert_all_state t as in
if (𝒰 s) then Unsat (the (𝒰⇩c s)) else Sat (𝒱 s)"
end
text‹The ‹assert_all_state› function can be implemented by first
applying the ‹init› function that creates an initial state based
on the starting tableau, and then by iteratively applying the ‹assert› function for each atom in the starting atoms list.›
text‹
\begin{small}
‹assert_loop as s ≡ foldl (λ s' a. if (𝒰 s') then s' else assert a s') s as›
‹assert_all_state t as ≡ assert_loop ats (init t)›
\end{small}
›
locale Init' =
fixes init :: "tableau ⇒ ('i,'a::lrv) state"
assumes init'_tableau_normalized: "△ t ⟹ △ (𝒯 (init t))"
assumes init'_tableau_equiv: "△ t ⟹ (v::'a valuation) ⊨⇩t t = v ⊨⇩t 𝒯 (init t)"
assumes init'_sat: "△ t ⟹ ¬ 𝒰 (init t) ⟶ ⊨ (init t)"
assumes init'_unsat: "△ t ⟹ 𝒰 (init t) ⟶ minimal_unsat_state_core (init t)"
assumes init'_atoms_equiv_bounds: "△ t ⟹ {} ≐ ℬ (init t)"
assumes init'_atoms_imply_bounds_index: "△ t ⟹ {} ⊨⇩i ℬℐ (init t)"
text‹Specification for ‹init› can be obtained from the
specification of ‹asser_all_state› since all its assumptions
must also hold for ‹init› (when the list of atoms is
empty). Also, since ‹init› is the first step in the ‹assert_all_state› implementation, the precondition for ‹init›
the same as for the ‹assert_all_state›. However,
unsatisfiability is never going to be detected during initialization
and @{term 𝒰} flag is never going to be raised. Also, the tableau in
the initial state can just be initialized with the starting
tableau. The condition @{term "{} ≐ ℬ (init t)"} is equivalent to
asking that initial bounds are empty. Therefore, specification for
‹init› can be refined to:›
locale Init = fixes init::"tableau ⇒ ('i,'a::lrv) state"
assumes
init_tableau_id: "𝒯 (init t) = t" and
init_unsat_flag: "¬ 𝒰 (init t)" and
init_satisfies_tableau: "⟨𝒱 (init t)⟩ ⊨⇩t t" and
init_bounds: "ℬ⇩i⇩l (init t) = Mapping.empty" "ℬ⇩i⇩u (init t) = Mapping.empty" and
init_tableau_valuated: "∇ (init t)"
begin
lemma init_satisfies_bounds:
"⟨𝒱 (init t)⟩ ⊨⇩b ℬ (init t)"
using init_bounds
unfolding satisfies_bounds.simps in_bounds.simps bound_compare_defs
by (auto simp: boundsl_def boundsu_def)
lemma init_satisfies:
"⊨ (init t)"
using init_satisfies_tableau init_satisfies_bounds init_tableau_id
unfolding curr_val_satisfies_state_def satisfies_state_def
by simp
lemma init_atoms_equiv_bounds:
"{} ≐ ℬ (init t)"
using init_bounds
unfolding atoms_equiv_bounds.simps satisfies_bounds.simps in_bounds.simps satisfies_atom_set_def
unfolding bound_compare_defs
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
lemma init_atoms_imply_bounds_index:
"{} ⊨⇩i ℬℐ (init t)"
using init_bounds
unfolding atoms_imply_bounds_index.simps satisfies_bounds_index.simps in_bounds.simps
i_satisfies_atom_set.simps satisfies_atom_set_def
unfolding bound_compare_defs
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
lemma init_tableau_normalized:
"△ t ⟹ △ (𝒯 (init t))"
using init_tableau_id
by simp
lemma init_index_valid: "index_valid as (init t)"
using init_bounds unfolding index_valid_def by auto
lemma init_indices: "indices_state (init t) = {}"
unfolding indices_state_def init_bounds by auto
end
sublocale Init < Init' init
using init_tableau_id init_satisfies init_unsat_flag init_atoms_equiv_bounds init_atoms_imply_bounds_index
by unfold_locales auto
abbreviation vars_list where
"vars_list t ≡ remdups (map lhs t @ (concat (map (Abstract_Linear_Poly.vars_list ∘ rhs) t)))"
lemma "tvars t = set (vars_list t)"
by (auto simp add: set_vars_list lvars_def rvars_def)
text‹\smallskip The ‹assert› function asserts a single
atom. Since the ‹init› function does not raise the ‹𝒰›
flag, from the definition of ‹assert_loop›, it is clear that the
flag is not raised when the ‹assert› function is
called. Moreover, the assumptions about the ‹assert_all_state›
imply that the loop invariant must be that if the ‹𝒰› flag is
not raised, then the current valuation must satisfy the state (i.e.,
‹⊨ s›). The ‹assert› function will be more easily
implemented if it is always applied to a state with a normalized and
valuated tableau, so we make this another loop invariant. Therefore,
the precondition for the ‹assert a s› function call is that
‹¬ 𝒰 s›, ‹⊨ s›, ‹△ (𝒯 s)› and ‹∇ s›
hold. The specification for ‹assert› directly follows from the
specification of ‹assert_all_state› (except that it is
additionally required that bounds reflect asserted atoms also when
unsatisfiability is detected, and that it is required that ‹assert› keeps the tableau normalized and valuated).›
locale Assert = fixes assert::"('i,'a::lrv) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state"
assumes
assert_tableau: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ let s' = assert a s in
((v::'a valuation) ⊨⇩t 𝒯 s ⟷ v ⊨⇩t 𝒯 s') ∧ △ (𝒯 s') ∧ ∇ s'" and
assert_sat: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ ¬ 𝒰 (assert a s) ⟹ ⊨ (assert a s)" and
assert_atoms_equiv_bounds: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ flat ats ≐ ℬ s ⟹ flat (ats ∪ {a}) ≐ ℬ (assert a s)" and
assert_atoms_imply_bounds_index: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ ats ⊨⇩i ℬℐ s ⟹
insert a ats ⊨⇩i ℬℐ (assert a s)" and
assert_unsat: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s; index_valid ats s⟧ ⟹ 𝒰 (assert a s) ⟹ minimal_unsat_state_core (assert a s)" and
assert_index_valid: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ index_valid ats s ⟹ index_valid (insert a ats) (assert a s)"
begin
lemma assert_tableau_equiv: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ (v::'a valuation) ⊨⇩t 𝒯 s ⟷ v ⊨⇩t 𝒯 (assert a s)"
using assert_tableau
by (simp add: Let_def)
lemma assert_tableau_normalized: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ △ (𝒯 (assert a s))"
using assert_tableau
by (simp add: Let_def)
lemma assert_tableau_valuated: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ ∇ (assert a s)"
using assert_tableau
by (simp add: Let_def)
end
locale AssertAllState' = Init init + Assert assert for
init :: "tableau ⇒ ('i,'a::lrv) state" and assert :: "('i,'a) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state"
begin
definition assert_loop where
"assert_loop as s ≡ foldl (λ s' a. if (𝒰 s') then s' else assert a s') s as"
definition assert_all_state where [simp]:
"assert_all_state t as ≡ assert_loop as (init t)"
lemma AssertAllState'_precond:
"△ t ⟹ △ (𝒯 (assert_all_state t as))
∧ (∇ (assert_all_state t as))
∧ (¬ 𝒰 (assert_all_state t as) ⟶ ⊨ (assert_all_state t as))"
unfolding assert_all_state_def assert_loop_def
using init_satisfies init_tableau_normalized init_index_valid
using assert_sat assert_tableau_normalized init_tableau_valuated assert_tableau_valuated
by (induct as rule: rev_induct) auto
lemma AssertAllState'Induct:
assumes
"△ t"
"P {} (init t)"
"⋀ as bs t. as ⊆ bs ⟹ P as t ⟹ P bs t"
"⋀ s a as. ⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s; P as s; index_valid as s⟧ ⟹ P (insert a as) (assert a s)"
shows "P (set as) (assert_all_state t as)"
proof -
have "P (set as) (assert_all_state t as) ∧ index_valid (set as) (assert_all_state t as)"
proof (induct as rule: rev_induct)
case Nil
then show ?case
unfolding assert_all_state_def assert_loop_def
using assms(2) init_index_valid by auto
next
case (snoc a as')
let ?f = "λs' a. if 𝒰 s' then s' else assert a s'"
let ?s = "foldl ?f (init t) as'"
show ?case
proof (cases "𝒰 ?s")
case True
from snoc index_valid_mono[of _ "set (a # as')" "(assert_all_state t as')"]
have index: "index_valid (set (a # as')) (assert_all_state t as')"
by auto
from snoc assms(3)[of "set as'" "set (a # as')"]
have P: "P (set (a # as')) (assert_all_state t as')" by auto
show ?thesis
using True P index
unfolding assert_all_state_def assert_loop_def
by simp
next
case False
then show ?thesis
using snoc
using assms(1) assms(4)
using AssertAllState'_precond assert_index_valid
unfolding assert_all_state_def assert_loop_def
by auto
qed
qed
then show ?thesis ..
qed
lemma AssertAllState_index_valid: "△ t ⟹ index_valid (set as) (assert_all_state t as)"
by (rule AssertAllState'Induct, auto intro: assert_index_valid init_index_valid index_valid_mono)
lemma AssertAllState'_sat_atoms_equiv_bounds:
"△ t ⟹ ¬ 𝒰 (assert_all_state t as) ⟹ flat (set as) ≐ ℬ (assert_all_state t as)"
using AssertAllState'_precond
using init_atoms_equiv_bounds assert_atoms_equiv_bounds
unfolding assert_all_state_def assert_loop_def
by (induct as rule: rev_induct) auto
lemma AssertAllState'_unsat_atoms_implies_bounds:
assumes "△ t"
shows "set as ⊨⇩i ℬℐ (assert_all_state t as)"
proof (induct as rule: rev_induct)
case Nil
then show ?case
using assms init_atoms_imply_bounds_index
unfolding assert_all_state_def assert_loop_def
by simp
next
case (snoc a as')
let ?s = "assert_all_state t as'"
show ?case
proof (cases "𝒰 ?s")
case True
then show ?thesis
using snoc atoms_imply_bounds_index_mono[of "set as'" "set (as' @ [a])"]
unfolding assert_all_state_def assert_loop_def
by auto
next
case False
then have id: "assert_all_state t (as' @ [a]) = assert a ?s"
unfolding assert_all_state_def assert_loop_def by simp
from snoc have as': "set as' ⊨⇩i ℬℐ ?s" by auto
from AssertAllState'_precond[of t as'] assms False
have "⊨ ?s" "△ (𝒯 ?s)" "∇ ?s" by auto
from assert_atoms_imply_bounds_index[OF False this as', of a]
show ?thesis unfolding id by auto
qed
qed
end
text‹Under these assumptions, it can easily be shown (mainly by
induction) that the previously shown implementation of ‹assert_all_state› satisfies its specification.›
sublocale AssertAllState' < AssertAllState assert_all_state
proof
fix v::"'a valuation" and t as s'
assume *: "△ t" and id: "assert_all_state t as = s'"
note idsym = id[symmetric]
show "v ⊨⇩t t = v ⊨⇩t 𝒯 s'" unfolding idsym
using init_tableau_id[of t] assert_tableau_equiv[of _ v]
by (induct rule: AssertAllState'Induct) (auto simp add: * )
show "¬ 𝒰 s' ⟹ ⊨ s'" unfolding idsym
using AssertAllState'_precond by (simp add: * )
show "¬ 𝒰 s' ⟹ flat (set as) ≐ ℬ s'"
unfolding idsym
using *
by (rule AssertAllState'_sat_atoms_equiv_bounds)
show "𝒰 s' ⟹ set as ⊨⇩i ℬℐ s'"
using * unfolding idsym
by (rule AssertAllState'_unsat_atoms_implies_bounds)
show "𝒰 s' ⟹ minimal_unsat_state_core s'"
using init_unsat_flag assert_unsat assert_index_valid unfolding idsym
by (induct rule: AssertAllState'Induct) (auto simp add: * )
show "indices_state s' ⊆ fst ` set as" unfolding idsym using *
by (intro index_valid_indices_state, induct rule: AssertAllState'Induct,
auto simp: init_index_valid index_valid_mono assert_index_valid)
show "index_valid (set as) s'" using "*" AssertAllState_index_valid idsym by blast
qed
subsection‹Asserting Single Atoms›
text‹The @{term assert} function is split in two phases. First,
@{term assert_bound} updates the bounds and checks only for conflicts
cheap to detect. Next, ‹check› performs the full simplex
algorithm. The ‹assert› function can be implemented as ‹assert a s = check (assert_bound a s)›. Note that it is also
possible to do the first phase for several asserted atoms, and only
then to let the expensive second phase work.
\medskip Asserting an atom ‹x ⨝ b› begins with the function
‹assert_bound›. If the atom is subsumed by the current bounds,
then no changes are performed. Otherwise, bounds for ‹x› are
changed to incorporate the atom. If the atom is inconsistent with the
previous bounds for ‹x›, the @{term 𝒰} flag is raised. If
‹x› is not a lhs variable in the current tableau and if the
value for ‹x› in the current valuation violates the new bound
‹b›, the value for ‹x› can be updated and set to
‹b›, meanwhile updating the values for lhs variables of
the tableau so that it remains satisfied. Otherwise, no changes to the
current valuation are performed.›
fun satisfies_bounds_set :: "'a::linorder valuation ⇒ 'a bounds × 'a bounds ⇒ var set ⇒ bool" where
"satisfies_bounds_set v (lb, ub) S ⟷ (∀ x ∈ S. in_bounds x v (lb, ub))"
declare satisfies_bounds_set.simps [simp del]
syntax
"_satisfies_bounds_set" :: "(var ⇒ 'a::linorder) ⇒ 'a bounds × 'a bounds ⇒ var set ⇒ bool" ("_ ⊨⇩b _ ∥/ _")
translations
"v ⊨⇩b b ∥ S" == "CONST satisfies_bounds_set v b S"
lemma satisfies_bounds_set_iff:
"v ⊨⇩b (lb, ub) ∥ S ≡ (∀ x ∈ S. v x ≥⇩l⇩b lb x ∧ v x ≤⇩u⇩b ub x)"
by (simp add: satisfies_bounds_set.simps)
definition curr_val_satisfies_no_lhs ("⊨⇩n⇩o⇩l⇩h⇩s") where
"⊨⇩n⇩o⇩l⇩h⇩s s ≡ ⟨𝒱 s⟩ ⊨⇩t (𝒯 s) ∧ (⟨𝒱 s⟩ ⊨⇩b (ℬ s) ∥ (- lvars (𝒯 s)))"
lemma satisfies_satisfies_no_lhs:
"⊨ s ⟹ ⊨⇩n⇩o⇩l⇩h⇩s s"
by (simp add: curr_val_satisfies_state_def satisfies_state_def curr_val_satisfies_no_lhs_def satisfies_bounds.simps satisfies_bounds_set.simps)
definition bounds_consistent :: "('i,'a::linorder) state ⇒ bool" ("◇") where
"◇ s ≡
∀ x. if ℬ⇩l s x = None ∨ ℬ⇩u s x = None then True else the (ℬ⇩l s x) ≤ the (ℬ⇩u s x)"
text‹So, the ‹assert_bound› function must ensure that the
given atom is included in the bounds, that the tableau remains
satisfied by the valuation and that all variables except the lhs variables
in the tableau are within their
bounds. To formalize this, we introduce the notation ‹v
⊨⇩b (lb, ub) ∥ S›, and define @{thm
satisfies_bounds_set_iff[no_vars]}, and @{thm
curr_val_satisfies_no_lhs_def[no_vars]}. The ‹assert_bound›
function raises the ‹𝒰› flag if and only if lower and upper
bounds overlap. This is formalized as @{thm
bounds_consistent_def[no_vars]}.›
lemma satisfies_bounds_consistent:
"(v::'a::linorder valuation) ⊨⇩b ℬ s ⟶ ◇ s"
unfolding satisfies_bounds.simps in_bounds.simps bounds_consistent_def bound_compare_defs
by (auto split: option.split) force
lemma satisfies_consistent:
"⊨ s ⟶ ◇ s"
by (auto simp add: curr_val_satisfies_state_def satisfies_state_def satisfies_bounds_consistent)
lemma bounds_consistent_geq_lb:
"⟦◇ s; ℬ⇩u s x⇩i = Some c⟧
⟹ c ≥⇩l⇩b ℬ⇩l s x⇩i"
unfolding bounds_consistent_def
by (cases "ℬ⇩l s x⇩i", auto simp add: bound_compare_defs split: if_splits)
(erule_tac x="x⇩i" in allE, auto)
lemma bounds_consistent_leq_ub:
"⟦◇ s; ℬ⇩l s x⇩i = Some c⟧
⟹ c ≤⇩u⇩b ℬ⇩u s x⇩i"
unfolding bounds_consistent_def
by (cases "ℬ⇩u s x⇩i", auto simp add: bound_compare_defs split: if_splits)
(erule_tac x="x⇩i" in allE, auto)
lemma bounds_consistent_gt_ub:
"⟦◇ s; c <⇩l⇩b ℬ⇩l s x ⟧ ⟹ ¬ c >⇩u⇩b ℬ⇩u s x"
unfolding bounds_consistent_def
by (case_tac[!] "ℬ⇩l s x", case_tac[!] "ℬ⇩u s x")
(auto simp add: bound_compare_defs, erule_tac x="x" in allE, simp)
lemma bounds_consistent_lt_lb:
"⟦◇ s; c >⇩u⇩b ℬ⇩u s x ⟧ ⟹ ¬ c <⇩l⇩b ℬ⇩l s x"
unfolding bounds_consistent_def
by (case_tac[!] "ℬ⇩l s x", case_tac[!] "ℬ⇩u s x")
(auto simp add: bound_compare_defs, erule_tac x="x" in allE, simp)
text‹Since the ‹assert_bound› is the first step in the ‹assert› function implementation, the preconditions for ‹assert_bound› are the same as preconditions for the ‹assert›
function. The specifiction for the ‹assert_bound› is:›
locale AssertBound = fixes assert_bound::"('i,'a::lrv) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state"
assumes
assert_bound_tableau: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ assert_bound a s = s' ⟹ 𝒯 s' = 𝒯 s ∧ ∇ s'" and
assert_bound_sat: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ assert_bound a s = s' ⟹ ¬ 𝒰 s' ⟹ ⊨⇩n⇩o⇩l⇩h⇩s s' ∧ ◇ s'" and
assert_bound_atoms_equiv_bounds: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹
flat ats ≐ ℬ s ⟹ flat (ats ∪ {a}) ≐ ℬ (assert_bound a s)" and
assert_bound_atoms_imply_bounds_index: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹
ats ⊨⇩i ℬℐ s ⟹ insert a ats ⊨⇩i ℬℐ (assert_bound a s)" and
assert_bound_unsat: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ index_valid as s ⟹ assert_bound a s = s' ⟹ 𝒰 s' ⟹ minimal_unsat_state_core s'" and
assert_bound_index_valid: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ index_valid as s ⟹ index_valid (insert a as) (assert_bound a s)"
begin
lemma assert_bound_tableau_id: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ 𝒯 (assert_bound a s) = 𝒯 s"
using assert_bound_tableau by blast
lemma assert_bound_tableau_valuated: "⟦¬ 𝒰 s; ⊨ s; △ (𝒯 s); ∇ s⟧ ⟹ ∇ (assert_bound a s)"
using assert_bound_tableau by blast
end
locale AssertBoundNoLhs =
fixes assert_bound :: "('i,'a::lrv) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state"
assumes assert_bound_nolhs_tableau_id: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s⟧ ⟹ 𝒯 (assert_bound a s) = 𝒯 s"
assumes assert_bound_nolhs_sat: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s⟧ ⟹
¬ 𝒰 (assert_bound a s) ⟹ ⊨⇩n⇩o⇩l⇩h⇩s (assert_bound a s) ∧ ◇ (assert_bound a s)"
assumes assert_bound_nolhs_atoms_equiv_bounds: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s⟧ ⟹
flat ats ≐ ℬ s ⟹ flat (ats ∪ {a}) ≐ ℬ (assert_bound a s)"
assumes assert_bound_nolhs_atoms_imply_bounds_index: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s⟧ ⟹
ats ⊨⇩i ℬℐ s ⟹ insert a ats ⊨⇩i ℬℐ (assert_bound a s)"
assumes assert_bound_nolhs_unsat: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s⟧ ⟹
index_valid as s ⟹ 𝒰 (assert_bound a s) ⟹ minimal_unsat_state_core (assert_bound a s)"
assumes assert_bound_nolhs_tableau_valuated: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s⟧ ⟹
∇ (assert_bound a s)"
assumes assert_bound_nolhs_index_valid: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s⟧ ⟹
index_valid as s ⟹ index_valid (insert a as) (assert_bound a s)"
sublocale AssertBoundNoLhs < AssertBound
by unfold_locales
((metis satisfies_satisfies_no_lhs satisfies_consistent
assert_bound_nolhs_tableau_id assert_bound_nolhs_sat assert_bound_nolhs_atoms_equiv_bounds
assert_bound_nolhs_index_valid assert_bound_nolhs_atoms_imply_bounds_index
assert_bound_nolhs_unsat assert_bound_nolhs_tableau_valuated)+)
text‹The second phase of ‹assert›, the ‹check› function,
is the heart of the Simplex algorithm. It is always called after
‹assert_bound›, but in two different situations. In the first
case ‹assert_bound› raised the ‹𝒰› flag and then
‹check› should retain the flag and should not perform any changes.
In the second case ‹assert_bound› did not raise the
‹𝒰› flag, so ‹⊨⇩n⇩o⇩l⇩h⇩s s›, ‹◇ s›, ‹△ (𝒯
s)›, and ‹∇ s› hold.›
locale Check = fixes check::"('i,'a::lrv) state ⇒ ('i,'a) state"
assumes
check_unsat_id: "𝒰 s ⟹ check s = s" and
check_tableau: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; △ (𝒯 s); ∇ s⟧ ⟹
let s' = check s in ((v::'a valuation) ⊨⇩t 𝒯 s ⟷ v ⊨⇩t 𝒯 s') ∧ △ (𝒯 s') ∧ ∇ s'" and
check_bounds_id: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; △ (𝒯 s); ∇ s⟧ ⟹ ℬ⇩i (check s) = ℬ⇩i s" and
check_sat: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; △ (𝒯 s); ∇ s⟧ ⟹ ¬ 𝒰 (check s) ⟹ ⊨ (check s)" and
check_unsat: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; △ (𝒯 s); ∇ s⟧ ⟹ 𝒰 (check s) ⟹ minimal_unsat_state_core (check s)"
begin
lemma check_tableau_equiv: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; △ (𝒯 s); ∇ s⟧ ⟹
(v::'a valuation) ⊨⇩t 𝒯 s ⟷ v ⊨⇩t 𝒯 (check s)"
using check_tableau
by (simp add: Let_def)
lemma check_tableau_index_valid: assumes "¬ 𝒰 s" " ⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s" "△ (𝒯 s)" "∇ s"
shows "index_valid as (check s) = index_valid as s"
unfolding index_valid_def using check_bounds_id[OF assms]
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
lemma check_tableau_normalized: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; △ (𝒯 s); ∇ s⟧ ⟹ △ (𝒯 (check s))"
using check_tableau
by (simp add: Let_def)
lemma check_tableau_valuated: "⟦¬ 𝒰 s; ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; △ (𝒯 s); ∇ s⟧ ⟹ ∇ (check s)"
using check_tableau
by (simp add: Let_def)
lemma check_indices_state: assumes "¬ 𝒰 s ⟹ ⊨⇩n⇩o⇩l⇩h⇩s s" "¬ 𝒰 s ⟹ ◇ s" "¬ 𝒰 s ⟹ △ (𝒯 s)" "¬ 𝒰 s ⟹ ∇ s"
shows "indices_state (check s) = indices_state s"
using check_bounds_id[OF _ assms] check_unsat_id[of s]
unfolding indices_state_def by (cases "𝒰 s", auto)
lemma check_distinct_indices_state: assumes "¬ 𝒰 s ⟹ ⊨⇩n⇩o⇩l⇩h⇩s s" "¬ 𝒰 s ⟹ ◇ s" "¬ 𝒰 s ⟹ △ (𝒯 s)" "¬ 𝒰 s ⟹ ∇ s"
shows "distinct_indices_state (check s) = distinct_indices_state s"
using check_bounds_id[OF _ assms] check_unsat_id[of s]
unfolding distinct_indices_state_def by (cases "𝒰 s", auto)
end
locale Assert' = AssertBound assert_bound + Check check for
assert_bound :: "('i,'a::lrv) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state" and
check :: "('i,'a::lrv) state ⇒ ('i,'a) state"
begin
definition assert :: "('i,'a) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state" where
"assert a s ≡ check (assert_bound a s)"
lemma Assert'Precond:
assumes "¬ 𝒰 s" "⊨ s" "△ (𝒯 s)" "∇ s"
shows
"△ (𝒯 (assert_bound a s))"
"¬ 𝒰 (assert_bound a s) ⟹ ⊨⇩n⇩o⇩l⇩h⇩s (assert_bound a s) ∧ ◇ (assert_bound a s)"
"∇ (assert_bound a s)"
using assms
using assert_bound_tableau_id assert_bound_sat assert_bound_tableau_valuated
by (auto simp add: satisfies_bounds_consistent Let_def)
end
sublocale Assert' < Assert assert
proof
fix s::"('i,'a) state" and v::"'a valuation" and a::"('i,'a) i_atom"
assume *: "¬ 𝒰 s" "⊨ s" "△ (𝒯 s)" "∇ s"
have "△ (𝒯 (assert a s))"
using check_tableau_normalized[of "assert_bound a s"] check_unsat_id[of "assert_bound a s"] *
using assert_bound_sat[of s a] Assert'Precond[of s a]
by (force simp add: assert_def)
moreover
have "v ⊨⇩t 𝒯 s = v ⊨⇩t 𝒯 (assert a s)"
using check_tableau_equiv[of "assert_bound a s" v] *
using check_unsat_id[of "assert_bound a s"]
by (force simp add: assert_def Assert'Precond assert_bound_sat assert_bound_tableau_id)
moreover
have "∇ (assert a s)"
using assert_bound_tableau_valuated[of s a] *
using check_tableau_valuated[of "assert_bound a s"]
using check_unsat_id[of "assert_bound a s"]
by (cases "𝒰 (assert_bound a s)") (auto simp add: Assert'Precond assert_def)
ultimately
show "let s' = assert a s in (v ⊨⇩t 𝒯 s = v ⊨⇩t 𝒯 s') ∧ △ (𝒯 s') ∧ ∇ s'"
by (simp add: Let_def)
next
fix s::"('i,'a) state" and a::"('i,'a) i_atom"
assume "¬ 𝒰 s" "⊨ s" "△ (𝒯 s)" "∇ s"
then show "¬ 𝒰 (assert a s) ⟹ ⊨ (assert a s)"
unfolding assert_def
using check_unsat_id[of "assert_bound a s"]
using check_sat[of "assert_bound a s"]
by (force simp add: Assert'Precond)
next
fix s::"('i,'a) state" and a::"('i,'a) i_atom" and ats::"('i,'a) i_atom set"
assume "¬ 𝒰 s" "⊨ s" "△ (𝒯 s)" "∇ s"
then show "flat ats ≐ ℬ s ⟹ flat (ats ∪ {a}) ≐ ℬ (assert a s)"
using assert_bound_atoms_equiv_bounds
using check_unsat_id[of "assert_bound a s"] check_bounds_id
by (cases "𝒰 (assert_bound a s)") (auto simp add: Assert'Precond assert_def
simp: indexl_def indexu_def boundsl_def boundsu_def)
next
fix s::"('i,'a) state" and a::"('i,'a) i_atom" and ats
assume *: "¬ 𝒰 s" "⊨ s" "△ (𝒯 s)" "∇ s" "𝒰 (assert a s)" "index_valid ats s"
show "minimal_unsat_state_core (assert a s)"
proof (cases "𝒰 (assert_bound a s)")
case True
then show ?thesis
unfolding assert_def
using * assert_bound_unsat check_tableau_equiv[of "assert_bound a s"] check_bounds_id
using check_unsat_id[of "assert_bound a s"]
by (auto simp add: Assert'Precond satisfies_state_def Let_def)
next
case False
then show ?thesis
unfolding assert_def
using * assert_bound_sat[of s a] check_unsat Assert'Precond
by (metis assert_def)
qed
next
fix ats
fix s::"('i,'a) state" and a::"('i,'a) i_atom"
assume *: "index_valid ats s"
and **: "¬ 𝒰 s" "⊨ s" "△ (𝒯 s)" "∇ s"
have *: "index_valid (insert a ats) (assert_bound a s)"
using assert_bound_index_valid[OF ** *] .
show "index_valid (insert a ats) (assert a s)"
proof (cases "𝒰 (assert_bound a s)")
case True
show ?thesis unfolding assert_def check_unsat_id[OF True] using * .
next
case False
show ?thesis unfolding assert_def using Assert'Precond[OF **, of a] False *
by (subst check_tableau_index_valid[OF False], auto)
qed
next
fix s ats a
let ?s = "assert_bound a s"
assume *: "¬ 𝒰 s" "⊨ s" "△ (𝒯 s)" "∇ s" "ats ⊨⇩i ℬℐ s"
from assert_bound_atoms_imply_bounds_index[OF this, of a]
have as: "insert a ats ⊨⇩i ℬℐ (assert_bound a s)" by auto
show "insert a ats ⊨⇩i ℬℐ (assert a s)"
proof (cases "𝒰 ?s")
case True
from check_unsat_id[OF True] as show ?thesis unfolding assert_def by auto
next
case False
from assert_bound_tableau_id[OF *(1-4)] *
have t: "△ (𝒯 ?s)" by simp
from assert_bound_tableau_valuated[OF *(1-4)]
have v: "∇ ?s" .
from assert_bound_sat[OF *(1-4) refl False]
have "⊨⇩n⇩o⇩l⇩h⇩s ?s" "◇ ?s" by auto
from check_bounds_id[OF False this t v] as
show ?thesis unfolding assert_def
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
qed
qed
text‹Under these assumptions for ‹assert_bound› and ‹check›, it can be easily shown that the implementation of ‹assert› (previously given) satisfies its specification.›
locale AssertAllState'' = Init init + AssertBoundNoLhs assert_bound + Check check for
init :: "tableau ⇒ ('i,'a::lrv) state" and
assert_bound :: "('i,'a::lrv) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state" and
check :: "('i,'a::lrv) state ⇒ ('i,'a) state"
begin
definition assert_bound_loop where
"assert_bound_loop ats s ≡ foldl (λ s' a. if (𝒰 s') then s' else assert_bound a s') s ats"
definition assert_all_state where [simp]:
"assert_all_state t ats ≡ check (assert_bound_loop ats (init t))"
text‹However, for efficiency reasons, we want to allow
implementations that delay the ‹check› function call and call it
after several ‹assert_bound› calls. For example:
\smallskip
\begin{small}
@{thm assert_bound_loop_def[no_vars]}
@{thm assert_all_state_def[no_vars]}
\end{small}
\smallskip
Then, the loop consists only of ‹assert_bound› calls, so ‹assert_bound› postcondition must imply its precondition. This is not
the case, since variables on the lhs may be out of their
bounds. Therefore, we make a refinement and specify weaker
preconditions (replace ‹⊨ s›, by ‹⊨⇩n⇩o⇩l⇩h⇩s s› and ‹◇ s›) for ‹assert_bound›, and show that these
preconditions are still good enough to prove the correctnes of this
alternative ‹assert_all_state› definition.›
lemma AssertAllState''_precond':
assumes "△ (𝒯 s)" "∇ s" "¬ 𝒰 s ⟶ ⊨⇩n⇩o⇩l⇩h⇩s s ∧ ◇ s"
shows "let s' = assert_bound_loop ats s in
△ (𝒯 s') ∧ ∇ s' ∧ (¬ 𝒰 s' ⟶ ⊨⇩n⇩o⇩l⇩h⇩s s' ∧ ◇ s')"
using assms
using assert_bound_nolhs_tableau_id assert_bound_nolhs_sat assert_bound_nolhs_tableau_valuated
by (induct ats rule: rev_induct) (auto simp add: assert_bound_loop_def Let_def)
lemma AssertAllState''_precond:
assumes "△ t"
shows "let s' = assert_bound_loop ats (init t) in
△ (𝒯 s') ∧ ∇ s' ∧ (¬ 𝒰 s' ⟶ ⊨⇩n⇩o⇩l⇩h⇩s s' ∧ ◇ s')"
using assms
using AssertAllState''_precond'[of "init t" ats]
by (simp add: Let_def init_tableau_id init_unsat_flag init_satisfies satisfies_consistent
satisfies_satisfies_no_lhs init_tableau_valuated)
lemma AssertAllState''Induct:
assumes
"△ t"
"P {} (init t)"
"⋀ as bs t. as ⊆ bs ⟹ P as t ⟹ P bs t"
"⋀ s a ats. ⟦¬ 𝒰 s; ⟨𝒱 s⟩ ⊨⇩t 𝒯 s; ⊨⇩n⇩o⇩l⇩h⇩s s; △ (𝒯 s); ∇ s; ◇ s; P (set ats) s; index_valid (set ats) s⟧
⟹ P (insert a (set ats)) (assert_bound a s)"
shows "P (set ats) (assert_bound_loop ats (init t))"
proof -
have "P (set ats) (assert_bound_loop ats (init t)) ∧ index_valid (set ats) (assert_bound_loop ats (init t))"
proof (induct ats rule: rev_induct)
case Nil
then show ?case
unfolding assert_bound_loop_def
using assms(2) init_index_valid
by simp
next
case (snoc a as')
let ?s = "assert_bound_loop as' (init t)"
from snoc index_valid_mono[of _ "set (a # as')" "assert_bound_loop as' (init t)"]
have index: "index_valid (set (a # as')) (assert_bound_loop as' (init t))"
by auto
from snoc assms(3)[of "set as'" "set (a # as')"]
have P: "P (set (a # as')) (assert_bound_loop as' (init t))" by auto
show ?case
proof (cases "𝒰 ?s")
case True
then show ?thesis
using P index
unfolding assert_bound_loop_def
by simp
next
case False
have id': "set (as' @ [a]) = insert a (set as')" by simp
have id: "assert_bound_loop (as' @ [a]) (init t) =
assert_bound a (assert_bound_loop as' (init t))"
using False unfolding assert_bound_loop_def by auto
from snoc have index: "index_valid (set as') ?s" by simp
show ?thesis unfolding id unfolding id' using False snoc AssertAllState''_precond[OF assms(1)]
by (intro conjI assert_bound_nolhs_index_valid index assms(4); (force simp: Let_def curr_val_satisfies_no_lhs_def)?)
qed
qed
then show ?thesis ..
qed
lemma AssertAllState''_tableau_id:
"△ t ⟹ 𝒯 (assert_bound_loop ats (init t)) = 𝒯 (init t)"
by (rule AssertAllState''Induct) (auto simp add: init_tableau_id assert_bound_nolhs_tableau_id)
lemma AssertAllState''_sat:
"△ t ⟹
¬ 𝒰 (assert_bound_loop ats (init t)) ⟶ ⊨⇩n⇩o⇩l⇩h⇩s (assert_bound_loop ats (init t)) ∧ ◇ (assert_bound_loop ats (init t))"
by (rule AssertAllState''Induct) (auto simp add: init_unsat_flag init_satisfies satisfies_consistent satisfies_satisfies_no_lhs assert_bound_nolhs_sat)
lemma AssertAllState''_unsat:
"△ t ⟹ 𝒰 (assert_bound_loop ats (init t)) ⟶ minimal_unsat_state_core (assert_bound_loop ats (init t))"
by (rule AssertAllState''Induct)
(auto simp add: init_tableau_id assert_bound_nolhs_unsat init_unsat_flag)
lemma AssertAllState''_sat_atoms_equiv_bounds:
"△ t ⟹ ¬ 𝒰 (assert_bound_loop ats (init t)) ⟶ flat (set ats) ≐ ℬ (assert_bound_loop ats (init t))"
using AssertAllState''_precond
using assert_bound_nolhs_atoms_equiv_bounds init_atoms_equiv_bounds
by (induct ats rule: rev_induct) (auto simp add: Let_def assert_bound_loop_def)
lemma AssertAllState''_atoms_imply_bounds_index:
assumes "△ t"
shows "set ats ⊨⇩i ℬℐ (assert_bound_loop ats (init t))"
proof (induct ats rule: rev_induct)
case Nil
then show ?case
unfolding assert_bound_loop_def
using init_atoms_imply_bounds_index assms
by simp
next
case (snoc a ats')
let ?s = "assert_bound_loop ats' (init t)"
show ?case
proof (cases "𝒰 ?s")
case True
then show ?thesis
using snoc atoms_imply_bounds_index_mono[of "set ats'" "set (ats' @ [a])"]
unfolding assert_bound_loop_def
by auto
next
case False
then have id: "assert_bound_loop (ats' @ [a]) (init t) = assert_bound a ?s"
unfolding assert_bound_loop_def by auto
from snoc have ats: "set ats' ⊨⇩i ℬℐ ?s" by auto
from AssertAllState''_precond[of t ats', OF assms, unfolded Let_def] False
have *: "⊨⇩n⇩o⇩l⇩h⇩s ?s" "△ (𝒯 ?s)" "∇ ?s" "◇ ?s" by auto
show ?thesis unfolding id using assert_bound_nolhs_atoms_imply_bounds_index[OF False * ats, of a] by auto
qed
qed
lemma AssertAllState''_index_valid:
"△ t ⟹ index_valid (set ats) (assert_bound_loop ats (init t))"
by (rule AssertAllState''Induct, auto simp: init_index_valid index_valid_mono assert_bound_nolhs_index_valid)
end
sublocale AssertAllState'' < AssertAllState assert_all_state
proof
fix v::"'a valuation" and t ats s'
assume *: "△ t" and "assert_all_state t ats = s'"
then have s': "s' = assert_all_state t ats" by simp
let ?s' = "assert_bound_loop ats (init t)"
show "v ⊨⇩t t = v ⊨⇩t 𝒯 s'"
unfolding assert_all_state_def s'
using * check_tableau_equiv[of ?s' v] AssertAllState''_tableau_id[of t ats] init_tableau_id[of t]
using AssertAllState''_sat[of t ats] check_unsat_id[of ?s']
using AssertAllState''_precond[of t ats]
by force
show "¬ 𝒰 s' ⟹ ⊨ s'"
unfolding assert_all_state_def s'
using * AssertAllState''_precond[of t ats]
using check_sat check_unsat_id
by (force simp add: Let_def)
show "𝒰 s' ⟹ minimal_unsat_state_core s'"
using * check_unsat check_unsat_id[of ?s'] check_bounds_id
using AssertAllState''_unsat[of t ats] AssertAllState''_precond[of t ats] s'
by (force simp add: Let_def satisfies_state_def)
show "¬ 𝒰 s' ⟹ flat (set ats) ≐ ℬ s'"
unfolding assert_all_state_def s'
using * AssertAllState''_precond[of t ats]
using check_bounds_id[of ?s'] check_unsat_id[of ?s']
using AssertAllState''_sat_atoms_equiv_bounds[of t ats]
by (force simp add: Let_def simp: indexl_def indexu_def boundsl_def boundsu_def)
show "𝒰 s' ⟹ set ats ⊨⇩i ℬℐ s'"
unfolding assert_all_state_def s'
using * AssertAllState''_precond[of t ats]
unfolding Let_def
using check_bounds_id[of ?s']
using AssertAllState''_atoms_imply_bounds_index[of t ats]
using check_unsat_id[of ?s']
by (cases "𝒰 ?s'") (auto simp add: Let_def simp: indexl_def indexu_def boundsl_def boundsu_def)
show "index_valid (set ats) s'"
unfolding assert_all_state_def s'
using * AssertAllState''_precond[of t ats] AssertAllState''_index_valid[OF *, of ats]
unfolding Let_def
using check_tableau_index_valid[of ?s']
using check_unsat_id[of ?s']
by (cases "𝒰 ?s'", auto)
show "indices_state s' ⊆ fst ` set ats"
by (intro index_valid_indices_state, fact)
qed
subsection‹Update and Pivot›
text‹Both ‹assert_bound› and ‹check› need to update
the valuation so that the tableau remains satisfied. If the value for
a variable not on the lhs of the tableau is changed, this
can be done rather easily (once the value of that variable is changed,
one should recalculate and change the values for all lhs
variables of the tableau). The ‹update› function does this, and
it is specified by:›
locale Update = fixes update::"var ⇒ 'a::lrv ⇒ ('i,'a) state ⇒ ('i,'a) state"
assumes
update_id: "⟦△ (𝒯 s); ∇ s; x ∉ lvars (𝒯 s)⟧ ⟹
let s' = update x c s in 𝒯 s' = 𝒯 s ∧ ℬ⇩i s' = ℬ⇩i s ∧ 𝒰 s' = 𝒰 s ∧ 𝒰⇩c s' = 𝒰⇩c s" and
update_tableau_valuated: "⟦△ (𝒯 s); ∇ s; x ∉ lvars (𝒯 s)⟧ ⟹ ∇ (update x v s)" and
update_valuation_nonlhs: "⟦△ (𝒯 s); ∇ s; x ∉ lvars (𝒯 s)⟧ ⟹ x' ∉ lvars (𝒯 s) ⟶
look (𝒱 (update x v s)) x' = (if x = x' then Some v else look (𝒱 s) x')" and
update_satisfies_tableau: "⟦△ (𝒯 s); ∇ s; x ∉ lvars (𝒯 s)⟧ ⟹ ⟨𝒱 s⟩ ⊨⇩t 𝒯 s ⟶ ⟨𝒱 (update x c s)⟩ ⊨⇩t 𝒯 s"
begin
lemma update_bounds_id:
assumes "△ (𝒯 s)" "∇ s" "x ∉ lvars (𝒯 s)"
shows "ℬ⇩i (update x c s) = ℬ⇩i s"
"ℬℐ (update x c s) = ℬℐ s"
"ℬ⇩l (update x c s) = ℬ⇩l s"
"ℬ⇩u (update x c s) = ℬ⇩u s"
using update_id assms
by (auto simp add: Let_def simp: indexl_def indexu_def boundsl_def boundsu_def)
lemma update_indices_state_id:
assumes "△ (𝒯 s)" "∇ s" "x ∉ lvars (𝒯 s)"
shows "indices_state (update x c s) = indices_state s"
using update_bounds_id[OF assms] unfolding indices_state_def by auto
lemma update_tableau_id: "⟦△ (𝒯 s); ∇ s; x ∉ lvars (𝒯 s)⟧ ⟹ 𝒯 (update x c s) = 𝒯 s"
using update_id
by (auto simp add: Let_def)
lemma update_unsat_id: "⟦△ (𝒯 s); ∇ s; x ∉ lvars (𝒯 s)⟧ ⟹ 𝒰 (update x c s) = 𝒰 s"
using update_id
by (auto simp add: Let_def)
lemma update_unsat_core_id: "⟦△ (𝒯 s); ∇ s; x ∉ lvars (𝒯 s)⟧ ⟹ 𝒰⇩c (update x c s) = 𝒰⇩c s"
using update_id
by (auto simp add: Let_def)
definition assert_bound' where
[simp]: "assert_bound' dir i x c s ≡
(if (⊵⇩u⇩b (lt dir)) c (UB dir s x) then s
else let s' = updateℬℐ (UBI_upd dir) i x c s in
if (⊲⇩l⇩b (lt dir)) c ((LB dir) s x) then
set_unsat [i, ((LI dir) s x)] s'
else if x ∉ lvars (𝒯 s') ∧ (lt dir) c (⟨𝒱 s⟩ x) then
update x c s'
else
s')"
fun assert_bound :: "('i,'a::lrv) i_atom ⇒ ('i,'a) state ⇒ ('i,'a) state" where
"assert_bound (i,Leq x c) s = assert_bound' Positive i x c s"
| "assert_bound (i,Geq x c) s = assert_bound' Negative i x c s"
lemma assert_bound'_cases:
assumes "⟦⊵⇩u⇩b (lt dir) c ((UB dir) s x)⟧ ⟹ P s"
assumes "⟦¬ (⊵⇩u⇩b (lt dir) c ((UB dir) s x)); ⊲⇩l⇩b (lt dir) c ((LB dir) s x)⟧ ⟹
P (set_unsat [i, ((LI dir) s x)] (updateℬℐ (UBI_upd dir) i x c s))"
assumes "⟦x ∉ lvars (𝒯 s); (lt dir) c (⟨𝒱 s⟩ x); ¬ (⊵⇩u⇩b (lt dir) c ((UB dir) s x)); ¬ (⊲⇩l⇩b (lt dir) c ((LB dir) s x))⟧ ⟹
P (update x c (updateℬℐ (UBI_upd dir) i x c s))"
assumes "⟦¬ (⊵⇩u⇩b (lt dir) c ((UB dir) s x)); ¬ (⊲⇩l⇩b (lt dir) c ((LB dir) s x)); x ∈ lvars (𝒯 s)⟧ ⟹
P (updateℬℐ (UBI_upd dir) i x c s)"
assumes "⟦¬ (⊵⇩u⇩b (lt dir) c ((UB dir) s x)); ¬ (⊲⇩l⇩b (lt dir) c ((LB dir) s x)); ¬ ((lt dir) c (⟨𝒱 s⟩ x))⟧ ⟹
P (updateℬℐ (UBI_upd dir) i x c s)"
assumes "dir = Positive ∨ dir = Negative"
shows "P (assert_bound' dir i x c s)"
proof (cases "⊵⇩u⇩b (lt dir) c ((UB dir) s x)")
case True
then show ?thesis
using assms(1)
by simp
next
case False
show ?thesis
proof (cases "⊲⇩l⇩b (lt dir) c ((LB dir) s x)")
case True
then show ?thesis
using ‹¬ ⊵⇩u⇩b (lt dir) c ((UB dir) s x)›
using assms(2)
by simp
next
case False
let ?s = "updateℬℐ (UBI_upd dir) i x c s"
show ?thesis
proof (cases "x ∉ lvars (𝒯 ?s) ∧ (lt dir) c (⟨𝒱 s⟩ x)")
case True
then show ?thesis
using ‹¬ ⊵⇩u⇩b (lt dir) c ((UB dir) s x)› ‹¬ ⊲⇩l⇩b (lt dir) c ((LB dir) s x)›
using assms(3) assms(6)
by auto
next
case False
then have "x ∈ lvars (𝒯 ?s) ∨ ¬ ((lt dir) c (⟨𝒱 s⟩ x))"
by simp
then show ?thesis
proof
assume "x ∈ lvars (𝒯 ?s)"
then show ?thesis
using ‹¬ ⊵⇩u⇩b (lt dir) c ((UB dir) s x)› ‹¬ ⊲⇩l⇩b (lt dir) c ((LB dir) s x)›
using assms(4) assms(6)
by auto
next
assume "¬ (lt dir) c (⟨𝒱 s⟩ x)"
then show ?thesis
using ‹¬ ⊵⇩u⇩b (lt dir) c ((UB dir) s x)› ‹¬ ⊲⇩l⇩b (lt dir) c ((LB dir) s x)›
using assms(5) assms(6)
by simp
qed
qed
qed
qed
lemma assert_bound_cases:
assumes "⋀ c x dir.
⟦ dir = Positive ∨ dir = Negative;
a = LE dir x c;
⊵⇩u⇩b (lt dir) c ((UB dir) s x)
⟧ ⟹
P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) s"
assumes "⋀ c x dir.
⟦dir = Positive ∨ dir = Negative;
a = LE dir x c;
¬ ⊵⇩u⇩b (lt dir) c ((UB dir) s x); ⊲⇩l⇩b (lt dir) c ((LB dir) s x)
⟧ ⟹
P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)
(set_unsat [i, ((LI dir) s x)] (updateℬℐ (UBI_upd dir) i x c s))"
assumes "⋀ c x dir.
⟦ dir = Positive ∨ dir = Negative;
a = LE dir x c;
x ∉ lvars (𝒯 s); (lt dir) c (⟨𝒱 s⟩ x);
¬ (⊵⇩u⇩b (lt dir) c ((UB dir) s x)); ¬ (⊲⇩l⇩b (lt dir) c ((LB dir) s x))
⟧ ⟹
P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)
(update x c ((updateℬℐ (UBI_upd dir) i x c s)))"
assumes "⋀ c x dir.
⟦ dir = Positive ∨ dir = Negative;
a = LE dir x c;
x ∈ lvars (𝒯 s); ¬ (⊵⇩u⇩b (lt dir) c ((UB dir) s x));
¬ (⊲⇩l⇩b (lt dir) c ((LB dir) s x))
⟧ ⟹
P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)
((updateℬℐ (UBI_upd dir) i x c s))"
assumes "⋀ c x dir.
⟦ dir = Positive ∨ dir = Negative;
a = LE dir x c;
¬ (⊵⇩u⇩b (lt dir) c ((UB dir) s x)); ¬ (⊲⇩l⇩b (lt dir) c ((LB dir) s x));
¬ ((lt dir) c (⟨𝒱 s⟩ x))
⟧ ⟹
P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)
((updateℬℐ (UBI_upd dir) i x c s))"
assumes "⋀ s. P s = P' (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℬ⇩i⇩l_update ℐ⇩l ℐ⇩u Geq Leq s"
assumes "⋀ s. P s = P' (<) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l ℬ⇩i⇩u_update ℐ⇩u ℐ⇩l Leq Geq s"
shows "P (assert_bound (i,a) s)"
proof (cases a)
case (Leq x c)
then show ?thesis
apply (simp del: assert_bound'_def)
apply (rule assert_bound'_cases, simp_all)
using assms(1)[of Positive x c]
using assms(2)[of Positive x c]
using assms(3)[of Positive x c]
using assms(4)[of Positive x c]
using assms(5)[of Positive x c]
using assms(7)
by auto
next
case (Geq x c)
then show ?thesis
apply (simp del: assert_bound'_def)
apply (rule assert_bound'_cases)
using assms(1)[of Negative x c]
using assms(2)[of Negative x c]
using assms(3)[of Negative x c]
using assms(4)[of Negative x c]
using assms(5)[of Negative x c]
using assms(6)
by auto
qed
end
lemma set_unsat_bounds_id: "ℬ (set_unsat I s) = ℬ s"
unfolding boundsl_def boundsu_def by auto
lemma decrease_ub_satisfied_inverse:
assumes lt: "⊲⇩u⇩b (lt dir) c (UB dir s x)" and dir: "dir = Positive ∨ dir = Negative"
assumes v: "v ⊨⇩b ℬ (updateℬℐ (UBI_upd dir) i x c s)"
shows "v ⊨⇩b ℬ s"
unfolding satisfies_bounds.simps
proof
fix x'
show "in_bounds x' v (ℬ s)"
proof (cases "x = x'")
case False
then show ?thesis
using v dir
unfolding satisfies_bounds.simps
by (auto split: if_splits simp: boundsl_def boundsu_def)
next
case True
then show ?thesis
using v dir
unfolding satisfies_bounds.simps
using lt
by (erule_tac x="x'" in allE)
(auto simp add: lt_ubound_def[THEN sym] gt_lbound_def[THEN sym] compare_strict_nonstrict
boundsl_def boundsu_def)
qed
qed
lemma atoms_equiv_bounds_extend:
fixes x c dir
assumes "dir = Positive ∨ dir = Negative" "¬ ⊵⇩u⇩b (lt dir) c (UB dir s x)" "ats ≐ ℬ s"
shows "(ats ∪ {LE dir x c}) ≐ ℬ (updateℬℐ (UBI_upd dir) i x c s)"
unfolding atoms_equiv_bounds.simps
proof
fix v
let ?s = "updateℬℐ (UBI_upd dir) i x c s"
show "v ⊨⇩a⇩s (ats ∪ {LE dir x c}) = v ⊨⇩b ℬ ?s"
proof
assume "v ⊨⇩a⇩s (ats ∪ {LE dir x c})"
then have "v ⊨⇩a⇩s ats" "le (lt dir) (v x) c"
using ‹dir = Positive ∨ dir = Negative›
unfolding satisfies_atom_set_def
by auto
show "v ⊨⇩b ℬ ?s"
unfolding satisfies_bounds.simps
proof
fix x'
show "in_bounds x' v (ℬ ?s)"
using ‹v ⊨⇩a⇩s ats› ‹le (lt dir) (v x) c› ‹ats ≐ ℬ s›
using ‹dir = Positive ∨ dir = Negative›
unfolding atoms_equiv_bounds.simps satisfies_bounds.simps
by (cases "x = x'") (auto simp: boundsl_def boundsu_def)
qed
next
assume "v ⊨⇩b ℬ ?s"
then have "v ⊨⇩b ℬ s"
using ‹¬ ⊵⇩u⇩b (lt dir) c (UB dir s x)›
using ‹dir = Positive ∨ dir = Negative›
using decrease_ub_satisfied_inverse[of dir c s x v]
using neg_bounds_compare(1)[of c "ℬ⇩u s x"]
using neg_bounds_compare(5)[of c "ℬ⇩l s x"]
by (auto simp add: lt_ubound_def[THEN sym] ge_ubound_def[THEN sym] le_lbound_def[THEN sym] gt_lbound_def[THEN sym])
show "v ⊨⇩a⇩s (ats ∪ {LE dir x c})"
unfolding satisfies_atom_set_def
proof
fix a
assume "a ∈ ats ∪ {LE dir x c}"
then show "v ⊨⇩a a"
proof
assume "a ∈ {LE dir x c}"
then show ?thesis
using ‹v ⊨⇩b ℬ ?s›
using ‹dir = Positive ∨ dir = Negative›
unfolding satisfies_bounds.simps
by (auto split: if_splits simp: boundsl_def boundsu_def)
next
assume "a ∈ ats"
then show ?thesis
using ‹ats ≐ ℬ s›
using ‹v ⊨⇩b ℬ s›
unfolding atoms_equiv_bounds.simps satisfies_atom_set_def
by auto
qed
qed
qed
qed
lemma bounds_updates: "ℬ⇩l (ℬ⇩i⇩u_update u s) = ℬ⇩l s"
"ℬ⇩u (ℬ⇩i⇩l_update u s) = ℬ⇩u s"
"ℬ⇩u (ℬ⇩i⇩u_update (upd x (i,c)) s) = (ℬ⇩u s) (x ↦ c)"
"ℬ⇩l (ℬ⇩i⇩l_update (upd x (i,c)) s) = (ℬ⇩l s) (x ↦ c)"
by (auto simp: boundsl_def boundsu_def)
locale EqForLVar =
fixes eq_idx_for_lvar :: "tableau ⇒ var ⇒ nat"
assumes eq_idx_for_lvar:
"⟦x ∈ lvars t⟧ ⟹ eq_idx_for_lvar t x < length t ∧ lhs (t ! eq_idx_for_lvar t x) = x"
begin
definition eq_for_lvar :: "tableau ⇒ var ⇒ eq" where
"eq_for_lvar t v ≡ t ! (eq_idx_for_lvar t v)"
lemma eq_for_lvar:
"⟦x ∈ lvars t⟧ ⟹ eq_for_lvar t x ∈ set t ∧ lhs (eq_for_lvar t x) = x"
unfolding eq_for_lvar_def
using eq_idx_for_lvar
by auto
abbreviation rvars_of_lvar where
"rvars_of_lvar t x ≡ rvars_eq (eq_for_lvar t x)"
lemma rvars_of_lvar_rvars:
assumes "x ∈ lvars t"
shows "rvars_of_lvar t x ⊆ rvars t"
using assms eq_for_lvar[of x t]
unfolding rvars_def
by auto
end
text ‹Updating changes the value of ‹x› and then updates
values of all lhs variables so that the tableau remains
satisfied. This can be based on a function that recalculates rhs
polynomial values in the changed valuation:›
locale RhsEqVal = fixes rhs_eq_val::"(var, 'a::lrv) mapping ⇒ var ⇒ 'a ⇒ eq ⇒ 'a"
assumes rhs_eq_val: "⟨v⟩ ⊨⇩e e ⟹ rhs_eq_val v x c e = rhs e ⦃ ⟨v⟩ (x := c) ⦄"
begin
text‹\noindent Then, the next implementation of ‹update›
satisfies its specification:›
abbreviation update_eq where
"update_eq v x c v' e ≡ upd (lhs e) (rhs_eq_val v x c e) v'"
definition update :: "var ⇒ 'a ⇒ ('i,'a) state ⇒ ('i,'a) state" where
"update x c s ≡ 𝒱_update (upd x c (foldl (update_eq (𝒱 s) x c) (𝒱 s) (𝒯 s))) s"
lemma update_no_set_none:
shows "look (𝒱 s) y ≠ None ⟹
look (foldl (update_eq (𝒱 s) x v) (𝒱 s) t) y ≠ None"
by (induct t rule: rev_induct, auto simp: lookup_update')
lemma update_no_left:
assumes "y ∉ lvars t"
shows "look (𝒱 s) y = look (foldl (update_eq (𝒱 s) x v) (𝒱 s) t) y"
using assms
by (induct t rule: rev_induct) (auto simp add: lvars_def lookup_update')
lemma update_left:
assumes "y ∈ lvars t"
shows "∃ eq ∈ set t. lhs eq = y ∧
look (foldl (update_eq (𝒱 s) x v) (𝒱 s) t) y = Some (rhs_eq_val (𝒱 s) x v eq)"
using assms
by (induct t rule: rev_induct) (auto simp add: lvars_def lookup_update')
lemma update_valuate_rhs:
assumes "e ∈ set (𝒯 s)" "△ (𝒯 s)"
shows "rhs e ⦃ ⟨𝒱 (update x c s)⟩ ⦄ = rhs e ⦃ ⟨𝒱 s⟩ (x := c) ⦄"
proof (rule valuate_depend, safe)
fix y
assume "y ∈ rvars_eq e"
then have "y ∉ lvars (𝒯 s)"
using ‹△ (𝒯 s)› ‹e ∈ set (𝒯 s)›
by (auto simp add: normalized_tableau_def rvars_def)
then show "⟨𝒱 (update x c s)⟩ y = (⟨𝒱 s⟩(x := c)) y"
using update_no_left[of y "𝒯 s" s x c]
by (auto simp add: update_def map2fun_def lookup_update')
qed
end
sublocale RhsEqVal < Update update
proof
fix s::"('i,'a) state" and x c
show "let s' = update x c s in 𝒯 s' = 𝒯 s ∧ ℬ⇩i s' = ℬ⇩i s ∧ 𝒰 s' = 𝒰 s ∧ 𝒰⇩c s' = 𝒰⇩c s"
by (simp add: Let_def update_def add: boundsl_def boundsu_def indexl_def indexu_def)
next
fix s::"('i,'a) state" and x c
assume "△ (𝒯 s)" "∇ s" "x ∉ lvars (𝒯 s)"
then show "∇ (update x c s)"
using update_no_set_none[of s]
by (simp add: Let_def update_def tableau_valuated_def lookup_update')
next
fix s::"('i,'a) state" and x x' c
assume "△ (𝒯 s)" "∇ s" "x ∉ lvars (𝒯 s)"
show "x' ∉ lvars (𝒯 s) ⟶
look (𝒱 (update x c s)) x' =
(if x = x' then Some c else look (𝒱 s) x')"
using update_no_left[of x' "𝒯 s" s x c]
unfolding update_def lvars_def Let_def
by (auto simp: lookup_update')
next
fix s::"('i,'a) state" and x c
assume "△ (𝒯 s)" "∇ s" "x ∉ lvars (𝒯 s)"
have "⟨𝒱 s⟩ ⊨⇩t 𝒯 s ⟹ ∀e ∈ set (𝒯 s). ⟨𝒱 (update x c s)⟩ ⊨⇩e e"
proof
fix e
assume "e ∈ set (𝒯 s)" "⟨𝒱 s⟩ ⊨⇩t 𝒯 s"
then have "⟨𝒱 s⟩ ⊨⇩e e"
by (simp add: satisfies_tableau_def)
have "x ≠ lhs e"
using ‹x ∉ lvars (𝒯 s)› ‹e ∈ set (𝒯 s)›
by (auto simp add: lvars_def)
then have "⟨𝒱 (update x c s)⟩ (lhs e) = rhs_eq_val (𝒱 s) x c e"
using update_left[of "lhs e" "𝒯 s" s x c] ‹e ∈ set (𝒯 s)› ‹△ (𝒯 s)›
by (auto simp add: lvars_def lookup_update' update_def Let_def map2fun_def normalized_tableau_def distinct_map inj_on_def)
then show "⟨𝒱 (update x c s)⟩ ⊨⇩e e"
using ‹⟨𝒱 s⟩ ⊨⇩e e› ‹e ∈ set (𝒯 s)› ‹x ∉ lvars (𝒯 s)› ‹△ (𝒯 s)›
using rhs_eq_val
by (simp add: satisfies_eq_def update_valuate_rhs)
qed
then show "⟨𝒱 s⟩ ⊨⇩t 𝒯 s ⟶ ⟨𝒱 (update x c s)⟩ ⊨⇩t 𝒯 s"
by(simp add: satisfies_tableau_def update_def)
qed
text‹To update the valuation for a variable that is on the lhs of
the tableau it should first be swapped with some rhs variable of its
equation, in an operation called \emph{pivoting}. Pivoting has the
precondition that the tableau is normalized and that it is always
called for a lhs variable of the tableau, and a rhs variable in the
equation with that lhs variable. The set of rhs variables for the
given lhs variable is found using the ‹rvars_of_lvar› function
(specified in a very simple locale ‹EqForLVar›, that we do not
print).›
locale Pivot = EqForLVar + fixes pivot::"var ⇒ var ⇒ ('i,'a::lrv) state ⇒ ('i,'a) state"
assumes
pivot_id: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
let s' = pivot x⇩i x⇩j s in 𝒱 s' = 𝒱 s ∧ ℬ⇩i s' = ℬ⇩i s ∧ 𝒰 s' = 𝒰 s ∧ 𝒰⇩c s' = 𝒰⇩c s" and
pivot_tableau: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
let s' = pivot x⇩i x⇩j s in ((v::'a valuation) ⊨⇩t 𝒯 s ⟷ v ⊨⇩t 𝒯 s') ∧ △ (𝒯 s') " and
pivot_vars': "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ let s' = pivot x⇩i x⇩j s in
rvars(𝒯 s') = rvars(𝒯 s)-{x⇩j}∪{x⇩i} ∧ lvars(𝒯 s') = lvars(𝒯 s)-{x⇩i}∪{x⇩j}"
begin
lemma pivot_bounds_id: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
ℬ⇩i (pivot x⇩i x⇩j s) = ℬ⇩i s"
using pivot_id
by (simp add: Let_def)
lemma pivot_bounds_id': assumes "△ (𝒯 s)" "x⇩i ∈ lvars (𝒯 s)" "x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i"
shows "ℬℐ (pivot x⇩i x⇩j s) = ℬℐ s" "ℬ (pivot x⇩i x⇩j s) = ℬ s" "ℐ (pivot x⇩i x⇩j s) = ℐ s"
using pivot_bounds_id[OF assms]
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
lemma pivot_valuation_id: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ 𝒱 (pivot x⇩i x⇩j s) = 𝒱 s"
using pivot_id
by (simp add: Let_def)
lemma pivot_unsat_id: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ 𝒰 (pivot x⇩i x⇩j s) = 𝒰 s"
using pivot_id
by (simp add: Let_def)
lemma pivot_unsat_core_id: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ 𝒰⇩c (pivot x⇩i x⇩j s) = 𝒰⇩c s"
using pivot_id
by (simp add: Let_def)
lemma pivot_tableau_equiv: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
(v::'a valuation) ⊨⇩t 𝒯 s = v ⊨⇩t 𝒯 (pivot x⇩i x⇩j s)"
using pivot_tableau
by (simp add: Let_def)
lemma pivot_tableau_normalized: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ △ (𝒯 (pivot x⇩i x⇩j s))"
using pivot_tableau
by (simp add: Let_def)
lemma pivot_rvars: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ rvars (𝒯 (pivot x⇩i x⇩j s)) = rvars (𝒯 s) - {x⇩j} ∪ {x⇩i}"
using pivot_vars'
by (simp add: Let_def)
lemma pivot_lvars: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ lvars (𝒯 (pivot x⇩i x⇩j s)) = lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
using pivot_vars'
by (simp add: Let_def)
lemma pivot_vars:
"⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ tvars (𝒯 (pivot x⇩i x⇩j s)) = tvars (𝒯 s) "
using pivot_lvars[of s x⇩i x⇩j] pivot_rvars[of s x⇩i x⇩j]
using rvars_of_lvar_rvars[of x⇩i "𝒯 s"]
by auto
lemma
pivot_tableau_valuated: "⟦△ (𝒯 s); x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i; ∇ s⟧ ⟹ ∇ (pivot x⇩i x⇩j s)"
using pivot_valuation_id pivot_vars
by (auto simp add: tableau_valuated_def)
end
text‹Functions ‹pivot› and ‹update› can be used to
implement the ‹check› function. In its context, ‹pivot›
and ‹update› functions are always called together, so the
following definition can be used: @{prop "pivot_and_update x⇩i x⇩j c s =
update x⇩i c (pivot x⇩i x⇩j s)"}. It is possible to make a more efficient
implementation of ‹pivot_and_update› that does not use separate
implementations of ‹pivot› and ‹update›. To allow this, a
separate specification for ‹pivot_and_update› can be given. It can be
easily shown that the ‹pivot_and_update› definition above
satisfies this specification.›
locale PivotAndUpdate = EqForLVar +
fixes pivot_and_update :: "var ⇒ var ⇒ 'a::lrv ⇒ ('i,'a) state ⇒ ('i,'a) state"
assumes pivotandupdate_unsat_id: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
𝒰 (pivot_and_update x⇩i x⇩j c s) = 𝒰 s"
assumes pivotandupdate_unsat_core_id: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
𝒰⇩c (pivot_and_update x⇩i x⇩j c s) = 𝒰⇩c s"
assumes pivotandupdate_bounds_id: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
ℬ⇩i (pivot_and_update x⇩i x⇩j c s) = ℬ⇩i s"
assumes pivotandupdate_tableau_normalized: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
△ (𝒯 (pivot_and_update x⇩i x⇩j c s))"
assumes pivotandupdate_tableau_equiv: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
(v::'a valuation) ⊨⇩t 𝒯 s ⟷ v ⊨⇩t 𝒯 (pivot_and_update x⇩i x⇩j c s)"
assumes pivotandupdate_satisfies_tableau: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
⟨𝒱 s⟩ ⊨⇩t 𝒯 s ⟶ ⟨𝒱 (pivot_and_update x⇩i x⇩j c s)⟩ ⊨⇩t 𝒯 s"
assumes pivotandupdate_rvars: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
rvars (𝒯 (pivot_and_update x⇩i x⇩j c s)) = rvars (𝒯 s) - {x⇩j} ∪ {x⇩i}"
assumes pivotandupdate_lvars: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
lvars (𝒯 (pivot_and_update x⇩i x⇩j c s)) = lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
assumes pivotandupdate_valuation_nonlhs: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
x ∉ lvars (𝒯 s) - {x⇩i} ∪ {x⇩j} ⟶ look (𝒱 (pivot_and_update x⇩i x⇩j c s)) x = (if x = x⇩i then Some c else look (𝒱 s) x)"
assumes pivotandupdate_tableau_valuated: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹
∇ (pivot_and_update x⇩i x⇩j c s)"
begin
lemma pivotandupdate_bounds_id': assumes "△ (𝒯 s)" "∇ s" "x⇩i ∈ lvars (𝒯 s)" "x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i"
shows "ℬℐ (pivot_and_update x⇩i x⇩j c s) = ℬℐ s"
"ℬ (pivot_and_update x⇩i x⇩j c s) = ℬ s"
"ℐ (pivot_and_update x⇩i x⇩j c s) = ℐ s"
using pivotandupdate_bounds_id[OF assms]
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
lemma pivotandupdate_valuation_xi: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i⟧ ⟹ look (𝒱 (pivot_and_update x⇩i x⇩j c s)) x⇩i = Some c"
using pivotandupdate_valuation_nonlhs[of s x⇩i x⇩j x⇩i c]
using rvars_of_lvar_rvars
by (auto simp add: normalized_tableau_def)
lemma pivotandupdate_valuation_other_nolhs: "⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i; x ∉ lvars (𝒯 s); x ≠ x⇩j⟧ ⟹ look (𝒱 (pivot_and_update x⇩i x⇩j c s)) x = look (𝒱 s) x"
using pivotandupdate_valuation_nonlhs[of s x⇩i x⇩j x c]
by auto
lemma pivotandupdate_nolhs:
"⟦ △ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i;
⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; ℬ⇩l s x⇩i = Some c ∨ ℬ⇩u s x⇩i = Some c⟧ ⟹
⊨⇩n⇩o⇩l⇩h⇩s (pivot_and_update x⇩i x⇩j c s)"
using pivotandupdate_satisfies_tableau[of s x⇩i x⇩j c]
using pivotandupdate_tableau_equiv[of s x⇩i x⇩j _ c]
using pivotandupdate_valuation_xi[of s x⇩i x⇩j c]
using pivotandupdate_valuation_other_nolhs[of s x⇩i x⇩j _ c]
using pivotandupdate_lvars[of s x⇩i x⇩j c]
by (auto simp add: curr_val_satisfies_no_lhs_def satisfies_bounds.simps satisfies_bounds_set.simps
bounds_consistent_geq_lb bounds_consistent_leq_ub map2fun_def pivotandupdate_bounds_id')
lemma pivotandupdate_bounds_consistent:
assumes "△ (𝒯 s)" "∇ s" "x⇩i ∈ lvars (𝒯 s)" "x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i"
shows "◇ (pivot_and_update x⇩i x⇩j c s) = ◇ s"
using assms pivotandupdate_bounds_id'[of s x⇩i x⇩j c]
by (simp add: bounds_consistent_def)
end
locale PivotUpdate = Pivot eq_idx_for_lvar pivot + Update update for
eq_idx_for_lvar :: "tableau ⇒ var ⇒ nat" and
pivot :: "var ⇒ var ⇒ ('i,'a::lrv) state ⇒ ('i,'a) state" and
update :: "var ⇒ 'a ⇒ ('i,'a) state ⇒ ('i,'a) state"
begin
definition pivot_and_update :: "var ⇒ var ⇒ 'a ⇒ ('i,'a) state ⇒ ('i,'a) state" where [simp]:
"pivot_and_update x⇩i x⇩j c s ≡ update x⇩i c (pivot x⇩i x⇩j s)"
lemma pivot_update_precond:
assumes "△ (𝒯 s)" "x⇩i ∈ lvars (𝒯 s)" "x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i"
shows "△ (𝒯 (pivot x⇩i x⇩j s))" "x⇩i ∉ lvars (𝒯 (pivot x⇩i x⇩j s))"
proof-
from assms have "x⇩i ≠ x⇩j"
using rvars_of_lvar_rvars[of x⇩i "𝒯 s"]
by (auto simp add: normalized_tableau_def)
then show "△ (𝒯 (pivot x⇩i x⇩j s))" "x⇩i ∉ lvars (𝒯 (pivot x⇩i x⇩j s))"
using assms
using pivot_tableau_normalized[of s x⇩i x⇩j]
using pivot_lvars[of s x⇩i x⇩j]
by auto
qed
end
sublocale PivotUpdate < PivotAndUpdate eq_idx_for_lvar pivot_and_update
using pivot_update_precond
using update_unsat_id pivot_unsat_id pivot_unsat_core_id update_bounds_id pivot_bounds_id
update_tableau_id pivot_tableau_normalized pivot_tableau_equiv update_satisfies_tableau
pivot_valuation_id pivot_lvars pivot_rvars update_valuation_nonlhs update_valuation_nonlhs
pivot_tableau_valuated update_tableau_valuated update_unsat_core_id
by (unfold_locales, auto)
text‹Given the @{term update} function, ‹assert_bound› can be
implemented as follows.
\vspace{-2mm}
@{text[display]
"assert_bound (Leq x c) s ≡
if c ≥⇩u⇩b ℬ⇩u s x then s
else let s' = s ⦇ ℬ⇩u := (ℬ⇩u s) (x := Some c) ⦈
in if c <⇩l⇩b ℬ⇩l s x then s' ⦇ 𝒰 := True ⦈
else if x ∉ lvars (𝒯 s') ∧ c < ⟨𝒱 s⟩ x then update x c s' else s'"
}
\vspace{-2mm}
\noindent The case of ‹Geq x c› atoms is analogous (a systematic way to
avoid symmetries is discussed in Section \ref{sec:symmetries}). This
implementation satisfies both its specifications.
›
lemma indices_state_set_unsat: "indices_state (set_unsat I s) = indices_state s"
by (cases s, auto simp: indices_state_def)
lemma ℬℐ_set_unsat: "ℬℐ (set_unsat I s) = ℬℐ s"
by (cases s, auto simp: boundsl_def boundsu_def indexl_def indexu_def)
lemma satisfies_tableau_cong: assumes "⋀ x. x ∈ tvars t ⟹ v x = w x"
shows "(v ⊨⇩t t) = (w ⊨⇩t t)"
unfolding satisfies_tableau_def satisfies_eq_def
by (intro ball_cong[OF refl] arg_cong2[of _ _ _ _ "(=)"] valuate_depend,
insert assms, auto simp: lvars_def rvars_def)
lemma satisfying_state_valuation_to_atom_tabl: assumes J: "J ⊆ indices_state s"
and model: "(J, v) ⊨⇩i⇩s⇩e s"
and ivalid: "index_valid as s"
and dist: "distinct_indices_atoms as"
shows "(J, v) ⊨⇩i⇩a⇩e⇩s as" "v ⊨⇩t 𝒯 s"
unfolding i_satisfies_atom_set'.simps
proof (intro ballI)
from model[unfolded satisfies_state_index'.simps]
have model: "v ⊨⇩t 𝒯 s" "(J, v) ⊨⇩i⇩b⇩e ℬℐ s" by auto
show "v ⊨⇩t 𝒯 s" by fact
fix a
assume "a ∈ restrict_to J as"
then obtain i where iJ: "i ∈ J" and mem: "(i,a) ∈ as" by auto
with J have "i ∈ indices_state s" by auto
from this[unfolded indices_state_def] obtain x c where
look: "look (ℬ⇩i⇩l s) x = Some (i, c) ∨ look (ℬ⇩i⇩u s) x = Some (i, c)" by auto
with ivalid[unfolded index_valid_def]
obtain b where "(i,b) ∈ as" "atom_var b = x" "atom_const b = c" by force
with dist[unfolded distinct_indices_atoms_def, rule_format, OF this(1) mem]
have a: "atom_var a = x" "atom_const a = c" by auto
from model(2)[unfolded satisfies_bounds_index'.simps] look iJ have "v x = c"
by (auto simp: boundsu_def boundsl_def indexu_def indexl_def)
thus "v ⊨⇩a⇩e a" unfolding satisfies_atom'_def a .
qed
text ‹Note that in order to ensure minimality of the unsat cores, pivoting is required.›
sublocale AssertAllState < AssertAll assert_all
proof
fix t as v I
assume D: "△ t"
from D show "assert_all t as = Sat v ⟹ ⟨v⟩ ⊨⇩t t ∧ ⟨v⟩ ⊨⇩a⇩s flat (set as)"
unfolding Let_def assert_all_def
using assert_all_state_tableau_equiv[OF D refl]
using assert_all_state_sat[OF D refl]
using assert_all_state_sat_atoms_equiv_bounds[OF D refl, of as]
unfolding atoms_equiv_bounds.simps curr_val_satisfies_state_def satisfies_state_def satisfies_atom_set_def
by (auto simp: Let_def split: if_splits)
let ?s = "assert_all_state t as"
assume "assert_all t as = Unsat I"
then have i: "I = the (𝒰⇩c ?s)" and U: "𝒰 ?s"
unfolding assert_all_def Let_def by (auto split: if_splits)
from assert_all_index_valid[OF D refl, of as] have ivalid: "index_valid (set as) ?s" .
note unsat = assert_all_state_unsat[OF D refl U, unfolded minimal_unsat_state_core_def unsat_state_core_def i[symmetric]]
from unsat have "set I ⊆ indices_state ?s" by auto
also have "… ⊆ fst ` set as" using assert_all_state_indices[OF D refl] .
finally have indices: "set I ⊆ fst ` set as" .
show "minimal_unsat_core_tabl_atoms (set I) t (set as)"
unfolding minimal_unsat_core_tabl_atoms_def
proof (intro conjI impI allI indices, clarify)
fix v
assume model: "v ⊨⇩t t" "(set I, v) ⊨⇩i⇩a⇩s set as"
from unsat have no_model: "¬ ((set I, v) ⊨⇩i⇩s ?s)" by auto
from assert_all_state_unsat_atoms_equiv_bounds[OF D refl U]
have equiv: "set as ⊨⇩i ℬℐ ?s" by auto
from assert_all_state_tableau_equiv[OF D refl, of v] model
have model_t: "v ⊨⇩t 𝒯 ?s" by auto
have model_as': "(set I, v) ⊨⇩i⇩a⇩s set as"
using model(2) by (auto simp: satisfies_atom_set_def)
with equiv model_t have "(set I, v) ⊨⇩i⇩s ?s"
unfolding satisfies_state_index.simps atoms_imply_bounds_index.simps by simp
with no_model show False by simp
next
fix J
assume dist: "distinct_indices_atoms (set as)" and J: "J ⊂ set I"
from J unsat[unfolded subsets_sat_core_def, folded i]
have J': "J ⊆ indices_state ?s" by auto
from index_valid_distinct_indices[OF ivalid dist] J unsat[unfolded subsets_sat_core_def, folded i]
obtain v where model: "(J, v) ⊨⇩i⇩s⇩e ?s" by blast
have "(J, v) ⊨⇩i⇩a⇩e⇩s set as" "v ⊨⇩t t"
using satisfying_state_valuation_to_atom_tabl[OF J' model ivalid dist]
assert_all_state_tableau_equiv[OF D refl] by auto
then show "∃ v. v ⊨⇩t t ∧ (J, v) ⊨⇩i⇩a⇩e⇩s set as" by blast
qed
qed
lemma (in Update) update_to_assert_bound_no_lhs: assumes pivot: "Pivot eqlvar (pivot :: var ⇒ var ⇒ ('i,'a) state ⇒ ('i,'a) state)"
shows "AssertBoundNoLhs assert_bound"
proof
fix s::"('i,'a) state" and a
assume "¬ 𝒰 s" "△ (𝒯 s)" "∇ s"
then show "𝒯 (assert_bound a s) = 𝒯 s"
by (cases a, cases "snd a") (auto simp add: Let_def update_tableau_id tableau_valuated_def)
next
fix s::"('i,'a) state" and ia and as
assume *: "¬ 𝒰 s" "△ (𝒯 s)" "∇ s" and **: "𝒰 (assert_bound ia s)"
and index: "index_valid as s"
and consistent: "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s"
obtain i a where ia: "ia = (i,a)" by force
let ?modelU = "λ lt UB UI s v x c i. UB s x = Some c ⟶ UI s x = i ⟶ i ∈ set (the (𝒰⇩c s)) ⟶ (lt (v x) c ∨ v x = c)"
let ?modelL = "λ lt LB LI s v x c i. LB s x = Some c ⟶ LI s x = i ⟶ i ∈ set (the (𝒰⇩c s)) ⟶ (lt c (v x) ∨ c = v x)"
let ?modelIU = "λ I lt UB UI s v x c i. UB s x = Some c ⟶ UI s x = i ⟶ i ∈ I ⟶ (v x = c)"
let ?modelIL = "λ I lt LB LI s v x c i. LB s x = Some c ⟶ LI s x = i ⟶ i ∈ I ⟶ (v x = c)"
let ?P' = "λ lt UBI LBI UB LB UBI_upd UI LI LE GE s.
𝒰 s ⟶ (set (the (𝒰⇩c s)) ⊆ indices_state s ∧ ¬ (∃v. (v ⊨⇩t 𝒯 s
∧ (∀ x c i. ?modelU lt UB UI s v x c i)
∧ (∀ x c i. ?modelL lt LB LI s v x c i))))
∧ (distinct_indices_state s ⟶ (∀ I. I ⊂ set (the (𝒰⇩c s)) ⟶ (∃ v. v ⊨⇩t 𝒯 s ∧
(∀ x c i. ?modelIU I lt UB UI s v x c i) ∧ (∀ x c i. ?modelIL I lt LB LI s v x c i))))"
have "𝒰 (assert_bound ia s) ⟶ (unsat_state_core (assert_bound ia s) ∧
(distinct_indices_state (assert_bound ia s) ⟶ subsets_sat_core (assert_bound ia s)))" (is "?P (assert_bound ia s)") unfolding ia
proof (rule assert_bound_cases[of _ _ ?P'])
fix s' :: "('i,'a) state"
have id: "((x :: 'a) < y ∨ x = y) ⟷ x ≤ y" "((x :: 'a) > y ∨ x = y) ⟷ x ≥ y" for x y by auto
have id': "?P' (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u undefined ℐ⇩l ℐ⇩u Geq Leq s' = ?P' (<) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l undefined ℐ⇩u ℐ⇩l Leq Geq s'"
by (intro arg_cong[of _ _ "λ y. _ ⟶ y"] arg_cong[of _ _ "λ x. _ ∧ x"],
intro arg_cong2[of _ _ _ _ "(∧)"] arg_cong[of _ _ "λ y. _ ⟶ y"] arg_cong[of _ _ "λ y. ∀ x ⊂ set (the (𝒰⇩c s')). y x"] ext arg_cong[of _ _ Not],
unfold id, auto)
show "?P s' = ?P' (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u undefined ℐ⇩l ℐ⇩u Geq Leq s'"
unfolding satisfies_state_def satisfies_bounds_index.simps satisfies_bounds.simps
in_bounds.simps unsat_state_core_def satisfies_state_index.simps subsets_sat_core_def
satisfies_state_index'.simps satisfies_bounds_index'.simps
unfolding bound_compare''_defs id
by ((intro arg_cong[of _ _ "λ x. _ ⟶ x"] arg_cong[of _ _ "λ x. _ ∧ x"],
intro arg_cong2[of _ _ _ _ "(∧)"] refl arg_cong[of _ _ "λ x. _ ⟶ x"] arg_cong[of _ _ Not]
arg_cong[of _ _ "λ y. ∀ x ⊂ set (the (𝒰⇩c s')). y x"] ext; intro arg_cong[of _ _ Ex] ext), auto)
then show "?P s' = ?P' (<) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l undefined ℐ⇩u ℐ⇩l Leq Geq s'" unfolding id' .
next
fix c::'a and x::nat and dir
assume "⊲⇩l⇩b (lt dir) c (LB dir s x)" and dir: "dir = Positive ∨ dir = Negative"
then obtain d where some: "LB dir s x = Some d" and lt: "lt dir c d"
by (auto simp: bound_compare'_defs split: option.splits)
from index[unfolded index_valid_def, rule_format, of x _ d]
some dir obtain j where ind: "LI dir s x = j" "look (LBI dir s) x = Some (j,d)" and ge: "(j, GE dir x d) ∈ as"
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
let ?s = "set_unsat [i, ((LI dir) s x)] (updateℬℐ (UBI_upd dir) i x c s)"
let ?ss = "updateℬℐ (UBI_upd dir) i x c s"
show "?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) ?s"
proof (intro conjI impI allI, goal_cases)
case 1
thus ?case using dir ind ge lt some by (force simp: indices_state_def split: if_splits)
next
case 2
{
fix v
assume vU: "∀ x c i. ?modelU (lt dir) (UB dir) (UI dir) ?s v x c i"
assume vL: "∀ x c i. ?modelL (lt dir) (LB dir) (LI dir) ?s v x c i"
from dir have "UB dir ?s x = Some c" "UI dir ?s x = i" by (auto simp: boundsl_def boundsu_def indexl_def indexu_def)
from vU[rule_format, OF this] have vx_le_c: "lt dir (v x) c ∨ v x = c" by auto
from dir ind some have *: "LB dir ?s x = Some d" "LI dir ?s x = j" by (auto simp: boundsl_def boundsu_def indexl_def indexu_def)
have d_le_vx: "lt dir d (v x) ∨ d = v x" by (intro vL[rule_format, OF *], insert some ind, auto)
from dir d_le_vx vx_le_c lt
have False by auto
}
thus ?case by blast
next
case (3 I)
then obtain j where I: "I ⊆ {j}" by (auto split: if_splits)
from 3 have dist: "distinct_indices_state ?ss" unfolding distinct_indices_state_def by auto
have id1: "UB dir ?s y = UB dir ?ss y" "LB dir ?s y = LB dir ?ss y"
"UI dir ?s y = UI dir ?ss y" "LI dir ?s y = LI dir ?ss y"
"𝒯 ?s = 𝒯 s"
"set (the (𝒰⇩c ?s)) = {i,LI dir s x}" for y
using dir by (auto simp: boundsu_def boundsl_def indexu_def indexl_def)
from I have id: "(∀ k. P1 k ⟶ P2 k ⟶ k ∈ I ⟶ Q k) ⟷ (I = {} ∨ (P1 j ⟶ P2 j ⟶ Q j))" for P1 P2 Q by auto
have id2: "(UB dir s xa = Some ca ⟶ UI dir s xa = j ⟶ P) = (look (UBI dir s) xa = Some (j,ca) ⟶ P)"
"(LB dir s xa = Some ca ⟶ LI dir s xa = j ⟶ P) = (look (LBI dir s) xa = Some (j,ca) ⟶ P)" for xa ca P s
using dir by (auto simp: boundsu_def indexu_def boundsl_def indexl_def)
have "∃v. v ⊨⇩t 𝒯 s ∧
(∀xa ca ia.
UB dir ?ss xa = Some ca ⟶ UI dir ?ss xa = ia ⟶ ia ∈ I ⟶ v xa = ca) ∧
(∀xa ca ia.
LB dir ?ss xa = Some ca ⟶ LI dir ?ss xa = ia ⟶ ia ∈ I ⟶ v xa = ca)"
proof (cases "∃ xa ca. look (UBI dir ?ss) xa = Some (j,ca) ∨ look (LBI dir ?ss) xa = Some (j,ca)")
case False
thus ?thesis unfolding id id2 using consistent unfolding curr_val_satisfies_no_lhs_def
by (intro exI[of _ "⟨𝒱 s⟩"], auto)
next
case True
from consistent have val: " ⟨𝒱 s⟩ ⊨⇩t 𝒯 s" unfolding curr_val_satisfies_no_lhs_def by auto
define ss where ss: "ss = ?ss"
from True obtain y b where "look (UBI dir ?ss) y = Some (j,b) ∨ look (LBI dir ?ss) y = Some (j,b)" by force
then have id3: "(look (LBI dir ss) yy = Some (j,bb) ∨ look (UBI dir ss) yy = Some (j,bb)) ⟷ (yy = y ∧ bb = b)" for yy bb
using distinct_indices_stateD(1)[OF dist, of y j b yy bb] using dir
unfolding ss[symmetric]
by (auto simp: boundsu_def boundsl_def indexu_def indexl_def)
have "∃v. v ⊨⇩t 𝒯 s ∧ v y = b"
proof (cases "y ∈ lvars (𝒯 s)")
case False
let ?v = "⟨𝒱 (update y b s)⟩"
show ?thesis
proof (intro exI[of _ ?v] conjI)
from update_satisfies_tableau[OF *(2,3) False] val
show "?v ⊨⇩t 𝒯 s" by simp
from update_valuation_nonlhs[OF *(2,3) False, of y b] False
show "?v y = b" by (simp add: map2fun_def')
qed
next
case True
from *(2)[unfolded normalized_tableau_def]
have zero: "0 ∉ rhs ` set (𝒯 s)" by auto
interpret Pivot eqlvar pivot by fact
interpret PivotUpdate eqlvar pivot update ..
let ?eq = "eq_for_lvar (𝒯 s) y"
from eq_for_lvar[OF True] have "?eq ∈ set (𝒯 s)" "lhs ?eq = y" by auto
with zero have rhs: "rhs ?eq ≠ 0" by force
hence "rvars_eq ?eq ≠ {}"
by (simp add: vars_empty_zero)
then obtain z where z: "z ∈ rvars_eq ?eq" by auto
let ?v = "𝒱 (pivot_and_update y z b s)"
let ?vv = "⟨?v⟩"
from pivotandupdate_valuation_xi[OF *(2,3) True z]
have "look ?v y = Some b" .
hence vv: "?vv y = b" unfolding map2fun_def' by auto
show ?thesis
proof (intro exI[of _ ?vv] conjI vv)
show "?vv ⊨⇩t 𝒯 s" using pivotandupdate_satisfies_tableau[OF *(2,3) True z] val by auto
qed
qed
thus ?thesis unfolding id id2 ss[symmetric] using id3 by metis
qed
thus ?case unfolding id1 .
qed
next
fix c::'a and x::nat and dir
assume **: "dir = Positive ∨ dir = Negative" "a = LE dir x c" "x ∉ lvars (𝒯 s)" "lt dir c (⟨𝒱 s⟩ x)"
"¬ ⊵⇩u⇩b (lt dir) c (UB dir s x)" "¬ ⊲⇩l⇩b (lt dir) c (LB dir s x)"
let ?s = "updateℬℐ (UBI_upd dir) i x c s"
show "?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)
(update x c ?s)"
using * **
by (auto simp add: update_unsat_id tableau_valuated_def)
qed (auto simp add: * update_unsat_id tableau_valuated_def)
with ** show "minimal_unsat_state_core (assert_bound ia s)" by (auto simp: minimal_unsat_state_core_def)
next
fix s::"('i,'a) state" and ia
assume *: "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s" "△ (𝒯 s)" "∇ s"
and **: "¬ 𝒰 (assert_bound ia s)" (is ?lhs)
obtain i a where ia: "ia = (i,a)" by force
have "⟨𝒱 (assert_bound ia s)⟩ ⊨⇩t 𝒯 (assert_bound ia s)"
proof-
let ?P = "λ lt UBI LBI UB LB UBI_upd UI LI LE GE s. ⟨𝒱 s⟩ ⊨⇩t 𝒯 s"
show ?thesis unfolding ia
proof (rule assert_bound_cases[of _ _ ?P])
fix c x and dir :: "('i,'a) Direction"
let ?s' = "updateℬℐ (UBI_upd dir) i x c s"
assume "x ∉ lvars (𝒯 s)" "(lt dir) c (⟨𝒱 s⟩ x)"
"dir = Positive ∨ dir = Negative"
then show "⟨𝒱 (update x c ?s')⟩ ⊨⇩t 𝒯 (update x c ?s')"
using *
using update_satisfies_tableau[of ?s' x c] update_tableau_id
by (auto simp add: curr_val_satisfies_no_lhs_def tableau_valuated_def)
qed (insert *, auto simp add: curr_val_satisfies_no_lhs_def)
qed
moreover
have "¬ 𝒰 (assert_bound ia s) ⟶ ⟨𝒱 (assert_bound ia s)⟩ ⊨⇩b ℬ (assert_bound ia s) ∥ - lvars (𝒯 (assert_bound ia s))" (is "?P (assert_bound ia s)")
proof-
let ?P' = "λ lt UBI LBI UB LB UB_upd UI LI LE GE s.
¬ 𝒰 s ⟶ (∀x∈- lvars (𝒯 s). ⊵⇩l⇩b lt (⟨𝒱 s⟩ x) (LB s x) ∧ ⊴⇩u⇩b lt (⟨𝒱 s⟩ x) (UB s x))"
let ?P'' = "λ dir. ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)"
have x: "⋀ s'. ?P s' = ?P' (<) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩u ℬ⇩l ℬ⇩i⇩u_update ℐ⇩u ℐ⇩l Leq Geq s'"
and xx: "⋀ s'. ?P s' = ?P' (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℬ⇩i⇩l_update ℐ⇩l ℐ⇩u Geq Leq s'"
unfolding satisfies_bounds_set.simps in_bounds.simps bound_compare_defs
by (auto split: option.split)
show ?thesis unfolding ia
proof (rule assert_bound_cases[of _ _ ?P'])
fix dir :: "('i,'a) Direction"
assume "dir = Positive ∨ dir = Negative"
then show "?P'' dir s"
using x[of s] xx[of s] ‹⊨⇩n⇩o⇩l⇩h⇩s s›
by (auto simp add: curr_val_satisfies_no_lhs_def)
next
fix x c and dir :: "('i,'a) Direction"
let ?s' = "updateℬℐ (UBI_upd dir) i x c s"
assume "x ∈ lvars (𝒯 s)" "dir = Positive ∨ dir = Negative"
then have "?P ?s'"
using ‹⊨⇩n⇩o⇩l⇩h⇩s s›
by (auto simp add: satisfies_bounds_set.simps curr_val_satisfies_no_lhs_def
boundsl_def boundsu_def indexl_def indexu_def)
then show "?P'' dir ?s'"
using x[of ?s'] xx[of ?s'] ‹dir = Positive ∨ dir = Negative›
by auto
next
fix c x and dir :: "('i,'a) Direction"
let ?s' = "updateℬℐ (UBI_upd dir) i x c s"
assume "¬ lt dir c (⟨𝒱 s⟩ x)" "dir = Positive ∨ dir = Negative"
then show "?P'' dir ?s'"
using ‹⊨⇩n⇩o⇩l⇩h⇩s s›
by (auto simp add: satisfies_bounds_set.simps curr_val_satisfies_no_lhs_def
simp: boundsl_def boundsu_def indexl_def indexu_def)
(auto simp add: bound_compare_defs)
next
fix c x and dir :: "('i,'a) Direction"
let ?s' = "update x c (updateℬℐ (UBI_upd dir) i x c s)"
assume "x ∉ lvars (𝒯 s)" "¬ ⊲⇩l⇩b (lt dir) c (LB dir s x)"
"dir = Positive ∨ dir = Negative"
show "?P'' dir ?s'"
proof (rule impI, rule ballI)
fix y
assume "¬ 𝒰 ?s'" "y ∈ - lvars (𝒯 ?s')"
show "⊵⇩l⇩b (lt dir) (⟨𝒱 ?s'⟩ y) (LB dir ?s' y) ∧ ⊴⇩u⇩b (lt dir) (⟨𝒱 ?s'⟩ y) (UB dir ?s' y)"
proof (cases "x = y")
case True
then show ?thesis
using ‹x ∉ lvars (𝒯 s)›
using ‹y ∈ - lvars (𝒯 ?s')›
using ‹¬ ⊲⇩l⇩b (lt dir) c (LB dir s x)›
using ‹dir = Positive ∨ dir = Negative›
using neg_bounds_compare(7) neg_bounds_compare(3)
using *
by (auto simp add: update_valuation_nonlhs update_tableau_id update_bounds_id bound_compare''_defs map2fun_def tableau_valuated_def bounds_updates) (force simp add: bound_compare'_defs)+
next
case False
then show ?thesis
using ‹x ∉ lvars (𝒯 s)› ‹y ∈ - lvars (𝒯 ?s')›
using ‹dir = Positive ∨ dir = Negative› *
by (auto simp add: update_valuation_nonlhs update_tableau_id update_bounds_id bound_compare''_defs satisfies_bounds_set.simps curr_val_satisfies_no_lhs_def map2fun_def
tableau_valuated_def bounds_updates)
qed
qed
qed (auto simp add: x xx)
qed
moreover
have "¬ 𝒰 (assert_bound ia s) ⟶ ◇ (assert_bound ia s)" (is "?P (assert_bound ia s)")
proof-
let ?P' = "λ lt UBI LBI UB LB UBI_upd UI LI LE GE s.
¬ 𝒰 s ⟶
(∀x. if LB s x = None ∨ UB s x = None then True
else lt (the (LB s x)) (the (UB s x)) ∨ (the (LB s x) = the (UB s x)))"
let ?P'' = "λ dir. ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)"
have x: "⋀ s'. ?P s' = ?P' (<) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩u ℬ⇩l ℬ⇩i⇩u_update ℐ⇩u ℐ⇩l Leq Geq s'" and
xx: "⋀ s'. ?P s' = ?P' (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℬ⇩i⇩l_update ℐ⇩l ℐ⇩u Geq Leq s'"
unfolding bounds_consistent_def
by auto
show ?thesis unfolding ia
proof (rule assert_bound_cases[of _ _ ?P'])
fix dir :: "('i,'a) Direction"
assume "dir = Positive ∨ dir = Negative"
then show "?P'' dir s"
using ‹◇ s›
by (auto simp add: bounds_consistent_def) (erule_tac x=x in allE, auto)+
next
fix x c and dir :: "('i,'a) Direction"
let ?s' = "update x c (updateℬℐ (UBI_upd dir) i x c s)"
assume "dir = Positive ∨ dir = Negative" "x ∉ lvars (𝒯 s)"
"¬ ⊵⇩u⇩b (lt dir) c (UB dir s x)" "¬ ⊲⇩l⇩b (lt dir) c (LB dir s x)"
then show "?P'' dir ?s'"
using ‹◇ s› *
unfolding bounds_consistent_def
by (auto simp add: update_bounds_id tableau_valuated_def bounds_updates split: if_splits)
(force simp add: bound_compare'_defs, erule_tac x=xa in allE, simp)+
next
fix x c and dir :: "('i,'a) Direction"
let ?s' = "updateℬℐ (UBI_upd dir) i x c s"
assume "¬ ⊵⇩u⇩b (lt dir) c (UB dir s x)" "¬ ⊲⇩l⇩b (lt dir) c (LB dir s x)"
"dir = Positive ∨ dir = Negative"
then have "?P'' dir ?s'"
using ‹◇ s›
unfolding bounds_consistent_def
by (auto split: if_splits simp: bounds_updates)
(force simp add: bound_compare'_defs, erule_tac x=xa in allE, simp)+
then show "?P'' dir ?s'" "?P'' dir ?s'"
by simp_all
qed (auto simp add: x xx)
qed
ultimately
show "⊨⇩n⇩o⇩l⇩h⇩s (assert_bound ia s) ∧ ◇ (assert_bound ia s)"
using ‹?lhs›
unfolding curr_val_satisfies_no_lhs_def
by simp
next
fix s :: "('i,'a) state" and ats and ia :: "('i,'a) i_atom"
assume "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "△ (𝒯 s)" "∇ s"
obtain i a where ia: "ia = (i,a)" by force
{
fix ats
let ?P' = "λ lt UBI LBI UB LB UB_upd UI LI LE GE s'. ats ≐ ℬ s ⟶ (ats ∪ {a}) ≐ ℬ s'"
let ?P'' = "λ dir. ?P' (lt dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)"
have "ats ≐ ℬ s ⟶ (ats ∪ {a}) ≐ ℬ (assert_bound ia s)" (is "?P (assert_bound ia s)")
unfolding ia
proof (rule assert_bound_cases[of _ _ ?P'])
fix x c and dir :: "('i,'a) Direction"
assume "dir = Positive ∨ dir = Negative" "a = LE dir x c" "⊵⇩u⇩b (lt dir) c (UB dir s x)"
then show "?P s"
unfolding atoms_equiv_bounds.simps satisfies_atom_set_def satisfies_bounds.simps
by auto (erule_tac x=x in allE, force simp add: bound_compare_defs)+
next
fix x c and dir :: "('i,'a) Direction"
let ?s' = "set_unsat [i, ((LI dir) s x)] (updateℬℐ (UBI_upd dir) i x c s)"
assume "dir = Positive ∨ dir = Negative" "a = LE dir x c" "¬ (⊵⇩u⇩b (lt dir) c (UB dir s x))"
then show "?P ?s'" unfolding set_unsat_bounds_id
using atoms_equiv_bounds_extend[of dir c s x ats i]
by auto
next
fix x c and dir :: "('i,'a) Direction"
let ?s' = "updateℬℐ (UBI_upd dir) i x c s"
assume "dir = Positive ∨ dir = Negative" "a = LE dir x c" "¬ (⊵⇩u⇩b (lt dir) c (UB dir s x))"
then have "?P ?s'"
using atoms_equiv_bounds_extend[of dir c s x ats i]
by auto
then show "?P ?s'" "?P ?s'"
by simp_all
next
fix x c and dir :: "('i,'a) Direction"
let ?s = "updateℬℐ (UBI_upd dir) i x c s"
let ?s' = "update x c ?s"
assume *: "dir = Positive ∨ dir = Negative" "a = LE dir x c" "¬ (⊵⇩u⇩b (lt dir) c (UB dir s x))" "x ∉ lvars (𝒯 s)"
then have "△ (𝒯 ?s)" "∇ ?s" "x ∉ lvars (𝒯 ?s)"
using ‹△ (𝒯 s)› ‹⊨⇩n⇩o⇩l⇩h⇩s s› ‹∇ s›
by (auto simp: tableau_valuated_def)
from update_bounds_id[OF this, of c]
have "ℬ⇩i ?s' = ℬ⇩i ?s" by blast
then have id: "ℬ ?s' = ℬ ?s" unfolding boundsl_def boundsu_def by auto
show "?P ?s'" unfolding id ‹a = LE dir x c›
by (intro impI atoms_equiv_bounds_extend[rule_format] *(1,3))
qed simp_all
}
then show "flat ats ≐ ℬ s ⟹ flat (ats ∪ {ia}) ≐ ℬ (assert_bound ia s)" unfolding ia by auto
next
fix s :: "('i,'a) state" and ats and ia :: "('i,'a) i_atom"
obtain i a where ia: "ia = (i,a)" by force
assume "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "△ (𝒯 s)" "∇ s"
have *: "⋀ dir x c s. dir = Positive ∨ dir = Negative ⟹
∇ (updateℬℐ (UBI_upd dir) i x c s) = ∇ s"
"⋀ s y I . ∇ (set_unsat I s) = ∇ s"
by (auto simp add: tableau_valuated_def)
show "∇ (assert_bound ia s)" (is "?P (assert_bound ia s)")
proof-
let ?P' = "λ lt UBI LBI UB LB UB_upd UI LI LE GE s'. ∇ s'"
let ?P'' = "λ dir. ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)"
show ?thesis unfolding ia
proof (rule assert_bound_cases[of _ _ ?P'])
fix x c and dir :: "('i,'a) Direction"
let ?s' = "updateℬℐ (UBI_upd dir) i x c s"
assume "dir = Positive ∨ dir = Negative"
then have "∇ ?s'"
using *(1)[of dir x c s] ‹∇ s›
by simp
then show "∇ (set_unsat [i, ((LI dir) s x)] ?s')"
using *(2) by auto
next
fix x c and dir :: "('i,'a) Direction"
assume *: "x ∉ lvars (𝒯 s)" "dir = Positive ∨ dir = Negative"
let ?s = "updateℬℐ (UBI_upd dir) i x c s"
let ?s' = "update x c ?s"
from * show "∇ ?s'"
using ‹△ (𝒯 s)› ‹∇ s›
using update_tableau_valuated[of ?s x c]
by (auto simp add: tableau_valuated_def)
qed (insert ‹∇ s› *(1), auto)
qed
next
fix s :: "('i,'a) state" and as and ia :: "('i,'a) i_atom"
obtain i a where ia: "ia = (i,a)" by force
assume *: "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "△ (𝒯 s)" "∇ s"
and valid: "index_valid as s"
have id: "⋀ dir x c s. dir = Positive ∨ dir = Negative ⟹
∇ (updateℬℐ (UBI_upd dir) i x c s) = ∇ s"
"⋀ s y I. ∇ (set_unsat I s) = ∇ s"
by (auto simp add: tableau_valuated_def)
let ?I = "insert (i,a) as"
define I where "I = ?I"
from index_valid_mono[OF _ valid] have valid: "index_valid I s" unfolding I_def by auto
have I: "(i,a) ∈ I" unfolding I_def by auto
let ?P = "λ s. index_valid I s"
let ?P' = "λ (lt :: 'a ⇒ 'a ⇒ bool)
(UBI :: ('i,'a) state ⇒ ('i,'a) bounds_index) (LBI :: ('i,'a) state ⇒ ('i,'a) bounds_index)
(UB :: ('i,'a) state ⇒ 'a bounds) (LB :: ('i,'a) state ⇒ 'a bounds)
(UBI_upd :: (('i,'a) bounds_index ⇒ ('i,'a) bounds_index) ⇒ ('i,'a) state ⇒ ('i,'a) state)
(UI :: ('i,'a) state ⇒ 'i bound_index) (LI :: ('i,'a) state ⇒ 'i bound_index)
LE GE s'.
(∀ x c i. look (UBI s') x = Some (i,c) ⟶ (i,LE (x :: var) c) ∈ I) ∧
(∀ x c i. look (LBI s') x = Some (i,c) ⟶ (i,GE (x :: nat) c) ∈ I)"
define P where "P = ?P'"
let ?P'' = "λ (dir :: ('i,'a) Direction).
P (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)"
have x: "⋀ s'. ?P s' = P (<) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l ℬ⇩i⇩u_update ℐ⇩u ℐ⇩l Leq Geq s'"
and xx: "⋀ s'. ?P s' = P (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℬ⇩i⇩l_update ℐ⇩l ℐ⇩u Geq Leq s'"
unfolding satisfies_bounds_set.simps in_bounds.simps bound_compare_defs index_valid_def P_def
by (auto split: option.split simp: indexl_def indexu_def boundsl_def boundsu_def)
from valid have P'': "dir = Positive ∨ dir = Negative ⟹ ?P'' dir s" for dir
using x[of s] xx[of s] by auto
have UTrue: "dir = Positive ∨ dir = Negative ⟹ ?P'' dir s ⟹ ?P'' dir (set_unsat I s)" for dir s I
unfolding P_def by (auto simp: boundsl_def boundsu_def indexl_def indexu_def)
have updateIB: "a = LE dir x c ⟹ dir = Positive ∨ dir = Negative ⟹ ?P'' dir s ⟹ ?P'' dir
(updateℬℐ (UBI_upd dir) i x c s)" for dir x c s
unfolding P_def using I by (auto split: if_splits simp: simp: boundsl_def boundsu_def indexl_def indexu_def)
show "index_valid (insert ia as) (assert_bound ia s)" unfolding ia I_def[symmetric]
proof ((rule assert_bound_cases[of _ _ P]; (intro UTrue x xx updateIB P'')?))
fix x c and dir :: "('i,'a) Direction"
assume **: "dir = Positive ∨ dir = Negative"
"a = LE dir x c"
"x ∉ lvars (𝒯 s)"
let ?s = "(updateℬℐ (UBI_upd dir) i x c s)"
define s' where "s' = ?s"
have 1: "△ (𝒯 ?s)" using * ** by auto
have 2: "∇ ?s" using id(1) ** * ‹∇ s› by auto
have 3: "x ∉ lvars (𝒯 ?s)" using id(1) ** * ‹∇ s› by auto
have "?P'' dir ?s" using ** by (intro updateIB P'') auto
with update_id[of ?s x c, OF 1 2 3, unfolded Let_def] **(1)
show "P (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)
(update x c (updateℬℐ (UBI_upd dir) i x c s))"
unfolding P_def s'_def[symmetric] by auto
qed auto
next
fix s and ia :: "('i,'a) i_atom" and ats :: "('i,'a) i_atom set"
assume *: "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "△ (𝒯 s)" "∇ s" "◇ s" and ats: "ats ⊨⇩i ℬℐ s"
obtain i a where ia: "ia = (i,a)" by force
have id: "⋀ dir x c s. dir = Positive ∨ dir = Negative ⟹
∇ (updateℬℐ (UBI_upd dir) i x c s) = ∇ s"
"⋀ s I. ∇ (set_unsat I s) = ∇ s"
by (auto simp add: tableau_valuated_def)
have idlt: "(c < (a :: 'a) ∨ c = a) = (c ≤ a)"
"(a < c ∨ c = a) = (c ≥ a)" for a c by auto
define A where "A = insert (i,a) ats"
let ?P = "λ (s :: ('i,'a) state). A ⊨⇩i ℬℐ s"
let ?Q = "λ bs (lt :: 'a ⇒ 'a ⇒ bool)
(UBI :: ('i,'a) state ⇒ ('i,'a) bounds_index) (LBI :: ('i,'a) state ⇒ ('i,'a) bounds_index)
(UB :: ('i,'a) state ⇒ 'a bounds) (LB :: ('i,'a) state ⇒ 'a bounds)
(UBI_upd :: (('i,'a) bounds_index ⇒ ('i,'a) bounds_index) ⇒ ('i,'a) state ⇒ ('i,'a) state)
UI LI
(LE :: nat ⇒ 'a ⇒ 'a atom) (GE :: nat ⇒ 'a ⇒ 'a atom) s'.
(∀ I v. (I :: 'i set,v) ⊨⇩i⇩a⇩s bs ⟶
((∀ x c. LB s' x = Some c ⟶ LI s' x ∈ I ⟶ lt c (v x) ∨ c = v x)
∧ (∀ x c. UB s' x = Some c ⟶ UI s' x ∈ I ⟶ lt (v x) c ∨ v x = c)))"
define Q where "Q = ?Q"
let ?P' = "Q A"
have equiv:
"bs ⊨⇩i ℬℐ s' ⟷ Q bs (<) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l ℬ⇩i⇩u_update ℐ⇩u ℐ⇩l Leq Geq s'"
"bs ⊨⇩i ℬℐ s' ⟷ Q bs (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℬ⇩i⇩l_update ℐ⇩l ℐ⇩u Geq Leq s'"
for bs s'
unfolding satisfies_bounds_set.simps in_bounds.simps bound_compare_defs index_valid_def Q_def
atoms_imply_bounds_index.simps
by (atomize(full), (intro conjI iff_exI allI arg_cong2[of _ _ _ _ "(∧)"] refl iff_allI
arg_cong2[of _ _ _ _ "(=)"]; unfold satisfies_bounds_index.simps idlt), auto)
have x: "⋀ s'. ?P s' = ?P' (<) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l ℬ⇩i⇩u_update ℐ⇩u ℐ⇩l Leq Geq s'"
and xx: "⋀ s'. ?P s' = ?P' (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℬ⇩i⇩l_update ℐ⇩l ℐ⇩u Geq Leq s'"
using equiv by blast+
from ats equiv[of ats s]
have Q_ats:
"Q ats (<) ℬ⇩i⇩u ℬ⇩i⇩l ℬ⇩u ℬ⇩l ℬ⇩i⇩u_update ℐ⇩u ℐ⇩l Leq Geq s"
"Q ats (>) ℬ⇩i⇩l ℬ⇩i⇩u ℬ⇩l ℬ⇩u ℬ⇩i⇩l_update ℐ⇩l ℐ⇩u Geq Leq s"
by auto
let ?P'' = "λ (dir :: ('i,'a) Direction). ?P' (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir)"
have P_upd: "dir = Positive ∨ dir = Negative ⟹ ?P'' dir (set_unsat I s) = ?P'' dir s" for s I dir
unfolding Q_def
by (intro iff_exI arg_cong2[of _ _ _ _ "(∧)"] refl iff_allI arg_cong2[of _ _ _ _ "(=)"]
arg_cong2[of _ _ _ _ "(⟶)"], auto simp: boundsl_def boundsu_def indexl_def indexu_def)
have P_upd: "dir = Positive ∨ dir = Negative ⟹ ?P'' dir s ⟹ ?P'' dir (set_unsat I s)" for s I dir
using P_upd[of dir] by blast
have ats_sub: "ats ⊆ A" unfolding A_def by auto
{
fix x c and dir :: "('i,'a) Direction"
assume dir: "dir = Positive ∨ dir = Negative"
and a: "a = LE dir x c"
from Q_ats dir
have Q_ats: "Q ats (lt dir) (UBI dir) (LBI dir) (UB dir) (LB dir) (UBI_upd dir) (UI dir) (LI dir) (LE dir) (GE dir) s"
by auto
have "?P'' dir (updateℬℐ (UBI_upd dir) i x c s)"
unfolding Q_def
proof (intro allI impI conjI)
fix I v y d
assume IvA: "(I, v) ⊨⇩i⇩a⇩s A"
from i_satisfies_atom_set_mono[OF ats_sub this]
have "(I, v) ⊨⇩i⇩a⇩s ats" by auto
from Q_ats[unfolded Q_def, rule_format, OF this]
have s_bnds:
"LB dir s x = Some c ⟹ LI dir s x ∈ I ⟹ lt dir c (v x) ∨ c = v x"
"UB dir s x = Some c ⟹ UI dir s x ∈ I ⟹ lt dir (v x) c ∨ v x = c" for x c by auto
from IvA[unfolded A_def, unfolded i_satisfies_atom_set.simps satisfies_atom_set_def, simplified]
have va: "i ∈ I ⟹ v ⊨⇩a a" by auto
with a dir have vc: "i ∈ I ⟹ lt dir (v x) c ∨ v x = c"
by auto
let ?s = "(updateℬℐ (UBI_upd dir) i x c s)"
show "LB dir ?s y = Some d ⟹ LI dir ?s y ∈ I ⟹ lt dir d (v y) ∨ d = v y"
"UB dir ?s y = Some d ⟹ UI dir ?s y ∈ I ⟹ lt dir (v y) d ∨ v y = d"
proof (atomize(full), goal_cases)
case 1
consider (main) "y = x" "UI dir ?s x = i" |
(easy1) "x ≠ y" | (easy2) "x = y" "UI dir ?s y ≠ i"
by blast
then show ?case
proof cases
case easy1
then show ?thesis using s_bnds[of y d] dir by (fastforce simp: boundsl_def boundsu_def indexl_def indexu_def)
next
case easy2
then show ?thesis using s_bnds[of y d] dir by (fastforce simp: boundsl_def boundsu_def indexl_def indexu_def)
next
case main
note s_bnds = s_bnds[of x]
show ?thesis unfolding main using s_bnds dir vc by (auto simp: boundsl_def boundsu_def indexl_def indexu_def) blast+
qed
qed
qed
} note main = this
have Ps: "dir = Positive ∨ dir = Negative ⟹ ?P'' dir s" for dir
using Q_ats unfolding Q_def using i_satisfies_atom_set_mono[OF ats_sub] by fastforce
have "?P (assert_bound (i,a) s)"
proof ((rule assert_bound_cases[of _ _ ?P']; (intro x xx P_upd main Ps)?))
fix c x and dir :: "('i,'a) Direction"
assume **: "dir = Positive ∨ dir = Negative"
"a = LE dir x c"
"x ∉ lvars (𝒯 s)"
let ?s = "updateℬℐ (UBI_upd dir) i x c s"
define s' where "s' = ?s"
from main[OF **(1-2)] have P: "?P'' dir s'" unfolding s'_def .
have 1: "△ (𝒯 ?s)" using * ** by auto
have 2: "∇ ?s" using id(1) ** * ‹∇ s› by auto
have 3: "x ∉ lvars (𝒯 ?s)" using id(1) ** * ‹∇ s› by auto
have "△ (𝒯 s')" "∇ s'" "x ∉ lvars (𝒯 s')" using 1 2 3 unfolding s'_def by auto
from update_bounds_id[OF this, of c] **(1)
have "?P'' dir (update x c s') = ?P'' dir s'"
unfolding Q_def
by (intro iff_allI arg_cong2[of _ _ _ _ "(⟶)"] arg_cong2[of _ _ _ _ "(∧)"] refl, auto)
with P
show "?P'' dir (update x c ?s)" unfolding s'_def by blast
qed auto
then show "insert ia ats ⊨⇩i ℬℐ (assert_bound ia s)" unfolding ia A_def by blast
qed
text ‹Pivoting the tableau can be reduced to pivoting single equations,
and substituting variable by polynomials. These operations are specified
by:›
locale PivotEq =
fixes pivot_eq::"eq ⇒ var ⇒ eq"
assumes
vars_pivot_eq:"
⟦x⇩j ∈ rvars_eq eq; lhs eq ∉ rvars_eq eq ⟧ ⟹ let eq' = pivot_eq eq x⇩j in
lhs eq' = x⇩j ∧ rvars_eq eq' = {lhs eq} ∪ (rvars_eq eq - {x⇩j})" and
equiv_pivot_eq:
"⟦x⇩j ∈ rvars_eq eq; lhs eq ∉ rvars_eq eq ⟧ ⟹
(v::'a::lrv valuation) ⊨⇩e pivot_eq eq x⇩j ⟷ v ⊨⇩e eq"
begin
lemma lhs_pivot_eq:
"⟦x⇩j ∈ rvars_eq eq; lhs eq ∉ rvars_eq eq ⟧ ⟹ lhs (pivot_eq eq x⇩j) = x⇩j"
using vars_pivot_eq
by (simp add: Let_def)
lemma rvars_pivot_eq:
"⟦x⇩j ∈ rvars_eq eq; lhs eq ∉ rvars_eq eq ⟧ ⟹ rvars_eq (pivot_eq eq x⇩j) = {lhs eq} ∪ (rvars_eq eq - {x⇩j})"
using vars_pivot_eq
by (simp add: Let_def)
end
abbreviation doublesub (" _ ⊆s _ ⊆s _" [50,51,51] 50) where
"doublesub a b c ≡ a ⊆ b ∧ b ⊆ c"
locale SubstVar =
fixes subst_var::"var ⇒ linear_poly ⇒ linear_poly ⇒ linear_poly"
assumes
vars_subst_var':
"(vars lp - {x⇩j}) - vars lp' ⊆s vars (subst_var x⇩j lp' lp) ⊆s (vars lp - {x⇩j}) ∪ vars lp'"and
subst_no_effect: "x⇩j ∉ vars lp ⟹ subst_var x⇩j lp' lp = lp" and
subst_with_effect: "x⇩j ∈ vars lp ⟹ x ∈ vars lp' - vars lp ⟹ x ∈ vars (subst_var x⇩j lp' lp)" and
equiv_subst_var:
"(v::'a :: lrv valuation) x⇩j = lp' ⦃v⦄ ⟶ lp ⦃v⦄ = (subst_var x⇩j lp' lp) ⦃v⦄"
begin
lemma vars_subst_var:
"vars (subst_var x⇩j lp' lp) ⊆ (vars lp - {x⇩j}) ∪ vars lp'"
using vars_subst_var'
by simp
lemma vars_subst_var_supset:
"vars (subst_var x⇩j lp' lp) ⊇ (vars lp - {x⇩j}) - vars lp'"
using vars_subst_var'
by simp
definition subst_var_eq :: "var ⇒ linear_poly ⇒ eq ⇒ eq" where
"subst_var_eq v lp' eq ≡ (lhs eq, subst_var v lp' (rhs eq))"
lemma rvars_eq_subst_var_eq:
shows "rvars_eq (subst_var_eq x⇩j lp eq) ⊆ (rvars_eq eq - {x⇩j}) ∪ vars lp"
unfolding subst_var_eq_def
by (auto simp add: vars_subst_var)
lemma rvars_eq_subst_var_eq_supset:
"rvars_eq (subst_var_eq x⇩j lp eq) ⊇ (rvars_eq eq) - {x⇩j} - (vars lp)"
unfolding subst_var_eq_def
by (simp add: vars_subst_var_supset)
lemma equiv_subst_var_eq:
assumes "(v::'a valuation) ⊨⇩e (x⇩j, lp')"
shows "v ⊨⇩e eq ⟷ v ⊨⇩e subst_var_eq x⇩j lp' eq"
using assms
unfolding subst_var_eq_def
unfolding satisfies_eq_def
using equiv_subst_var[of v x⇩j lp' "rhs eq"]
by auto
end
locale Pivot' = EqForLVar + PivotEq + SubstVar
begin
definition pivot_tableau' :: "var ⇒ var ⇒ tableau ⇒ tableau" where
"pivot_tableau' x⇩i x⇩j t ≡
let x⇩i_idx = eq_idx_for_lvar t x⇩i; eq = t ! x⇩i_idx; eq' = pivot_eq eq x⇩j in
map (λ idx. if idx = x⇩i_idx then
eq'
else
subst_var_eq x⇩j (rhs eq') (t ! idx)
) [0..<length t]"
definition pivot' :: "var ⇒ var ⇒ ('i,'a::lrv) state ⇒ ('i,'a) state" where
"pivot' x⇩i x⇩j s ≡ 𝒯_update (pivot_tableau' x⇩i x⇩j (𝒯 s)) s"
text‹\noindent Then, the next implementation of ‹pivot› satisfies its specification:›
definition pivot_tableau :: "var ⇒ var ⇒ tableau ⇒ tableau" where
"pivot_tableau x⇩i x⇩j t ≡ let eq = eq_for_lvar t x⇩i; eq' = pivot_eq eq x⇩j in
map (λ e. if lhs e = lhs eq then eq' else subst_var_eq x⇩j (rhs eq') e) t"
definition pivot :: "var ⇒ var ⇒ ('i,'a::lrv) state ⇒ ('i,'a) state" where
"pivot x⇩i x⇩j s ≡ 𝒯_update (pivot_tableau x⇩i x⇩j (𝒯 s)) s"
lemma pivot_tableau'pivot_tableau:
assumes "△ t" "x⇩i ∈ lvars t"
shows "pivot_tableau' x⇩i x⇩j t = pivot_tableau x⇩i x⇩j t"
proof-
let ?f = "λidx. if idx = eq_idx_for_lvar t x⇩i then pivot_eq (t ! eq_idx_for_lvar t x⇩i) x⇩j
else subst_var_eq x⇩j (rhs (pivot_eq (t ! eq_idx_for_lvar t x⇩i) x⇩j)) (t ! idx)"
let ?f' = "λe. if lhs e = lhs (eq_for_lvar t x⇩i) then pivot_eq (eq_for_lvar t x⇩i) x⇩j else subst_var_eq x⇩j (rhs (pivot_eq (eq_for_lvar t x⇩i) x⇩j)) e"
have "∀ i < length t. ?f' (t ! i) = ?f i"
proof(safe)
fix i
assume "i < length t"
then have "t ! i ∈ set t" "i < length t"
by auto
moreover
have "t ! eq_idx_for_lvar t x⇩i ∈ set t" "eq_idx_for_lvar t x⇩i < length t"
using eq_for_lvar[of x⇩i t] ‹x⇩i ∈ lvars t› eq_idx_for_lvar[of x⇩i t]
by (auto simp add: eq_for_lvar_def)
ultimately
have "lhs (t ! i) = lhs (t ! eq_idx_for_lvar t x⇩i) ⟹ t ! i = t ! (eq_idx_for_lvar t x⇩i)" "distinct t"
using ‹△ t›
unfolding normalized_tableau_def
by (auto simp add: distinct_map inj_on_def)
then have "lhs (t ! i) = lhs (t ! eq_idx_for_lvar t x⇩i) ⟹ i = eq_idx_for_lvar t x⇩i"
using ‹i < length t› ‹eq_idx_for_lvar t x⇩i < length t›
by (auto simp add: distinct_conv_nth)
then show "?f' (t ! i) = ?f i"
by (auto simp add: eq_for_lvar_def)
qed
then show "pivot_tableau' x⇩i x⇩j t = pivot_tableau x⇩i x⇩j t"
unfolding pivot_tableau'_def pivot_tableau_def
unfolding Let_def
by (auto simp add: map_reindex)
qed
lemma pivot'pivot: fixes s :: "('i,'a::lrv)state"
assumes "△ (𝒯 s)" "x⇩i ∈ lvars (𝒯 s)"
shows "pivot' x⇩i x⇩j s = pivot x⇩i x⇩j s"
using pivot_tableau'pivot_tableau[OF assms]
unfolding pivot_def pivot'_def by auto
end
sublocale Pivot' < Pivot eq_idx_for_lvar pivot
proof
fix s::"('i,'a) state" and x⇩i x⇩j and v::"'a valuation"
assume "△ (𝒯 s)" "x⇩i ∈ lvars (𝒯 s)"
"x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)"
show "let s' = pivot x⇩i x⇩j s in 𝒱 s' = 𝒱 s ∧ ℬ⇩i s' = ℬ⇩i s ∧ 𝒰 s' = 𝒰 s ∧ 𝒰⇩c s' = 𝒰⇩c s"
unfolding pivot_def
by (auto simp add: Let_def simp: boundsl_def boundsu_def indexl_def indexu_def)
let ?t = "𝒯 s"
let ?idx = "eq_idx_for_lvar ?t x⇩i"
let ?eq = "?t ! ?idx"
let ?eq' = "pivot_eq ?eq x⇩j"
have "?idx < length ?t" "lhs (?t ! ?idx) = x⇩i"
using ‹x⇩i ∈ lvars ?t›
using eq_idx_for_lvar
by auto
have "distinct (map lhs ?t)"
using ‹△ ?t›
unfolding normalized_tableau_def
by simp
have "x⇩j ∈ rvars_eq ?eq"
using ‹x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)›
unfolding eq_for_lvar_def
by simp
then have "x⇩j ∈ rvars ?t"
using ‹?idx < length ?t›
using in_set_conv_nth[of ?eq ?t]
by (auto simp add: rvars_def)
then have "x⇩j ∉ lvars ?t"
using ‹△ ?t›
unfolding normalized_tableau_def
by auto
have "x⇩i ∉ rvars ?t"
using ‹x⇩i ∈ lvars ?t› ‹△ ?t›
unfolding normalized_tableau_def rvars_def
by auto
then have "x⇩i ∉ rvars_eq ?eq"
unfolding rvars_def
using ‹?idx < length ?t›
using in_set_conv_nth[of ?eq ?t]
by auto
have "x⇩i ≠ x⇩j"
using ‹x⇩j ∈ rvars_eq ?eq› ‹x⇩i ∉ rvars_eq ?eq›
by auto
have "?eq' = (x⇩j, rhs ?eq')"
using lhs_pivot_eq[of x⇩j ?eq]
using ‹x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)› ‹lhs (?t ! ?idx) = x⇩i› ‹x⇩i ∉ rvars_eq ?eq›
by (auto simp add: eq_for_lvar_def) (cases "?eq'", simp)+
let ?I1 = "[0..<?idx]"
let ?I2 = "[?idx + 1..<length ?t]"
have "[0..<length ?t] = ?I1 @ [?idx] @ ?I2"
using ‹?idx < length ?t›
by (rule interval_3split)
then have map_lhs_pivot:
"map lhs (𝒯 (pivot' x⇩i x⇩j s)) =
map (λidx. lhs (?t ! idx)) ?I1 @ [x⇩j] @ map (λidx. lhs (?t ! idx)) ?I2"
using ‹x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)› ‹lhs (?t ! ?idx) = x⇩i› ‹x⇩i ∉ rvars_eq ?eq›
by (auto simp add: Let_def subst_var_eq_def eq_for_lvar_def lhs_pivot_eq pivot'_def pivot_tableau'_def)
have lvars_pivot: "lvars (𝒯 (pivot' x⇩i x⇩j s)) =
lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
proof-
have "lvars (𝒯 (pivot' x⇩i x⇩j s)) =
{x⇩j} ∪ (λidx. lhs (?t ! idx)) ` ({0..<length?t} - {?idx})"
using ‹?idx < length ?t› ‹?eq' = (x⇩j, rhs ?eq')›
by (cases ?eq', auto simp add: Let_def pivot'_def pivot_tableau'_def lvars_def subst_var_eq_def)+
also have "... = {x⇩j} ∪ (((λidx. lhs (?t ! idx)) ` {0..<length?t}) - {lhs (?t ! ?idx)})"
using ‹?idx < length ?t› ‹distinct (map lhs ?t)›
by (auto simp add: distinct_conv_nth)
also have "... = {x⇩j} ∪ (set (map lhs ?t) - {x⇩i})"
using ‹lhs (?t ! ?idx) = x⇩i›
by (auto simp add: in_set_conv_nth rev_image_eqI) (auto simp add: image_def)
finally show "lvars (𝒯 (pivot' x⇩i x⇩j s)) =
lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
by (simp add: lvars_def)
qed
moreover
have rvars_pivot: "rvars (𝒯 (pivot' x⇩i x⇩j s)) =
rvars (𝒯 s) - {x⇩j} ∪ {x⇩i}"
proof-
have "rvars_eq ?eq' = {x⇩i} ∪ (rvars_eq ?eq - {x⇩j})"
using rvars_pivot_eq[of x⇩j ?eq]
using ‹lhs (?t ! ?idx) = x⇩i›
using ‹x⇩j ∈ rvars_eq ?eq› ‹x⇩i ∉ rvars_eq ?eq›
by simp
let ?S1 = "rvars_eq ?eq'"
let ?S2 = "⋃idx∈({0..<length ?t} - {?idx}).
rvars_eq (subst_var_eq x⇩j (rhs ?eq') (?t ! idx))"
have "rvars (𝒯 (pivot' x⇩i x⇩j s)) = ?S1 ∪ ?S2"
unfolding pivot'_def pivot_tableau'_def rvars_def
using ‹?idx < length ?t›
by (auto simp add: Let_def split: if_splits)
also have "... = {x⇩i} ∪ (rvars ?t - {x⇩j})" (is "?S1 ∪ ?S2 = ?rhs")
proof
show "?S1 ∪ ?S2 ⊆ ?rhs"
proof-
have "?S1 ⊆ ?rhs"
using ‹?idx < length ?t›
unfolding rvars_def
using ‹rvars_eq ?eq' = {x⇩i} ∪ (rvars_eq ?eq - {x⇩j})›
by (force simp add: in_set_conv_nth)
moreover
have "?S2 ⊆ ?rhs"
proof-
have "?S2 ⊆ (⋃idx∈{0..<length ?t}. (rvars_eq (?t ! idx) - {x⇩j}) ∪ rvars_eq ?eq')"
apply (rule UN_mono)
using rvars_eq_subst_var_eq
by auto
also have "... ⊆ rvars_eq ?eq' ∪ (⋃idx∈{0..<length ?t}. rvars_eq (?t ! idx) - {x⇩j})"
by auto
also have "... = rvars_eq ?eq' ∪ (rvars ?t - {x⇩j})"
unfolding rvars_def
by (force simp add: in_set_conv_nth)
finally show ?thesis
using ‹rvars_eq ?eq' = {x⇩i} ∪ (rvars_eq ?eq - {x⇩j})›
unfolding rvars_def
using ‹?idx < length ?t›
by auto
qed
ultimately
show ?thesis
by simp
qed
next
show "?rhs ⊆ ?S1 ∪ ?S2"
proof
fix x
assume "x ∈ ?rhs"
show "x ∈ ?S1 ∪ ?S2"
proof (cases "x ∈ rvars_eq ?eq'")
case True
then show ?thesis
by auto
next
case False
let ?S2' = "⋃idx∈({0..<length ?t} - {?idx}).
(rvars_eq (?t ! idx) - {x⇩j}) - rvars_eq ?eq'"
have "x ∈ ?S2'"
using False ‹x ∈ ?rhs›
using ‹rvars_eq ?eq' = {x⇩i} ∪ (rvars_eq ?eq - {x⇩j})›
unfolding rvars_def
by (force simp add: in_set_conv_nth)
moreover
have "?S2 ⊇ ?S2'"
apply (rule UN_mono)
using rvars_eq_subst_var_eq_supset[of _ x⇩j "rhs ?eq'" ]
by auto
ultimately
show ?thesis
by auto
qed
qed
qed
ultimately
show ?thesis
by simp
qed
ultimately
show "let s' = pivot x⇩i x⇩j s in rvars (𝒯 s') = rvars (𝒯 s) - {x⇩j} ∪ {x⇩i} ∧ lvars (𝒯 s') = lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
using pivot'pivot[where ?'i = 'i]
using ‹△ (𝒯 s)› ‹x⇩i ∈ lvars (𝒯 s)›
by (simp add: Let_def)
have "△ (𝒯 (pivot' x⇩i x⇩j s))"
unfolding normalized_tableau_def
proof
have "lvars (𝒯 (pivot' x⇩i x⇩j s)) ∩ rvars (𝒯 (pivot' x⇩i x⇩j s)) = {}" (is ?g1)
using ‹△ (𝒯 s)›
unfolding normalized_tableau_def
using lvars_pivot rvars_pivot
using ‹x⇩i ≠ x⇩j›
by auto
moreover have "0 ∉ rhs ` set (𝒯 (pivot' x⇩i x⇩j s))" (is ?g2)
proof
let ?eq = "eq_for_lvar (𝒯 s) x⇩i"
from eq_for_lvar[OF ‹x⇩i ∈ lvars (𝒯 s)›]
have "?eq ∈ set (𝒯 s)" and var: "lhs ?eq = x⇩i" by auto
have "lhs ?eq ∉ rvars_eq ?eq" using ‹△ (𝒯 s)› ‹?eq ∈ set (𝒯 s)›
using ‹x⇩i ∉ rvars_eq (𝒯 s ! eq_idx_for_lvar (𝒯 s) x⇩i)› eq_for_lvar_def var by auto
from vars_pivot_eq[OF ‹x⇩j ∈ rvars_eq ?eq› this]
have vars_pivot: "lhs (pivot_eq ?eq x⇩j) = x⇩j" "rvars_eq (pivot_eq ?eq x⇩j) = {lhs (eq_for_lvar (𝒯 s) x⇩i)} ∪ (rvars_eq (eq_for_lvar (𝒯 s) x⇩i) - {x⇩j})"
unfolding Let_def by auto
from vars_pivot(2) have rhs_pivot0: "rhs (pivot_eq ?eq x⇩j) ≠ 0" using vars_zero by auto
assume "0 ∈ rhs ` set (𝒯 (pivot' x⇩i x⇩j s))"
from this[unfolded pivot'pivot[OF ‹△ (𝒯 s)› ‹x⇩i ∈ lvars (𝒯 s)›] pivot_def]
have "0 ∈ rhs ` set (pivot_tableau x⇩i x⇩j (𝒯 s))" by simp
from this[unfolded pivot_tableau_def Let_def var, unfolded var] rhs_pivot0
obtain e where "e ∈ set (𝒯 s)" "lhs e ≠ x⇩i" and rvars_eq: "rvars_eq (subst_var_eq x⇩j (rhs (pivot_eq ?eq x⇩j)) e) = {}"
by (auto simp: vars_zero)
from rvars_eq[unfolded subst_var_eq_def]
have empty: "vars (subst_var x⇩j (rhs (pivot_eq ?eq x⇩j)) (rhs e)) = {}" by auto
show False
proof (cases "x⇩j ∈ vars (rhs e)")
case False
from empty[unfolded subst_no_effect[OF False]]
have "rvars_eq e = {}" by auto
hence "rhs e = 0" using zero_coeff_zero coeff_zero by auto
with ‹e ∈ set (𝒯 s)› ‹△ (𝒯 s)› show False unfolding normalized_tableau_def by auto
next
case True
from ‹e ∈ set (𝒯 s)› have "rvars_eq e ⊆ rvars (𝒯 s)" unfolding rvars_def by auto
hence "x⇩i ∈ vars (rhs (pivot_eq ?eq x⇩j)) - rvars_eq e"
unfolding vars_pivot(2) var
using ‹△ (𝒯 s)›[unfolded normalized_tableau_def] ‹x⇩i ∈ lvars (𝒯 s)› by auto
from subst_with_effect[OF True this] rvars_eq
show ?thesis by (simp add: subst_var_eq_def)
qed
qed
ultimately show "?g1 ∧ ?g2" ..
show "distinct (map lhs (𝒯 (pivot' x⇩i x⇩j s)))"
using map_parametrize_idx[of lhs ?t]
using map_lhs_pivot
using ‹distinct (map lhs ?t)›
using interval_3split[of ?idx "length ?t"] ‹?idx < length ?t›
using ‹x⇩j ∉ lvars ?t›
unfolding lvars_def
by auto
qed
moreover
have "v ⊨⇩t ?t = v ⊨⇩t 𝒯 (pivot' x⇩i x⇩j s)"
unfolding satisfies_tableau_def
proof
assume "∀e∈set (?t). v ⊨⇩e e"
show "∀e∈set (𝒯 (pivot' x⇩i x⇩j s)). v ⊨⇩e e"
proof-
have "v ⊨⇩e ?eq'"
using ‹x⇩i ∉ rvars_eq ?eq›
using ‹?idx < length ?t› ‹∀e∈set (?t). v ⊨⇩e e›
using ‹x⇩j ∈ rvars_eq ?eq› ‹x⇩i ∈ lvars ?t›
by (simp add: equiv_pivot_eq eq_idx_for_lvar)
moreover
{
fix idx
assume "idx < length ?t" "idx ≠ ?idx"
have "v ⊨⇩e subst_var_eq x⇩j (rhs ?eq') (?t ! idx)"
using ‹?eq' = (x⇩j, rhs ?eq')›
using ‹v ⊨⇩e ?eq'› ‹idx < length ?t› ‹∀e∈set (?t). v ⊨⇩e e›
using equiv_subst_var_eq[of v x⇩j "rhs ?eq'" "?t ! idx"]
by auto
}
ultimately
show ?thesis
by (auto simp add: pivot'_def pivot_tableau'_def Let_def)
qed
next
assume "∀e∈set (𝒯 (pivot' x⇩i x⇩j s)). v ⊨⇩e e"
then have "v ⊨⇩e ?eq'"
"⋀ idx. ⟦idx < length ?t; idx ≠ ?idx ⟧ ⟹ v ⊨⇩e subst_var_eq x⇩j (rhs ?eq') (?t ! idx)"
using ‹?idx < length ?t›
unfolding pivot'_def pivot_tableau'_def
by (auto simp add: Let_def)
show "∀e∈set (𝒯 s). v ⊨⇩e e"
proof-
{
fix idx
assume "idx < length ?t"
have "v ⊨⇩e (?t ! idx)"
proof (cases "idx = ?idx")
case True
then show ?thesis
using ‹v ⊨⇩e ?eq'›
using ‹x⇩j ∈ rvars_eq ?eq› ‹x⇩i ∈ lvars ?t› ‹x⇩i ∉ rvars_eq ?eq›
by (simp add: eq_idx_for_lvar equiv_pivot_eq)
next
case False
then show ?thesis
using ‹idx < length ?t›
using ‹⟦idx < length ?t; idx ≠ ?idx ⟧ ⟹ v ⊨⇩e subst_var_eq x⇩j (rhs ?eq') (?t ! idx)›
using ‹v ⊨⇩e ?eq'› ‹?eq' = (x⇩j, rhs ?eq')›
using equiv_subst_var_eq[of v x⇩j "rhs ?eq'" "?t ! idx"]
by auto
qed
}
then show ?thesis
by (force simp add: in_set_conv_nth)
qed
qed
ultimately
show "let s' = pivot x⇩i x⇩j s in v ⊨⇩t 𝒯 s = v ⊨⇩t 𝒯 s' ∧ △ (𝒯 s')"
using pivot'pivot[where ?'i = 'i]
using ‹△ (𝒯 s)› ‹x⇩i ∈ lvars (𝒯 s)›
by (simp add: Let_def)
qed
subsection‹Check implementation›
text‹The ‹check› function is called when all rhs variables are
in bounds, and it checks if there is a lhs variable that is not. If
there is no such variable, then satisfiability is detected and ‹check› succeeds. If there is a lhs variable ‹x⇩i› out of its
bounds, a rhs variable ‹x⇩j› is sought which allows pivoting
with ‹x⇩i› and updating ‹x⇩i› to its violated bound. If
‹x⇩i› is under its lower bound it must be increased, and if
‹x⇩j› has a positive coefficient it must be increased so it
must be under its upper bound and if it has a negative coefficient it
must be decreased so it must be above its lower bound. The case when
‹x⇩i› is above its upper bound is symmetric (avoiding
symmetries is discussed in Section \ref{sec:symmetries}). If there is
no such ‹x⇩j›, unsatisfiability is detected and ‹check›
fails. The procedure is recursively repeated, until it either succeeds
or fails. To ensure termination, variables ‹x⇩i› and ‹x⇩j› must be chosen with respect to a fixed variable ordering. For
choosing these variables auxiliary functions ‹min_lvar_not_in_bounds›, ‹min_rvar_inc› and ‹min_rvar_dec› are specified (each in its own locale). For, example:
›
locale MinLVarNotInBounds = fixes min_lvar_not_in_bounds::"('i,'a::lrv) state ⇒ var option"
assumes
min_lvar_not_in_bounds_None: "min_lvar_not_in_bounds s = None ⟶ (∀x∈lvars (𝒯 s). in_bounds x ⟨𝒱 s⟩ (ℬ s))" and
min_lvar_not_in_bounds_Some': "min_lvar_not_in_bounds s = Some x⇩i ⟶ x⇩i∈lvars (𝒯 s) ∧ ¬in_bounds x⇩i ⟨𝒱 s⟩ (ℬ s)
∧ (∀x∈lvars (𝒯 s). x < x⇩i ⟶ in_bounds x ⟨𝒱 s⟩ (ℬ s))"
begin
lemma min_lvar_not_in_bounds_None':
"min_lvar_not_in_bounds s = None ⟶ (⟨𝒱 s⟩ ⊨⇩b ℬ s ∥ lvars (𝒯 s))"
unfolding satisfies_bounds_set.simps
by (rule min_lvar_not_in_bounds_None)
lemma min_lvar_not_in_bounds_lvars:"min_lvar_not_in_bounds s = Some x⇩i ⟶ x⇩i ∈ lvars (𝒯 s)"
using min_lvar_not_in_bounds_Some'
by simp
lemma min_lvar_not_in_bounds_Some: "min_lvar_not_in_bounds s = Some x⇩i ⟶ ¬ in_bounds x⇩i ⟨𝒱 s⟩ (ℬ s)"
using min_lvar_not_in_bounds_Some'
by simp
lemma min_lvar_not_in_bounds_Some_min: "min_lvar_not_in_bounds s = Some x⇩i ⟶ (∀ x ∈ lvars (𝒯 s). x < x⇩i ⟶ in_bounds x ⟨𝒱 s⟩ (ℬ s))"
using min_lvar_not_in_bounds_Some'
by simp
end
abbreviation reasable_var where
"reasable_var dir x eq s ≡
(coeff (rhs eq) x > 0 ∧ ⊲⇩u⇩b (lt dir) (⟨𝒱 s⟩ x) (UB dir s x)) ∨
(coeff (rhs eq) x < 0 ∧ ⊳⇩l⇩b (lt dir) (⟨𝒱 s⟩ x) (LB dir s x))"
locale MinRVarsEq =
fixes min_rvar_incdec_eq :: "('i,'a) Direction ⇒ ('i,'a::lrv) state ⇒ eq ⇒ 'i list + var"
assumes min_rvar_incdec_eq_None:
"min_rvar_incdec_eq dir s eq = Inl is ⟹
(∀ x ∈ rvars_eq eq. ¬ reasable_var dir x eq s) ∧
(set is = {LI dir s (lhs eq)} ∪ {LI dir s x | x. x ∈ rvars_eq eq ∧ coeff (rhs eq) x < 0}
∪ {UI dir s x | x. x ∈ rvars_eq eq ∧ coeff (rhs eq) x > 0}) ∧
((dir = Positive ∨ dir = Negative) ⟶ LI dir s (lhs eq) ∈ indices_state s ⟶ set is ⊆ indices_state s)"
assumes min_rvar_incdec_eq_Some_rvars:
"min_rvar_incdec_eq dir s eq = Inr x⇩j ⟹ x⇩j ∈ rvars_eq eq"
assumes min_rvar_incdec_eq_Some_incdec:
"min_rvar_incdec_eq dir s eq = Inr x⇩j ⟹ reasable_var dir x⇩j eq s"
assumes min_rvar_incdec_eq_Some_min:
"min_rvar_incdec_eq dir s eq = Inr x⇩j ⟹
(∀ x ∈ rvars_eq eq. x < x⇩j ⟶ ¬ reasable_var dir x eq s)"
begin
lemma min_rvar_incdec_eq_None':
assumes *: "dir = Positive ∨ dir = Negative"
and min: "min_rvar_incdec_eq dir s eq = Inl is"
and sub: "I = set is"
and Iv: "(I,v) ⊨⇩i⇩b ℬℐ s"
shows "le (lt dir) ((rhs eq) ⦃v⦄) ((rhs eq) ⦃⟨𝒱 s⟩⦄)"
proof -
have "∀ x ∈ rvars_eq eq. ¬ reasable_var dir x eq s"
using min
using min_rvar_incdec_eq_None
by simp
have "∀ x ∈ rvars_eq eq. (0 < coeff (rhs eq) x ⟶ le (lt dir) 0 (⟨𝒱 s⟩ x - v x)) ∧ (coeff (rhs eq) x < 0 ⟶ le (lt dir) (⟨𝒱 s⟩ x - v x) 0)"
proof (safe)
fix x
assume x: "x ∈ rvars_eq eq" "0 < coeff (rhs eq) x" "0 ≠ ⟨𝒱 s⟩ x - v x"
then have "¬ (⊲⇩u⇩b (lt dir) (⟨𝒱 s⟩ x) (UB dir s x))"
using ‹∀ x ∈ rvars_eq eq. ¬ reasable_var dir x eq s›
by auto
then have "⊵⇩u⇩b (lt dir) (⟨𝒱 s⟩ x) (UB dir s x)"
using *
by (cases "UB dir s x") (auto simp add: bound_compare_defs)
moreover
from min_rvar_incdec_eq_None[OF min] x sub have "UI dir s x ∈ I" by auto
from Iv * this
have "⊴⇩u⇩b (lt dir) (v x) (UB dir s x)"
unfolding satisfies_bounds_index.simps
by (cases "UB dir s x", auto simp: indexl_def indexu_def boundsl_def boundsu_def bound_compare'_defs)
(fastforce)+
ultimately
have "le (lt dir) (v x) (⟨𝒱 s⟩ x)"
using *
by (cases "UB dir s x") (auto simp add: bound_compare_defs)
then show "lt dir 0 (⟨𝒱 s⟩ x - v x)"
using ‹0 ≠ ⟨𝒱 s⟩ x - v x› *
using minus_gt[of "v x" "⟨𝒱 s⟩ x"] minus_lt[of "⟨𝒱 s⟩ x" "v x"]
by auto
next
fix x
assume x: "x ∈ rvars_eq eq" "0 > coeff (rhs eq) x" "⟨𝒱 s⟩ x - v x ≠ 0"
then have "¬ (⊳⇩l⇩b (lt dir) (⟨𝒱 s⟩ x) (LB dir s x))"
using ‹∀ x ∈ rvars_eq eq. ¬ reasable_var dir x eq s›
by auto
then have "⊴⇩l⇩b (lt dir) (⟨𝒱 s⟩ x) (LB dir s x)"
using *
by (cases "LB dir s x") (auto simp add: bound_compare_defs)
moreover
from min_rvar_incdec_eq_None[OF min] x sub have "LI dir s x ∈ I" by auto
from Iv * this
have "⊵⇩l⇩b (lt dir) (v x) (LB dir s x)"
unfolding satisfies_bounds_index.simps
by (cases "LB dir s x", auto simp: indexl_def indexu_def boundsl_def boundsu_def bound_compare'_defs)
(fastforce)+
ultimately
have "le (lt dir) (⟨𝒱 s⟩ x) (v x)"
using *
by (cases "LB dir s x") (auto simp add: bound_compare_defs)
then show "lt dir (⟨𝒱 s⟩ x - v x) 0"
using ‹⟨𝒱 s⟩ x - v x ≠ 0› *
using minus_lt[of "⟨𝒱 s⟩ x" "v x"] minus_gt[of "v x" "⟨𝒱 s⟩ x"]
by auto
qed
then have "le (lt dir) 0 (rhs eq ⦃ λ x. ⟨𝒱 s⟩ x - v x⦄)"
using *
apply auto
using valuate_nonneg[of "rhs eq" "λx. ⟨𝒱 s⟩ x - v x"]
apply force
using valuate_nonpos[of "rhs eq" "λx. ⟨𝒱 s⟩ x - v x"]
apply force
done
then show "le (lt dir) rhs eq ⦃ v ⦄ rhs eq ⦃ ⟨𝒱 s⟩ ⦄"
using ‹dir = Positive ∨ dir = Negative›
using minus_gt[of "rhs eq ⦃ v ⦄" "rhs eq ⦃ ⟨𝒱 s⟩ ⦄"]
by (auto simp add: valuate_diff[THEN sym])
qed
end
locale MinRVars = EqForLVar + MinRVarsEq min_rvar_incdec_eq
for min_rvar_incdec_eq :: "('i, 'a :: lrv) Direction ⇒ _"
begin
abbreviation min_rvar_incdec :: "('i,'a) Direction ⇒ ('i,'a) state ⇒ var ⇒ 'i list + var" where
"min_rvar_incdec dir s x⇩i ≡ min_rvar_incdec_eq dir s (eq_for_lvar (𝒯 s) x⇩i)"
end
locale MinVars = MinLVarNotInBounds min_lvar_not_in_bounds + MinRVars eq_idx_for_lvar min_rvar_incdec_eq
for min_lvar_not_in_bounds :: "('i,'a::lrv) state ⇒ _" and
eq_idx_for_lvar and
min_rvar_incdec_eq :: "('i, 'a :: lrv) Direction ⇒ _"
locale PivotUpdateMinVars =
PivotAndUpdate eq_idx_for_lvar pivot_and_update +
MinVars min_lvar_not_in_bounds eq_idx_for_lvar min_rvar_incdec_eq for
eq_idx_for_lvar :: "tableau ⇒ var ⇒ nat" and
min_lvar_not_in_bounds :: "('i,'a::lrv) state ⇒ var option" and
min_rvar_incdec_eq :: "('i,'a) Direction ⇒ ('i,'a) state ⇒ eq ⇒ 'i list + var" and
pivot_and_update :: "var ⇒ var ⇒ 'a ⇒ ('i,'a) state ⇒ ('i,'a) state"
begin
definition check' where
"check' dir x⇩i s ≡
let l⇩i = the (LB dir s x⇩i);
x⇩j' = min_rvar_incdec dir s x⇩i
in case x⇩j' of
Inl I ⇒ set_unsat I s
| Inr x⇩j ⇒ pivot_and_update x⇩i x⇩j l⇩i s"
lemma check'_cases:
assumes "⋀ I. ⟦min_rvar_incdec dir s x⇩i = Inl I; check' dir x⇩i s = set_unsat I s⟧ ⟹ P (set_unsat I s)"
assumes "⋀ x⇩j l⇩i. ⟦min_rvar_incdec dir s x⇩i = Inr x⇩j;
l⇩i = the (LB dir s x⇩i);
check' dir x⇩i s = pivot_and_update x⇩i x⇩j l⇩i s⟧ ⟹
P (pivot_and_update x⇩i x⇩j l⇩i s)"
shows "P (check' dir x⇩i s)"
using assms
unfolding check'_def
by (cases "min_rvar_incdec dir s x⇩i", auto)
partial_function (tailrec) check where
"check s =
(if 𝒰 s then s
else let x⇩i' = min_lvar_not_in_bounds s
in case x⇩i' of
None ⇒ s
| Some x⇩i ⇒ let dir = if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then Positive
else Negative
in check (check' dir x⇩i s))"
declare check.simps[code]
inductive check_dom where
step: "⟦⋀x⇩i. ⟦¬ 𝒰 s; Some x⇩i = min_lvar_not_in_bounds s; ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i⟧
⟹ check_dom (check' Positive x⇩i s);
⋀x⇩i. ⟦¬ 𝒰 s; Some x⇩i = min_lvar_not_in_bounds s; ¬ ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i⟧
⟹ check_dom (check' Negative x⇩i s)⟧
⟹ check_dom s"
text‹
The definition of ‹check› can be given by:
@{text[display]
"check s ≡ if 𝒰 s then s
else let x⇩i' = min_lvar_not_in_bounds s in
case x⇩i' of None ⇒ s
| Some x⇩i ⇒ if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then check (check_inc x⇩i s)
else check (check_dec x⇩i s)"
}
@{text[display]
"check_inc x⇩i s ≡ let l⇩i = the (ℬ⇩l s x⇩i); x⇩j' = min_rvar_inc s x⇩i in
case x⇩j' of None ⇒ s ⦇ 𝒰 := True ⦈ | Some x⇩j ⇒ pivot_and_update x⇩i x⇩j l⇩i s"
}
The definition of ‹check_dec› is analogous. It is shown (mainly
by induction) that this definition satisfies the ‹check›
specification. Note that this definition uses general recursion, so
its termination is non-trivial. It has been shown that it terminates
for all states satisfying the check preconditions. The proof is based
on the proof outline given in \cite{simplex-rad}. It is very
technically involved, but conceptually uninteresting so we do not
discuss it in more details.›
lemma pivotandupdate_check_precond:
assumes
"dir = (if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then Positive else Negative)"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec dir s x⇩i = Inr x⇩j"
"l⇩i = the (LB dir s x⇩i)"
"∇ s" "△ (𝒯 s)" "⊨⇩n⇩o⇩l⇩h⇩s s" " ◇ s"
shows "△ (𝒯 (pivot_and_update x⇩i x⇩j l⇩i s)) ∧ ⊨⇩n⇩o⇩l⇩h⇩s (pivot_and_update x⇩i x⇩j l⇩i s) ∧ ◇ (pivot_and_update x⇩i x⇩j l⇩i s) ∧ ∇ (pivot_and_update x⇩i x⇩j l⇩i s)"
proof-
have "ℬ⇩l s x⇩i = Some l⇩i ∨ ℬ⇩u s x⇩i = Some l⇩i"
using ‹l⇩i = the (LB dir s x⇩i)› ‹dir = (if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then Positive else Negative)›
using ‹min_lvar_not_in_bounds s = Some x⇩i› min_lvar_not_in_bounds_Some[of s x⇩i]
using ‹◇ s›
by (case_tac[!] "ℬ⇩l s x⇩i", case_tac[!] "ℬ⇩u s x⇩i") (auto simp add: bounds_consistent_def bound_compare_defs)
then show ?thesis
using assms
using pivotandupdate_tableau_normalized[of s x⇩i x⇩j l⇩i]
using pivotandupdate_nolhs[of s x⇩i x⇩j l⇩i]
using pivotandupdate_bounds_consistent[of s x⇩i x⇩j l⇩i]
using pivotandupdate_tableau_valuated[of s x⇩i x⇩j l⇩i]
by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars)
qed
abbreviation gt_state' where
"gt_state' dir s s' x⇩i x⇩j l⇩i ≡
min_lvar_not_in_bounds s = Some x⇩i ∧
l⇩i = the (LB dir s x⇩i) ∧
min_rvar_incdec dir s x⇩i = Inr x⇩j ∧
s' = pivot_and_update x⇩i x⇩j l⇩i s"
definition gt_state :: "('i,'a) state ⇒ ('i,'a) state ⇒ bool" (infixl "≻⇩x" 100) where
"s ≻⇩x s' ≡
∃ x⇩i x⇩j l⇩i.
let dir = if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then Positive else Negative in
gt_state' dir s s' x⇩i x⇩j l⇩i"
abbreviation succ :: "('i,'a) state ⇒ ('i,'a) state ⇒ bool" (infixl "≻" 100) where
"s ≻ s' ≡ △ (𝒯 s) ∧ ◇ s ∧ ⊨⇩n⇩o⇩l⇩h⇩s s ∧ ∇ s ∧ s ≻⇩x s' ∧ ℬ⇩i s' = ℬ⇩i s ∧ 𝒰⇩c s' = 𝒰⇩c s"
abbreviation succ_rel :: "('i,'a) state rel" where
"succ_rel ≡ {(s, s'). s ≻ s'}"
abbreviation succ_rel_trancl :: "('i,'a) state ⇒ ('i,'a) state ⇒ bool" (infixl "≻⇧+" 100) where
"s ≻⇧+ s' ≡ (s, s') ∈ succ_rel⇧+"
abbreviation succ_rel_rtrancl :: "('i,'a) state ⇒ ('i,'a) state ⇒ bool" (infixl "≻⇧*" 100) where
"s ≻⇧* s' ≡ (s, s') ∈ succ_rel⇧*"
lemma succ_vars:
assumes "s ≻ s'"
obtains x⇩i x⇩j where
"x⇩i ∈ lvars (𝒯 s)"
"x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i" "x⇩j ∈ rvars (𝒯 s)"
"lvars (𝒯 s') = lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
"rvars (𝒯 s') = rvars (𝒯 s) - {x⇩j} ∪ {x⇩i}"
proof-
from assms
obtain x⇩i x⇩j c
where *:
"△ (𝒯 s)" "∇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec Positive s x⇩i = Inr x⇩j ∨ min_rvar_incdec Negative s x⇩i = Inr x⇩j"
"s' = pivot_and_update x⇩i x⇩j c s"
unfolding gt_state_def
by (auto split: if_splits)
then have
"x⇩i ∈ lvars (𝒯 s)"
"x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)"
"lvars (𝒯 s') = lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
"rvars (𝒯 s') = rvars (𝒯 s) - {x⇩j} ∪ {x⇩i}"
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using pivotandupdate_rvars[of s x⇩i x⇩j]
using pivotandupdate_lvars[of s x⇩i x⇩j]
by auto
moreover
have "x⇩j ∈ rvars (𝒯 s)"
using ‹x⇩i ∈ lvars (𝒯 s)›
using ‹x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)›
using eq_for_lvar[of x⇩i "𝒯 s"]
unfolding rvars_def
by auto
ultimately
have
"x⇩i ∈ lvars (𝒯 s)"
"x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i" "x⇩j ∈ rvars (𝒯 s)"
"lvars (𝒯 s') = lvars (𝒯 s) - {x⇩i} ∪ {x⇩j}"
"rvars (𝒯 s') = rvars (𝒯 s) - {x⇩j} ∪ {x⇩i}"
by auto
then show thesis
..
qed
lemma succ_vars_id:
assumes "s ≻ s'"
shows "lvars (𝒯 s) ∪ rvars (𝒯 s) =
lvars (𝒯 s') ∪ rvars (𝒯 s')"
using assms
by (rule succ_vars) auto
lemma succ_inv:
assumes "s ≻ s'"
shows "△ (𝒯 s')" "∇ s'" "◇ s'" "ℬ⇩i s = ℬ⇩i s'"
"(v::'a valuation) ⊨⇩t (𝒯 s) ⟷ v ⊨⇩t (𝒯 s')"
proof-
from assms obtain x⇩i x⇩j c
where *:
"△ (𝒯 s)" "∇ s" "◇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec Positive s x⇩i = Inr x⇩j ∨ min_rvar_incdec Negative s x⇩i = Inr x⇩j"
"s' = pivot_and_update x⇩i x⇩j c s"
unfolding gt_state_def
by (auto split: if_splits)
then show "△ (𝒯 s')" "∇ s'" "◇ s'" "ℬ⇩i s = ℬ⇩i s'"
"(v::'a valuation) ⊨⇩t (𝒯 s) ⟷ v ⊨⇩t (𝒯 s')"
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using pivotandupdate_tableau_normalized[of s x⇩i x⇩j c]
using pivotandupdate_bounds_consistent[of s x⇩i x⇩j c]
using pivotandupdate_bounds_id[of s x⇩i x⇩j c]
using pivotandupdate_tableau_equiv
using pivotandupdate_tableau_valuated
by auto
qed
lemma succ_rvar_valuation_id:
assumes "s ≻ s'" "x ∈ rvars (𝒯 s)" "x ∈ rvars (𝒯 s')"
shows "⟨𝒱 s⟩ x = ⟨𝒱 s'⟩ x"
proof-
from assms obtain x⇩i x⇩j c
where *:
"△ (𝒯 s)" "∇ s" "◇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec Positive s x⇩i = Inr x⇩j ∨ min_rvar_incdec Negative s x⇩i = Inr x⇩j"
"s' = pivot_and_update x⇩i x⇩j c s"
unfolding gt_state_def
by (auto split: if_splits)
then show ?thesis
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using ‹x ∈ rvars (𝒯 s)› ‹x ∈ rvars (𝒯 s')›
using pivotandupdate_rvars[of s x⇩i x⇩j c]
using pivotandupdate_valuation_xi[of s x⇩i x⇩j c]
using pivotandupdate_valuation_other_nolhs[of s x⇩i x⇩j x c]
by (force simp add: normalized_tableau_def map2fun_def)
qed
lemma succ_min_lvar_not_in_bounds:
assumes "s ≻ s'"
"xr ∈ lvars (𝒯 s)" "xr ∈ rvars (𝒯 s')"
shows "¬ in_bounds xr (⟨𝒱 s⟩) (ℬ s)"
"∀ x ∈ lvars (𝒯 s). x < xr ⟶ in_bounds x (⟨𝒱 s⟩) (ℬ s)"
proof-
from assms obtain x⇩i x⇩j c
where *:
"△ (𝒯 s)" "∇ s" "◇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec Positive s x⇩i = Inr x⇩j ∨ min_rvar_incdec Negative s x⇩i = Inr x⇩j"
"s' = pivot_and_update x⇩i x⇩j c s"
unfolding gt_state_def
by (auto split: if_splits)
then have "x⇩i = xr"
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using ‹xr ∈ lvars (𝒯 s)› ‹xr ∈ rvars (𝒯 s')›
using pivotandupdate_rvars
by (auto simp add: normalized_tableau_def)
then show "¬ in_bounds xr (⟨𝒱 s⟩) (ℬ s)"
"∀ x ∈ lvars (𝒯 s). x < xr ⟶ in_bounds x (⟨𝒱 s⟩) (ℬ s)"
using ‹min_lvar_not_in_bounds s = Some x⇩i›
using min_lvar_not_in_bounds_Some min_lvar_not_in_bounds_Some_min
by simp_all
qed
lemma succ_min_rvar:
assumes "s ≻ s'"
"xs ∈ lvars (𝒯 s)" "xs ∈ rvars (𝒯 s')"
"xr ∈ rvars (𝒯 s)" "xr ∈ lvars (𝒯 s')"
"eq = eq_for_lvar (𝒯 s) xs" and
dir: "dir = Positive ∨ dir = Negative"
shows
"¬ ⊵⇩l⇩b (lt dir) (⟨𝒱 s⟩ xs) (LB dir s xs) ⟶
reasable_var dir xr eq s ∧ (∀ x ∈ rvars_eq eq. x < xr ⟶ ¬ reasable_var dir x eq s)"
proof-
from assms(1) obtain x⇩i x⇩j c
where"△ (𝒯 s) ∧ ∇ s ∧ ◇ s ∧ ⊨⇩n⇩o⇩l⇩h⇩s s"
"gt_state' (if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then Positive else Negative) s s' x⇩i x⇩j c"
by (auto simp add: gt_state_def Let_def)
then have
"△ (𝒯 s)" "∇ s" "◇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"s' = pivot_and_update x⇩i x⇩j c s" and
*: "(⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i ∧ min_rvar_incdec Positive s x⇩i = Inr x⇩j) ∨
(¬ ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i ∧ min_rvar_incdec Negative s x⇩i = Inr x⇩j)"
by (auto split: if_splits)
then have "xr = x⇩j" "xs = x⇩i"
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using ‹xr ∈ rvars (𝒯 s)› ‹xr ∈ lvars (𝒯 s')›
using ‹xs ∈ lvars (𝒯 s)› ‹xs ∈ rvars (𝒯 s')›
using pivotandupdate_lvars pivotandupdate_rvars
by (auto simp add: normalized_tableau_def)
show "¬ (⊵⇩l⇩b (lt dir) (⟨𝒱 s⟩ xs) (LB dir s xs)) ⟶
reasable_var dir xr eq s ∧ (∀ x ∈ rvars_eq eq. x < xr ⟶ ¬ reasable_var dir x eq s)"
proof
assume "¬ ⊵⇩l⇩b (lt dir) (⟨𝒱 s⟩ xs) (LB dir s xs)"
then have "⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ xs) (LB dir s xs)"
using dir
by (cases "LB dir s xs") (auto simp add: bound_compare_defs)
moreover
then have "¬ (⊳⇩u⇩b (lt dir) (⟨𝒱 s⟩ xs) (UB dir s xs))"
using ‹◇ s› dir
using bounds_consistent_gt_ub bounds_consistent_lt_lb
by (force simp add: bound_compare''_defs)
ultimately
have "min_rvar_incdec dir s xs = Inr xr"
using * ‹xr = x⇩j› ‹xs = x⇩i› dir
by (auto simp add: bound_compare''_defs)
then show "reasable_var dir xr eq s ∧ (∀ x ∈ rvars_eq eq. x < xr ⟶ ¬ reasable_var dir x eq s)"
using ‹eq = eq_for_lvar (𝒯 s) xs›
using min_rvar_incdec_eq_Some_min[of dir s eq xr]
using min_rvar_incdec_eq_Some_incdec[of dir s eq xr]
by simp
qed
qed
lemma succ_set_on_bound:
assumes
"s ≻ s'" "x⇩i ∈ lvars (𝒯 s)" "x⇩i ∈ rvars (𝒯 s')" and
dir: "dir = Positive ∨ dir = Negative"
shows
"¬ ⊵⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i) ⟶ ⟨𝒱 s'⟩ x⇩i = the (LB dir s x⇩i)"
"⟨𝒱 s'⟩ x⇩i = the (ℬ⇩l s x⇩i) ∨ ⟨𝒱 s'⟩ x⇩i = the (ℬ⇩u s x⇩i)"
proof-
from assms(1) obtain x⇩i' x⇩j c
where"△ (𝒯 s) ∧ ∇ s ∧ ◇ s ∧ ⊨⇩n⇩o⇩l⇩h⇩s s"
"gt_state' (if ⟨𝒱 s⟩ x⇩i' <⇩l⇩b ℬ⇩l s x⇩i' then Positive else Negative) s s' x⇩i' x⇩j c"
by (auto simp add: gt_state_def Let_def)
then have
"△ (𝒯 s)" "∇ s" "◇ s"
"min_lvar_not_in_bounds s = Some x⇩i'"
"s' = pivot_and_update x⇩i' x⇩j c s" and
*: "(⟨𝒱 s⟩ x⇩i' <⇩l⇩b ℬ⇩l s x⇩i' ∧ c = the (ℬ⇩l s x⇩i') ∧ min_rvar_incdec Positive s x⇩i' = Inr x⇩j) ∨
(¬ ⟨𝒱 s⟩ x⇩i' <⇩l⇩b ℬ⇩l s x⇩i' ∧ c = the (ℬ⇩u s x⇩i') ∧ min_rvar_incdec Negative s x⇩i' = Inr x⇩j)"
by (auto split: if_splits)
then have "x⇩i = x⇩i'" "x⇩i' ∈ lvars (𝒯 s)"
"x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i')"
using min_lvar_not_in_bounds_lvars[of s x⇩i']
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i'" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i'" x⇩j]
using ‹x⇩i ∈ lvars (𝒯 s)› ‹x⇩i ∈ rvars (𝒯 s')›
using pivotandupdate_rvars
by (auto simp add: normalized_tableau_def)
show "¬ ⊵⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i) ⟶ ⟨𝒱 s'⟩ x⇩i = the (LB dir s x⇩i)"
proof
assume "¬ ⊵⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i)"
then have "⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i)"
using dir
by (cases "LB dir s x⇩i") (auto simp add: bound_compare_defs)
moreover
then have "¬ ⊳⇩u⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (UB dir s x⇩i)"
using ‹◇ s› dir
using bounds_consistent_gt_ub bounds_consistent_lt_lb
by (force simp add: bound_compare''_defs)
ultimately
show "⟨𝒱 s'⟩ x⇩i = the (LB dir s x⇩i)"
using * ‹x⇩i = x⇩i'› ‹s' = pivot_and_update x⇩i' x⇩j c s›
using ‹△ (𝒯 s)› ‹∇ s› ‹x⇩i' ∈ lvars (𝒯 s)›
‹x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i')›
using pivotandupdate_valuation_xi[of s x⇩i x⇩j c] dir
by (case_tac[!] "ℬ⇩l s x⇩i'", case_tac[!] "ℬ⇩u s x⇩i'") (auto simp add: bound_compare_defs map2fun_def)
qed
have "¬ ⟨𝒱 s⟩ x⇩i' <⇩l⇩b ℬ⇩l s x⇩i' ⟶ ⟨𝒱 s⟩ x⇩i' >⇩u⇩b ℬ⇩u s x⇩i'"
using ‹min_lvar_not_in_bounds s = Some x⇩i'›
using min_lvar_not_in_bounds_Some[of s x⇩i']
using not_in_bounds[of x⇩i' "⟨𝒱 s⟩" "ℬ⇩l s" "ℬ⇩u s"]
by auto
then show "⟨𝒱 s'⟩ x⇩i = the (ℬ⇩l s x⇩i) ∨ ⟨𝒱 s'⟩ x⇩i = the (ℬ⇩u s x⇩i)"
using ‹△ (𝒯 s)› ‹∇ s› ‹x⇩i' ∈ lvars (𝒯 s)›
‹x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i')›
using ‹s' = pivot_and_update x⇩i' x⇩j c s› ‹x⇩i = x⇩i'›
using pivotandupdate_valuation_xi[of s x⇩i x⇩j c]
using *
by (case_tac[!] "ℬ⇩l s x⇩i'", case_tac[!] "ℬ⇩u s x⇩i'") (auto simp add: map2fun_def bound_compare_defs)
qed
lemma succ_rvar_valuation:
assumes
"s ≻ s'" "x ∈ rvars (𝒯 s')"
shows
"⟨𝒱 s'⟩ x = ⟨𝒱 s⟩ x ∨ ⟨𝒱 s'⟩ x = the (ℬ⇩l s x) ∨ ⟨𝒱 s'⟩ x = the (ℬ⇩u s x)"
proof-
from assms
obtain x⇩i x⇩j b where
"△ (𝒯 s)" "∇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec Positive s x⇩i = Inr x⇩j ∨ min_rvar_incdec Negative s x⇩i = Inr x⇩j"
"b = the (ℬ⇩l s x⇩i) ∨ b = the (ℬ⇩u s x⇩i)"
"s' = pivot_and_update x⇩i x⇩j b s"
unfolding gt_state_def
by (auto simp add: Let_def split: if_splits)
then have
"x⇩i ∈ lvars (𝒯 s)" "x⇩i ∉ rvars (𝒯 s)"
"x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)"
"x⇩j ∈ rvars (𝒯 s)" "x⇩j ∉ lvars (𝒯 s)" "x⇩i ≠ x⇩j"
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using rvars_of_lvar_rvars ‹△ (𝒯 s)›
by (auto simp add: normalized_tableau_def)
then have
"rvars (𝒯 s') = rvars (𝒯 s) - {x⇩j} ∪ {x⇩i}"
"x ∈ rvars (𝒯 s) ∨ x = x⇩i" "x ≠ x⇩j" "x ≠ x⇩i ⟶ x ∉ lvars (𝒯 s)"
using ‹x ∈ rvars (𝒯 s')›
using pivotandupdate_rvars[of s x⇩i x⇩j]
using ‹△ (𝒯 s)› ‹∇ s› ‹s' = pivot_and_update x⇩i x⇩j b s›
by (auto simp add: normalized_tableau_def)
then show ?thesis
using pivotandupdate_valuation_xi[of s x⇩i x⇩j b]
using pivotandupdate_valuation_other_nolhs[of s x⇩i x⇩j x b]
using ‹x⇩i ∈ lvars (𝒯 s)› ‹x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)›
using ‹△ (𝒯 s)› ‹∇ s› ‹s' = pivot_and_update x⇩i x⇩j b s› ‹b = the (ℬ⇩l s x⇩i) ∨ b = the (ℬ⇩u s x⇩i)›
by (auto simp add: map2fun_def)
qed
lemma succ_no_vars_valuation:
assumes
"s ≻ s'" "x ∉ tvars (𝒯 s')"
shows "look (𝒱 s') x = look (𝒱 s) x"
proof-
from assms
obtain x⇩i x⇩j b where
"△ (𝒯 s)" "∇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec Positive s x⇩i = Inr x⇩j ∨ min_rvar_incdec Negative s x⇩i = Inr x⇩j"
"b = the (ℬ⇩l s x⇩i) ∨ b = the (ℬ⇩u s x⇩i)"
"s' = pivot_and_update x⇩i x⇩j b s"
unfolding gt_state_def
by (auto simp add: Let_def split: if_splits)
then have
"x⇩i ∈ lvars (𝒯 s)" "x⇩i ∉ rvars (𝒯 s)"
"x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)"
"x⇩j ∈ rvars (𝒯 s)" "x⇩j ∉ lvars (𝒯 s)" "x⇩i ≠ x⇩j"
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using rvars_of_lvar_rvars ‹△ (𝒯 s)›
by (auto simp add: normalized_tableau_def)
then show ?thesis
using pivotandupdate_valuation_other_nolhs[of s x⇩i x⇩j x b]
using ‹△ (𝒯 s)› ‹∇ s› ‹s' = pivot_and_update x⇩i x⇩j b s›
using ‹x ∉ tvars (𝒯 s')›
using pivotandupdate_rvars[of s x⇩i x⇩j]
using pivotandupdate_lvars[of s x⇩i x⇩j]
by (auto simp add: map2fun_def)
qed
lemma succ_valuation_satisfies:
assumes "s ≻ s'" "⟨𝒱 s⟩ ⊨⇩t 𝒯 s"
shows "⟨𝒱 s'⟩ ⊨⇩t 𝒯 s'"
proof-
from ‹s ≻ s'›
obtain x⇩i x⇩j b where
"△ (𝒯 s)" "∇ s"
"min_lvar_not_in_bounds s = Some x⇩i"
"min_rvar_incdec Positive s x⇩i = Inr x⇩j ∨ min_rvar_incdec Negative s x⇩i = Inr x⇩j"
"b = the (ℬ⇩l s x⇩i) ∨ b = the (ℬ⇩u s x⇩i)"
"s' = pivot_and_update x⇩i x⇩j b s"
unfolding gt_state_def
by (auto simp add: Let_def split: if_splits)
then have
"x⇩i ∈ lvars (𝒯 s)"
"x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i"
using min_lvar_not_in_bounds_lvars[of s x⇩i]
using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (𝒯 s) x⇩i" x⇩j] ‹△ (𝒯 s)›
by (auto simp add: normalized_tableau_def)
then show ?thesis
using pivotandupdate_satisfies_tableau[of s x⇩i x⇩j b]
using pivotandupdate_tableau_equiv[of s x⇩i x⇩j ]
using ‹△ (𝒯 s)› ‹∇ s› ‹⟨𝒱 s⟩ ⊨⇩t 𝒯 s› ‹s' = pivot_and_update x⇩i x⇩j b s›
by auto
qed
lemma succ_tableau_valuated:
assumes "s ≻ s'" "∇ s"
shows "∇ s'"
using succ_inv(2) assms by blast
abbreviation succ_chain where
"succ_chain l ≡ rel_chain l succ_rel"
lemma succ_chain_induct:
assumes *: "succ_chain l" "i ≤ j" "j < length l"
assumes base: "⋀ i. P i i"
assumes step: "⋀ i. l ! i ≻ (l ! (i + 1)) ⟹ P i (i + 1)"
assumes trans: "⋀ i j k. ⟦P i j; P j k; i < j; j ≤ k⟧ ⟹ P i k"
shows "P i j"
using *
proof (induct "j - i" arbitrary: i)
case 0
then show ?case
by (simp add: base)
next
case (Suc k)
have "P (i + 1) j"
using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5)
by auto
moreover
have "P i (i + 1)"
proof (rule step)
show "l ! i ≻ (l ! (i + 1))"
using Suc(2) Suc(3) Suc(5)
unfolding rel_chain_def
by auto
qed
ultimately
show ?case
using trans[of i "i + 1" j] Suc(2)
by simp
qed
lemma succ_chain_bounds_id:
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "ℬ⇩i (l ! i) = ℬ⇩i (l ! j)"
using assms
proof (rule succ_chain_induct)
fix i
assume "l ! i ≻ (l ! (i + 1))"
then show "ℬ⇩i (l ! i) = ℬ⇩i (l ! (i + 1))"
by (rule succ_inv(4))
qed simp_all
lemma succ_chain_vars_id':
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "lvars (𝒯 (l ! i)) ∪ rvars (𝒯 (l ! i)) =
lvars (𝒯 (l ! j)) ∪ rvars (𝒯 (l ! j))"
using assms
proof (rule succ_chain_induct)
fix i
assume "l ! i ≻ (l ! (i + 1))"
then show "tvars (𝒯 (l ! i)) = tvars (𝒯 (l ! (i + 1)))"
by (rule succ_vars_id)
qed simp_all
lemma succ_chain_vars_id:
assumes "succ_chain l" "i < length l" "j < length l"
shows "lvars (𝒯 (l ! i)) ∪ rvars (𝒯 (l ! i)) =
lvars (𝒯 (l ! j)) ∪ rvars (𝒯 (l ! j))"
proof (cases "i ≤ j")
case True
then show ?thesis
using assms succ_chain_vars_id'[of l i j]
by simp
next
case False
then have "j ≤ i"
by simp
then show ?thesis
using assms succ_chain_vars_id'[of l j i]
by simp
qed
lemma succ_chain_tableau_equiv':
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "(v::'a valuation) ⊨⇩t 𝒯 (l ! i) ⟷ v ⊨⇩t 𝒯 (l ! j)"
using assms
proof (rule succ_chain_induct)
fix i
assume "l ! i ≻ (l ! (i + 1))"
then show "v ⊨⇩t 𝒯 (l ! i) = v ⊨⇩t 𝒯 (l ! (i + 1))"
by (rule succ_inv(5))
qed simp_all
lemma succ_chain_tableau_equiv:
assumes "succ_chain l" "i < length l" "j < length l"
shows "(v::'a valuation) ⊨⇩t 𝒯 (l ! i) ⟷ v ⊨⇩t 𝒯 (l ! j)"
proof (cases "i ≤ j")
case True
then show ?thesis
using assms succ_chain_tableau_equiv'[of l i j v]
by simp
next
case False
then have "j ≤ i"
by auto
then show ?thesis
using assms succ_chain_tableau_equiv'[of l j i v]
by simp
qed
lemma succ_chain_no_vars_valuation:
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "∀ x. x ∉ tvars (𝒯 (l ! i)) ⟶ look (𝒱 (l ! i)) x = look (𝒱 (l ! j)) x" (is "?P i j")
using assms
proof (induct "j - i" arbitrary: i)
case 0
then show ?case
by simp
next
case (Suc k)
have "?P (i + 1) j"
using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5)
by auto
moreover
have "?P (i + 1) i"
proof (rule+, rule succ_no_vars_valuation)
show "l ! i ≻ (l ! (i + 1))"
using Suc(2) Suc(3) Suc(5)
unfolding rel_chain_def
by auto
qed
moreover
have "tvars (𝒯 (l ! i)) = tvars (𝒯 (l ! (i + 1)))"
proof (rule succ_vars_id)
show "l ! i ≻ (l ! (i + 1))"
using Suc(2) Suc(3) Suc(5)
unfolding rel_chain_def
by simp
qed
ultimately
show ?case
by simp
qed
lemma succ_chain_rvar_valuation:
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "∀x∈rvars (𝒯 (l ! j)).
⟨𝒱 (l ! j)⟩ x = ⟨𝒱 (l ! i)⟩ x ∨
⟨𝒱 (l ! j)⟩ x = the (ℬ⇩l (l ! i) x) ∨
⟨𝒱 (l ! j)⟩ x = the (ℬ⇩u (l ! i) x)" (is "?P i j")
using assms
proof (induct "j - i" arbitrary: j)
case 0
then show ?case
by simp
next
case (Suc k)
have "k = j - 1 - i" "succ_chain l" "i ≤ j - 1" "j - 1 < length l" "j > 0"
using Suc(2) Suc(3) Suc(4) Suc(5)
by auto
then have ji: "?P i (j - 1)"
using Suc(1)
by simp
have "l ! (j - 1) ≻ (l ! j)"
using Suc(3) ‹j < length l› ‹j > 0›
unfolding rel_chain_def
by (erule_tac x="j - 1" in allE) simp
then have
jj: "?P (j - 1) j"
using succ_rvar_valuation
by auto
obtain x⇩i x⇩j where
vars: "x⇩i ∈ lvars (𝒯 (l ! (j - 1)))" "x⇩j ∈ rvars (𝒯 (l ! (j - 1)))"
"rvars (𝒯 (l ! j)) = rvars (𝒯 (l ! (j - 1))) - {x⇩j} ∪ {x⇩i}"
using ‹l ! (j - 1) ≻ (l ! j)›
by (rule succ_vars) simp
then have bounds:
"ℬ⇩l (l ! (j - 1)) = ℬ⇩l (l ! i)" "ℬ⇩l (l ! j) = ℬ⇩l (l ! i)"
"ℬ⇩u (l ! (j - 1)) = ℬ⇩u (l ! i)" "ℬ⇩u (l ! j) = ℬ⇩u (l ! i)"
using ‹succ_chain l›
using succ_chain_bounds_id[of l i "j - 1", THEN sym] ‹j - 1 < length l› ‹i ≤ j - 1›
using succ_chain_bounds_id[of l "j - 1" j, THEN sym] ‹j < length l›
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
show ?case
proof
fix x
assume "x ∈ rvars (𝒯 (l ! j))"
then have "x ≠ x⇩j ∧ x ∈ rvars (𝒯 (l ! (j - 1))) ∨ x = x⇩i"
using vars
by auto
then show "⟨𝒱 (l ! j)⟩ x = ⟨𝒱 (l ! i)⟩ x ∨
⟨𝒱 (l ! j)⟩ x = the (ℬ⇩l (l ! i) x) ∨
⟨𝒱 (l ! j)⟩ x = the (ℬ⇩u (l ! i) x)"
proof
assume "x ≠ x⇩j ∧ x ∈ rvars (𝒯 (l ! (j - 1)))"
then show ?thesis
using jj ‹x ∈ rvars (𝒯 (l ! j))› ji
using bounds
by force
next
assume "x = x⇩i"
then show ?thesis
using succ_set_on_bound(2)[of "l ! (j - 1)" "l ! j" x⇩i] ‹l ! (j - 1) ≻ (l ! j)›
using vars bounds
by auto
qed
qed
qed
lemma succ_chain_valuation_satisfies:
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "⟨𝒱 (l ! i)⟩ ⊨⇩t 𝒯 (l ! i) ⟶ ⟨𝒱 (l ! j)⟩ ⊨⇩t 𝒯 (l ! j)"
using assms
proof (rule succ_chain_induct)
fix i
assume "l ! i ≻ (l ! (i + 1))"
then show "⟨𝒱 (l ! i)⟩ ⊨⇩t 𝒯 (l ! i) ⟶ ⟨𝒱 (l ! (i + 1))⟩ ⊨⇩t 𝒯 (l ! (i + 1))"
using succ_valuation_satisfies
by auto
qed simp_all
lemma succ_chain_tableau_valuated:
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "∇ (l ! i) ⟶ ∇ (l ! j)"
using assms
proof(rule succ_chain_induct)
fix i
assume "l ! i ≻ (l ! (i + 1))"
then show "∇ (l ! i) ⟶ ∇ (l ! (i + 1))"
using succ_tableau_valuated
by auto
qed simp_all
abbreviation swap_lr where
"swap_lr l i x ≡ i + 1 < length l ∧ x ∈ lvars (𝒯 (l ! i)) ∧ x ∈ rvars (𝒯 (l ! (i + 1)))"
abbreviation swap_rl where
"swap_rl l i x ≡ i + 1 < length l ∧ x ∈ rvars (𝒯 (l ! i)) ∧ x ∈ lvars (𝒯 (l ! (i + 1)))"
abbreviation always_r where
"always_r l i j x ≡ ∀ k. i ≤ k ∧ k ≤ j ⟶ x ∈ rvars (𝒯 (l ! k))"
lemma succ_chain_always_r_valuation_id:
assumes "succ_chain l" "i ≤ j" "j < length l"
shows "always_r l i j x ⟶ ⟨𝒱 (l ! i)⟩ x = ⟨𝒱 (l ! j)⟩ x" (is "?P i j")
using assms
proof (rule succ_chain_induct)
fix i
assume "l ! i ≻ (l ! (i + 1))"
then show "?P i (i + 1)"
using succ_rvar_valuation_id
by simp
qed simp_all
lemma succ_chain_swap_rl_exists:
assumes "succ_chain l" "i < j" "j < length l"
"x ∈ rvars (𝒯 (l ! i))" "x ∈ lvars (𝒯 (l ! j))"
shows "∃ k. i ≤ k ∧ k < j ∧ swap_rl l k x"
using assms
proof (induct "j - i" arbitrary: i)
case 0
then show ?case
by simp
next
case (Suc k)
have "l ! i ≻ (l ! (i + 1))"
using Suc(3) Suc(4) Suc(5)
unfolding rel_chain_def
by auto
then have "△ (𝒯 (l ! (i + 1)))"
by (rule succ_inv)
show ?case
proof (cases "x ∈ rvars (𝒯 (l ! (i + 1)))")
case True
then have "j ≠ i + 1"
using Suc(7) ‹△ (𝒯 (l ! (i + 1)))›
by (auto simp add: normalized_tableau_def)
have "k = j - Suc i"
using Suc(2)
by simp
then obtain k where "k ≥ i + 1" "k < j" "swap_rl l k x"
using ‹x ∈ rvars (𝒯 (l ! (i + 1)))› ‹j ≠ i + 1›
using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5) Suc(6) Suc(7)
by auto
then show ?thesis
by (rule_tac x="k" in exI) simp
next
case False
then have "x ∈ lvars (𝒯 (l ! (i + 1)))"
using Suc(6)
using ‹l ! i ≻ (l ! (i + 1))› succ_vars_id
by auto
then show ?thesis
using Suc(4) Suc(5) Suc(6)
by force
qed
qed
lemma succ_chain_swap_lr_exists:
assumes "succ_chain l" "i < j" "j < length l"
"x ∈ lvars (𝒯 (l ! i))" "x ∈ rvars (𝒯 (l ! j))"
shows "∃ k. i ≤ k ∧ k < j ∧ swap_lr l k x"
using assms
proof (induct "j - i" arbitrary: i)
case 0
then show ?case
by simp
next
case (Suc k)
have "l ! i ≻ (l ! (i + 1))"
using Suc(3) Suc(4) Suc(5)
unfolding rel_chain_def
by auto
then have "△ (𝒯 (l ! (i + 1)))"
by (rule succ_inv)
show ?case
proof (cases "x ∈ lvars (𝒯 (l ! (i + 1)))")
case True
then have "j ≠ i + 1"
using Suc(7) ‹△ (𝒯 (l ! (i + 1)))›
by (auto simp add: normalized_tableau_def)
have "k = j - Suc i"
using Suc(2)
by simp
then obtain k where "k ≥ i + 1" "k < j" "swap_lr l k x"
using ‹x ∈ lvars (𝒯 (l ! (i + 1)))› ‹j ≠ i + 1›
using Suc(1)[of "i + 1"] Suc(2) Suc(3) Suc(4) Suc(5) Suc(6) Suc(7)
by auto
then show ?thesis
by (rule_tac x="k" in exI) simp
next
case False
then have "x ∈ rvars (𝒯 (l ! (i + 1)))"
using Suc(6)
using ‹l ! i ≻ (l ! (i + 1))› succ_vars_id
by auto
then show ?thesis
using Suc(4) Suc(5) Suc(6)
by force
qed
qed
lemma finite_tableaus_aux:
shows "finite {t. lvars t = L ∧ rvars t = V - L ∧ △ t ∧ (∀ v::'a valuation. v ⊨⇩t t = v ⊨⇩t t0)}" (is "finite (?Al L)")
proof (cases "?Al L = {}")
case True
show ?thesis
by (subst True) simp
next
case False
then have "∃ t. t ∈ ?Al L"
by auto
let ?t = "SOME t. t ∈ ?Al L"
have "?t ∈ ?Al L"
using ‹∃ t. t ∈ ?Al L›
by (rule someI_ex)
have "?Al L ⊆ {t. t <~~> ?t}"
proof
fix x
assume "x ∈ ?Al L"
have "x <~~> ?t"
apply (rule tableau_perm)
using ‹?t ∈ ?Al L› ‹x ∈ ?Al L›
by auto
then show "x ∈ {t. t <~~> ?t}"
by simp
qed
moreover
have "finite {t. t <~~> ?t}"
by (rule perm_finite)
ultimately
show ?thesis
by (rule finite_subset)
qed
lemma finite_tableaus:
assumes "finite V"
shows "finite {t. tvars t = V ∧ △ t ∧ (∀ v::'a valuation. v ⊨⇩t t = v ⊨⇩t t0)}" (is "finite ?A")
proof-
let ?Al = "λ L. {t. lvars t = L ∧ rvars t = V - L ∧ △ t ∧ (∀ v::'a valuation. v ⊨⇩t t = v ⊨⇩t t0)}"
have "?A = ⋃ (?Al ` {L. L ⊆ V})"
by (auto simp add: normalized_tableau_def)
then show ?thesis
using ‹finite V›
using finite_tableaus_aux
by auto
qed
lemma finite_accessible_tableaus:
shows "finite (𝒯 ` {s'. s ≻⇧* s'})"
proof-
have "{s'. s ≻⇧* s'} = {s'. s ≻⇧+ s'} ∪ {s}"
by (auto simp add: rtrancl_eq_or_trancl)
moreover
have "finite (𝒯 ` {s'. s ≻⇧+ s'})" (is "finite ?A")
proof-
let ?T = "{t. tvars t = tvars (𝒯 s) ∧ △ t ∧ (∀ v::'a valuation. v ⊨⇩t t = v ⊨⇩t(𝒯 s))}"
have "?A ⊆ ?T"
proof
fix t
assume "t ∈ ?A"
then obtain s' where "s ≻⇧+ s'" "t = 𝒯 s'"
by auto
then obtain l where *: "l ≠ []" "1 < length l" "hd l = s" "last l = s'" "succ_chain l"
using trancl_rel_chain[of s s' succ_rel]
by auto
show "t ∈ ?T"
proof-
have "tvars (𝒯 s') = tvars (𝒯 s)"
using succ_chain_vars_id[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l]
by simp
moreover
have "△ (𝒯 s')"
using ‹s ≻⇧+ s'›
using succ_inv(1)[of _ s']
by (auto dest: tranclD2)
moreover
have "∀v::'a valuation. v ⊨⇩t 𝒯 s' = v ⊨⇩t 𝒯 s"
using succ_chain_tableau_equiv[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l]
by auto
ultimately
show ?thesis
using ‹t = 𝒯 s'›
by simp
qed
qed
moreover
have "finite (tvars (𝒯 s))"
by (auto simp add: lvars_def rvars_def finite_vars)
ultimately
show ?thesis
using finite_tableaus[of "tvars (𝒯 s)" "𝒯 s"]
by (auto simp add: finite_subset)
qed
ultimately
show ?thesis
by simp
qed
abbreviation check_valuation where
"check_valuation (v::'a valuation) v0 bl0 bu0 t0 V ≡
∃ t. tvars t = V ∧ △ t ∧ (∀ v::'a valuation. v ⊨⇩t t = v ⊨⇩t t0) ∧ v ⊨⇩t t ∧
(∀ x ∈ rvars t. v x = v0 x ∨ v x = bl0 x ∨ v x = bu0 x) ∧
(∀ x. x ∉ V ⟶ v x = v0 x)"
lemma finite_valuations:
assumes "finite V"
shows "finite {v::'a valuation. check_valuation v v0 bl0 bu0 t0 V}" (is "finite ?A")
proof-
let ?Al = "λ L. {t. lvars t = L ∧ rvars t = V - L ∧ △ t ∧ (∀ v::'a valuation. v ⊨⇩t t = v ⊨⇩t t0)}"
let ?Vt = "λ t. {v::'a valuation. v ⊨⇩t t ∧ (∀ x ∈ rvars t. v x = v0 x ∨ v x = bl0 x ∨ v x = bu0 x) ∧ (∀ x. x ∉ V ⟶ v x = v0 x)}"
have "finite {L. L ⊆ V}"
using ‹finite V›
by auto
have "∀ L. L ⊆ V ⟶ finite (?Al L)"
using finite_tableaus_aux
by auto
have "∀ L t. L ⊆ V ∧ t ∈ ?Al L ⟶ finite (?Vt t)"
proof (safe)
fix L t
assume "lvars t ⊆ V" "rvars t = V - lvars t" "△ t" "∀v. v ⊨⇩t t = v ⊨⇩t t0"
then have "rvars t ∪ lvars t = V"
by auto
let ?f = "λ v x. if x ∈ rvars t then v x else 0"
have "inj_on ?f (?Vt t)"
unfolding inj_on_def
proof (safe, rule ext)
fix v1 v2 x
assume "(λx. if x ∈ rvars t then v1 x else (0 :: 'a)) =
(λx. if x ∈ rvars t then v2 x else (0 :: 'a))" (is "?f1 = ?f2")
have "∀x∈rvars t. v1 x = v2 x"
proof
fix x
assume "x ∈ rvars t"
then show "v1 x = v2 x"
using ‹?f1 = ?f2› fun_cong[of ?f1 ?f2 x]
by auto
qed
assume *: "v1 ⊨⇩t t" "v2 ⊨⇩t t"
"∀x. x ∉ V ⟶ v1 x = v0 x" "∀x. x ∉ V ⟶ v2 x = v0 x"
show "v1 x = v2 x"
proof (cases "x ∈ lvars t")
case False
then show ?thesis
using * ‹∀x∈rvars t. v1 x = v2 x› ‹rvars t ∪ lvars t = V›
by auto
next
case True
let ?eq = "eq_for_lvar t x"
have "?eq ∈ set t ∧ lhs ?eq = x"
using eq_for_lvar ‹x ∈ lvars t›
by simp
then have "v1 x = rhs ?eq ⦃ v1 ⦄" "v2 x = rhs ?eq ⦃ v2 ⦄"
using ‹v1 ⊨⇩t t› ‹v2 ⊨⇩t t›
unfolding satisfies_tableau_def satisfies_eq_def
by auto
moreover
have "rhs ?eq ⦃ v1 ⦄ = rhs ?eq ⦃ v2 ⦄"
apply (rule valuate_depend)
using ‹∀x∈rvars t. v1 x = v2 x› ‹?eq ∈ set t ∧ lhs ?eq = x›
unfolding rvars_def
by auto
ultimately
show ?thesis
by simp
qed
qed
let ?R = "{v. ∀ x. if x ∈ rvars t then v x = v0 x ∨ v x = bl0 x ∨ v x = bu0 x else v x = 0 }"
have "?f ` (?Vt t) ⊆ ?R"
by auto
moreover
have "finite ?R"
proof-
have "finite (rvars t)"
using ‹finite V› ‹rvars t ∪ lvars t = V›
using finite_subset[of "rvars t" V]
by auto
moreover
let ?R' = "{v. ∀ x. if x ∈ rvars t then v x ∈ {v0 x, bl0 x, bu0 x} else v x = 0}"
have "?R = ?R'"
by auto
ultimately
show ?thesis
using finite_fun_args[of "rvars t" "λ x. {v0 x, bl0 x, bu0 x}" "λ x. 0"]
by auto
qed
ultimately
have "finite (?f ` (?Vt t))"
by (simp add: finite_subset)
then show "finite (?Vt t)"
using ‹inj_on ?f (?Vt t)›
by (auto dest: finite_imageD)
qed
have "?A = ⋃ (⋃ (((`) ?Vt) ` (?Al ` {L. L ⊆ V})))" (is "?A = ?A'")
by (auto simp add: normalized_tableau_def cong del: image_cong_simp)
moreover
have "finite ?A'"
proof (rule finite_Union)
show "finite (⋃ (((`) ?Vt) ` (?Al ` {L. L ⊆ V})))"
using ‹finite {L. L ⊆ V}› ‹∀ L. L ⊆ V ⟶ finite (?Al L)›
by auto
next
fix M
assume "M ∈ ⋃ (((`) ?Vt) ` (?Al ` {L. L ⊆ V}))"
then obtain L t where "L ⊆ V" "t ∈ ?Al L" "M = ?Vt t"
by blast
then show "finite M"
using ‹∀ L t. L ⊆ V ∧ t ∈ ?Al L ⟶ finite (?Vt t)›
by blast
qed
ultimately
show ?thesis
by simp
qed
lemma finite_accessible_valuations:
shows "finite (𝒱 ` {s'. s ≻⇧* s'})"
proof-
have "{s'. s ≻⇧* s'} = {s'. s ≻⇧+ s'} ∪ {s}"
by (auto simp add: rtrancl_eq_or_trancl)
moreover
have "finite (𝒱 ` {s'. s ≻⇧+ s'})" (is "finite ?A")
proof-
let ?P = "λ v. check_valuation v (⟨𝒱 s⟩) (λ x. the (ℬ⇩l s x)) (λ x. the (ℬ⇩u s x)) (𝒯 s) (tvars (𝒯 s))"
let ?P' = "λ v::(var, 'a) mapping.
∃ t. tvars t = tvars (𝒯 s) ∧ △ t ∧ (∀ v::'a valuation. v ⊨⇩t t = v ⊨⇩t 𝒯 s) ∧ ⟨v⟩ ⊨⇩t t ∧
(∀ x ∈ rvars t. ⟨v⟩ x = ⟨𝒱 s⟩ x ∨
⟨v⟩ x = the (ℬ⇩l s x) ∨
⟨v⟩ x = the (ℬ⇩u s x)) ∧
(∀ x. x ∉ tvars (𝒯 s) ⟶ look v x = look (𝒱 s) x) ∧
(∀ x. x ∈ tvars (𝒯 s) ⟶ look v x ≠ None)"
have "finite (tvars (𝒯 s))"
by (auto simp add: lvars_def rvars_def finite_vars)
then have "finite {v. ?P v}"
using finite_valuations[of "tvars (𝒯 s)" "𝒯 s" "⟨𝒱 s⟩" "λ x. the (ℬ⇩l s x)" "λ x. the (ℬ⇩u s x)"]
by auto
moreover
have "map2fun ` {v. ?P' v} ⊆ {v. ?P v}"
by (auto simp add: map2fun_def)
ultimately
have "finite (map2fun ` {v. ?P' v})"
by (auto simp add: finite_subset)
moreover
have "inj_on map2fun {v. ?P' v}"
unfolding inj_on_def
proof (safe)
fix x y
assume "⟨x⟩ = ⟨y⟩" and *:
"∀x. x ∉ Simplex.tvars (𝒯 s) ⟶ look y x = look (𝒱 s) x"
"∀xa. xa ∉ Simplex.tvars (𝒯 s) ⟶ look x xa = look (𝒱 s) xa"
"∀x. x ∈ Simplex.tvars (𝒯 s) ⟶ look y x ≠ None"
"∀xa. xa ∈ Simplex.tvars (𝒯 s) ⟶ look x xa ≠ None"
show "x = y"
proof (rule mapping_eqI)
fix k
have "⟨x⟩ k = ⟨y⟩ k"
using ‹⟨x⟩ = ⟨y⟩›
by simp
then show "look x k = look y k"
using *
by (cases "k ∈ tvars (𝒯 s)") (auto simp add: map2fun_def split: option.split)
qed
qed
ultimately
have "finite {v. ?P' v}"
by (rule finite_imageD)
moreover
have "?A ⊆ {v. ?P' v}"
proof (safe)
fix s'
assume "s ≻⇧+ s'"
then obtain l where *: "l ≠ []" "1 < length l" "hd l = s" "last l = s'" "succ_chain l"
using trancl_rel_chain[of s s' succ_rel]
by auto
show "?P' (𝒱 s')"
proof-
have "∇ s" "△ (𝒯 s)" "⟨𝒱 s⟩ ⊨⇩t 𝒯 s"
using ‹s ≻⇧+ s'›
using tranclD[of s s' succ_rel]
by (auto simp add: curr_val_satisfies_no_lhs_def)
have "tvars (𝒯 s') = tvars (𝒯 s)"
using succ_chain_vars_id[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l]
by simp
moreover
have "△(𝒯 s')"
using ‹s ≻⇧+ s'›
using succ_inv(1)[of _ s']
by (auto dest: tranclD2)
moreover
have "∀v::'a valuation. v ⊨⇩t 𝒯 s' = v ⊨⇩t 𝒯 s"
using succ_chain_tableau_equiv[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l]
by auto
moreover
have "⟨𝒱 s'⟩ ⊨⇩t 𝒯 s'"
using succ_chain_valuation_satisfies[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l] ‹⟨𝒱 s⟩ ⊨⇩t 𝒯 s›
by simp
moreover
have "∀x∈rvars (𝒯 s'). ⟨𝒱 s'⟩ x = ⟨𝒱 s⟩ x ∨ ⟨𝒱 s'⟩ x = the (ℬ⇩l s x) ∨ ⟨𝒱 s'⟩ x = the (ℬ⇩u s x)"
using succ_chain_rvar_valuation[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l]
by auto
moreover
have "∀x. x ∉ tvars (𝒯 s) ⟶ look (𝒱 s') x = look (𝒱 s) x"
using succ_chain_no_vars_valuation[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l]
by auto
moreover
have "∀x. x ∈ Simplex.tvars (𝒯 s') ⟶ look (𝒱 s') x ≠ None"
using succ_chain_tableau_valuated[of l 0 "length l - 1"]
using * hd_conv_nth[of l] last_conv_nth[of l]
using ‹tvars (𝒯 s') = tvars (𝒯 s)› ‹∇ s›
by (auto simp add: tableau_valuated_def)
ultimately
show ?thesis
by (rule_tac x="𝒯 s'" in exI) auto
qed
qed
ultimately
show ?thesis
by (auto simp add: finite_subset)
qed
ultimately
show ?thesis
by simp
qed
lemma accessible_bounds:
shows "ℬ⇩i ` {s'. s ≻⇧* s'} = {ℬ⇩i s}"
proof -
have "s ≻⇧* s' ⟹ ℬ⇩i s' = ℬ⇩i s" for s'
by (induct s s' rule: rtrancl.induct, auto)
then show ?thesis by blast
qed
lemma accessible_unsat_core:
shows "𝒰⇩c ` {s'. s ≻⇧* s'} = {𝒰⇩c s}"
proof -
have "s ≻⇧* s' ⟹ 𝒰⇩c s' = 𝒰⇩c s" for s'
by (induct s s' rule: rtrancl.induct, auto)
then show ?thesis by blast
qed
lemma state_eqI:
"ℬ⇩i⇩l s = ℬ⇩i⇩l s' ⟹ ℬ⇩i⇩u s = ℬ⇩i⇩u s' ⟹
𝒯 s = 𝒯 s' ⟹ 𝒱 s = 𝒱 s' ⟹
𝒰 s = 𝒰 s' ⟹ 𝒰⇩c s = 𝒰⇩c s' ⟹
s = s'"
by (cases s, cases s', auto)
lemma finite_accessible_states:
shows "finite {s'. s ≻⇧* s'}" (is "finite ?A")
proof-
let ?V = "𝒱 ` ?A"
let ?T = "𝒯 ` ?A"
let ?P = "?V × ?T × {ℬ⇩i s} × {True, False} × {𝒰⇩c s}"
have "finite ?P"
using finite_accessible_valuations finite_accessible_tableaus
by auto
moreover
let ?f = "λ s. (𝒱 s, 𝒯 s, ℬ⇩i s, 𝒰 s, 𝒰⇩c s)"
have "?f ` ?A ⊆ ?P"
using accessible_bounds[of s] accessible_unsat_core[of s]
by auto
moreover
have "inj_on ?f ?A"
unfolding inj_on_def by (auto intro: state_eqI)
ultimately
show ?thesis
using finite_imageD [of ?f ?A]
using finite_subset
by auto
qed
lemma acyclic_suc_rel: "acyclic succ_rel"
proof (rule acyclicI, rule allI)
fix s
show "(s, s) ∉ succ_rel⇧+"
proof
assume "s ≻⇧+ s"
then obtain l where
"l ≠ []" "length l > 1" "hd l = s" "last l = s" "succ_chain l"
using trancl_rel_chain[of s s succ_rel]
by auto
have "l ! 0 = s"
using ‹l ≠ []› ‹hd l = s›
by (simp add: hd_conv_nth)
then have "s ≻ (l ! 1)"
using ‹succ_chain l›
unfolding rel_chain_def
using ‹length l > 1›
by auto
then have "△ (𝒯 s)"
by simp
let ?enter_rvars =
"{x. ∃ sl. swap_lr l sl x}"
have "finite ?enter_rvars"
proof-
let ?all_vars = "⋃ (set (map (λ t. lvars t ∪ rvars t) (map 𝒯 l)))"
have "finite ?all_vars"
by (auto simp add: lvars_def rvars_def finite_vars)
moreover
have "?enter_rvars ⊆ ?all_vars"
by force
ultimately
show ?thesis
by (simp add: finite_subset)
qed
let ?xr = "Max ?enter_rvars"
have "?xr ∈ ?enter_rvars"
proof (rule Max_in)
show "?enter_rvars ≠ {}"
proof-
from ‹s ≻ (l ! 1)›
obtain x⇩i x⇩j :: var where
"x⇩i ∈ lvars (𝒯 s)" "x⇩i ∈ rvars (𝒯 (l ! 1))"
by (rule succ_vars) auto
then have "x⇩i ∈ ?enter_rvars"
using ‹hd l = s› ‹l ≠ []› ‹length l > 1›
by (auto simp add: hd_conv_nth)
then show ?thesis
by auto
qed
next
show "finite ?enter_rvars"
using ‹finite ?enter_rvars›
.
qed
then obtain xr sl where
"xr = ?xr" "swap_lr l sl xr"
by auto
then have "sl + 1 < length l"
by simp
have "(l ! sl) ≻ (l ! (sl + 1))"
using ‹sl + 1 < length l› ‹succ_chain l›
unfolding rel_chain_def
by auto
have "length l > 2"
proof (rule ccontr)
assume "¬ ?thesis"
with ‹length l > 1›
have "length l = 2"
by auto
then have "last l = l ! 1"
by (cases l) (auto simp add: last_conv_nth nth_Cons split: nat.split)
then have "xr ∈ lvars (𝒯 s)" "xr ∈ rvars (𝒯 s)"
using ‹length l = 2›
using ‹swap_lr l sl xr›
using ‹hd l = s› ‹last l = s› ‹l ≠ []›
by (auto simp add: hd_conv_nth)
then show False
using ‹△ (𝒯 s)›
unfolding normalized_tableau_def
by auto
qed
obtain l' where
"hd l' = l ! (sl + 1)" "last l' = l ! sl" "length l' = length l - 1" "succ_chain l'" and
l'_l: "∀ i. i + 1 < length l' ⟶
(∃ j. j + 1 < length l ∧ l' ! i = l ! j ∧ l' ! (i + 1) = l ! (j + 1))"
using ‹length l > 2› ‹sl + 1 < length l› ‹hd l = s› ‹last l = s› ‹succ_chain l›
using reorder_cyclic_list[of l s sl]
by blast
then have "xr ∈ rvars (𝒯 (hd l'))" "xr ∈ lvars (𝒯 (last l'))" "length l' > 1" "l' ≠ []"
using ‹swap_lr l sl xr› ‹length l > 2›
by auto
then have "∃ sp. swap_rl l' sp xr"
using ‹succ_chain l'›
using succ_chain_swap_rl_exists[of l' 0 "length l' - 1" xr]
by (auto simp add: hd_conv_nth last_conv_nth)
then have "∃ sp. swap_rl l' sp xr ∧ (∀ sp'. sp' < sp ⟶ ¬ swap_rl l' sp' xr)"
by (rule min_element)
then obtain sp where
"swap_rl l' sp xr" "∀ sp'. sp' < sp ⟶ ¬ swap_rl l' sp' xr"
by blast
then have "sp + 1 < length l'"
by simp
have "⟨𝒱 (l' ! 0)⟩ xr = ⟨𝒱 (l' ! sp)⟩ xr"
proof-
have "always_r l' 0 sp xr"
using ‹xr ∈ rvars (𝒯 (hd l'))› ‹sp + 1 < length l'›
‹∀ sp'. sp' < sp ⟶ ¬ swap_rl l' sp' xr›
proof (induct sp)
case 0
then have "l' ≠ []"
by auto
then show ?case
using 0(1)
by (auto simp add: hd_conv_nth)
next
case (Suc sp')
show ?case
proof (safe)
fix k
assume "k ≤ Suc sp'"
show "xr ∈ rvars (𝒯 (l' ! k))"
proof (cases "k = sp' + 1")
case False
then show ?thesis
using Suc ‹k ≤ Suc sp'›
by auto
next
case True
then have "xr ∈ rvars (𝒯 (l' ! (k - 1)))"
using Suc
by auto
moreover
then have "xr ∉ lvars (𝒯 (l' ! k))"
using True Suc(3) Suc(4)
by auto
moreover
have "(l' ! (k - 1)) ≻ (l' ! k)"
using ‹succ_chain l'›
using Suc(3) True
by (simp add: rel_chain_def)
ultimately
show ?thesis
using succ_vars_id[of "l' ! (k - 1)" "l' ! k"]
by auto
qed
qed
qed
then show ?thesis
using ‹sp + 1 < length l'›
using ‹succ_chain l'›
using succ_chain_always_r_valuation_id
by simp
qed
have "(l' ! sp) ≻ (l' ! (sp+1))"
using ‹sp + 1 < length l'› ‹succ_chain l'›
unfolding rel_chain_def
by simp
then obtain xs xr' :: var where
"xs ∈ lvars (𝒯 (l' ! sp))"
"xr ∈ rvars (𝒯 (l' ! sp))"
"swap_lr l' sp xs"
apply (rule succ_vars)
using ‹swap_rl l' sp xr› ‹sp + 1 < length l'›
by auto
then have "xs ≠ xr"
using ‹(l' ! sp) ≻ (l' ! (sp+1))›
by (auto simp add: normalized_tableau_def)
obtain sp' where
"l' ! sp = l ! sp'" "l' ! (sp + 1) = l ! (sp' + 1)"
"sp' + 1 < length l"
using ‹sp + 1 < length l'› l'_l
by auto
have "xs ∈ ?enter_rvars"
using ‹swap_lr l' sp xs› l'_l
by force
have "xs < xr"
proof-
have "xs ≤ ?xr"
using ‹finite ?enter_rvars› ‹xs ∈ ?enter_rvars›
by (rule Max_ge)
then show ?thesis
using ‹xr = ?xr› ‹xs ≠ xr›
by simp
qed
let ?sl = "l ! sl"
let ?sp = "l' ! sp"
let ?eq = "eq_for_lvar (𝒯 ?sp) xs"
let ?bl = "𝒱 ?sl"
let ?bp = "𝒱 ?sp"
have "⊨⇩n⇩o⇩l⇩h⇩s ?sl" "⊨⇩n⇩o⇩l⇩h⇩s ?sp"
using ‹l ! sl ≻ (l ! (sl + 1))›
using ‹l' ! sp ≻ (l' ! (sp+ 1))›
by simp_all
have "ℬ⇩i ?sp = ℬ⇩i ?sl"
proof-
have "ℬ⇩i (l' ! sp) = ℬ⇩i (l' ! (length l' - 1))"
using ‹sp + 1 < length l'› ‹succ_chain l'›
using succ_chain_bounds_id
by auto
then have "ℬ⇩i (last l') = ℬ⇩i (l' ! sp)"
using ‹l' ≠ []›
by (simp add: last_conv_nth)
then show ?thesis
using ‹last l' = l ! sl›
by simp
qed
have diff_satified: "⟨?bl⟩ xs - ⟨?bp⟩ xs = ((rhs ?eq) ⦃ ⟨?bl⟩ ⦄) - ((rhs ?eq) ⦃ ⟨?bp⟩ ⦄)"
proof-
have "⟨?bp⟩ ⊨⇩e ?eq"
using ‹⊨⇩n⇩o⇩l⇩h⇩s ?sp›
using eq_for_lvar[of xs "𝒯 ?sp"]
using ‹xs ∈ lvars (𝒯 (l' ! sp))›
unfolding curr_val_satisfies_no_lhs_def satisfies_tableau_def
by auto
moreover
have "⟨?bl⟩ ⊨⇩e ?eq"
proof-
have "⟨𝒱 (l ! sl)⟩ ⊨⇩t 𝒯 (l' ! sp)"
using ‹l' ! sp = l ! sp'› ‹sp' + 1 < length l› ‹sl + 1 < length l›
using ‹succ_chain l›
using succ_chain_tableau_equiv[of l sl sp']
using ‹⊨⇩n⇩o⇩l⇩h⇩s ?sl›
unfolding curr_val_satisfies_no_lhs_def
by simp
then show ?thesis
unfolding satisfies_tableau_def
using eq_for_lvar
using ‹xs ∈ lvars (𝒯 (l' ! sp))›
by simp
qed
moreover
have "lhs ?eq = xs"
using ‹xs ∈ lvars (𝒯 (l' ! sp))›
using eq_for_lvar
by simp
ultimately
show ?thesis
unfolding satisfies_eq_def
by auto
qed
have "¬ in_bounds xr ⟨?bl⟩ (ℬ ?sl)"
using ‹l ! sl ≻ (l ! (sl + 1))› ‹swap_lr l sl xr›
using succ_min_lvar_not_in_bounds(1)[of ?sl "l ! (sl + 1)" xr]
by simp
have "∀ x. x < xr ⟶ in_bounds x ⟨?bl⟩ (ℬ ?sl)"
proof (safe)
fix x
assume "x < xr"
show "in_bounds x ⟨?bl⟩ (ℬ ?sl)"
proof (cases "x ∈ lvars (𝒯 ?sl)")
case True
then show ?thesis
using succ_min_lvar_not_in_bounds(2)[of ?sl "l ! (sl + 1)" xr]
using ‹l ! sl ≻ (l ! (sl + 1))› ‹swap_lr l sl xr› ‹x < xr›
by simp
next
case False
then show ?thesis
using ‹⊨⇩n⇩o⇩l⇩h⇩s ?sl›
unfolding curr_val_satisfies_no_lhs_def
by (simp add: satisfies_bounds_set.simps)
qed
qed
then have "in_bounds xs ⟨?bl⟩ (ℬ ?sl)"
using ‹xs < xr›
by simp
have "¬ in_bounds xs ⟨?bp⟩ (ℬ ?sp)"
using ‹l' ! sp ≻ (l' ! (sp + 1))› ‹swap_lr l' sp xs›
using succ_min_lvar_not_in_bounds(1)[of ?sp "l' ! (sp + 1)" xs]
by simp
have "∀ x ∈ rvars_eq ?eq. x > xr ⟶ ⟨?bp⟩ x = ⟨?bl⟩ x"
proof (safe)
fix x
assume "x ∈ rvars_eq ?eq" "x > xr"
then have "always_r l' 0 (length l' - 1) x"
proof (safe)
fix k
assume "x ∈ rvars_eq ?eq" "x > xr" "0 ≤ k" "k ≤ length l' - 1"
obtain k' where "l ! k' = l' ! k" "k' < length l"
using l'_l ‹k ≤ length l' - 1› ‹length l' > 1›
apply (cases "k > 0")
apply (erule_tac x="k - 1" in allE)
apply (drule mp)
by auto
let ?eq' = "eq_for_lvar (𝒯 (l ! sp')) xs"
have "∀ x ∈ rvars_eq ?eq'. x > xr ⟶ always_r l 0 (length l - 1) x"
proof (safe)
fix x k
assume "x ∈ rvars_eq ?eq'" "xr < x" "0 ≤ k" "k ≤ length l - 1"
then have "x ∈ rvars (𝒯 (l ! sp'))"
using eq_for_lvar[of xs "𝒯 (l ! sp')"]
using ‹swap_lr l' sp xs› ‹l' ! sp = l ! sp'›
by (auto simp add: rvars_def)
have *: "∀ i. i < sp' ⟶ x ∈ rvars (𝒯 (l ! i))"
proof (safe, rule ccontr)
fix i
assume "i < sp'" "x ∉ rvars (𝒯 (l ! i))"
then have "x ∈ lvars (𝒯 (l ! i))"
using ‹x ∈ rvars (𝒯 (l ! sp'))›
using ‹sp' + 1 < length l›
using ‹succ_chain l›
using succ_chain_vars_id[of l i sp']
by auto
obtain i' where "swap_lr l i' x"
using ‹x ∈ lvars (𝒯 (l ! i))›
using ‹x ∈ rvars (𝒯 (l ! sp'))›
using ‹i < sp'› ‹sp' + 1 < length l›
using ‹succ_chain l›
using succ_chain_swap_lr_exists[of l i sp' x]
by auto
then have "x ∈ ?enter_rvars"
by auto
then have "x ≤ ?xr"
using ‹finite ?enter_rvars›
using Max_ge[of ?enter_rvars x]
by simp
then show False
using ‹x > xr›
using ‹xr = ?xr›
by simp
qed
then have "x ∈ rvars (𝒯 (last l))"
using ‹hd l = s› ‹last l = s› ‹l ≠ []›
using ‹x ∈ rvars (𝒯 (l ! sp'))›
by (auto simp add: hd_conv_nth)
show "x ∈ rvars (𝒯 (l ! k))"
proof (cases "k = length l - 1")
case True
then show ?thesis
using ‹x ∈ rvars (𝒯 (last l))›
using ‹l ≠ []›
by (simp add: last_conv_nth)
next
case False
then have "k < length l - 1"
using ‹k ≤ length l - 1›
by simp
then have "k < length l"
using ‹length l > 1›
by auto
show ?thesis
proof (rule ccontr)
assume "¬ ?thesis"
then have "x ∈ lvars (𝒯 (l ! k))"
using ‹x ∈ rvars (𝒯 (l ! sp'))›
using ‹sp' + 1 < length l› ‹k < length l›
using succ_chain_vars_id[of l k sp']
using ‹succ_chain l› ‹l ≠ []›
by auto
obtain i' where "swap_lr l i' x"
using ‹succ_chain l›
using ‹x ∈ lvars (𝒯 (l ! k))›
using ‹x ∈ rvars (𝒯 (last l))›
using ‹k < length l - 1› ‹l ≠ []›
using succ_chain_swap_lr_exists[of l k "length l - 1" x]
by (auto simp add: last_conv_nth)
then have "x ∈ ?enter_rvars"
by auto
then have "x ≤ ?xr"
using ‹finite ?enter_rvars›
using Max_ge[of ?enter_rvars x]
by simp
then show False
using ‹x > xr›
using ‹xr = ?xr›
by simp
qed
qed
qed
then have "x ∈ rvars (𝒯 (l ! k'))"
using ‹x ∈ rvars_eq ?eq› ‹x > xr› ‹k' < length l›
using ‹l' ! sp = l ! sp'›
by simp
then show "x ∈ rvars (𝒯 (l' ! k))"
using ‹l ! k' = l' ! k›
by simp
qed
then have "⟨?bp⟩ x = ⟨𝒱 (l' ! (length l' - 1))⟩ x"
using ‹succ_chain l'› ‹sp + 1 < length l'›
by (auto intro!: succ_chain_always_r_valuation_id[rule_format])
then have "⟨?bp⟩ x = ⟨𝒱 (last l')⟩ x"
using ‹l' ≠ []›
by (simp add: last_conv_nth)
then show "⟨?bp⟩ x = ⟨?bl⟩ x"
using ‹last l' = l ! sl›
by simp
qed
have "⟨?bp⟩ xr = ⟨𝒱 (l ! (sl + 1))⟩ xr"
using ‹⟨𝒱 (l' ! 0)⟩ xr = ⟨𝒱 (l' ! sp)⟩ xr›
using ‹hd l' = l ! (sl + 1)› ‹l' ≠ []›
by (simp add: hd_conv_nth)
{
fix dir1 dir2 :: "('i,'a) Direction"
assume dir1: "dir1 = (if ⟨?bl⟩ xr <⇩l⇩b ℬ⇩l ?sl xr then Positive else Negative)"
then have "⊲⇩l⇩b (lt dir1) (⟨?bl⟩ xr) (LB dir1 ?sl xr)"
using ‹¬ in_bounds xr ⟨?bl⟩ (ℬ ?sl)›
using neg_bounds_compare(7) neg_bounds_compare(3)
by (auto simp add: bound_compare''_defs)
then have "¬ ⊵⇩l⇩b (lt dir1) (⟨?bl⟩ xr) (LB dir1 ?sl xr)"
using bounds_compare_contradictory(7) bounds_compare_contradictory(3) neg_bounds_compare(6) dir1
unfolding bound_compare''_defs
by auto force
have "LB dir1 ?sl xr ≠ None"
using ‹⊲⇩l⇩b (lt dir1) (⟨?bl⟩ xr) (LB dir1 ?sl xr)›
by (cases "LB dir1 ?sl xr") (auto simp add: bound_compare_defs)
assume dir2: "dir2 = (if ⟨?bp⟩ xs <⇩l⇩b ℬ⇩l ?sp xs then Positive else Negative)"
then have "⊲⇩l⇩b (lt dir2) (⟨?bp⟩ xs) (LB dir2 ?sp xs)"
using ‹¬ in_bounds xs ⟨?bp⟩ (ℬ ?sp)›
using neg_bounds_compare(2) neg_bounds_compare(6)
by (auto simp add: bound_compare''_defs)
then have "¬ ⊵⇩l⇩b (lt dir2) (⟨?bp⟩ xs) (LB dir2 ?sp xs)"
using bounds_compare_contradictory(3) bounds_compare_contradictory(7) neg_bounds_compare(6) dir2
unfolding bound_compare''_defs
by auto force
then have "∀ x ∈ rvars_eq ?eq. x < xr ⟶ ¬ reasable_var dir2 x ?eq ?sp"
using succ_min_rvar[of ?sp "l' ! (sp + 1)" xs xr ?eq]
using ‹l' ! sp ≻ (l' ! (sp + 1))›
using ‹swap_lr l' sp xs› ‹swap_rl l' sp xr› dir2
unfolding bound_compare''_defs
by auto
have "LB dir2 ?sp xs ≠ None"
using ‹⊲⇩l⇩b (lt dir2) (⟨?bp⟩ xs) (LB dir2 ?sp xs)›
by (cases "LB dir2 ?sp xs") (auto simp add: bound_compare_defs)
have *: "∀ x ∈ rvars_eq ?eq. x < xr ⟶
((coeff (rhs ?eq) x > 0 ⟶ ⊵⇩u⇩b (lt dir2) (⟨?bp⟩ x) (UB dir2 ?sp x)) ∧
(coeff (rhs ?eq) x < 0 ⟶ ⊴⇩l⇩b (lt dir2) (⟨?bp⟩ x) (LB dir2 ?sp x)))"
proof (safe)
fix x
assume "x ∈ rvars_eq ?eq" "x < xr" "coeff (rhs ?eq) x > 0"
then have "¬ ⊲⇩u⇩b (lt dir2) (⟨?bp⟩ x) (UB dir2 ?sp x)"
using ‹∀ x ∈ rvars_eq ?eq. x < xr ⟶ ¬ reasable_var dir2 x ?eq ?sp›
by simp
then show "⊵⇩u⇩b (lt dir2) (⟨?bp⟩ x) (UB dir2 ?sp x)"
using dir2 neg_bounds_compare(4) neg_bounds_compare(8)
unfolding bound_compare''_defs
by force
next
fix x
assume "x ∈ rvars_eq ?eq" "x < xr" "coeff (rhs ?eq) x < 0"
then have "¬ ⊳⇩l⇩b (lt dir2) (⟨?bp⟩ x) (LB dir2 ?sp x)"
using ‹∀ x ∈ rvars_eq ?eq. x < xr ⟶ ¬ reasable_var dir2 x ?eq ?sp›
by simp
then show "⊴⇩l⇩b (lt dir2) (⟨?bp⟩ x) (LB dir2 ?sp x)"
using dir2 neg_bounds_compare(4) neg_bounds_compare(8) dir2
unfolding bound_compare''_defs
by force
qed
have "(lt dir2) (⟨?bp⟩ xs) (⟨?bl⟩ xs)"
using ‹⊲⇩l⇩b (lt dir2) (⟨?bp⟩ xs) (LB dir2 ?sp xs)›
using ‹ℬ⇩i ?sp = ℬ⇩i ?sl› dir2
using ‹in_bounds xs ⟨?bl⟩ (ℬ ?sl)›
by (auto simp add: bound_compare''_defs
simp: indexl_def indexu_def boundsl_def boundsu_def)
then have "(lt dir2) 0 (⟨?bl⟩ xs - ⟨?bp⟩ xs)"
using dir2
by (auto simp add: minus_gt[THEN sym] minus_lt[THEN sym])
moreover
have "le (lt dir2) ((rhs ?eq) ⦃ ⟨?bl⟩ ⦄ - (rhs ?eq) ⦃ ⟨?bp⟩ ⦄) 0"
proof-
have "∀ x ∈ rvars_eq ?eq. (0 < coeff (rhs ?eq) x ⟶ le (lt dir2) 0 (⟨?bp⟩ x - ⟨?bl⟩ x)) ∧
(coeff (rhs ?eq) x < 0 ⟶ le (lt dir2) (⟨?bp⟩ x - ⟨?bl⟩ x) 0)"
proof
fix x
assume "x ∈ rvars_eq ?eq"
show "(0 < coeff (rhs ?eq) x ⟶ le (lt dir2) 0 (⟨?bp⟩ x - ⟨?bl⟩ x)) ∧
(coeff (rhs ?eq) x < 0 ⟶ le (lt dir2) (⟨?bp⟩ x - ⟨?bl⟩ x) 0)"
proof (cases "x < xr")
case True
then have "in_bounds x ⟨?bl⟩ (ℬ ?sl)"
using ‹∀ x. x < xr ⟶ in_bounds x ⟨?bl⟩ (ℬ ?sl)›
by simp
show ?thesis
proof (safe)
assume "coeff (rhs ?eq) x > 0" "0 ≠ ⟨?bp⟩ x - ⟨?bl⟩ x"
then have "⊵⇩u⇩b (lt dir2) (⟨𝒱 (l' ! sp)⟩ x) (UB dir2 (l' ! sp) x)"
using * ‹x < xr› ‹x ∈ rvars_eq ?eq›
by simp
then have "le (lt dir2) (⟨?bl⟩ x) (⟨?bp⟩ x)"
using ‹in_bounds x ⟨?bl⟩ (ℬ ?sl)› ‹ℬ⇩i ?sp = ℬ⇩i ?sl› dir2
apply (auto simp add: bound_compare''_defs)
using bounds_lg(3)[of "⟨?bp⟩ x" "ℬ⇩u (l ! sl) x" "⟨?bl⟩ x"]
using bounds_lg(6)[of "⟨?bp⟩ x" "ℬ⇩l (l ! sl) x" "⟨?bl⟩ x"]
unfolding bound_compare''_defs
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
then show "lt dir2 0 (⟨?bp⟩ x - ⟨?bl⟩ x)"
using ‹0 ≠ ⟨?bp⟩ x - ⟨?bl⟩ x›
using minus_gt[of "⟨?bl⟩ x" "⟨?bp⟩ x"] minus_lt[of "⟨?bp⟩ x" "⟨?bl⟩ x"] dir2
by auto
next
assume "coeff (rhs ?eq) x < 0" "⟨?bp⟩ x - ⟨?bl⟩ x ≠ 0"
then have "⊴⇩l⇩b (lt dir2) (⟨𝒱 (l' ! sp)⟩ x) (LB dir2 (l' ! sp) x)"
using * ‹x < xr› ‹x ∈ rvars_eq ?eq›
by simp
then have "le (lt dir2) (⟨?bp⟩ x) (⟨?bl⟩ x)"
using ‹in_bounds x ⟨?bl⟩ (ℬ ?sl)› ‹ℬ⇩i ?sp = ℬ⇩i ?sl› dir2
apply (auto simp add: bound_compare''_defs)
using bounds_lg(3)[of "⟨?bp⟩ x" "ℬ⇩u (l ! sl) x" "⟨?bl⟩ x"]
using bounds_lg(6)[of "⟨?bp⟩ x" "ℬ⇩l (l ! sl) x" "⟨?bl⟩ x"]
unfolding bound_compare''_defs
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
then show "lt dir2 (⟨?bp⟩ x - ⟨?bl⟩ x) 0"
using ‹⟨?bp⟩ x - ⟨?bl⟩ x ≠ 0›
using minus_gt[of "⟨?bl⟩ x" "⟨?bp⟩ x"] minus_lt[of "⟨?bp⟩ x" "⟨?bl⟩ x"] dir2
by auto
qed
next
case False
show ?thesis
proof (cases "x = xr")
case True
have "⟨𝒱 (l ! (sl + 1))⟩ xr = the (LB dir1 ?sl xr)"
using ‹l ! sl ≻ (l ! (sl + 1))›
using ‹swap_lr l sl xr›
using succ_set_on_bound(1)[of "l ! sl" "l ! (sl + 1)" xr]
using ‹¬ ⊵⇩l⇩b (lt dir1) (⟨?bl⟩ xr) (LB dir1 ?sl xr)› dir1
unfolding bound_compare''_defs
by auto
then have "⟨?bp⟩ xr = the (LB dir1 ?sl xr)"
using ‹⟨?bp⟩ xr = ⟨𝒱 (l ! (sl + 1))⟩ xr›
by simp
then have "lt dir1 (⟨?bl⟩ xr) (⟨?bp⟩ xr)"
using ‹LB dir1 ?sl xr ≠ None›
using ‹⊲⇩l⇩b (lt dir1) (⟨?bl⟩ xr) (LB dir1 ?sl xr)› dir1
by (auto simp add: bound_compare_defs)
moreover
have "reasable_var dir2 xr ?eq ?sp"
using ‹¬ ⊵⇩l⇩b (lt dir2) (⟨?bp⟩ xs) (LB dir2 ?sp xs)›
using ‹l' ! sp ≻ (l' ! (sp + 1))›
using ‹swap_lr l' sp xs› ‹swap_rl l' sp xr›
using succ_min_rvar[of "l' ! sp" "l' ! (sp + 1)"xs xr ?eq] dir2
unfolding bound_compare''_defs
by auto
then have "if dir1 = dir2 then coeff (rhs ?eq) xr > 0 else coeff (rhs ?eq) xr < 0"
using ‹⟨?bp⟩ xr = the (LB dir1 ?sl xr)›
using ‹ℬ⇩i ?sp = ℬ⇩i ?sl›[THEN sym] dir1
using ‹LB dir1 ?sl xr ≠ None› dir1 dir2
by (auto split: if_splits simp add: bound_compare_defs
indexl_def indexu_def boundsl_def boundsu_def)
moreover
have "dir1 = Positive ∨ dir1 = Negative" "dir2 = Positive ∨ dir2 = Negative"
using dir1 dir2
by auto
ultimately
show ?thesis
using ‹x = xr›
using minus_lt[of "⟨?bp⟩ xr" "⟨?bl⟩ xr"] minus_gt[of "⟨?bl⟩ xr" "⟨?bp⟩ xr"]
by (auto split: if_splits)
next
case False
then have "x > xr"
using ‹¬ x < xr›
by simp
then have "⟨?bp⟩ x = ⟨?bl⟩ x"
using ‹∀ x ∈ rvars_eq ?eq. x > xr ⟶ ⟨?bp⟩ x = ⟨?bl⟩ x›
using ‹x ∈ rvars_eq ?eq›
by simp
then show ?thesis
by simp
qed
qed
qed
then have "le (lt dir2) 0 (rhs ?eq ⦃ λ x. ⟨?bp⟩ x - ⟨?bl⟩ x ⦄)"
using dir2
apply auto
using valuate_nonneg[of "rhs ?eq" "λ x. ⟨?bp⟩ x - ⟨?bl⟩ x"]
apply force
using valuate_nonpos[of "rhs ?eq" "λ x. ⟨?bp⟩ x - ⟨?bl⟩ x"]
apply force
done
then have "le (lt dir2) 0 ((rhs ?eq) ⦃ ⟨?bp⟩ ⦄ - (rhs ?eq) ⦃ ⟨?bl⟩ ⦄)"
by (subst valuate_diff)+ simp
then have "le (lt dir2) ((rhs ?eq) ⦃ ⟨?bl⟩ ⦄) ((rhs ?eq) ⦃ ⟨?bp⟩ ⦄)"
using minus_lt[of "(rhs ?eq) ⦃ ⟨?bp⟩ ⦄" "(rhs ?eq) ⦃ ⟨?bl⟩ ⦄"] dir2
by auto
then show ?thesis
using dir2
using minus_lt[of "(rhs ?eq) ⦃ ⟨?bl⟩ ⦄" "(rhs ?eq) ⦃ ⟨?bp⟩ ⦄"]
using minus_gt[of "(rhs ?eq) ⦃ ⟨?bp⟩ ⦄" "(rhs ?eq) ⦃ ⟨?bl⟩ ⦄"]
by auto
qed
ultimately
have False
using diff_satified dir2
by (auto split: if_splits)
}
then show False
by auto
qed
qed
lemma check_unsat_terminates:
assumes "𝒰 s"
shows "check_dom s"
by (rule check_dom.intros) (auto simp add: assms)
lemma check_sat_terminates'_aux:
assumes
dir: "dir = (if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then Positive else Negative)" and
*: "⋀ s'. ⟦s ≻ s'; ∇ s'; △ (𝒯 s'); ◇ s'; ⊨⇩n⇩o⇩l⇩h⇩s s' ⟧ ⟹ check_dom s'" and
"∇ s" "△ (𝒯 s)" "◇ s" "⊨⇩n⇩o⇩l⇩h⇩s s"
"¬ 𝒰 s" "min_lvar_not_in_bounds s = Some x⇩i"
"⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i)"
shows "check_dom
(case min_rvar_incdec dir s x⇩i of Inl I ⇒ set_unsat I s
| Inr x⇩j ⇒ pivot_and_update x⇩i x⇩j (the (LB dir s x⇩i)) s)"
proof (cases "min_rvar_incdec dir s x⇩i")
case Inl
then show ?thesis
using check_unsat_terminates by simp
next
case (Inr x⇩j)
then have xj: "x⇩j ∈ rvars_of_lvar (𝒯 s) x⇩i"
using min_rvar_incdec_eq_Some_rvars[of _ s "eq_for_lvar (𝒯 s) x⇩i" x⇩j]
using dir
by simp
let ?s' = "pivot_and_update x⇩i x⇩j (the (LB dir s x⇩i)) s"
have "check_dom ?s'"
proof (rule * )
show **: "∇ ?s'" "△ (𝒯 ?s')" "◇ ?s'" "⊨⇩n⇩o⇩l⇩h⇩s ?s'"
using ‹min_lvar_not_in_bounds s = Some x⇩i› Inr
using ‹∇ s› ‹△ (𝒯 s)› ‹◇ s› ‹⊨⇩n⇩o⇩l⇩h⇩s s› dir
using pivotandupdate_check_precond
by auto
have xi: "x⇩i ∈ lvars (𝒯 s)"
using assms(8) min_lvar_not_in_bounds_lvars by blast
show "s ≻ ?s'"
unfolding gt_state_def
using ‹△ (𝒯 s)› ‹◇ s› ‹⊨⇩n⇩o⇩l⇩h⇩s s› ‹∇ s›
using ‹min_lvar_not_in_bounds s = Some x⇩i› ‹⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i)›
Inr dir
by (intro conjI pivotandupdate_bounds_id pivotandupdate_unsat_core_id,
auto intro!: xj xi)
qed
then show ?thesis using Inr by simp
qed
lemma check_sat_terminates':
assumes "∇ s" "△ (𝒯 s)" "◇ s" "⊨⇩n⇩o⇩l⇩h⇩s s" "s⇩0 ≻⇧* s"
shows "check_dom s"
using assms
proof (induct s rule: wf_induct[of "{(y, x). s⇩0 ≻⇧* x ∧ x ≻ y}"])
show "wf {(y, x). s⇩0 ≻⇧* x ∧ x ≻ y}"
proof (rule finite_acyclic_wf)
let ?A = "{(s', s). s⇩0 ≻⇧* s ∧ s ≻ s'}"
let ?B = "{s. s⇩0 ≻⇧* s}"
have "?A ⊆ ?B × ?B"
proof
fix p
assume "p ∈ ?A"
then have "fst p ∈ ?B" "snd p ∈ ?B"
using rtrancl_into_trancl1[of s⇩0 "snd p" succ_rel "fst p"]
by auto
then show "p ∈ ?B × ?B"
using mem_Sigma_iff[of "fst p" "snd p"]
by auto
qed
then show "finite ?A"
using finite_accessible_states[of s⇩0]
using finite_subset[of ?A "?B × ?B"]
by simp
show "acyclic ?A"
proof-
have "?A ⊆ succ_rel¯"
by auto
then show ?thesis
using acyclic_converse acyclic_subset
using acyclic_suc_rel
by auto
qed
qed
next
fix s
assume "∀ s'. (s', s) ∈ {(y, x). s⇩0 ≻⇧* x ∧ x ≻ y} ⟶ ∇ s' ⟶ △ (𝒯 s') ⟶ ◇ s' ⟶ ⊨⇩n⇩o⇩l⇩h⇩s s' ⟶ s⇩0 ≻⇧* s' ⟶ check_dom s'"
"∇ s" "△ (𝒯 s)" "◇ s" " ⊨⇩n⇩o⇩l⇩h⇩s s" "s⇩0 ≻⇧* s"
then have *: "⋀ s'. ⟦s ≻ s'; ∇ s'; △ (𝒯 s'); ◇ s'; ⊨⇩n⇩o⇩l⇩h⇩s s' ⟧ ⟹ check_dom s'"
using rtrancl_into_trancl1[of s⇩0 s succ_rel]
using trancl_into_rtrancl[of s⇩0 _ succ_rel]
by auto
show "check_dom s"
proof (rule check_dom.intros, simp_all add: check'_def, unfold Positive_def[symmetric], unfold Negative_def[symmetric])
fix x⇩i
assume "¬ 𝒰 s" "Some x⇩i = min_lvar_not_in_bounds s" "⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i"
have "ℬ⇩l s x⇩i = LB Positive s x⇩i"
by simp
show "check_dom
(case min_rvar_incdec Positive s x⇩i of
Inl I ⇒ set_unsat I s
| Inr x⇩j ⇒ pivot_and_update x⇩i x⇩j (the (ℬ⇩l s x⇩i)) s)"
apply (subst ‹ℬ⇩l s x⇩i = LB Positive s x⇩i›)
apply (rule check_sat_terminates'_aux[of Positive s x⇩i])
using ‹∇ s› ‹△ (𝒯 s)› ‹◇ s› ‹⊨⇩n⇩o⇩l⇩h⇩s s› *
using ‹¬ 𝒰 s› ‹Some x⇩i = min_lvar_not_in_bounds s› ‹⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i›
by (simp_all add: bound_compare''_defs)
next
fix x⇩i
assume "¬ 𝒰 s" "Some x⇩i = min_lvar_not_in_bounds s" "¬ ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i"
then have "⟨𝒱 s⟩ x⇩i >⇩u⇩b ℬ⇩u s x⇩i"
using min_lvar_not_in_bounds_Some[of s x⇩i]
using neg_bounds_compare(7) neg_bounds_compare(2)
by auto
have "ℬ⇩u s x⇩i = LB Negative s x⇩i"
by simp
show "check_dom
(case min_rvar_incdec Negative s x⇩i of
Inl I ⇒ set_unsat I s
| Inr x⇩j ⇒ pivot_and_update x⇩i x⇩j (the (ℬ⇩u s x⇩i)) s)"
apply (subst ‹ℬ⇩u s x⇩i = LB Negative s x⇩i›)
apply (rule check_sat_terminates'_aux)
using ‹∇ s› ‹△ (𝒯 s)› ‹◇ s› ‹⊨⇩n⇩o⇩l⇩h⇩s s› *
using ‹¬ 𝒰 s› ‹Some x⇩i = min_lvar_not_in_bounds s› ‹⟨𝒱 s⟩ x⇩i >⇩u⇩b ℬ⇩u s x⇩i› ‹¬ ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i›
by (simp_all add: bound_compare''_defs)
qed
qed
lemma check_sat_terminates:
assumes "∇ s" "△ (𝒯 s)" "◇ s" "⊨⇩n⇩o⇩l⇩h⇩s s"
shows "check_dom s"
using assms
using check_sat_terminates'[of s s]
by simp
lemma check_cases:
assumes "𝒰 s ⟹ P s"
assumes "⟦¬ 𝒰 s; min_lvar_not_in_bounds s = None⟧ ⟹ P s"
assumes "⋀ x⇩i dir I.
⟦dir = Positive ∨ dir = Negative;
¬ 𝒰 s; min_lvar_not_in_bounds s = Some x⇩i;
⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i);
min_rvar_incdec dir s x⇩i = Inl I⟧ ⟹
P (set_unsat I s)"
assumes "⋀ x⇩i x⇩j l⇩i dir.
⟦dir = (if ⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i then Positive else Negative);
¬ 𝒰 s; min_lvar_not_in_bounds s = Some x⇩i;
⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i);
min_rvar_incdec dir s x⇩i = Inr x⇩j;
l⇩i = the (LB dir s x⇩i);
check' dir x⇩i s = pivot_and_update x⇩i x⇩j l⇩i s⟧ ⟹
P (check (pivot_and_update x⇩i x⇩j l⇩i s))"
assumes "△ (𝒯 s)" "◇ s" "⊨⇩n⇩o⇩l⇩h⇩s s"
shows "P (check s)"
proof (cases "𝒰 s")
case True
then show ?thesis
using assms(1)
using check.simps[of s]
by simp
next
case False
show ?thesis
proof (cases "min_lvar_not_in_bounds s")
case None
then show ?thesis
using ‹¬ 𝒰 s›
using assms(2) ‹△ (𝒯 s)› ‹◇ s› ‹⊨⇩n⇩o⇩l⇩h⇩s s›
using check.simps[of s]
by simp
next
case (Some x⇩i)
let ?dir = "if (⟨𝒱 s⟩ x⇩i <⇩l⇩b ℬ⇩l s x⇩i) then (Positive :: ('i,'a)Direction) else Negative"
let ?s' = "check' ?dir x⇩i s"
have "⊲⇩l⇩b (lt ?dir) (⟨𝒱 s⟩ x⇩i) (LB ?dir s x⇩i)"
using ‹min_lvar_not_in_bounds s = Some x⇩i›
using min_lvar_not_in_bounds_Some[of s x⇩i]
using not_in_bounds[of x⇩i "⟨𝒱 s⟩" "ℬ⇩l s" "ℬ⇩u s"]
by (auto split: if_splits simp add: bound_compare''_defs)
have "P (check ?s')"
apply (rule check'_cases)
using ‹¬ 𝒰 s› ‹min_lvar_not_in_bounds s = Some x⇩i› ‹⊲⇩l⇩b (lt ?dir) (⟨𝒱 s⟩ x⇩i) (LB ?dir s x⇩i)›
using assms(3)[of ?dir x⇩i]
using assms(4)[of ?dir x⇩i]
using check.simps[of "set_unsat (_ :: 'i list) s"]
using ‹△ (𝒯 s)› ‹◇ s› ‹⊨⇩n⇩o⇩l⇩h⇩s s›
by (auto simp add: bounds_consistent_def curr_val_satisfies_no_lhs_def)
then show ?thesis
using ‹¬ 𝒰 s› ‹min_lvar_not_in_bounds s = Some x⇩i›
using check.simps[of s]
using ‹△ (𝒯 s)› ‹◇ s› ‹⊨⇩n⇩o⇩l⇩h⇩s s›
by auto
qed
qed
lemma check_induct:
fixes s :: "('i,'a) state"
assumes *: "∇ s" "△ (𝒯 s)" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s"
assumes **:
"⋀ s. 𝒰 s ⟹ P s s"
"⋀ s. ⟦¬ 𝒰 s; min_lvar_not_in_bounds s = None⟧ ⟹ P s s"
"⋀ s x⇩i dir I. ⟦dir = Positive ∨ dir = Negative; ¬ 𝒰 s; min_lvar_not_in_bounds s = Some x⇩i;
⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i); min_rvar_incdec dir s x⇩i = Inl I⟧
⟹ P s (set_unsat I s)"
assumes step': "⋀ s x⇩i x⇩j l⇩i. ⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i)⟧ ⟹ P s (pivot_and_update x⇩i x⇩j l⇩i s)"
assumes trans': "⋀ si sj sk. ⟦P si sj; P sj sk⟧ ⟹ P si sk"
shows "P s (check s)"
proof-
have "check_dom s"
using *
by (simp add: check_sat_terminates)
then show ?thesis
using *
proof (induct s rule: check_dom.induct)
case (step s')
show ?case
proof (rule check_cases)
fix x⇩i x⇩j l⇩i dir
let ?dir = "if ⟨𝒱 s'⟩ x⇩i <⇩l⇩b ℬ⇩l s' x⇩i then Positive else Negative"
let ?s' = "check' dir x⇩i s'"
assume "¬ 𝒰 s'" "min_lvar_not_in_bounds s' = Some x⇩i" "min_rvar_incdec dir s' x⇩i = Inr x⇩j" "l⇩i = the (LB dir s' x⇩i)"
"?s' = pivot_and_update x⇩i x⇩j l⇩i s'" "dir = ?dir"
moreover
then have "∇ ?s'" "△ (𝒯 ?s')" "⊨⇩n⇩o⇩l⇩h⇩s ?s'" "◇ ?s'"
using ‹∇ s'› ‹△ (𝒯 s')› ‹⊨⇩n⇩o⇩l⇩h⇩s s'› ‹◇ s'›
using ‹?s' = pivot_and_update x⇩i x⇩j l⇩i s'›
using pivotandupdate_check_precond[of dir s' x⇩i x⇩j l⇩i]
by auto
ultimately
have "P (check' dir x⇩i s') (check (check' dir x⇩i s'))"
using step(2)[of x⇩i] step(4)[of x⇩i] ‹△ (𝒯 s')› ‹∇ s'›
by auto
then show "P s' (check (pivot_and_update x⇩i x⇩j l⇩i s'))"
using ‹?s' = pivot_and_update x⇩i x⇩j l⇩i s'› ‹△ (𝒯 s')› ‹∇ s'›
using ‹min_lvar_not_in_bounds s' = Some x⇩i› ‹min_rvar_incdec dir s' x⇩i = Inr x⇩j›
using step'[of s' x⇩i x⇩j l⇩i]
using trans'[of s' ?s' "check ?s'"]
by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars)
qed (simp_all add: ‹∇ s'› ‹△ (𝒯 s')› ‹⊨⇩n⇩o⇩l⇩h⇩s s'› ‹◇ s'› **)
qed
qed
lemma check_induct':
fixes s :: "('i,'a) state"
assumes "∇ s" "△ (𝒯 s)" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s"
assumes "⋀ s x⇩i dir I. ⟦dir = Positive ∨ dir = Negative; ¬ 𝒰 s; min_lvar_not_in_bounds s = Some x⇩i;
⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i); min_rvar_incdec dir s x⇩i = Inl I; P s⟧
⟹ P (set_unsat I s)"
assumes "⋀ s x⇩i x⇩j l⇩i. ⟦△ (𝒯 s); ∇ s; x⇩i ∈ lvars (𝒯 s); x⇩j ∈ rvars_eq (eq_for_lvar (𝒯 s) x⇩i); P s⟧ ⟹ P (pivot_and_update x⇩i x⇩j l⇩i s)"
assumes "P s"
shows "P (check s)"
proof-
have "P s ⟶ P (check s)"
by (rule check_induct) (simp_all add: assms)
then show ?thesis
using ‹P s›
by simp
qed
lemma check_induct'':
fixes s :: "('i,'a) state"
assumes *: "∇ s" "△ (𝒯 s)" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s"
assumes **:
"𝒰 s ⟹ P s"
"⋀ s. ⟦∇ s; △ (𝒯 s); ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; ¬ 𝒰 s; min_lvar_not_in_bounds s = None⟧ ⟹ P s"
"⋀ s x⇩i dir I. ⟦dir = Positive ∨ dir = Negative; ∇ s; △ (𝒯 s); ⊨⇩n⇩o⇩l⇩h⇩s s; ◇ s; ¬ 𝒰 s;
min_lvar_not_in_bounds s = Some x⇩i; ⊲⇩l⇩b (lt dir) (⟨𝒱 s⟩ x⇩i) (LB dir s x⇩i);
min_rvar_incdec dir s x⇩i = Inl I⟧
⟹ P (set_unsat I s)"
shows "P (check s)"
proof (cases "𝒰 s")
case True
then show ?thesis
using ‹𝒰 s ⟹ P s›
by (simp add: check.simps)
next
case False
have "check_dom s"
using *
by (simp add: check_sat_terminates)
then show ?thesis
using * False
proof (induct s rule: check_dom.induct)
case (step s')
show ?case
proof (rule check_cases)
fix x⇩i x⇩j l⇩i dir
let ?dir = "if ⟨𝒱 s'⟩ x⇩i <⇩l⇩b ℬ⇩l s' x⇩i then Positive else Negative"
let ?s' = "check' dir x⇩i s'"
assume "¬ 𝒰 s'" "min_lvar_not_in_bounds s' = Some x⇩i" "min_rvar_incdec dir s' x⇩i = Inr x⇩j" "l⇩i = the (LB dir s' x⇩i)"
"?s' = pivot_and_update x⇩i x⇩j l⇩i s'" "dir = ?dir"
moreover
then have "∇ ?s'" "△ (𝒯 ?s')" "⊨⇩n⇩o⇩l⇩h⇩s ?s'" "◇ ?s'" "¬ 𝒰 ?s'"
using ‹∇ s'› ‹△ (𝒯 s')› ‹⊨⇩n⇩o⇩l⇩h⇩s s'› ‹◇ s'›
using ‹?s' = pivot_and_update x⇩i x⇩j l⇩i s'›
using pivotandupdate_check_precond[of dir s' x⇩i x⇩j l⇩i]
using pivotandupdate_unsat_id[of s' x⇩i x⇩j l⇩i]
by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars)
ultimately
have "P (check (check' dir x⇩i s'))"
using step(2)[of x⇩i] step(4)[of x⇩i] ‹△ (𝒯 s')› ‹∇ s'›
by auto
then show "P (check (pivot_and_update x⇩i x⇩j l⇩i s'))"
using ‹?s' = pivot_and_update x⇩i x⇩j l⇩i s'›
by simp
qed (simp_all add: ‹∇ s'› ‹△ (𝒯 s')› ‹⊨⇩n⇩o⇩l⇩h⇩s s'› ‹◇ s'› ‹¬ 𝒰 s'› ** )
qed
qed
end
lemma poly_eval_update: "(p ⦃ v ( x := c :: 'a :: lrv) ⦄) = (p ⦃ v ⦄) + coeff p x *R (c - v x)"
proof (transfer, simp, goal_cases)
case (1 p v x c)
hence fin: "finite {v. p v ≠ 0}" by simp
have "(∑y∈{v. p v ≠ 0}. p y *R (if y = x then c else v y)) =
(∑y∈{v. p v ≠ 0} ∩ {x}. p y *R (if y = x then c else v y))
+ (∑y∈{v. p v ≠ 0} ∩ (UNIV - {x}). p y *R (if y = x then c else v y))" (is "?l = ?a + ?b")
by (subst sum.union_disjoint[symmetric], auto intro: sum.cong fin)
also have "?a = (if p x = 0 then 0 else p x *R c)" by auto
also have "… = p x *R c" by auto
also have "?b = (∑y∈{v. p v ≠ 0} ∩ (UNIV - {x}). p y *R v y)" (is "_ = ?c") by (rule sum.cong, auto)
finally have l: "?l = p x *R c + ?c" .
define r where "r = (∑y∈{v. p v ≠ 0}. p y *R v y) + p x *R (c - v x)"
have "r = (∑y∈{v. p v ≠ 0}. p y *R v y) + p x *R (c - v x)" by (simp add: r_def)
also have "(∑y∈{v. p v ≠ 0}. p y *R v y) =
(∑y∈{v. p v ≠ 0} ∩ {x}. p y *R v y) + ?c" (is "_ = ?d + _")
by (subst sum.union_disjoint[symmetric], auto intro: sum.cong fin)
also have "?d = (if p x = 0 then 0 else p x *R v x)" by auto
also have "… = p x *R v x" by auto
finally have "(p x *R (c - v x) + p x *R v x) + ?c = r" by simp
also have "(p x *R (c - v x) + p x *R v x) = p x *R c" unfolding scaleRat_right_distrib[symmetric] by simp
finally have r: "p x *R c + ?c = r" .
show ?case unfolding l r r_def ..
qed
context PivotUpdateMinVars
begin
context
fixes rhs_eq_val :: "(var, 'a::lrv) mapping ⇒ var ⇒ 'a ⇒ eq ⇒ 'a"
assumes "RhsEqVal rhs_eq_val"
begin
lemma check_minimal_unsat_state_core:
assumes *: "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s" "△ (𝒯 s)" "∇ s"
shows "𝒰 (check s) ⟶ minimal_unsat_state_core (check s)"
(is "?P (check s)")
proof (rule check_induct'')
fix s' :: "('i,'a) state" and x⇩i dir I
assume nolhs: "⊨⇩n⇩o⇩l⇩h⇩s s'"
and min_rvar: "min_rvar_incdec dir s' x⇩i = Inl I"
and sat: "¬ 𝒰 s'"
and min_lvar: "min_lvar_not_in_bounds s' = Some x⇩i"
and dir: "dir = Positive ∨ dir = Negative"
and lt: "⊲⇩l⇩b (lt dir) (⟨𝒱 s'⟩ x⇩i) (LB dir s' x⇩i)"
and norm: "△ (𝒯 s')"
and valuated: "∇ s'"
let ?eq = "eq_for_lvar (𝒯 s') x⇩i"
have unsat_core: "set (the (𝒰⇩c (set_unsat I s'))) = set I"
by auto
obtain l⇩i where LB_Some: "LB dir s' x⇩i = Some l⇩i" and lt: "lt dir (⟨𝒱 s'⟩ x⇩i) l⇩i"
using lt by (cases "LB dir s' x⇩i") (auto simp add: bound_compare_defs)
from LB_Some dir obtain i where LBI: "look (LBI dir s') x⇩i = Some (i,l⇩i)" and LI: "LI dir s' x⇩i = i"
by (auto simp: simp: indexl_def indexu_def boundsl_def boundsu_def)
from min_rvar_incdec_eq_None[OF min_rvar] dir
have Is': "LI dir s' (lhs (eq_for_lvar (𝒯 s') x⇩i)) ∈ indices_state s' ⟹ set I ⊆ indices_state s'" and
reasable: "⋀ x. x ∈ rvars_eq ?eq ⟹ ¬ reasable_var dir x ?eq s'" and
setI: "set I =
{LI dir s' (lhs ?eq)} ∪
{LI dir s' x |x. x ∈ rvars_eq ?eq ∧ coeff (rhs ?eq) x < 0} ∪
{UI dir s' x |x. x ∈ rvars_eq ?eq ∧ 0 < coeff (rhs ?eq) x}" (is "_ = ?L ∪ ?R1 ∪ ?R2") by auto
note setI also have id: "lhs ?eq = x⇩i"
by (simp add: EqForLVar.eq_for_lvar EqForLVar_axioms min_lvar min_lvar_not_in_bounds_lvars)
finally have iI: "i ∈ set I" unfolding LI by auto
note setI = setI[unfolded id]
have "LI dir s' x⇩i ∈ indices_state s'" using LBI LI
unfolding indices_state_def using dir by force
from Is'[unfolded id, OF this]
have Is': "set I ⊆ indices_state s'" .
have "x⇩i ∈ lvars (𝒯 s')"
using min_lvar
by (simp add: min_lvar_not_in_bounds_lvars)
then have **: "?eq ∈ set (𝒯 s')" "lhs ?eq = x⇩i"
by (auto simp add: eq_for_lvar)
have Is': "set I ⊆ indices_state (set_unsat I s')"
using Is' * unfolding indices_state_def by auto
have "⟨𝒱 s'⟩ ⊨⇩t 𝒯 s'" and b: "⟨𝒱 s'⟩ ⊨⇩b ℬ s' ∥ - lvars (𝒯 s')"
using nolhs[unfolded curr_val_satisfies_no_lhs_def] by auto
from norm[unfolded normalized_tableau_def]
have lvars_rvars: "lvars (𝒯 s') ∩ rvars (𝒯 s') = {}" by auto
hence in_bnds: "x ∈ rvars (𝒯 s') ⟹ in_bounds x ⟨𝒱 s'⟩ (ℬ s')" for x
by (intro b[unfolded satisfies_bounds_set.simps, rule_format, of x], auto)
{
assume dist: "distinct_indices_state (set_unsat I s')"
hence "distinct_indices_state s'" unfolding distinct_indices_state_def by auto
note dist = this[unfolded distinct_indices_state_def, rule_format]
{
fix x c i y
assume c: "look (ℬ⇩i⇩l s') x = Some (i,c) ∨ look (ℬ⇩i⇩u s') x = Some (i,c)"
and y: "y ∈ rvars_eq ?eq" and
coeff: "coeff (rhs ?eq) y < 0 ∧ i = LI dir s' y ∨ coeff (rhs ?eq) y > 0 ∧ i = UI dir s' y"
{
assume coeff: "coeff (rhs ?eq) y < 0" and i: "i = LI dir s' y"
from reasable[OF y] coeff have not_gt: "¬ (⊳⇩l⇩b (lt dir) (⟨𝒱 s'⟩ y) (LB dir s' y))" by auto
then obtain d where LB: "LB dir s' y = Some d" using dir by (cases "LB dir s' y", auto simp: bound_compare_defs)
with not_gt have le: "le (lt dir) (⟨𝒱 s'⟩ y) d" using dir by (auto simp: bound_compare_defs)
from LB have "look (LBI dir s') y = Some (i, d)" unfolding i using dir
by (auto simp: boundsl_def boundsu_def indexl_def indexu_def)
with c dist[of x i c y d] dir
have yx: "y = x" "d = c" by auto
from y[unfolded yx] have "x ∈ rvars (𝒯 s')" using **(1) unfolding rvars_def by force
from in_bnds[OF this] le LB not_gt i have "⟨𝒱 s'⟩ x = c" unfolding yx using dir by auto
note yx(1) this
}
moreover
{
assume coeff: "coeff (rhs ?eq) y > 0" and i: "i = UI dir s' y"
from reasable[OF y] coeff have not_gt: "¬ (⊲⇩u⇩b (lt dir) (⟨𝒱 s'⟩ y) (UB dir s' y))" by auto
then obtain d where UB: "UB dir s' y = Some d" using dir by (cases "UB dir s' y", auto simp: bound_compare_defs)
with not_gt have le: "le (lt dir) d (⟨𝒱 s'⟩ y)" using dir by (auto simp: bound_compare_defs)
from UB have "look (UBI dir s') y = Some (i, d)" unfolding i using dir
by (auto simp: boundsl_def boundsu_def indexl_def indexu_def)
with c dist[of x i c y d] dir
have yx: "y = x" "d = c" by auto
from y[unfolded yx] have "x ∈ rvars (𝒯 s')" using **(1) unfolding rvars_def by force
from in_bnds[OF this] le UB not_gt i have "⟨𝒱 s'⟩ x = c" unfolding yx using dir by auto
note yx(1) this
}
ultimately have "y = x" "⟨𝒱 s'⟩ x = c" using coeff by blast+
} note x_vars_main = this
{
fix x c i
assume c: "look (ℬ⇩i⇩l s') x = Some (i,c) ∨ look (ℬ⇩i⇩u s') x = Some (i,c)" and i: "i ∈ ?R1 ∪ ?R2"
from i obtain y where y: "y ∈ rvars_eq ?eq" and
coeff: "coeff (rhs ?eq) y < 0 ∧ i = LI dir s' y ∨ coeff (rhs ?eq) y > 0 ∧ i = UI dir s' y"
by auto
from x_vars_main[OF c y coeff]
have "y = x" "⟨𝒱 s'⟩ x = c" using coeff by blast+
with y have "x ∈ rvars_eq ?eq" "x ∈ rvars (𝒯 s')" "⟨𝒱 s'⟩ x = c" using **(1) unfolding rvars_def by force+
} note x_rvars = this
have R1R2: "(?R1 ∪ ?R2, ⟨𝒱 s'⟩) ⊨⇩i⇩s⇩e s'"
unfolding satisfies_state_index'.simps
proof (intro conjI)
show "⟨𝒱 s'⟩ ⊨⇩t 𝒯 s'" by fact
show "(?R1 ∪ ?R2, ⟨𝒱 s'⟩) ⊨⇩i⇩b⇩e ℬℐ s'"
unfolding satisfies_bounds_index'.simps
proof (intro conjI impI allI)
fix x c
assume c: "ℬ⇩l s' x = Some c" and i: "ℐ⇩l s' x ∈ ?R1 ∪ ?R2"
from c have ci: "look (ℬ⇩i⇩l s') x = Some (ℐ⇩l s' x, c)" unfolding boundsl_def indexl_def by auto
from x_rvars[OF _ i] ci show "⟨𝒱 s'⟩ x = c" by auto
next
fix x c
assume c: "ℬ⇩u s' x = Some c" and i: "ℐ⇩u s' x ∈ ?R1 ∪ ?R2"
from c have ci: "look (ℬ⇩i⇩u s') x = Some (ℐ⇩u s' x, c)" unfolding boundsu_def indexu_def by auto
from x_rvars[OF _ i] ci show "⟨𝒱 s'⟩ x = c" by auto
qed
qed
have id1: "set (the (𝒰⇩c (set_unsat I s'))) = set I"
"⋀ x. x ⊨⇩i⇩s⇩e set_unsat I s' ⟷ x ⊨⇩i⇩s⇩e s'"
by (auto simp: satisfies_state_index'.simps boundsl_def boundsu_def indexl_def indexu_def)
have "subsets_sat_core (set_unsat I s')" unfolding subsets_sat_core_def id1
proof (intro allI impI)
fix J
assume sub: "J ⊂ set I"
show "∃v. (J, v) ⊨⇩i⇩s⇩e s'"
proof (cases "J ⊆ ?R1 ∪ ?R2")
case True
with R1R2 have "(J, ⟨𝒱 s'⟩) ⊨⇩i⇩s⇩e s'"
unfolding satisfies_state_index'.simps satisfies_bounds_index'.simps by blast
thus ?thesis by blast
next
case False
with sub obtain k where k: "k ∈ ?R1 ∪ ?R2" "k ∉ J" "k ∈ set I" unfolding setI by auto
from k(1) obtain y where y: "y ∈ rvars_eq ?eq"
and coeff: "coeff (rhs ?eq) y < 0 ∧ k = LI dir s' y ∨ coeff (rhs ?eq) y > 0 ∧ k = UI dir s' y" by auto
hence cy0: "coeff (rhs ?eq) y ≠ 0" by auto
from y **(1) have ry: "y ∈ rvars (𝒯 s')" unfolding rvars_def by force
hence yl: "y ∉ lvars (𝒯 s')" using lvars_rvars by blast
interpret rev: RhsEqVal rhs_eq_val by fact
note update = rev.update_valuation_nonlhs[THEN mp, OF norm valuated yl]
define diff where "diff = l⇩i - ⟨𝒱 s'⟩ x⇩i"
have "⟨𝒱 s'⟩ x⇩i < l⇩i ⟹ 0 < l⇩i - ⟨𝒱 s'⟩ x⇩i" "l⇩i < ⟨𝒱 s'⟩ x⇩i ⟹ l⇩i - ⟨𝒱 s'⟩ x⇩i < 0"
using minus_gt by (blast, insert minus_lt, blast)
with lt dir have diff: "lt dir 0 diff" by (auto simp: diff_def)
define up where "up = inverse (coeff (rhs ?eq) y) *R diff"
define v where "v = ⟨𝒱 (rev.update y (⟨𝒱 s'⟩ y + up) s')⟩"
show ?thesis unfolding satisfies_state_index'.simps
proof (intro exI[of _ v] conjI)
show "v ⊨⇩t 𝒯 s'" unfolding v_def
using rev.update_satisfies_tableau[OF norm valuated yl] ‹⟨𝒱 s'⟩ ⊨⇩t 𝒯 s'› by auto
with **(1) have "v ⊨⇩e ?eq" unfolding satisfies_tableau_def by auto
from this[unfolded satisfies_eq_def id]
have v_xi: "v x⇩i = (rhs ?eq ⦃ v ⦄)" .
from ‹⟨𝒱 s'⟩ ⊨⇩t 𝒯 s'› **(1) have "⟨𝒱 s'⟩ ⊨⇩e ?eq" unfolding satisfies_tableau_def by auto
hence V_xi: "⟨𝒱 s'⟩ x⇩i = (rhs ?eq ⦃ ⟨𝒱 s'⟩ ⦄)" unfolding satisfies_eq_def id .
have "v x⇩i = ⟨𝒱 s'⟩ x⇩i + coeff (rhs ?eq) y *R up"
unfolding v_xi unfolding v_def rev.update_valuate_rhs[OF **(1) norm] poly_eval_update V_xi by simp
also have "… = l⇩i" unfolding up_def diff_def scaleRat_scaleRat using cy0 by simp
finally have v_xi_l: "v x⇩i = l⇩i" .
{
assume both: "ℐ⇩u s' y ∈ ?R1 ∪ ?R2" "ℬ⇩u s' y ≠ None" "ℐ⇩l s' y ∈ ?R1 ∪ ?R2" "ℬ⇩l s' y ≠ None"
and diff: "ℐ⇩l s' y ≠ ℐ⇩u s' y"
from both(1) dir obtain xu cu where
looku: "look (ℬ⇩i⇩l s') xu = Some (ℐ⇩u s' y, cu) ∨ look (ℬ⇩i⇩u s') xu = Some (ℐ⇩u s' y,cu)"
by (smt Is' indices_state_def le_sup_iff mem_Collect_eq setI set_unsat_simps subsetCE)
from both(1) obtain xu' where "xu' ∈ rvars_eq ?eq" "coeff (rhs ?eq) xu' < 0 ∧ ℐ⇩u s' y = LI dir s' xu' ∨
coeff (rhs ?eq) xu' > 0 ∧ ℐ⇩u s' y = UI dir s' xu'" by blast
with x_vars_main(1)[OF looku this]
have xu: "xu ∈ rvars_eq ?eq" "coeff (rhs ?eq) xu < 0 ∧ ℐ⇩u s' y = LI dir s' xu ∨
coeff (rhs ?eq) xu > 0 ∧ ℐ⇩u s' y = UI dir s' xu" by auto
{
assume "xu ≠ y"
with dist[OF looku, of y] have "look (ℬ⇩i⇩u s') y = None"
by (cases "look (ℬ⇩i⇩u s') y", auto simp: boundsu_def indexu_def, blast)
with both(2) have False by (simp add: boundsu_def)
}
hence xu_y: "xu = y" by blast
from both(3) dir obtain xl cl where
lookl: "look (ℬ⇩i⇩l s') xl = Some (ℐ⇩l s' y, cl) ∨ look (ℬ⇩i⇩u s') xl = Some (ℐ⇩l s' y,cl)"
by (smt Is' indices_state_def le_sup_iff mem_Collect_eq setI set_unsat_simps subsetCE)
from both(3) obtain xl' where "xl' ∈ rvars_eq ?eq" "coeff (rhs ?eq) xl' < 0 ∧ ℐ⇩l s' y = LI dir s' xl' ∨
coeff (rhs ?eq) xl' > 0 ∧ ℐ⇩l s' y = UI dir s' xl'" by blast
with x_vars_main(1)[OF lookl this]
have xl: "xl ∈ rvars_eq ?eq" "coeff (rhs ?eq) xl < 0 ∧ ℐ⇩l s' y = LI dir s' xl ∨
coeff (rhs ?eq) xl > 0 ∧ ℐ⇩l s' y = UI dir s' xl" by auto
{
assume "xl ≠ y"
with dist[OF lookl, of y] have "look (ℬ⇩i⇩l s') y = None"
by (cases "look (ℬ⇩i⇩l s') y", auto simp: boundsl_def indexl_def, blast)
with both(4) have False by (simp add: boundsl_def)
}
hence xl_y: "xl = y" by blast
from xu(2) xl(2) diff have diff: "xu ≠ xl" by auto
with xu_y xl_y have False by simp
} note both_y_False = this
show "(J, v) ⊨⇩i⇩b⇩e ℬℐ s'" unfolding satisfies_bounds_index'.simps
proof (intro conjI allI impI)
fix x c
assume x: "ℬ⇩l s' x = Some c" "ℐ⇩l s' x ∈ J"
with k have not_k: "ℐ⇩l s' x ≠ k" by auto
from x have ci: "look (ℬ⇩i⇩l s') x = Some (ℐ⇩l s' x, c)" unfolding boundsl_def indexl_def by auto
show "v x = c"
proof (cases "ℐ⇩l s' x = i")
case False
hence iR12: "ℐ⇩l s' x ∈ ?R1 ∪ ?R2" using sub x unfolding setI LI by blast
from x_rvars(2-3)[OF _ iR12] ci have xr: "x ∈ rvars (𝒯 s')" and val: "⟨𝒱 s'⟩ x = c" by auto
with lvars_rvars have xl: "x ∉ lvars (𝒯 s')" by auto
show ?thesis
proof (cases "x = y")
case False
thus ?thesis using val unfolding v_def map2fun_def' update[OF xl] using val by auto
next
case True
note coeff = coeff[folded True]
from coeff not_k dir ci have Iu: "ℐ⇩u s' x = k" by auto
with ci Iu x(2) k sub False True
have both: "ℐ⇩u s' y ∈ ?R1 ∪ ?R2" "ℐ⇩l s' y ∈ ?R1 ∪ ?R2" and diff: "ℐ⇩l s' y ≠ ℐ⇩u s' y"
unfolding setI LI by auto
have "ℬ⇩l s' y ≠ None" using x True by simp
from both_y_False[OF both(1) _ both(2) this diff]
have "ℬ⇩u s' y = None" by metis
with reasable[OF y] dir coeff True
have "dir = Negative ⟹ 0 < coeff (rhs ?eq) y" "dir = Positive ⟹ 0 > coeff (rhs ?eq) y" by (auto simp: bound_compare_defs)
with dir coeff[unfolded True] have "k = ℐ⇩l s' y" by auto
with diff Iu False True
have False by auto
thus ?thesis ..
qed
next
case True
from LBI ci[unfolded True] dir
dist[unfolded distinct_indices_state_def, rule_format, of x i c x⇩i l⇩i]
have xxi: "x = x⇩i" and c: "c = l⇩i" by auto
have vxi: "v x = l⇩i" unfolding xxi v_xi_l ..
thus ?thesis unfolding c by simp
qed
next
fix x c
assume x: "ℬ⇩u s' x = Some c" "ℐ⇩u s' x ∈ J"
with k have not_k: "ℐ⇩u s' x ≠ k" by auto
from x have ci: "look (ℬ⇩i⇩u s') x = Some (ℐ⇩u s' x, c)" unfolding boundsu_def indexu_def by auto
show "v x = c"
proof (cases "ℐ⇩u s' x = i")
case False
hence iR12: "ℐ⇩u s' x ∈ ?R1 ∪ ?R2" using sub x unfolding setI LI by blast
from x_rvars(2-3)[OF _ iR12] ci have xr: "x ∈ rvars (𝒯 s')" and val: "⟨𝒱 s'⟩ x = c" by auto
with lvars_rvars have xl: "x ∉ lvars (𝒯 s')" by auto
show ?thesis
proof (cases "x = y")
case False
thus ?thesis using val unfolding v_def map2fun_def' update[OF xl] using val by auto
next
case True
note coeff = coeff[folded True]
from coeff not_k dir ci have Iu: "ℐ⇩l s' x = k" by auto
with ci Iu x(2) k sub False True
have both: "ℐ⇩u s' y ∈ ?R1 ∪ ?R2" "ℐ⇩l s' y ∈ ?R1 ∪ ?R2" and diff: "ℐ⇩l s' y ≠ ℐ⇩u s' y"
unfolding setI LI by auto
have "ℬ⇩u s' y ≠ None" using x True by simp
from both_y_False[OF both(1) this both(2) _ diff]
have "ℬ⇩l s' y = None" by metis
with reasable[OF y] dir coeff True
have "dir = Negative ⟹ 0 > coeff (rhs ?eq) y" "dir = Positive ⟹ 0 < coeff (rhs ?eq) y" by (auto simp: bound_compare_defs)
with dir coeff[unfolded True] have "k = ℐ⇩u s' y" by auto
with diff Iu False True
have False by auto
thus ?thesis ..
qed
next
case True
from LBI ci[unfolded True] dir
dist[unfolded distinct_indices_state_def, rule_format, of x i c x⇩i l⇩i]
have xxi: "x = x⇩i" and c: "c = l⇩i" by auto
have vxi: "v x = l⇩i" unfolding xxi v_xi_l ..
thus ?thesis unfolding c by simp
qed
qed
qed
qed
qed
} note minimal_core = this
have unsat_core: "unsat_state_core (set_unsat I s')"
unfolding unsat_state_core_def unsat_core
proof (intro impI conjI Is', clarify)
fix v
assume "(set I, v) ⊨⇩i⇩s set_unsat I s'"
then have Iv: "(set I, v) ⊨⇩i⇩s s'"
unfolding satisfies_state_index.simps
by (auto simp: indexl_def indexu_def boundsl_def boundsu_def)
from Iv have vt: "v ⊨⇩t 𝒯 s'" and Iv: "(set I, v) ⊨⇩i⇩b ℬℐ s'"
unfolding satisfies_state_index.simps by auto
have lt_le_eq: "⋀ x y :: 'a. (x < y) ⟷ (x ≤ y ∧ x ≠ y)" by auto
from Iv dir
have lb: "⋀ x i c l. look (LBI dir s') x = Some (i,l) ⟹ i ∈ set I ⟹ le (lt dir) l (v x)"
unfolding satisfies_bounds_index.simps
by (auto simp: lt_le_eq indexl_def indexu_def boundsl_def boundsu_def)
from lb[OF LBI iI] have li_x: "le (lt dir) l⇩i (v x⇩i)" .
have "⟨𝒱 s'⟩ ⊨⇩e ?eq"
using nolhs ‹?eq ∈ set (𝒯 s')›
unfolding curr_val_satisfies_no_lhs_def
by (simp add: satisfies_tableau_def)
then have "⟨𝒱 s'⟩ x⇩i = (rhs ?eq) ⦃ ⟨𝒱 s'⟩ ⦄"
using ‹lhs ?eq = x⇩i›
by (simp add: satisfies_eq_def)
moreover
have "v ⊨⇩e ?eq"
using vt ‹?eq ∈ set (𝒯 s')›
by (simp add: satisfies_state_def satisfies_tableau_def)
then have "v x⇩i = (rhs ?eq) ⦃ v ⦄"
using ‹lhs ?eq = x⇩i›
by (simp add: satisfies_eq_def)
moreover
have "⊵⇩l⇩b (lt dir) (v x⇩i) (LB dir s' x⇩i)"
using li_x dir unfolding LB_Some by (auto simp: bound_compare'_defs)
moreover
from min_rvar_incdec_eq_None'[rule_format, OF dir min_rvar refl Iv]
have "le (lt dir) (rhs (?eq) ⦃v⦄) (rhs (?eq) ⦃ ⟨𝒱 s'⟩ ⦄)" .
ultimately
show False
using dir lt LB_Some
by (auto simp add: bound_compare_defs)
qed
thus "𝒰 (set_unsat I s') ⟶ minimal_unsat_state_core (set_unsat I s')" using minimal_core
by (auto simp: minimal_unsat_state_core_def)
qed (simp_all add: *)
lemma Check_check: "Check check"
proof
fix s :: "('i,'a) state"
assume "𝒰 s"
then show "check s = s"
by (simp add: check.simps)
next
fix s :: "('i,'a) state" and v :: "'a valuation"
assume *: "∇ s" "△ (𝒯 s)" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s"
then have "v ⊨⇩t 𝒯 s = v ⊨⇩t 𝒯 (check s)"
by (rule check_induct, simp_all add: pivotandupdate_tableau_equiv)
moreover
have "△ (𝒯 (check s))"
by (rule check_induct', simp_all add: * pivotandupdate_tableau_normalized)
moreover
have "∇ (check s)"
proof (rule check_induct', simp_all add: * pivotandupdate_tableau_valuated)
fix s I
show "∇ s ⟹ ∇ (set_unsat I s)"
by (simp add: tableau_valuated_def)
qed
ultimately
show "let s' = check s in v ⊨⇩t 𝒯 s = v ⊨⇩t 𝒯 s' ∧ △ (𝒯 s') ∧ ∇ s'"
by (simp add: Let_def)
next
fix s :: "('i,'a) state"
assume *: "∇ s" "△ (𝒯 s)" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s"
from * show "ℬ⇩i (check s) = ℬ⇩i s"
by (rule check_induct, simp_all add: pivotandupdate_bounds_id)
next
fix s :: "('i,'a) state"
assume *: "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s" "△ (𝒯 s)" "∇ s"
have "¬ 𝒰 (check s) ⟶ ⊨ (check s)"
proof (rule check_induct'', simp_all add: *)
fix s
assume "min_lvar_not_in_bounds s = None" "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s"
then show " ⊨ s"
using min_lvar_not_in_bounds_None[of s]
unfolding curr_val_satisfies_state_def satisfies_state_def
unfolding curr_val_satisfies_no_lhs_def
by (auto simp add: satisfies_bounds_set.simps satisfies_bounds.simps)
qed
then show "¬ 𝒰 (check s) ⟹ ⊨ (check s)" by blast
next
fix s :: "('i,'a) state"
assume *: "¬ 𝒰 s" "⊨⇩n⇩o⇩l⇩h⇩s s" "◇ s" "△ (𝒯 s)" "∇ s"
have "𝒰 (check s) ⟶ minimal_unsat_state_core (check s)"
by (rule check_minimal_unsat_state_core[OF *])
then show "𝒰 (check s) ⟹ minimal_unsat_state_core (check s)" by blast
qed
end
end
subsection‹Symmetries›
text‹\label{sec:symmetries} Simplex algorithm exhibits many
symmetric cases. For example, ‹assert_bound› treats atoms
‹Leq x c› and ‹Geq x c› in a symmetric manner, ‹check_inc› and ‹check_dec› are symmetric, etc. These
symmetric cases differ only in several aspects: order relations
between numbers (‹<› vs ‹>› and ‹≤› vs ‹≥›), the role of lower and upper bounds (‹ℬ⇩l› vs
‹ℬ⇩u›) and their updating functions, comparisons with bounds
(e.g., ‹≥⇩u⇩b› vs ‹≤⇩l⇩b› or
‹<⇩l⇩b› vs ‹>⇩u⇩b›), and atom constructors (‹Leq›
and ‹Geq›). These can be attributed to two different
orientations (positive and negative) of rational axis. To avoid
duplicating definitions and proofs, ‹assert_bound› definition
cases for ‹Leq› and ‹Geq› are replaced by a call to a
newly introduced function parametrized by a ‹Direction› --- a
record containing minimal set of aspects listed above that differ in
two definition cases such that other aspects can be derived from them
(e.g., only ‹<› need to be stored while ‹≤› can be
derived from it). Two constants of the type ‹Direction› are
defined: ‹Positive› (with ‹<›, ‹≤› orders,
‹ℬ⇩l› for lower and ‹ℬ⇩u› for upper bounds and their
corresponding updating functions, and ‹Leq› constructor) and
‹Negative› (completely opposite from the previous
one). Similarly, ‹check_inc› and ‹check_dec› are
replaced by a new function ‹check_incdec› parametrized by a
‹Direction›. All lemmas, previously repeated for each
symmetric instance, were replaced by a more abstract one, again
parametrized by a ‹Direction› parameter.
\vspace{-3mm}
›
subsection‹Concrete implementation›
text‹It is easy to give a concrete implementation of the initial
state constructor, which satisfies the specification of the @{term
Init} locale. For example:›
definition init_state :: "_ ⇒ ('i,'a :: zero)state" where
"init_state t = State t Mapping.empty Mapping.empty (Mapping.tabulate (vars_list t) (λ v. 0)) False None"
interpretation Init "init_state :: _ ⇒ ('i,'a :: lrv)state"
proof
fix t
let ?init = "init_state t :: ('i,'a)state"
show "⟨𝒱 ?init⟩ ⊨⇩t t"
unfolding satisfies_tableau_def satisfies_eq_def
proof (safe)
fix l r
assume "(l, r) ∈ set t"
then have "l ∈ set (vars_list t)" "vars r ⊆ set (vars_list t)"
by (auto simp: set_vars_list) (transfer, force)
then have *: "vars r ⊆ lhs ` set t ∪ (⋃x∈set t. rvars_eq x)" by (auto simp: set_vars_list)
have "⟨𝒱 ?init⟩ l = (0::'a)"
using ‹l ∈ set (vars_list t)›
unfolding init_state_def by (auto simp: map2fun_def lookup_tabulate)
moreover
have "r ⦃ ⟨𝒱 ?init⟩ ⦄ = (0::'a)" using *
proof (transfer fixing: t, goal_cases)
case (1 r)
{
fix x
assume "x∈{v. r v ≠ 0}"
then have "r x *R ⟨𝒱 ?init⟩ x = (0::'a)"
using 1
unfolding init_state_def
by (auto simp add: map2fun_def lookup_tabulate comp_def restrict_map_def set_vars_list Abstract_Linear_Poly.vars_def)
}
then show ?case by auto
qed
ultimately
show "⟨𝒱 ?init⟩ (lhs (l, r)) = rhs (l, r) ⦃ ⟨𝒱 ?init⟩ ⦄"
by auto
qed
next
fix t
show "∇ (init_state t)"
unfolding init_state_def
by (auto simp add: lookup_tabulate tableau_valuated_def comp_def restrict_map_def set_vars_list lvars_def rvars_def)
qed (simp_all add: init_state_def add: boundsl_def boundsu_def indexl_def indexu_def)
definition min_lvar_not_in_bounds :: "('i,'a::{linorder,zero}) state ⇒ var option" where
"min_lvar_not_in_bounds s ≡
min_satisfying (λ x. ¬ in_bounds x (⟨𝒱 s⟩) (ℬ s)) (map lhs (𝒯 s))"
interpretation MinLVarNotInBounds "min_lvar_not_in_bounds :: ('i,'a::lrv) state ⇒ _"
proof
fix s::"('i,'a) state"
show "min_lvar_not_in_bounds s = None ⟶
(∀x∈lvars (𝒯 s). in_bounds x (⟨𝒱 s⟩) (ℬ s))"
unfolding min_lvar_not_in_bounds_def lvars_def
using min_satisfying_None
by blast
next
fix s x⇩i
show "min_lvar_not_in_bounds s = Some x⇩i ⟶
x⇩i ∈ lvars (𝒯 s) ∧
¬ in_bounds x⇩i ⟨𝒱 s⟩ (ℬ s) ∧
(∀x∈lvars (𝒯 s). x < x⇩i ⟶ in_bounds x ⟨𝒱 s⟩ (ℬ s))"
unfolding min_lvar_not_in_bounds_def lvars_def
using min_satisfying_Some
by blast+
qed
definition unsat_indices :: "('i,'a :: linorder) Direction ⇒ ('i,'a) state ⇒ var list ⇒ eq ⇒ 'i list" where
"unsat_indices dir s vs eq = (let r = rhs eq; li = LI dir s; ui = UI dir s in
remdups (li (lhs eq) # map (λ x. if coeff r x < 0 then li x else ui x) vs))"
definition min_rvar_incdec_eq :: "('i,'a) Direction ⇒ ('i,'a::lrv) state ⇒ eq ⇒ 'i list + var" where
"min_rvar_incdec_eq dir s eq = (let rvars = Abstract_Linear_Poly.vars_list (rhs eq)
in case min_satisfying (λ x. reasable_var dir x eq s) rvars of
None ⇒ Inl (unsat_indices dir s rvars eq)
| Some x⇩j ⇒ Inr x⇩j)"
interpretation MinRVarsEq "min_rvar_incdec_eq :: ('i,'a :: lrv) Direction ⇒ _"
proof
fix s eq "is" and dir :: "('i,'a) Direction"
let ?min = "min_satisfying (λ x. reasable_var dir x eq s) (Abstract_Linear_Poly.vars_list (rhs eq))"
let ?vars = "Abstract_Linear_Poly.vars_list (rhs eq)"
{
assume "min_rvar_incdec_eq dir s eq = Inl is"
from this[unfolded min_rvar_incdec_eq_def Let_def, simplified]
have "?min = None" and I: "set is = set (unsat_indices dir s ?vars eq)" by (cases ?min, auto)+
from this min_satisfying_None set_vars_list
have 1: "⋀ x. x ∈ rvars_eq eq ⟹ ¬ reasable_var dir x eq s" by blast
{
fix i
assume "i ∈ set is" and dir: "dir = Positive ∨ dir = Negative" and lhs_eq: "LI dir s (lhs eq) ∈ indices_state s"
from this[unfolded I unsat_indices_def Let_def]
consider (lhs) "i = LI dir s (lhs eq)"
| (LI_rhs) x where "i = LI dir s x" "x ∈ rvars_eq eq" "coeff (rhs eq) x < 0"
| (UI_rhs) x where "i = UI dir s x" "x ∈ rvars_eq eq" "coeff (rhs eq) x ≥ 0"
by (auto split: if_splits simp: set_vars_list)
then have "i ∈ indices_state s"
proof cases
case lhs
show ?thesis unfolding lhs using lhs_eq by auto
next
case LI_rhs
from 1[OF LI_rhs(2)] LI_rhs(3)
have "¬ (⊳⇩l⇩b (lt dir) (⟨𝒱 s⟩ x) (LB dir s x))" by auto
then show ?thesis unfolding LI_rhs(1) unfolding indices_state_def using dir
by (auto simp: bound_compare'_defs boundsl_def boundsu_def indexl_def indexu_def
split: option.splits intro!: exI[of _ x]) auto
next
case UI_rhs
from UI_rhs(2) have "coeff (rhs eq) x ≠ 0"
by (simp add: coeff_zero)
with UI_rhs(3) have "0 < coeff (rhs eq) x" by auto
from 1[OF UI_rhs(2)] this have "¬ (⊲⇩u⇩b (lt dir) (⟨𝒱 s⟩ x) (UB dir s x))" by auto
then show ?thesis unfolding UI_rhs(1) unfolding indices_state_def using dir
by (auto simp: bound_compare'_defs boundsl_def boundsu_def indexl_def indexu_def
split: option.splits intro!: exI[of _ x]) auto
qed
}
then have 2: "dir = Positive ∨ dir = Negative ⟹ LI dir s (lhs eq) ∈ indices_state s ⟹
set is ⊆ indices_state s" by auto
show "
(∀ x ∈ rvars_eq eq. ¬ reasable_var dir x eq s) ∧ set is =
{LI dir s (lhs eq)} ∪ {LI dir s x |x. x ∈ rvars_eq eq ∧
coeff (rhs eq) x < 0} ∪ {UI dir s x |x. x ∈ rvars_eq eq ∧ 0 < coeff (rhs eq) x} ∧
(dir = Positive ∨ dir = Negative ⟶ LI dir s (lhs eq) ∈ indices_state s ⟶ set is ⊆ indices_state s)"
proof (intro conjI impI 2, goal_cases)
case 2
have "set is = {LI dir s (lhs eq)} ∪ LI dir s ` (rvars_eq eq ∩ {x. coeff (rhs eq) x < 0}) ∪ UI dir s ` (rvars_eq eq ∩ {x. ¬ coeff (rhs eq) x < 0})"
unfolding I unsat_indices_def Let_def
by (auto simp add: set_vars_list)
also have "… = {LI dir s (lhs eq)} ∪ LI dir s ` {x. x ∈ rvars_eq eq ∧ coeff (rhs eq) x < 0}
∪ UI dir s ` { x. x ∈ rvars_eq eq ∧ 0 < coeff (rhs eq) x}"
proof (intro arg_cong2[of _ _ _ _ "(∪)"] arg_cong[of _ _ "λ x. _ ` x"] refl, goal_cases)
case 2
{
fix x
assume "x ∈ rvars_eq eq"
hence "coeff (rhs eq) x ≠ 0"
by (simp add: coeff_zero)
hence or: "coeff (rhs eq) x < 0 ∨ coeff (rhs eq) x > 0" by auto
assume "¬ coeff (rhs eq) x < 0"
hence "coeff (rhs eq) x > 0" using or by simp
} note [dest] = this
show ?case by auto
qed auto
finally
show "set is = {LI dir s (lhs eq)} ∪ {LI dir s x |x. x ∈ rvars_eq eq ∧ coeff (rhs eq) x < 0}
∪ {UI dir s x |x. x ∈ rvars_eq eq ∧ 0 < coeff (rhs eq) x}" by auto
qed (insert 1, auto)
}
fix x⇩j
assume "min_rvar_incdec_eq dir s eq = Inr x⇩j"
from this[unfolded min_rvar_incdec_eq_def Let_def]
have "?min = Some x⇩j" by (cases ?min, auto)
then show "x⇩j ∈ rvars_eq eq" "reasable_var dir x⇩j eq s"
"(∀ x' ∈ rvars_eq eq. x' < x⇩j ⟶ ¬ reasable_var dir x' eq s)"
using min_satisfying_Some set_vars_list by blast+
qed
primrec eq_idx_for_lvar_aux :: "tableau ⇒ var ⇒ nat ⇒ nat" where
"eq_idx_for_lvar_aux [] x i = i"
| "eq_idx_for_lvar_aux (eq # t) x i =
(if lhs eq = x then i else eq_idx_for_lvar_aux t x (i+1))"
definition eq_idx_for_lvar where
"eq_idx_for_lvar t x ≡ eq_idx_for_lvar_aux t x 0"
lemma eq_idx_for_lvar_aux:
assumes "x ∈ lvars t"
shows "let idx = eq_idx_for_lvar_aux t x i in
i ≤ idx ∧ idx < i + length t ∧ lhs (t ! (idx - i)) = x"
using assms
proof (induct t arbitrary: i)
case Nil
then show ?case
by (simp add: lvars_def)
next
case (Cons eq t)
show ?case
using Cons(1)[of "i+1"] Cons(2)
by (cases "x = lhs eq") (auto simp add: Let_def lvars_def nth_Cons')
qed
global_interpretation EqForLVarDefault: EqForLVar eq_idx_for_lvar
defines eq_for_lvar_code = EqForLVarDefault.eq_for_lvar
proof (unfold_locales)
fix x t
assume "x ∈ lvars t"
then show "eq_idx_for_lvar t x < length t ∧
lhs (t ! eq_idx_for_lvar t x) = x"
using eq_idx_for_lvar_aux[of x t 0]
by (simp add: Let_def eq_idx_for_lvar_def)
qed
definition pivot_eq :: "eq ⇒ var ⇒ eq" where
"pivot_eq e y ≡ let cy = coeff (rhs e) y in
(y, (-1/cy) *R ((rhs e) - cy *R (Var y)) + (1/cy) *R (Var (lhs e)))"
lemma pivot_eq_satisfies_eq:
assumes "y ∈ rvars_eq e"
shows "v ⊨⇩e e = v ⊨⇩e pivot_eq e y"
using assms
using scaleRat_right_distrib[of "1 / Rep_linear_poly (rhs e) y" "- (rhs e ⦃ v ⦄)" "v (lhs e)"]
using Groups.group_add_class.minus_unique[of "- ((rhs e) ⦃ v ⦄)" "v (lhs e)"]
unfolding coeff_def vars_def
by (simp add: coeff_def vars_def Let_def pivot_eq_def satisfies_eq_def)
(auto simp add: rational_vector.scale_right_diff_distrib valuate_add valuate_minus valuate_uminus valuate_scaleRat valuate_Var)
lemma pivot_eq_rvars:
assumes "x ∈ vars (rhs (pivot_eq e v))" "x ≠ lhs e" "coeff (rhs e) v ≠ 0" "v ≠ lhs e"
shows "x ∈ vars (rhs e)"
proof-
have "v ∉ vars ((1 / coeff (rhs e) v) *R (rhs e - coeff (rhs e) v *R Var v))"
using coeff_zero
by force
then have "x ≠ v"
using assms(1) assms(3) assms(4)
using vars_plus[of "(-1 / coeff (rhs e) v) *R (rhs e - coeff (rhs e) v *R Var v)" "(1 / coeff (rhs e) v) *R Var (lhs e)"]
by (auto simp add: Let_def vars_scaleRat pivot_eq_def)
then show ?thesis
using assms
using vars_plus[of "(-1 / coeff (rhs e) v) *R (rhs e - coeff (rhs e) v *R Var v)" "(1 / coeff (rhs e) v) *R Var (lhs e)"]
using vars_minus[of "rhs e" "coeff (rhs e) v *R Var v"]
by (auto simp add: vars_scaleRat Let_def pivot_eq_def)
qed
interpretation PivotEq pivot_eq
proof
fix eq x⇩j
assume "x⇩j ∈ rvars_eq eq" "lhs eq ∉ rvars_eq eq"
have "lhs (pivot_eq eq x⇩j) = x⇩j"
unfolding pivot_eq_def
by (simp add: Let_def)
moreover
have "rvars_eq (pivot_eq eq x⇩j) =
{lhs eq} ∪ (rvars_eq eq - {x⇩j})"
proof
show "rvars_eq (pivot_eq eq x⇩j) ⊆ {lhs eq} ∪ (rvars_eq eq - {x⇩j})"
proof
fix x
assume "x ∈ rvars_eq (pivot_eq eq x⇩j)"
have *: "coeff (rhs (pivot_eq eq x⇩j)) x⇩j = 0"
using ‹x⇩j ∈ rvars_eq eq› ‹lhs eq ∉ rvars_eq eq›
using coeff_Var2[of "lhs eq" x⇩j]
by (auto simp add: Let_def pivot_eq_def)
have "coeff (rhs eq) x⇩j ≠ 0"
using ‹x⇩j ∈ rvars_eq eq›
using coeff_zero
by (cases eq) (auto simp add:)
then show "x ∈ {lhs eq} ∪ (rvars_eq eq - {x⇩j})"
using pivot_eq_rvars[of x eq x⇩j]
using ‹x ∈ rvars_eq (pivot_eq eq x⇩j)› ‹x⇩j ∈ rvars_eq eq› ‹lhs eq ∉ rvars_eq eq›
using coeff_zero *
by auto
qed
show "{lhs eq} ∪ (rvars_eq eq - {x⇩j}) ⊆ rvars_eq (pivot_eq eq x⇩j)"
proof
fix x
assume "x ∈ {lhs eq} ∪ (rvars_eq eq - {x⇩j})"
have *: "coeff (rhs eq) (lhs eq) = 0"
using coeff_zero
using ‹lhs eq ∉ rvars_eq eq›
by auto
have **: "coeff (rhs eq) x⇩j ≠ 0"
using ‹x⇩j ∈ rvars_eq eq›
by (simp add: coeff_zero)
have ***: "x ∈ rvars_eq eq ⟹ coeff (Var (lhs eq)) x = 0"
using ‹lhs eq ∉ rvars_eq eq›
using coeff_Var2[of "lhs eq" x]
by auto
have "coeff (Var x⇩j) (lhs eq) = 0"
using ‹x⇩j ∈ rvars_eq eq› ‹lhs eq ∉ rvars_eq eq›
using coeff_Var2[of x⇩j "lhs eq"]
by auto
then have "coeff (rhs (pivot_eq eq x⇩j)) x ≠ 0"
using ‹x ∈ {lhs eq} ∪ (rvars_eq eq - {x⇩j})› * ** ***
using coeff_zero[of "rhs eq" x]
by (auto simp add: Let_def coeff_Var2 pivot_eq_def)
then show "x ∈ rvars_eq (pivot_eq eq x⇩j)"
by (simp add: coeff_zero)
qed
qed
ultimately
show "let eq' = pivot_eq eq x⇩j in lhs eq' = x⇩j ∧ rvars_eq eq' = {lhs eq} ∪ (rvars_eq eq - {x⇩j})"
by (simp add: Let_def)
next
fix v eq x⇩j
assume "x⇩j ∈ rvars_eq eq"
then show "v ⊨⇩e pivot_eq eq x⇩j = v ⊨⇩e eq"
using pivot_eq_satisfies_eq
by blast
qed
definition subst_var:: "var ⇒ linear_poly ⇒ linear_poly ⇒ linear_poly" where
"subst_var v lp' lp ≡ lp + (coeff lp v) *R lp' - (coeff lp v) *R (Var v)"
definition "subst_var_eq_code = SubstVar.subst_var_eq subst_var"
global_interpretation SubstVar subst_var rewrites
"SubstVar.subst_var_eq subst_var = subst_var_eq_code"
proof (unfold_locales)
fix x⇩j lp' lp
have *: "⋀x. ⟦x ∈ vars (lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j); x ∉ vars lp'⟧ ⟹ x ∈ vars lp"
proof-
fix x
assume "x ∈ vars (lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j)"
then have "coeff (lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j) x ≠ 0"
using coeff_zero
by force
assume "x ∉ vars lp'"
then have "coeff lp' x = 0"
using coeff_zero
by auto
show "x ∈ vars lp"
proof(rule ccontr)
assume "x ∉ vars lp"
then have "coeff lp x = 0"
using coeff_zero
by auto
then show False
using ‹coeff (lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j) x ≠ 0›
using ‹coeff lp' x = 0›
by (cases "x = x⇩j") (auto simp add: coeff_Var2)
qed
qed
have "vars (subst_var x⇩j lp' lp) ⊆ (vars lp - {x⇩j}) ∪ vars lp'"
unfolding subst_var_def
using coeff_zero[of "lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j" x⇩j]
using coeff_zero[of lp' x⇩j]
using *
by auto
moreover
have "⋀x. ⟦x ∉ vars (lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j); x ∈ vars lp; x ∉ vars lp'⟧ ⟹ x = x⇩j"
proof-
fix x
assume "x ∈ vars lp" "x ∉ vars lp'"
then have "coeff lp x ≠ 0" "coeff lp' x = 0"
using coeff_zero
by auto
assume "x ∉ vars (lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j)"
then have "coeff (lp + coeff lp x⇩j *R lp' - coeff lp x⇩j *R Var x⇩j) x = 0"
using coeff_zero
by force
then show "x = x⇩j"
using ‹coeff lp x ≠ 0› ‹coeff lp' x = 0›
by (cases "x = x⇩j") (auto simp add: coeff_Var2)
qed
then have "vars lp - {x⇩j} - vars lp' ⊆ vars (subst_var x⇩j lp' lp)"
by (auto simp add: subst_var_def)
ultimately show "vars lp - {x⇩j} - vars lp' ⊆s vars (subst_var x⇩j lp' lp)
⊆s vars lp - {x⇩j} ∪ vars lp'"
by simp
next
fix v x⇩j lp' lp
show "v x⇩j = lp' ⦃ v ⦄ ⟶ lp ⦃ v ⦄ = (subst_var x⇩j lp' lp) ⦃ v ⦄"
unfolding subst_var_def
using valuate_minus[of "lp + coeff lp x⇩j *R lp'" "coeff lp x⇩j *R Var x⇩j" v]
using valuate_add[of lp "coeff lp x⇩j *R lp'" v]
using valuate_scaleRat[of "coeff lp x⇩j" lp' v] valuate_scaleRat[of "coeff lp x⇩j" "Var x⇩j" v]
using valuate_Var[of x⇩j v]
by auto
next
fix x⇩j lp lp'
assume "x⇩j ∉ vars lp"
hence 0: "coeff lp x⇩j = 0" using coeff_zero by blast
show "subst_var x⇩j lp' lp = lp"
unfolding subst_var_def 0 by simp
next
fix x⇩j lp x lp'
assume "x⇩j ∈ vars lp" "x ∈ vars lp' - vars lp"
hence x: "x ≠ x⇩j" and 0: "coeff lp x = 0" and no0: "coeff lp x⇩j ≠ 0" "coeff lp' x ≠ 0"
using coeff_zero by blast+
from x have 00: "coeff (Var x⇩j) x = 0" using coeff_Var2 by auto
show "x ∈ vars (subst_var x⇩j lp' lp)"
unfolding subst_var_def coeff_zero[symmetric]
by (simp add: 0 00 no0)
qed (simp_all add: subst_var_eq_code_def)
definition rhs_eq_val where
"rhs_eq_val v x⇩i c e ≡ let x⇩j = lhs e; a⇩i⇩j = coeff (rhs e) x⇩i in
⟨v⟩ x⇩j + a⇩i⇩j *R (c - ⟨v⟩ x⇩i)"
definition "update_code = RhsEqVal.update rhs_eq_val"
definition "assert_bound'_code = Update.assert_bound' update_code"
definition "assert_bound_code = Update.assert_bound update_code"
global_interpretation RhsEqValDefault': RhsEqVal rhs_eq_val
rewrites
"RhsEqVal.update rhs_eq_val = update_code" and
"Update.assert_bound update_code = assert_bound_code" and
"Update.assert_bound' update_code = assert_bound'_code"
proof unfold_locales
fix v x c e
assume "⟨v⟩ ⊨⇩e e"
then show "rhs_eq_val v x c e = rhs e ⦃ ⟨v⟩(x := c) ⦄"
unfolding rhs_eq_val_def Let_def
using valuate_update_x[of "rhs e" x "⟨v⟩" "⟨v⟩(x := c)"]
by (auto simp add: satisfies_eq_def)
qed (auto simp: update_code_def assert_bound'_code_def assert_bound_code_def)
sublocale PivotUpdateMinVars < Check check
proof (rule Check_check)
show "RhsEqVal rhs_eq_val" ..
qed
definition "pivot_code = Pivot'.pivot eq_idx_for_lvar pivot_eq subst_var"
definition "pivot_tableau_code = Pivot'.pivot_tableau eq_idx_for_lvar pivot_eq subst_var"
global_interpretation Pivot'Default: Pivot' eq_idx_for_lvar pivot_eq subst_var
rewrites
"Pivot'.pivot eq_idx_for_lvar pivot_eq subst_var = pivot_code" and
"Pivot'.pivot_tableau eq_idx_for_lvar pivot_eq subst_var = pivot_tableau_code" and
"SubstVar.subst_var_eq subst_var = subst_var_eq_code"
by (unfold_locales, auto simp: pivot_tableau_code_def pivot_code_def subst_var_eq_code_def)
definition "pivot_and_update_code = PivotUpdate.pivot_and_update pivot_code update_code"
global_interpretation PivotUpdateDefault: PivotUpdate eq_idx_for_lvar pivot_code update_code
rewrites
"PivotUpdate.pivot_and_update pivot_code update_code = pivot_and_update_code"
by (unfold_locales, auto simp: pivot_and_update_code_def)
sublocale Update < AssertBoundNoLhs assert_bound
proof (rule update_to_assert_bound_no_lhs)
show "Pivot eq_idx_for_lvar pivot_code" ..
qed
definition "check_code = PivotUpdateMinVars.check eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code"
definition "check'_code = PivotUpdateMinVars.check' eq_idx_for_lvar min_rvar_incdec_eq pivot_and_update_code"
global_interpretation PivotUpdateMinVarsDefault: PivotUpdateMinVars eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code
rewrites
"PivotUpdateMinVars.check eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code = check_code" and
"PivotUpdateMinVars.check' eq_idx_for_lvar min_rvar_incdec_eq pivot_and_update_code = check'_code"
by (unfold_locales) (simp_all add: check_code_def check'_code_def)
definition "assert_code = Assert'.assert assert_bound_code check_code"
global_interpretation Assert'Default: Assert' assert_bound_code check_code
rewrites
"Assert'.assert assert_bound_code check_code = assert_code"
by (unfold_locales, auto simp: assert_code_def)
definition "assert_bound_loop_code = AssertAllState''.assert_bound_loop assert_bound_code"
definition "assert_all_state_code = AssertAllState''.assert_all_state init_state assert_bound_code check_code"
definition "assert_all_code = AssertAllState.assert_all assert_all_state_code"
global_interpretation AssertAllStateDefault: AssertAllState'' init_state assert_bound_code check_code
rewrites
"AssertAllState''.assert_bound_loop assert_bound_code = assert_bound_loop_code" and
"AssertAllState''.assert_all_state init_state assert_bound_code check_code = assert_all_state_code" and
"AssertAllState.assert_all assert_all_state_code = assert_all_code"
by unfold_locales (simp_all add: assert_bound_loop_code_def assert_all_state_code_def assert_all_code_def)
primrec
monom_to_atom:: "QDelta ns_constraint ⇒ QDelta atom" where
"monom_to_atom (LEQ_ns l r) = (if (monom_coeff l < 0) then
(Geq (monom_var l) (r /R monom_coeff l))
else
(Leq (monom_var l) (r /R monom_coeff l)))"
| "monom_to_atom (GEQ_ns l r) = (if (monom_coeff l < 0) then
(Leq (monom_var l) (r /R monom_coeff l))
else
(Geq (monom_var l) (r /R monom_coeff l)))"
primrec
qdelta_constraint_to_atom:: "QDelta ns_constraint ⇒ var ⇒ QDelta atom" where
"qdelta_constraint_to_atom (LEQ_ns l r) v = (if (is_monom l) then (monom_to_atom (LEQ_ns l r)) else (Leq v r))"
| "qdelta_constraint_to_atom (GEQ_ns l r) v = (if (is_monom l) then (monom_to_atom (GEQ_ns l r)) else (Geq v r))"
primrec
qdelta_constraint_to_atom':: "QDelta ns_constraint ⇒ var ⇒ QDelta atom" where
"qdelta_constraint_to_atom' (LEQ_ns l r) v = (Leq v r)"
| "qdelta_constraint_to_atom' (GEQ_ns l r) v = (Geq v r)"
fun linear_poly_to_eq:: "linear_poly ⇒ var ⇒ eq" where
"linear_poly_to_eq p v = (v, p)"
datatype 'i istate = IState
(FirstFreshVariable: var)
(Tableau: tableau)
(Atoms: "('i,QDelta) i_atom list")
(Poly_Mapping: "linear_poly ⇀ var")
(UnsatIndices: "'i list")
primrec zero_satisfies :: "'a :: lrv ns_constraint ⇒ bool" where
"zero_satisfies (LEQ_ns l r) ⟷ 0 ≤ r"
| "zero_satisfies (GEQ_ns l r) ⟷ 0 ≥ r"
lemma zero_satisfies: "poly c = 0 ⟹ zero_satisfies c ⟹ v ⊨⇩n⇩s c"
by (cases c, auto simp: valuate_zero)
lemma not_zero_satisfies: "poly c = 0 ⟹ ¬ zero_satisfies c ⟹ ¬ v ⊨⇩n⇩s c"
by (cases c, auto simp: valuate_zero)
fun
preprocess' :: "('i,QDelta) i_ns_constraint list ⇒ var ⇒ 'i istate" where
"preprocess' [] v = IState v [] [] (λ p. None) []"
| "preprocess' ((i,h) # t) v = (let s' = preprocess' t v; p = poly h; is_monom_h = is_monom p;
v' = FirstFreshVariable s';
t' = Tableau s';
a' = Atoms s';
m' = Poly_Mapping s';
u' = UnsatIndices s' in
if is_monom_h then IState v' t'
((i,qdelta_constraint_to_atom h v') # a') m' u'
else if p = 0 then
if zero_satisfies h then s' else
IState v' t' a' m' (i # u')
else (case m' p of Some v ⇒
IState v' t' ((i,qdelta_constraint_to_atom h v) # a') m' u'
| None ⇒ IState (v' + 1) (linear_poly_to_eq p v' # t')
((i,qdelta_constraint_to_atom h v') # a') (m' (p ↦ v')) u')
)"
lemma preprocess'_simps: "preprocess' ((i,h) # t) v = (let s' = preprocess' t v; p = poly h; is_monom_h = is_monom p;
v' = FirstFreshVariable s';
t' = Tableau s';
a' = Atoms s';
m' = Poly_Mapping s';
u' = UnsatIndices s' in
if is_monom_h then IState v' t'
((i,monom_to_atom h) # a') m' u'
else if p = 0 then
if zero_satisfies h then s' else
IState v' t' a' m' (i # u')
else (case m' p of Some v ⇒
IState v' t' ((i,qdelta_constraint_to_atom' h v) # a') m' u'
| None ⇒ IState (v' + 1) (linear_poly_to_eq p v' # t')
((i,qdelta_constraint_to_atom' h v') # a') (m' (p ↦ v')) u')
)" by (cases h, auto simp add: Let_def split: option.splits)
lemmas preprocess'_code = preprocess'.simps(1) preprocess'_simps
declare preprocess'_code[code]
text ‹Normalization of constraints helps to identify same polynomials, e.g.,
the constraints $x + y \leq 5$ and $-2x-2y \leq -12$ will be normalized
to $x + y \leq 5$ and $x + y \geq 6$, so that only one slack-variable will
be introduced for the polynomial $x+y$, and not another one for $-2x-2y$.
Normalization will take care that the max-var of the polynomial in the constraint
will have coefficient 1 (if the polynomial is non-zero)›
fun normalize_ns_constraint :: "'a :: lrv ns_constraint ⇒ 'a ns_constraint" where
"normalize_ns_constraint (LEQ_ns l r) = (let v = max_var l; c = coeff l v in
if c = 0 then LEQ_ns l r else
let ic = inverse c in if c < 0 then GEQ_ns (ic *R l) (scaleRat ic r) else LEQ_ns (ic *R l) (scaleRat ic r))"
| "normalize_ns_constraint (GEQ_ns l r) = (let v = max_var l; c = coeff l v in
if c = 0 then GEQ_ns l r else
let ic = inverse c in if c < 0 then LEQ_ns (ic *R l) (scaleRat ic r) else GEQ_ns (ic *R l) (scaleRat ic r))"
lemma normalize_ns_constraint[simp]: "v ⊨⇩n⇩s (normalize_ns_constraint c) ⟷ v ⊨⇩n⇩s (c :: 'a :: lrv ns_constraint)"
proof -
let ?c = "coeff (poly c) (max_var (poly c))"
consider (0) "?c = 0" | (pos) "?c > 0" | (neg) "?c < 0" by linarith
thus ?thesis
proof cases
case 0
thus ?thesis by (cases c, auto)
next
case pos
from pos have id: "a /R ?c ≤ b /R ?c ⟷ (a :: 'a) ≤ b" for a b
using scaleRat_leq1 by fastforce
show ?thesis using pos id by (cases c, auto simp: Let_def valuate_scaleRat id)
next
case neg
from neg have id: "a /R ?c ≤ b /R ?c ⟷ (a :: 'a) ≥ b" for a b
using scaleRat_leq2 by fastforce
show ?thesis using neg id by (cases c, auto simp: Let_def valuate_scaleRat id)
qed
qed
declare normalize_ns_constraint.simps[simp del]
lemma i_satisfies_normalize_ns_constraint[simp]: "Iv ⊨⇩i⇩n⇩s⇩s (map_prod id normalize_ns_constraint ` cs)
⟷ Iv ⊨⇩i⇩n⇩s⇩s cs"
by (cases Iv, force)
abbreviation max_var:: "QDelta ns_constraint ⇒ var" where
"max_var C ≡ Abstract_Linear_Poly.max_var (poly C)"
fun
start_fresh_variable :: "('i,QDelta) i_ns_constraint list ⇒ var" where
"start_fresh_variable [] = 0"
| "start_fresh_variable ((i,h)#t) = max (max_var h + 1) (start_fresh_variable t)"
definition
preprocess_part_1 :: "('i,QDelta) i_ns_constraint list ⇒ tableau × (('i,QDelta) i_atom list) × 'i list" where
"preprocess_part_1 l ≡ let start = start_fresh_variable l; is = preprocess' l start in (Tableau is, Atoms is, UnsatIndices is)"
lemma lhs_linear_poly_to_eq [simp]:
"lhs (linear_poly_to_eq h v) = v"
by (cases h) auto
lemma rvars_eq_linear_poly_to_eq [simp]:
"rvars_eq (linear_poly_to_eq h v) = vars h"
by simp
lemma fresh_var_monoinc:
"FirstFreshVariable (preprocess' cs start) ≥ start"
by (induct cs) (auto simp add: Let_def split: option.splits)
abbreviation vars_constraints where
"vars_constraints cs ≡ ⋃ (set (map vars (map poly cs)))"
lemma start_fresh_variable_fresh:
"∀ var ∈ vars_constraints (flat_list cs). var < start_fresh_variable cs"
using max_var_max
by (induct cs, auto simp add: max_def) force+
lemma vars_tableau_vars_constraints:
"rvars (Tableau (preprocess' cs start)) ⊆ vars_constraints (flat_list cs)"
by (induct cs start rule: preprocess'.induct) (auto simp add: rvars_def Let_def split: option.splits)
lemma lvars_tableau_ge_start:
"∀ var ∈ lvars (Tableau (preprocess' cs start)). var ≥ start"
by (induct cs start rule: preprocess'.induct) (auto simp add: Let_def lvars_def fresh_var_monoinc split: option.splits)
lemma rhs_no_zero_tableau_start:
"0 ∉ rhs ` set (Tableau (preprocess' cs start))"
by (induct cs start rule: preprocess'.induct, auto simp add: Let_def rvars_def fresh_var_monoinc split: option.splits)
lemma first_fresh_variable_not_in_lvars:
"∀var ∈ lvars (Tableau (preprocess' cs start)). FirstFreshVariable (preprocess' cs start) > var"
by (induct cs start rule: preprocess'.induct) (auto simp add: Let_def lvars_def split: option.splits)
lemma sat_atom_sat_eq_sat_constraint_non_monom:
assumes "v ⊨⇩a qdelta_constraint_to_atom h var" "v ⊨⇩e linear_poly_to_eq (poly h) var" "¬ is_monom (poly h)"
shows "v ⊨⇩n⇩s h"
using assms
by (cases h) (auto simp add: satisfies_eq_def split: if_splits)
lemma qdelta_constraint_to_atom_monom:
assumes "is_monom (poly h)"
shows "v ⊨⇩a qdelta_constraint_to_atom h var ⟷ v ⊨⇩n⇩s h"
proof (cases h)
case (LEQ_ns l a)
then show ?thesis
using assms
using monom_valuate[of _ v]
apply auto
using scaleRat_leq2[of "a /R monom_coeff l" "v (monom_var l)" "monom_coeff l"]
using divide_leq1[of "monom_coeff l" "v (monom_var l)" a]
apply (force, simp add: divide_rat_def)
using scaleRat_leq1[of "v (monom_var l)" "a /R monom_coeff l" "monom_coeff l"]
using is_monom_monom_coeff_not_zero[of l]
using divide_leq[of "monom_coeff l" "v (monom_var l)" a]
using is_monom_monom_coeff_not_zero[of l]
by (simp_all add: divide_rat_def)
next
case (GEQ_ns l a)
then show ?thesis
using assms
using monom_valuate[of _ v]
apply auto
using scaleRat_leq2[of "v (monom_var l)" "a /R monom_coeff l" "monom_coeff l"]
using divide_geq1[of a "monom_coeff l" "v (monom_var l)"]
apply (force, simp add: divide_rat_def)
using scaleRat_leq1[of "a /R monom_coeff l" "v (monom_var l)" "monom_coeff l"]
using is_monom_monom_coeff_not_zero[of l]
using divide_geq[of a "monom_coeff l" "v (monom_var l)"]
using is_monom_monom_coeff_not_zero[of l]
by (simp_all add: divide_rat_def)
qed
lemma preprocess'_Tableau_Poly_Mapping_None: "(Poly_Mapping (preprocess' cs start)) p = None
⟹ linear_poly_to_eq p v ∉ set (Tableau (preprocess' cs start))"
by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits if_splits)
lemma preprocess'_Tableau_Poly_Mapping_Some: "(Poly_Mapping (preprocess' cs start)) p = Some v
⟹ linear_poly_to_eq p v ∈ set (Tableau (preprocess' cs start))"
by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits if_splits)
lemma preprocess'_Tableau_Poly_Mapping_Some': "(Poly_Mapping (preprocess' cs start)) p = Some v
⟹ ∃ h. poly h = p ∧ ¬ is_monom (poly h) ∧ qdelta_constraint_to_atom h v ∈ flat (set (Atoms (preprocess' cs start)))"
by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits if_splits)
lemma not_one_le_zero_qdelta: "¬ (1 ≤ (0 :: QDelta))" by code_simp
lemma one_zero_contra[dest,consumes 2]: "1 ≤ x ⟹ (x :: QDelta) ≤ 0 ⟹ False"
using order.trans[of 1 x 0] not_one_le_zero_qdelta by simp
lemma i_preprocess'_sat:
assumes "(I,v) ⊨⇩i⇩a⇩s set (Atoms (preprocess' s start))" "v ⊨⇩t Tableau (preprocess' s start)"
"I ∩ set (UnsatIndices (preprocess' s start)) = {}"
shows "(I,v) ⊨⇩i⇩n⇩s⇩s set s"
using assms
by (induct s start rule: preprocess'.induct)
(auto simp add: Let_def satisfies_atom_set_def satisfies_tableau_def qdelta_constraint_to_atom_monom
sat_atom_sat_eq_sat_constraint_non_monom
split: if_splits option.splits dest!: preprocess'_Tableau_Poly_Mapping_Some zero_satisfies)
lemma preprocess'_sat:
assumes "v ⊨⇩a⇩s flat (set (Atoms (preprocess' s start)))" "v ⊨⇩t Tableau (preprocess' s start)" "set (UnsatIndices (preprocess' s start)) = {}"
shows "v ⊨⇩n⇩s⇩s flat (set s)"
using i_preprocess'_sat[of UNIV v s start] assms by simp
lemma sat_constraint_valuation:
assumes "∀ var ∈ vars (poly c). v1 var = v2 var"
shows "v1 ⊨⇩n⇩s c ⟷ v2 ⊨⇩n⇩s c"
using assms
using valuate_depend
by (cases c) (force)+
lemma atom_var_first:
assumes "a ∈ flat (set (Atoms (preprocess' cs start)))" "∀ var ∈ vars_constraints (flat_list cs). var < start"
shows "atom_var a < FirstFreshVariable (preprocess' cs start)"
using assms
proof(induct cs arbitrary: a)
case (Cons hh t a)
obtain i h where hh: "hh = (i,h)" by force
let ?s = "preprocess' t start"
show ?case
proof(cases "a ∈ flat (set (Atoms ?s))")
case True
then show ?thesis
using Cons(1)[of a] Cons(3) hh
by (auto simp add: Let_def split: option.splits)
next
case False
consider (monom) "is_monom (poly h)" | (normal) "¬ is_monom (poly h)" "poly h ≠ 0" "(Poly_Mapping ?s) (poly h) = None"
| (old) var where "¬ is_monom (poly h)" "poly h ≠ 0" "(Poly_Mapping ?s) (poly h) = Some var"
| (zero) "¬ is_monom (poly h)" "poly h = 0"
by auto
then show ?thesis
proof cases
case monom
from Cons(3) monom_var_in_vars hh monom
have "monom_var (poly h) < start" by auto
moreover from False have "a = qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))"
using Cons(2) hh monom by (auto simp: Let_def)
ultimately show ?thesis
using fresh_var_monoinc[of start t] hh monom
by (cases a; cases h) (auto simp add: Let_def )
next
case normal
have "a = qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))"
using False normal Cons(2) hh by (auto simp: Let_def)
then show ?thesis using hh normal
by (cases a; cases h) (auto simp add: Let_def )
next
case (old var)
from preprocess'_Tableau_Poly_Mapping_Some'[OF old(3)]
obtain h' where "poly h' = poly h" "qdelta_constraint_to_atom h' var ∈ flat (set (Atoms ?s))"
by blast
from Cons(1)[OF this(2)] Cons(3) this(1) old(1)
have var: "var < FirstFreshVariable ?s" by (cases h', auto)
have "a = qdelta_constraint_to_atom h var"
using False old Cons(2) hh by (auto simp: Let_def)
then have a: "atom_var a = var" using old by (cases a; cases h; auto simp: Let_def)
show ?thesis unfolding a hh by (simp add: old Let_def var)
next
case zero
from False show ?thesis using Cons(2) hh zero by (auto simp: Let_def split: if_splits)
qed
qed
qed simp
lemma satisfies_tableau_satisfies_tableau:
assumes "v1 ⊨⇩t t" "∀ var ∈ tvars t. v1 var = v2 var"
shows "v2 ⊨⇩t t"
using assms
using valuate_depend[of _ v1 v2]
by (force simp add: lvars_def rvars_def satisfies_eq_def satisfies_tableau_def)
lemma preprocess'_unsat_indices:
assumes "i ∈ set (UnsatIndices (preprocess' s start))"
shows "¬ ({i},v) ⊨⇩i⇩n⇩s⇩s set s"
using assms
proof (induct s start rule: preprocess'.induct)
case (2 j h t v)
then show ?case by (auto simp: Let_def not_zero_satisfies split: if_splits option.splits)
qed simp
lemma preprocess'_unsat:
assumes "(I,v) ⊨⇩i⇩n⇩s⇩s set s" "vars_constraints (flat_list s) ⊆ V" "∀var ∈ V. var < start"
shows "∃v'. (∀var ∈ V. v var = v' var)
∧ v' ⊨⇩a⇩s restrict_to I (set (Atoms (preprocess' s start)))
∧ v' ⊨⇩t Tableau (preprocess' s start)"
using assms
proof(induct s)
case Nil
show ?case
by (auto simp add: satisfies_atom_set_def satisfies_tableau_def)
next
case (Cons hh t)
obtain i h where hh: "hh = (i,h)" by force
from Cons hh obtain v' where
var: "(∀var∈V. v var = v' var)"
and v'_as: "v' ⊨⇩a⇩s restrict_to I (set (Atoms (preprocess' t start)))"
and v'_t: "v' ⊨⇩t Tableau (preprocess' t start)"
and vars_h: "vars_constraints [h] ⊆ V"
by auto
from Cons(2)[unfolded hh]
have i: "i ∈ I ⟹ v ⊨⇩n⇩s h" by auto
have "∀ var ∈ vars (poly h). v var = v' var"
using ‹(∀var∈V. v var = v' var)› Cons(3) hh
by auto
then have vh_v'h: "v ⊨⇩n⇩s h ⟷ v' ⊨⇩n⇩s h"
by (rule sat_constraint_valuation)
show ?case
proof(cases "is_monom (poly h)")
case True
then have id: "is_monom (poly h) = True" by simp
show ?thesis
unfolding hh preprocess'.simps Let_def id if_True istate.simps istate.sel
proof (intro exI[of _ v'] conjI v'_t var satisifies_atom_restrict_to_Cons[OF v'_as])
assume "i ∈ I"
from i[OF this] var vh_v'h
show "v' ⊨⇩a qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))"
unfolding qdelta_constraint_to_atom_monom[OF True] by auto
qed
next
case False
then have id: "is_monom (poly h) = False" by simp
let ?s = "preprocess' t start"
let ?x = "FirstFreshVariable ?s"
show ?thesis
proof (cases "poly h = 0")
case zero: False
hence id': "(poly h = 0) = False" by simp
let ?look = "(Poly_Mapping ?s) (poly h)"
show ?thesis
proof (cases ?look)
case None
let ?y = "poly h ⦃ v'⦄"
let ?v' = "v'(?x:=?y)"
show ?thesis unfolding preprocess'.simps hh Let_def id id' if_False istate.simps istate.sel None option.simps
proof (rule exI[of _ ?v'], intro conjI satisifies_atom_restrict_to_Cons satisfies_tableau_Cons)
show vars': "(∀var∈V. v var = ?v' var)"
using ‹(∀var∈V. v var = v' var)›
using fresh_var_monoinc[of start t]
using Cons(4)
by auto
{
assume "i ∈ I"
from vh_v'h i[OF this] False
show "?v' ⊨⇩a qdelta_constraint_to_atom h (FirstFreshVariable (preprocess' t start))"
by (cases h, auto)
}
let ?atoms = "restrict_to I (set (Atoms (preprocess' t start)))"
show "?v' ⊨⇩a⇩s ?atoms"
unfolding satisfies_atom_set_def
proof
fix a
assume "a ∈ ?atoms"
then have "v' ⊨⇩a a"
using ‹v' ⊨⇩a⇩s ?atoms› hh by (force simp add: satisfies_atom_set_def)
then show "?v' ⊨⇩a a"
using ‹a ∈ ?atoms› atom_var_first[of a t start]
using Cons(3) Cons(4)
by (cases a) auto
qed
show "?v' ⊨⇩e linear_poly_to_eq (poly h) (FirstFreshVariable (preprocess' t start))"
using Cons(3) Cons(4)
using valuate_depend[of "poly h" v' "v'(FirstFreshVariable (preprocess' t start) := (poly h) ⦃ v' ⦄)"]
using fresh_var_monoinc[of start t] hh
by (cases h) (force simp add: satisfies_eq_def)+
have "FirstFreshVariable (preprocess' t start) ∉ tvars (Tableau (preprocess' t start))"
using first_fresh_variable_not_in_lvars[of t start]
using Cons(3) Cons(4)
using vars_tableau_vars_constraints[of t start]
using fresh_var_monoinc[of start t]
by force
then show "?v' ⊨⇩t Tableau (preprocess' t start)"
using ‹v' ⊨⇩t Tableau (preprocess' t start)›
using satisfies_tableau_satisfies_tableau[of v' "Tableau (preprocess' t start)" ?v']
by auto
qed
next
case (Some var)
from preprocess'_Tableau_Poly_Mapping_Some[OF Some]
have "linear_poly_to_eq (poly h) var ∈ set (Tableau ?s)" by auto
with v'_t[unfolded satisfies_tableau_def]
have v'_h_var: "v' ⊨⇩e linear_poly_to_eq (poly h) var" by auto
show ?thesis unfolding preprocess'.simps hh Let_def id id' if_False istate.simps istate.sel Some option.simps
proof (intro exI[of _ v'] conjI var v'_t satisifies_atom_restrict_to_Cons satisfies_tableau_Cons v'_as)
assume "i ∈ I"
from vh_v'h i[OF this] False v'_h_var
show "v' ⊨⇩a qdelta_constraint_to_atom h var"
by (cases h, auto simp: satisfies_eq_iff)
qed
qed
next
case zero: True
hence id': "(poly h = 0) = True" by simp
show ?thesis
proof (cases "zero_satisfies h")
case True
hence id'': "zero_satisfies h = True" by simp
show ?thesis
unfolding hh preprocess'.simps Let_def id id' id'' if_True if_False istate.simps istate.sel
by (intro exI[of _ v'] conjI v'_t var v'_as)
next
case False
hence id'': "zero_satisfies h = False" by simp
{
assume "i ∈ I"
from i[OF this] not_zero_satisfies[OF zero False] have False by simp
} note no_I = this
show ?thesis
unfolding hh preprocess'.simps Let_def id id' id'' if_True if_False istate.simps istate.sel
proof (rule Cons(1)[OF _ _ Cons(4)])
show "(I, v) ⊨⇩i⇩n⇩s⇩s set t" using Cons(2) by auto
show "vars_constraints (map snd t) ⊆ V" using Cons(3) by force
qed
qed
qed
qed
qed
lemma lvars_distinct:
"distinct (map lhs (Tableau (preprocess' cs start)))"
using first_fresh_variable_not_in_lvars[where ?'a = 'a]
by (induct cs, auto simp add: Let_def lvars_def) (force split: option.splits)
lemma normalized_tableau_preprocess': "△ (Tableau (preprocess' cs (start_fresh_variable cs)))"
proof -
let ?s = "start_fresh_variable cs"
show ?thesis
using lvars_distinct[of cs ?s]
using lvars_tableau_ge_start[of cs ?s]
using vars_tableau_vars_constraints[of cs ?s]
using start_fresh_variable_fresh[of cs]
unfolding normalized_tableau_def Let_def
by (smt disjoint_iff_not_equal inf.absorb_iff2 inf.strict_order_iff rhs_no_zero_tableau_start subsetD)
qed
text ‹Improved preprocessing: Deletion. An equation x = p can be deleted from the tableau,
if x does not occur in the atoms.›
lemma delete_lhs_var: assumes norm: "△ t" and t: "t = t1 @ (x,p) # t2"
and t': "t' = t1 @ t2"
and tv: "tv = (λ v. upd x (p ⦃ ⟨v⟩ ⦄) v)"
and x_as: "x ∉ atom_var ` snd ` set as"
shows "△ t'"
"⟨w⟩ ⊨⇩t t' ⟹ ⟨tv w⟩ ⊨⇩t t"
"(I, ⟨w⟩) ⊨⇩i⇩a⇩s set as ⟹ (I, ⟨tv w⟩) ⊨⇩i⇩a⇩s set as"
"v ⊨⇩t t ⟹ v ⊨⇩t t'"
proof -
have tv: "⟨tv w⟩ = ⟨w⟩ (x := p ⦃ ⟨w⟩ ⦄)" unfolding tv map2fun_def' by auto
from norm
show "△ t'" unfolding t t' normalized_tableau_def by (auto simp: lvars_def rvars_def)
show "v ⊨⇩t t ⟹ v ⊨⇩t t'" unfolding t t' satisfies_tableau_def by auto
from norm have dist: "distinct (map lhs t)" "lvars t ∩ rvars t = {}"
unfolding normalized_tableau_def by auto
from x_as have x_as: "x ∉ atom_var ` snd ` (set as ∩ I × UNIV)" by auto
have "(I, ⟨tv w⟩) ⊨⇩i⇩a⇩s set as ⟷ (I, ⟨w⟩) ⊨⇩i⇩a⇩s set as" unfolding i_satisfies_atom_set.simps
satisfies_atom_set_def
proof (intro ball_cong[OF refl])
fix a
assume "a ∈ snd ` (set as ∩ I × UNIV)"
with x_as have "x ≠ atom_var a" by auto
then show "⟨tv w⟩ ⊨⇩a a = ⟨w⟩ ⊨⇩a a" unfolding tv
by (cases a, auto)
qed
then show "(I, ⟨w⟩) ⊨⇩i⇩a⇩s set as ⟹ (I, ⟨tv w⟩) ⊨⇩i⇩a⇩s set as" by blast
assume w: "⟨w⟩ ⊨⇩t t'"
from dist(2)[unfolded t] have xp: "x ∉ vars p"
unfolding lvars_def rvars_def by auto
{
fix eq
assume mem: "eq ∈ set t1 ∪ set t2"
then have "eq ∈ set t'" unfolding t' by auto
with w have w: "⟨w⟩ ⊨⇩e eq" unfolding satisfies_tableau_def by auto
obtain y q where eq: "eq = (y,q)" by force
from mem[unfolded eq] dist(1)[unfolded t] have xy: "x ≠ y" by force
from mem[unfolded eq] dist(2)[unfolded t] have xq: "x ∉ vars q"
unfolding lvars_def rvars_def by auto
from w have "⟨tv w⟩ ⊨⇩e eq" unfolding tv eq satisfies_eq_iff using xy xq
by (auto intro!: valuate_depend)
}
moreover
have "⟨tv w⟩ ⊨⇩e (x,p)" unfolding satisfies_eq_iff tv using xp
by (auto intro!: valuate_depend)
ultimately
show "⟨tv w⟩ ⊨⇩t t" unfolding t satisfies_tableau_def by auto
qed
definition pivot_tableau_eq :: "tableau ⇒ eq ⇒ tableau ⇒ var ⇒ tableau × eq × tableau" where
"pivot_tableau_eq t1 eq t2 x ≡ let eq' = pivot_eq eq x; m = map (λ e. subst_var_eq x (rhs eq') e) in
(m t1, eq', m t2)"
lemma pivot_tableau_eq: assumes t: "t = t1 @ eq # t2" "t' = t1' @ eq' # t2'"
and x: "x ∈ rvars_eq eq" and norm: "△ t" and pte: "pivot_tableau_eq t1 eq t2 x = (t1',eq',t2')"
shows "△ t'" "lhs eq' = x" "(v :: 'a :: lrv valuation) ⊨⇩t t' ⟷ v ⊨⇩t t"
proof -
let ?s = "λ t. State t undefined undefined undefined undefined undefined"
let ?y = "lhs eq"
have yl: "?y ∈ lvars t" unfolding t lvars_def by auto
from norm have eq_t12: "?y ∉ lhs ` (set t1 ∪ set t2)"
unfolding normalized_tableau_def t lvars_def by auto
have eq: "eq_for_lvar_code t ?y = eq"
by (metis (mono_tags, lifting) EqForLVarDefault.eq_for_lvar Un_insert_right eq_t12
image_iff insert_iff list.set(2) set_append t(1) yl)
have *: "(?y, b) ∈ set t1 ⟹ ?y ∈ lhs ` (set t1)" for b t1
by (metis image_eqI lhs.simps)
have pivot: "pivot_tableau_code ?y x t = t'"
unfolding Pivot'Default.pivot_tableau_def Let_def eq using pte[symmetric]
unfolding t pivot_tableau_eq_def Let_def using eq_t12 by (auto dest!: *)
note thms = Pivot'Default.pivot_vars' Pivot'Default.pivot_tableau
note thms = thms[unfolded Pivot'Default.pivot_def, of "?s t", simplified,
OF norm yl, unfolded eq, OF x, unfolded pivot]
from thms(1) thms(2)[of v] show "△ t'" "v ⊨⇩t t' ⟷ v ⊨⇩t t" by auto
show "lhs eq' = x" using pte[symmetric]
unfolding t pivot_tableau_eq_def Let_def pivot_eq_def by auto
qed
function preprocess_opt :: "var set ⇒ tableau ⇒ tableau ⇒ tableau × ((var,'a :: lrv)mapping ⇒ (var,'a)mapping)" where
"preprocess_opt X t1 [] = (t1,id)"
| "preprocess_opt X t1 ((x,p) # t2) = (if x ∉ X then
case preprocess_opt X t1 t2 of (t,tv) ⇒ (t, (λ v. upd x (p ⦃ ⟨v⟩ ⦄) v) o tv)
else case find (λ x. x ∉ X) (Abstract_Linear_Poly.vars_list p) of
None ⇒ preprocess_opt X ((x,p) # t1) t2
| Some y ⇒ case pivot_tableau_eq t1 (x,p) t2 y of
(tt1,(z,q),tt2) ⇒ case preprocess_opt X tt1 tt2 of (t,tv) ⇒ (t, (λ v. upd z (q ⦃ ⟨v⟩ ⦄) v) o tv))"
by pat_completeness auto
termination by (relation "measure (λ (X,t1,t2). length t2)", auto simp: pivot_tableau_eq_def Let_def)
lemma preprocess_opt: assumes "X = atom_var ` snd ` set as"
"preprocess_opt X t1 t2 = (t',tv)" "△ t" "t = rev t1 @ t2"
shows "△ t'"
"(⟨w⟩ :: 'a :: lrv valuation) ⊨⇩t t' ⟹ ⟨tv w⟩ ⊨⇩t t"
"(I, ⟨w⟩) ⊨⇩i⇩a⇩s set as ⟹ (I, ⟨tv w⟩) ⊨⇩i⇩a⇩s set as"
"v ⊨⇩t t ⟹ (v :: 'a valuation) ⊨⇩t t'"
using assms
proof (atomize(full), induct X t1 t2 arbitrary: t tv w rule: preprocess_opt.induct)
case (1 X t1 t tv)
then show ?case by (auto simp: normalized_tableau_def lvars_def rvars_def satisfies_tableau_def
simp flip: rev_map)
next
case (2 X t1 x p t2 t tv w)
note IH = 2(1-3)
note X = 2(4)
note res = 2(5)
have norm: "△ t" by fact
have t: "t = rev t1 @ (x, p) # t2" by fact
show ?case
proof (cases "x ∈ X")
case False
with res obtain tv' where res: "preprocess_opt X t1 t2 = (t', tv')" and
tv: "tv = (λv. Mapping.update x (p ⦃ ⟨v⟩ ⦄) v) o tv'"
by (auto split: prod.splits)
note delete = delete_lhs_var[OF norm t refl refl False[unfolded X]]
note IH = IH(1)[OF False X res delete(1) refl]
from delete(2)[of "tv' w"] delete(3)[of I "tv' w"] delete(4)[of v] IH[of w]
show ?thesis unfolding tv o_def
by auto
next
case True
then have "¬ x ∉ X" by simp
note IH = IH(2-3)[OF this]
show ?thesis
proof (cases "find (λx. x ∉ X) (Abstract_Linear_Poly.vars_list p)")
case None
with res True have pre: "preprocess_opt X ((x, p) # t1) t2 = (t', tv)" by auto
from t have t: "t = rev ((x, p) # t1) @ t2" by simp
from IH(1)[OF None X pre norm t]
show ?thesis .
next
case (Some z)
from Some[unfolded find_Some_iff] have zX: "z ∉ X" and "z ∈ set (Abstract_Linear_Poly.vars_list p)"
unfolding set_conv_nth by auto
then have z: "z ∈ rvars_eq (x, p)" by (simp add: set_vars_list)
obtain tt1 z' q tt2 where pte: "pivot_tableau_eq t1 (x, p) t2 z = (tt1,(z',q),tt2)"
by (cases "pivot_tableau_eq t1 (x, p) t2 z", auto)
then have pte_rev: "pivot_tableau_eq (rev t1) (x, p) t2 z = (rev tt1,(z',q),tt2)"
unfolding pivot_tableau_eq_def Let_def by (auto simp: rev_map)
note eq = pivot_tableau_eq[OF t refl z norm pte_rev]
then have z': "z' = z" by auto
note eq = eq(1,3)[unfolded z']
note pte = pte[unfolded z']
note pte_rev = pte_rev[unfolded z']
note delete = delete_lhs_var[OF eq(1) refl refl refl zX[unfolded X]]
from res[unfolded preprocess_opt.simps Some option.simps pte] True
obtain tv' where res: "preprocess_opt X tt1 tt2 = (t', tv')" and
tv: "tv = (λv. Mapping.update z (q ⦃ ⟨v⟩ ⦄) v) o tv'"
by (auto split: prod.splits)
note IH = IH(2)[OF Some, unfolded pte, OF refl refl refl X res delete(1) refl]
from IH[of w] delete(2)[of "tv' w"] delete(3)[of I "tv' w"] delete(4)[of v] show ?thesis
unfolding tv o_def eq(2) by auto
qed
qed
qed
definition "preprocess_part_2 as t = preprocess_opt (atom_var ` snd ` set as) [] t"
lemma preprocess_part_2: assumes "preprocess_part_2 as t = (t',tv)" "△ t"
shows "△ t'"
"(⟨w⟩ :: 'a :: lrv valuation) ⊨⇩t t' ⟹ ⟨tv w⟩ ⊨⇩t t"
"(I, ⟨w⟩) ⊨⇩i⇩a⇩s set as ⟹ (I, ⟨tv w⟩) ⊨⇩i⇩a⇩s set as"
"v ⊨⇩t t ⟹ (v :: 'a valuation) ⊨⇩t t'"
using preprocess_opt[OF refl assms(1)[unfolded preprocess_part_2_def] assms(2)] by auto
definition preprocess :: "('i,QDelta) i_ns_constraint list ⇒ _ × _ × (_ ⇒ (var,QDelta)mapping) × 'i list" where
"preprocess l = (case preprocess_part_1 (map (map_prod id normalize_ns_constraint) l) of
(t,as,ui) ⇒ case preprocess_part_2 as t of (t,tv) ⇒ (t,as,tv,ui))"
lemma preprocess:
assumes id: "preprocess cs = (t, as, trans_v, ui)"
shows "△ t"
"fst ` set as ∪ set ui ⊆ fst ` set cs"
"distinct_indices_ns (set cs) ⟹ distinct_indices_atoms (set as)"
"I ∩ set ui = {} ⟹ (I, ⟨v⟩) ⊨⇩i⇩a⇩s set as ⟹
⟨v⟩ ⊨⇩t t ⟹ (I, ⟨trans_v v⟩) ⊨⇩i⇩n⇩s⇩s set cs"
"i ∈ set ui ⟹ ∄v. ({i}, v) ⊨⇩i⇩n⇩s⇩s set cs"
"∃ v. (I,v) ⊨⇩i⇩n⇩s⇩s set cs ⟹ ∃v'. (I,v') ⊨⇩i⇩a⇩s set as ∧ v' ⊨⇩t t"
proof -
define ncs where "ncs = map (map_prod id normalize_ns_constraint) cs"
have ncs: "fst ` set ncs = fst ` set cs" "⋀ Iv. Iv ⊨⇩i⇩n⇩s⇩s set ncs ⟷ Iv ⊨⇩i⇩n⇩s⇩s set cs"
unfolding ncs_def by force auto
from id obtain t1 where part1: "preprocess_part_1 ncs = (t1,as,ui)"
unfolding preprocess_def by (auto simp: ncs_def split: prod.splits)
from id[unfolded preprocess_def part1 split ncs_def[symmetric]]
have part_2: "preprocess_part_2 as t1 = (t,trans_v)"
by (auto split: prod.splits)
have norm: "△ t1" using normalized_tableau_preprocess' part1
by (auto simp: preprocess_part_1_def Let_def)
note part_2 = preprocess_part_2[OF part_2 norm]
show "△ t" by fact
have unsat: "(I,⟨v⟩) ⊨⇩i⇩a⇩s set as ⟹ ⟨v⟩ ⊨⇩t t1 ⟹ I ∩ set ui = {} ⟹ (I,⟨v⟩) ⊨⇩i⇩n⇩s⇩s set ncs" for v
using part1[unfolded preprocess_part_1_def Let_def, simplified] i_preprocess'_sat[of I] by blast
with part_2(2,3) show "I ∩ set ui = {} ⟹ (I,⟨v⟩) ⊨⇩i⇩a⇩s set as ⟹ ⟨v⟩ ⊨⇩t t ⟹ (I,⟨trans_v v⟩) ⊨⇩i⇩n⇩s⇩s set cs"
by (auto simp: ncs)
from part1[unfolded preprocess_part_1_def Let_def] obtain var where
as: "as = Atoms (preprocess' ncs var)" and ui: "ui = UnsatIndices (preprocess' ncs var)" by auto
note min_defs = distinct_indices_atoms_def distinct_indices_ns_def
have min1: "(distinct_indices_ns (set ncs) ⟶ (∀ k a. (k,a) ∈ set as ⟶ (∃ v p. a = qdelta_constraint_to_atom p v ∧ (k,p) ∈ set ncs
∧ (¬ is_monom (poly p) ⟶ Poly_Mapping (preprocess' ncs var) (poly p) = Some v) )))
∧ fst ` set as ∪ set ui ⊆ fst ` set ncs"
unfolding as ui
proof (induct ncs var rule: preprocess'.induct)
case (2 i h t v)
hence sub: "fst ` set (Atoms (preprocess' t v)) ∪ set (UnsatIndices (preprocess' t v)) ⊆ fst ` set t" by auto
show ?case
proof (intro conjI impI allI, goal_cases)
show "fst ` set (Atoms (preprocess' ((i, h) # t) v)) ∪ set (UnsatIndices (preprocess' ((i,h) #t) v)) ⊆ fst ` set ((i, h) # t)"
using sub by (auto simp: Let_def split: option.splits)
next
case (1 k a)
hence min': "distinct_indices_ns (set t)" unfolding min_defs list.simps by blast
note IH = 2[THEN conjunct1, rule_format, OF min']
show ?case
proof (cases "(k,a) ∈ set (Atoms (preprocess' t v))")
case True
from IH[OF this] show ?thesis
by (force simp: Let_def split: option.splits if_split)
next
case new: False
with 1(2) have ki: "k = i" by (auto simp: Let_def split: if_splits option.splits)
show ?thesis
proof (cases "is_monom (poly h)")
case True
thus ?thesis using new 1(2) by (auto simp: Let_def True intro!: exI)
next
case no_monom: False
thus ?thesis using new 1(2) by (auto simp: Let_def no_monom split: option.splits if_splits intro!: exI)
qed
qed
qed
qed (auto simp: min_defs)
then show "fst ` set as ∪ set ui ⊆ fst ` set cs" by (auto simp: ncs)
{
assume mini: "distinct_indices_ns (set cs)"
have mini: "distinct_indices_ns (set ncs)" unfolding distinct_indices_ns_def
proof (intro impI allI, goal_cases)
case (1 n1 n2 i)
from 1(1) obtain c1 where c1: "(i,c1) ∈ set cs" and n1: "n1 = normalize_ns_constraint c1"
unfolding ncs_def by auto
from 1(2) obtain c2 where c2: "(i,c2) ∈ set cs" and n2: "n2 = normalize_ns_constraint c2"
unfolding ncs_def by auto
from mini[unfolded distinct_indices_ns_def, rule_format, OF c1 c2]
show ?case unfolding n1 n2
by (cases c1; cases c2; auto simp: normalize_ns_constraint.simps Let_def)
qed
note min = min1[THEN conjunct1, rule_format, OF this]
show "distinct_indices_atoms (set as)"
unfolding distinct_indices_atoms_def
proof (intro allI impI)
fix i a b
assume a: "(i,a) ∈ set as" and b: "(i,b) ∈ set as"
from min[OF a] obtain v p where aa: "a = qdelta_constraint_to_atom p v" "(i, p) ∈ set ncs"
"¬ is_monom (poly p) ⟹ Poly_Mapping (preprocess' ncs var) (poly p) = Some v"
by auto
from min[OF b] obtain w q where bb: "b = qdelta_constraint_to_atom q w" "(i, q) ∈ set ncs"
"¬ is_monom (poly q) ⟹ Poly_Mapping (preprocess' ncs var) (poly q) = Some w"
by auto
from mini[unfolded distinct_indices_ns_def, rule_format, OF aa(2) bb(2)]
have *: "poly p = poly q" "ns_constraint_const p = ns_constraint_const q" by auto
show "atom_var a = atom_var b ∧ atom_const a = atom_const b"
proof (cases "is_monom (poly q)")
case True
thus ?thesis unfolding aa(1) bb(1) using * by (cases p; cases q, auto)
next
case False
thus ?thesis unfolding aa(1) bb(1) using * aa(3) bb(3) by (cases p; cases q, auto)
qed
qed
}
show "i ∈ set ui ⟹ ∄v. ({i}, v) ⊨⇩i⇩n⇩s⇩s set cs"
using preprocess'_unsat_indices[of i ncs] part1 unfolding preprocess_part_1_def Let_def
by (auto simp: ncs)
assume "∃ w. (I,w) ⊨⇩i⇩n⇩s⇩s set cs"
then obtain w where "(I,w) ⊨⇩i⇩n⇩s⇩s set cs" by blast
hence "(I,w) ⊨⇩i⇩n⇩s⇩s set ncs" unfolding ncs .
from preprocess'_unsat[OF this _ start_fresh_variable_fresh, of ncs]
have "∃v'. (I,v') ⊨⇩i⇩a⇩s set as ∧ v' ⊨⇩t t1"
using part1
unfolding preprocess_part_1_def Let_def by auto
then show "∃v'. (I,v') ⊨⇩i⇩a⇩s set as ∧ v' ⊨⇩t t"
using part_2(4) by auto
qed
interpretation PreprocessDefault: Preprocess preprocess
by (unfold_locales; rule preprocess, auto)
global_interpretation Solve_exec_ns'Default: Solve_exec_ns' preprocess assert_all_code
defines solve_exec_ns_code = Solve_exec_ns'Default.solve_exec_ns
by unfold_locales
primrec
constraint_to_qdelta_constraint:: "constraint ⇒ QDelta ns_constraint list" where
"constraint_to_qdelta_constraint (LT l r) = [LEQ_ns l (QDelta.QDelta r (-1))]"
| "constraint_to_qdelta_constraint (GT l r) = [GEQ_ns l (QDelta.QDelta r 1)]"
| "constraint_to_qdelta_constraint (LEQ l r) = [LEQ_ns l (QDelta.QDelta r 0)]"
| "constraint_to_qdelta_constraint (GEQ l r) = [GEQ_ns l (QDelta.QDelta r 0)]"
| "constraint_to_qdelta_constraint (EQ l r) = [LEQ_ns l (QDelta.QDelta r 0), GEQ_ns l (QDelta.QDelta r 0)]"
| "constraint_to_qdelta_constraint (LTPP l1 l2) = [LEQ_ns (l1 - l2) (QDelta.QDelta 0 (-1))]"
| "constraint_to_qdelta_constraint (GTPP l1 l2) = [GEQ_ns (l1 - l2) (QDelta.QDelta 0 1)]"
| "constraint_to_qdelta_constraint (LEQPP l1 l2) = [LEQ_ns (l1 - l2) 0]"
| "constraint_to_qdelta_constraint (GEQPP l1 l2) = [GEQ_ns (l1 - l2) 0]"
| "constraint_to_qdelta_constraint (EQPP l1 l2) = [LEQ_ns (l1 - l2) 0, GEQ_ns (l1 - l2) 0]"
primrec
i_constraint_to_qdelta_constraint:: "'i i_constraint ⇒ ('i,QDelta) i_ns_constraint list" where
"i_constraint_to_qdelta_constraint (i,c) = map (Pair i) (constraint_to_qdelta_constraint c)"
definition
to_ns :: "'i i_constraint list ⇒ ('i,QDelta) i_ns_constraint list" where
"to_ns l ≡ concat (map i_constraint_to_qdelta_constraint l)"
primrec
δ0_val :: "QDelta ns_constraint ⇒ QDelta valuation ⇒ rat" where
"δ0_val (LEQ_ns lll rrr) vl = δ0 lll⦃vl⦄ rrr"
| "δ0_val (GEQ_ns lll rrr) vl = δ0 rrr lll⦃vl⦄"
primrec
δ0_val_min :: "QDelta ns_constraint list ⇒ QDelta valuation ⇒ rat" where
"δ0_val_min [] vl = 1"
| "δ0_val_min (h#t) vl = min (δ0_val_min t vl) (δ0_val h vl)"
abbreviation vars_list_constraints where
"vars_list_constraints cs ≡ remdups (concat (map Abstract_Linear_Poly.vars_list (map poly cs)))"
definition
from_ns ::"(var, QDelta) mapping ⇒ QDelta ns_constraint list ⇒ (var, rat) mapping" where
"from_ns vl cs ≡ let δ = δ0_val_min cs ⟨vl⟩ in
Mapping.tabulate (vars_list_constraints cs) (λ var. val (⟨vl⟩ var) δ)"
global_interpretation SolveExec'Default: SolveExec' to_ns from_ns solve_exec_ns_code
defines solve_exec_code = SolveExec'Default.solve_exec
and solve_code = SolveExec'Default.solve
proof unfold_locales
{
fix ics :: "'i i_constraint list" and v' and I
let ?to_ns = "to_ns ics"
let ?flat = "set ?to_ns"
assume sat: "(I,⟨v'⟩) ⊨⇩i⇩n⇩s⇩s ?flat"
define cs where "cs = map snd (filter (λ ic. fst ic ∈ I) ics)"
define to_ns' where to_ns: "to_ns' = (λ l. concat (map constraint_to_qdelta_constraint l))"
show "(I,⟨from_ns v' (flat_list ?to_ns)⟩) ⊨⇩i⇩c⇩s set ics" unfolding i_satisfies_cs.simps
proof
let ?listf = "map (λC. case C of (LEQ_ns l r) ⇒ (l⦃⟨v'⟩⦄, r)
| (GEQ_ns l r) ⇒ (r, l⦃⟨v'⟩⦄)
)"
let ?to_ns = "λ ics. to_ns' (map snd (filter (λic. fst ic ∈ I) ics))"
let ?list = "?listf (to_ns' cs)"
let ?f_list = "flat_list (to_ns ics)"
let ?flist = "?listf ?f_list"
obtain i_list where i_list: "?list = i_list" by force
obtain f_list where f_list: "?flist = f_list" by force
have if_list: "set i_list ⊆ set f_list" unfolding
i_list[symmetric] f_list[symmetric] to_ns_def to_ns set_map set_concat cs_def
by (intro image_mono, force)
have "⋀ qd1 qd2. (qd1, qd2) ∈ set ?list ⟹ qd1 ≤ qd2"
proof-
fix qd1 qd2
assume "(qd1, qd2) ∈ set ?list"
then show "qd1 ≤ qd2"
using sat unfolding cs_def
proof(induct ics)
case Nil
then show ?case
by (simp add: to_ns)
next
case (Cons h t)
obtain i c where h: "h = (i,c)" by force
from Cons(2) consider (ic) "(qd1,qd2) ∈ set (?listf (?to_ns [(i,c)]))"
| (t) "(qd1,qd2) ∈ set (?listf (?to_ns t))"
unfolding to_ns h set_map set_concat by fastforce
then show ?case
proof cases
case t
from Cons(1)[OF this] Cons(3) show ?thesis unfolding to_ns_def by auto
next
case ic
note ic = ic[unfolded to_ns, simplified]
from ic have i: "(i ∈ I) = True" by (cases "i ∈ I", auto)
note ic = ic[unfolded i if_True, simplified]
from Cons(3)[unfolded h] i have "⟨v'⟩ ⊨⇩n⇩s⇩s set (to_ns' [c])"
unfolding i_satisfies_ns_constraints.simps unfolding to_ns to_ns_def by force
with ic show ?thesis by (induct c) (auto simp add: to_ns)
qed
qed
qed
then have l1: "ε > 0 ⟹ ε ≤ (δ_min ?list) ⟹ ∀qd1 qd2. (qd1, qd2) ∈ set ?list ⟶ val qd1 ε ≤ val qd2 ε" for ε
unfolding i_list
by (simp add: delta_gt_zero delta_min[of i_list])
have "δ_min ?flist ≤ δ_min ?list" unfolding f_list i_list
by (rule delta_min_mono[OF if_list])
from l1[OF delta_gt_zero this]
have l1: "∀qd1 qd2. (qd1, qd2) ∈ set ?list ⟶ val qd1 (δ_min f_list) ≤ val qd2 (δ_min f_list)"
unfolding f_list .
have "δ0_val_min (flat_list (to_ns ics)) ⟨v'⟩ = δ_min f_list" unfolding f_list[symmetric]
proof(induct ics)
case Nil
show ?case
by (simp add: to_ns_def)
next
case (Cons h t)
then show ?case
by (cases h; cases "snd h") (auto simp add: to_ns_def)
qed
then have l2: "from_ns v' ?f_list = Mapping.tabulate (vars_list_constraints ?f_list) (λ var. val (⟨v'⟩ var) (δ_min f_list))"
by (auto simp add: from_ns_def)
fix c
assume "c ∈ restrict_to I (set ics)"
then obtain i where i: "i ∈ I" and mem: "(i,c) ∈ set ics" by auto
from mem show "⟨from_ns v' ?f_list⟩ ⊨⇩c c"
proof (induct c)
case (LT lll rrr)
then have "(lll⦃⟨v'⟩⦄, (QDelta.QDelta rrr (-1))) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns)
then have "val (lll⦃⟨v'⟩⦄) (δ_min f_list) ≤ val (QDelta.QDelta rrr (-1)) (δ_min f_list)"
using l1
by simp
moreover
have "lll⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
lll⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars lll"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using LT
by (auto simp add: comp_def lookup_tabulate restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "lll⦃⟨from_ns v' ?f_list⟩⦄ ≤ (val (QDelta.QDelta rrr (-1)) (δ_min f_list))"
by (auto simp add: valuate_rat_valuate)
then show ?case
using delta_gt_zero[of f_list]
by (simp add: val_def)
next
case (GT lll rrr)
then have "((QDelta.QDelta rrr 1), lll⦃⟨v'⟩⦄) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns)
then have "val (lll⦃⟨v'⟩⦄) (δ_min f_list) ≥ val (QDelta.QDelta rrr 1) (δ_min f_list)"
using l1
by simp
moreover
have "lll⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
lll⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars lll"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using GT
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "lll⦃⟨from_ns v' ?f_list⟩⦄ ≥ val (QDelta.QDelta rrr 1) (δ_min f_list)"
using l2
by (simp add: valuate_rat_valuate)
then show ?case
using delta_gt_zero[of f_list]
by (simp add: val_def)
next
case (LEQ lll rrr)
then have "(lll⦃⟨v'⟩⦄, (QDelta.QDelta rrr 0) ) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns)
then have "val (lll⦃⟨v'⟩⦄) (δ_min f_list) ≤ val (QDelta.QDelta rrr 0) (δ_min f_list)"
using l1
by simp
moreover
have "lll⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
lll⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars lll"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using LEQ
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "lll⦃⟨from_ns v' ?f_list⟩⦄ ≤ val (QDelta.QDelta rrr 0) (δ_min f_list)"
using l2
by (simp add: valuate_rat_valuate)
then show ?case
by (simp add: val_def)
next
case (GEQ lll rrr)
then have "((QDelta.QDelta rrr 0), lll⦃⟨v'⟩⦄) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns)
then have "val (lll⦃⟨v'⟩⦄) (δ_min f_list) ≥ val (QDelta.QDelta rrr 0) (δ_min f_list)"
using l1
by simp
moreover
have "lll⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
lll⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars lll"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using GEQ
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "lll⦃⟨from_ns v' ?f_list⟩⦄ ≥ val (QDelta.QDelta rrr 0) (δ_min f_list)"
using l2
by (simp add: valuate_rat_valuate)
then show ?case
by (simp add: val_def)
next
case (EQ lll rrr)
then have "((QDelta.QDelta rrr 0), lll⦃⟨v'⟩⦄) ∈ set ?list" and
"(lll⦃⟨v'⟩⦄, (QDelta.QDelta rrr 0) ) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns)+
then have "val (lll⦃⟨v'⟩⦄) (δ_min f_list) ≥ val (QDelta.QDelta rrr 0) (δ_min f_list)" and
"val (lll⦃⟨v'⟩⦄) (δ_min f_list) ≤ val (QDelta.QDelta rrr 0) (δ_min f_list)"
using l1
by simp_all
moreover
have "lll⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
lll⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars lll"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using EQ
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "lll⦃⟨from_ns v' ?f_list⟩⦄ ≥ val (QDelta.QDelta rrr 0) (δ_min f_list)" and
"lll⦃⟨from_ns v' ?f_list⟩⦄ ≤ val (QDelta.QDelta rrr 0) (δ_min f_list)"
using l1
by (auto simp add: valuate_rat_valuate)
then show ?case
by (simp add: val_def)
next
case (LTPP ll1 ll2)
then have "((ll1-ll2)⦃⟨v'⟩⦄, (QDelta.QDelta 0 (-1)) ) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns)
then have "val ((ll1-ll2)⦃⟨v'⟩⦄) (δ_min f_list) ≤ val (QDelta.QDelta 0 (-1)) (δ_min f_list)"
using l1
by simp
moreover
have "(ll1-ll2)⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars (ll1 - ll2)"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using LTPP
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄ ≤ val (QDelta.QDelta 0 (-1)) (δ_min f_list)"
using l1
by (simp add: valuate_rat_valuate)
then show ?case
using delta_gt_zero[of f_list]
by (simp add: val_def valuate_minus)
next
case (GTPP ll1 ll2)
then have "((QDelta.QDelta 0 1), (ll1-ll2)⦃⟨v'⟩⦄) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns)
then have "val ((ll1-ll2)⦃⟨v'⟩⦄) (δ_min f_list) ≥ val (QDelta.QDelta 0 1) (δ_min f_list)"
using l1
by simp
moreover
have "(ll1-ll2)⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars (ll1 - ll2)"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using GTPP
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄ ≥ val (QDelta.QDelta 0 1) (δ_min f_list)"
using l1
by (simp add: valuate_rat_valuate)
then show ?case
using delta_gt_zero[of f_list]
by (simp add: val_def valuate_minus)
next
case (LEQPP ll1 ll2)
then have "((ll1-ll2)⦃⟨v'⟩⦄, (QDelta.QDelta 0 0) ) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns zero_QDelta_def)
then have "val ((ll1-ll2)⦃⟨v'⟩⦄) (δ_min f_list) ≤ val (QDelta.QDelta 0 0) (δ_min f_list)"
using l1
by simp
moreover
have "(ll1-ll2)⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars (ll1 - ll2)"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using LEQPP
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄ ≤ val (QDelta.QDelta 0 0) (δ_min f_list)"
using l1
by (simp add: valuate_rat_valuate)
then show ?case
using delta_gt_zero[of f_list]
by (simp add: val_def valuate_minus)
next
case (GEQPP ll1 ll2)
then have "((QDelta.QDelta 0 0), (ll1-ll2)⦃⟨v'⟩⦄) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns zero_QDelta_def)
then have "val ((ll1-ll2)⦃⟨v'⟩⦄) (δ_min f_list) ≥ val (QDelta.QDelta 0 0) (δ_min f_list)"
using l1
by simp
moreover
have "(ll1-ll2)⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars (ll1 - ll2)"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using GEQPP
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄ ≥ val (QDelta.QDelta 0 0) (δ_min f_list)"
using l1
by (simp add: valuate_rat_valuate)
then show ?case
using delta_gt_zero[of f_list]
by (simp add: val_def valuate_minus)
next
case (EQPP ll1 ll2)
then have "((ll1-ll2)⦃⟨v'⟩⦄, (QDelta.QDelta 0 0) ) ∈ set ?list" and
"((QDelta.QDelta 0 0), (ll1-ll2)⦃⟨v'⟩⦄) ∈ set ?list" using i unfolding cs_def
by (force simp add: to_ns zero_QDelta_def)+
then have "val ((ll1-ll2)⦃⟨v'⟩⦄) (δ_min f_list) ≥ val (QDelta.QDelta 0 0) (δ_min f_list)" and
"val ((ll1-ll2)⦃⟨v'⟩⦄) (δ_min f_list) ≤ val (QDelta.QDelta 0 0) (δ_min f_list)"
using l1
by simp_all
moreover
have "(ll1-ll2)⦃(λx. val (⟨v'⟩ x) (δ_min f_list))⦄ =
(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄"
proof (rule valuate_depend, rule)
fix x
assume "x ∈ vars (ll1 - ll2)"
then show "val (⟨v'⟩ x) (δ_min f_list) = ⟨from_ns v' ?f_list⟩ x"
using l2
using EQPP
by (auto simp add: lookup_tabulate comp_def restrict_map_def set_vars_list to_ns_def map2fun_def')
qed
ultimately
have "(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄ ≥ val (QDelta.QDelta 0 0) (δ_min f_list)" and
"(ll1-ll2)⦃⟨from_ns v' ?f_list⟩⦄ ≤ val (QDelta.QDelta 0 0) (δ_min f_list)"
using l1
by (auto simp add: valuate_rat_valuate)
then show ?case
by (simp add: val_def valuate_minus)
qed
qed
} note sat = this
fix cs :: "('i × constraint) list"
have set_to_ns: "set (to_ns cs) = { (i,n) | i n c. (i,c) ∈ set cs ∧ n ∈ set (constraint_to_qdelta_constraint c)}"
unfolding to_ns_def by auto
show indices: "fst ` set (to_ns cs) = fst ` set cs"
proof
show "fst ` set (to_ns cs) ⊆ fst ` set cs"
unfolding set_to_ns by force
{
fix i
assume "i ∈ fst ` set cs"
then obtain c where "(i,c) ∈ set cs" by force
hence "i ∈ fst ` set (to_ns cs)" unfolding set_to_ns by (cases c; force)
}
thus "fst ` set cs ⊆ fst ` set (to_ns cs)" by blast
qed
{
assume dist: "distinct_indices cs"
show "distinct_indices_ns (set (to_ns cs))" unfolding distinct_indices_ns_def
proof (intro allI impI conjI notI)
fix n1 n2 i
assume "(i,n1) ∈ set (to_ns cs)" "(i,n2) ∈ set (to_ns cs)"
then obtain c1 c2 where i: "(i,c1) ∈ set cs" "(i,c2) ∈ set cs"
and n: "n1 ∈ set (constraint_to_qdelta_constraint c1)" "n2 ∈ set (constraint_to_qdelta_constraint c2)"
unfolding set_to_ns by auto
from dist
have "distinct (map fst cs)" unfolding distinct_indices_def by auto
with i have c12: "c1 = c2" by (metis eq_key_imp_eq_value)
note n = n[unfolded c12]
show "poly n1 = poly n2" using n by (cases c2, auto)
show "ns_constraint_const n1 = ns_constraint_const n2" using n by (cases c2, auto)
qed
} note mini = this
fix I mode
assume unsat: "minimal_unsat_core_ns I (set (to_ns cs))"
note unsat = unsat[unfolded minimal_unsat_core_ns_def indices]
hence indices: "I ⊆ fst ` set cs" by auto
show "minimal_unsat_core I cs"
unfolding minimal_unsat_core_def
proof (intro conjI indices impI allI, clarify)
fix v
assume v: "(I,v) ⊨⇩i⇩c⇩s set cs"
let ?v = "λvar. QDelta.QDelta (v var) 0"
have "(I,?v) ⊨⇩i⇩n⇩s⇩s (set (to_ns cs))" using v
proof(induct cs)
case (Cons ic cs)
obtain i c where ic: "ic = (i,c)" by force
from Cons(2-) ic
have rec: "(I,v) ⊨⇩i⇩c⇩s set cs" and c: "i ∈ I ⟹ v ⊨⇩c c" by auto
{
fix jn
assume i: "i ∈ I" and "jn ∈ set (to_ns [(i,c)])"
then have "jn ∈ set (i_constraint_to_qdelta_constraint (i,c))"
unfolding to_ns_def by auto
then obtain n where n: "n ∈ set (constraint_to_qdelta_constraint c)"
and jn: "jn = (i,n)" by force
from c[OF i] have c: "v ⊨⇩c c" by force
from c n jn have "?v ⊨⇩n⇩s snd jn"
by (cases c) (auto simp add: less_eq_QDelta_def to_ns_def valuate_valuate_rat valuate_minus zero_QDelta_def)
} note main = this
from Cons(1)[OF rec] have IH: "(I,?v) ⊨⇩i⇩n⇩s⇩s set (to_ns cs)" .
show ?case unfolding i_satisfies_ns_constraints.simps
proof (intro ballI)
fix x
assume "x ∈ snd ` (set (to_ns (ic # cs)) ∩ I × UNIV)"
then consider (1) "x ∈ snd ` (set (to_ns cs) ∩ I × UNIV)"
| (2) "x ∈ snd ` (set (to_ns [(i,c)]) ∩ I × UNIV)"
unfolding ic to_ns_def by auto
then show "?v ⊨⇩n⇩s x"
proof cases
case 1
then show ?thesis using IH by auto
next
case 2
then obtain jn where x: "snd jn = x" and "jn ∈ set (to_ns [(i,c)]) ∩ I × UNIV"
by auto
with main[of jn] show ?thesis unfolding to_ns_def by auto
qed
qed
qed (simp add: to_ns_def)
with unsat show False unfolding minimal_unsat_core_ns_def by simp blast
next
fix J
assume *: "distinct_indices cs" "J ⊂ I"
hence "distinct_indices_ns (set (to_ns cs))"
using mini by auto
with * unsat obtain v where model: "(J, v) ⊨⇩i⇩n⇩s⇩s set (to_ns cs)" by blast
define w where "w = Mapping.Mapping (λ x. Some (v x))"
have "v = ⟨w⟩" unfolding w_def map2fun_def
by (intro ext, transfer, auto)
with model have model: "(J, ⟨w⟩) ⊨⇩i⇩n⇩s⇩s set (to_ns cs)" by auto
from sat[OF this]
show " ∃v. (J, v) ⊨⇩i⇩c⇩s set cs" by blast
qed
qed
hide_const (open) le lt LE GE LB UB LI UI LBI UBI UBI_upd le_rat
inv zero Var add flat flat_list restrict_to look upd
text ‹Simplex version with indexed constraints as input›
definition simplex_index :: "'i i_constraint list ⇒ 'i list + (var, rat) mapping" where
"simplex_index = solve_exec_code"
lemma simplex_index:
"simplex_index cs = Unsat I ⟹ set I ⊆ fst ` set cs ∧ ¬ (∃ v. (set I, v) ⊨⇩i⇩c⇩s set cs) ∧
(distinct_indices cs ⟶ (∀ J ⊂ set I. (∃ v. (J, v) ⊨⇩i⇩c⇩s set cs)))"
"simplex_index cs = Sat v ⟹ ⟨v⟩ ⊨⇩c⇩s (snd ` set cs)"
proof (unfold simplex_index_def)
assume "solve_exec_code cs = Unsat I"
from SolveExec'Default.simplex_unsat0[OF this]
have core: "minimal_unsat_core (set I) cs" by auto
then show "set I ⊆ fst ` set cs ∧ ¬ (∃ v. (set I, v) ⊨⇩i⇩c⇩s set cs) ∧
(distinct_indices cs ⟶ (∀J⊂set I. ∃v. (J, v) ⊨⇩i⇩c⇩s set cs))"
unfolding minimal_unsat_core_def by auto
next
assume "solve_exec_code cs = Sat v"
from SolveExec'Default.simplex_sat0[OF this]
show "⟨v⟩ ⊨⇩c⇩s (snd ` set cs)" .
qed
text ‹Simplex version where indices will be created›
definition simplex where "simplex cs = simplex_index (zip [0..<length cs] cs)"
lemma simplex:
"simplex cs = Unsat I ⟹ ¬ (∃ v. v ⊨⇩c⇩s set cs)"
"simplex cs = Unsat I ⟹ set I ⊆ {0..<length cs} ∧ ¬ (∃ v. v ⊨⇩c⇩s {cs ! i | i. i ∈ set I})
∧ (∀J⊂set I. ∃v. v ⊨⇩c⇩s {cs ! i |i. i ∈ J})"
"simplex cs = Sat v ⟹ ⟨v⟩ ⊨⇩c⇩s set cs"
proof (unfold simplex_def)
let ?cs = "zip [0..<length cs] cs"
assume "simplex_index ?cs = Unsat I"
from simplex_index(1)[OF this]
have index: "set I ⊆ {0 ..< length cs}" and
core: "∄v. v ⊨⇩c⇩s (snd ` (set ?cs ∩ set I × UNIV))"
"(distinct_indices (zip [0..<length cs] cs) ⟶ (∀ J ⊂ set I. ∃v. v ⊨⇩c⇩s (snd ` (set ?cs ∩ J × UNIV))))"
by (auto simp flip: set_map)
note core(2)
also have "distinct_indices (zip [0..<length cs] cs)"
unfolding distinct_indices_def set_zip by (auto simp: set_conv_nth)
also have "(∀ J ⊂ set I. ∃v. v ⊨⇩c⇩s (snd ` (set ?cs ∩ J × UNIV))) =
(∀ J ⊂ set I. ∃v. v ⊨⇩c⇩s { cs ! i | i. i ∈ J})" using index
by (intro all_cong1 imp_cong ex_cong1 arg_cong[of _ _ "λ x. _ ⊨⇩c⇩s x"] refl, force simp: set_zip)
finally have core': "(∀J⊂set I. ∃v. v ⊨⇩c⇩s {cs ! i |i. i ∈ J}) " .
note unsat = unsat_mono[OF core(1)]
show "¬ (∃ v. v ⊨⇩c⇩s set cs)"
by (rule unsat, auto simp: set_zip)
show "set I ⊆ {0..<length cs} ∧ ¬ (∃ v. v ⊨⇩c⇩s {cs ! i | i. i ∈ set I})
∧ (∀J⊂set I. ∃v. v ⊨⇩c⇩s {cs ! i |i. i ∈ J})"
by (intro conjI index core', rule unsat, auto simp: set_zip)
next
assume "simplex_index (zip [0..<length cs] cs) = Sat v"
from simplex_index(2)[OF this]
show "⟨v⟩ ⊨⇩c⇩s set cs" by (auto simp flip: set_map)
qed
text ‹check executability›
lemma "case simplex [LTPP (lp_monom 2 1) (lp_monom 3 2 - lp_monom 3 0), GT (lp_monom 1 1) 5]
of Sat _ ⇒ True | Unsat _ ⇒ False"
by eval
text ‹check unsat core›
lemma
"case simplex_index [
(1 :: int, LT (lp_monom 1 1) 4),
(2, GTPP (lp_monom 2 1) (lp_monom 1 2)),
(3, EQPP (lp_monom 1 1) (lp_monom 2 2)),
(4, GT (lp_monom 2 2) 5),
(5, GT (lp_monom 3 0) 7)]
of Sat _ ⇒ False | Unsat I ⇒ set I = {1,3,4}"
by eval
end