diff --git a/thys/Farkas/Farkas.thy b/thys/Farkas/Farkas.thy --- a/thys/Farkas/Farkas.thy +++ b/thys/Farkas/Farkas.thy @@ -1,1633 +1,1624 @@ (* Authors: R. Bottesch, M. W. Haslbeck, R. Thiemann *) section \Farkas Coefficients via the Simplex Algorithm of Duterte and de~Moura\ text \Let $c_1,\dots,c_n$ be a finite list of linear inequalities. Let $C$ be a list of pairs $(r_i,c_i)$ where $r_i$ is a rational number. We say that $C$ is a list of \emph{Farkas coefficients} if the sum of all products $r_i \cdot c_i$ results in an inequality that is trivially unsatisfiable. Farkas' Lemma states that a finite set of non-strict linear inequalities is unsatisfiable if and only if Farkas coefficients exist. We will prove this lemma with the help of the simplex algorithm of Dutertre and de~Moura's. Note that the simplex implementation works on four layers, and we will formulate and prove a variant of Farkas' Lemma for each of these layers.\ theory Farkas imports Simplex.Simplex begin subsection \Linear Inequalities\ text \Both Farkas' Lemma and Motzkin's Transposition Theorem require linear combinations of inequalities. To this end we define one type that permits strict and non-strict inequalities which are always of the form ``polynomial R constant'' where R is either $\leq$ or $<$. On this type we can then define a commutative monoid.\ text \A type for the two relations: less-or-equal and less-than.\ datatype le_rel = Leq_Rel | Lt_Rel primrec rel_of :: "le_rel \ 'a :: lrv \ 'a \ bool" where "rel_of Leq_Rel = (\)" | "rel_of Lt_Rel = (<)" instantiation le_rel :: comm_monoid_add begin definition "zero_le_rel = Leq_Rel" fun plus_le_rel where "plus_le_rel Leq_Rel Leq_Rel = Leq_Rel" | "plus_le_rel _ _ = Lt_Rel" instance proof fix a b c :: le_rel show "a + b + c = a + (b + c)" by (cases a; cases b; cases c, auto) show "a + b = b + a" by (cases a; cases b, auto) show "0 + a = a" unfolding zero_le_rel_def by (cases a, auto) qed end lemma Leq_Rel_0: "Leq_Rel = 0" unfolding zero_le_rel_def by simp datatype 'a le_constraint = Le_Constraint (lec_rel: le_rel) (lec_poly: linear_poly) (lec_const: 'a) abbreviation (input) "Leqc \ Le_Constraint Leq_Rel" instantiation le_constraint :: (lrv) comm_monoid_add begin fun plus_le_constraint :: "'a le_constraint \ 'a le_constraint \ 'a le_constraint" where "plus_le_constraint (Le_Constraint r1 p1 c1) (Le_Constraint r2 p2 c2) = (Le_Constraint (r1 + r2) (p1 + p2) (c1 + c2))" definition zero_le_constraint :: "'a le_constraint" where "zero_le_constraint = Leqc 0 0" instance proof fix a b c :: "'a le_constraint" show "0 + a = a" by (cases a, auto simp: zero_le_constraint_def Leq_Rel_0) show "a + b = b + a" by (cases a; cases b, auto simp: ac_simps) show "a + b + c = a + (b + c)" by (cases a; cases b; cases c, auto simp: ac_simps) qed end primrec satisfiable_le_constraint :: "'a::lrv valuation \ 'a le_constraint \ bool" (infixl "\\<^sub>l\<^sub>e" 100) where "(v \\<^sub>l\<^sub>e (Le_Constraint rel l r)) \ (rel_of rel (l\v\) r)" lemma satisfies_zero_le_constraint: "v \\<^sub>l\<^sub>e 0" by (simp add: valuate_zero zero_le_constraint_def) lemma satisfies_sum_le_constraints: assumes "v \\<^sub>l\<^sub>e c" "v \\<^sub>l\<^sub>e d" shows "v \\<^sub>l\<^sub>e (c + d)" proof - obtain lc rc ld rd rel1 rel2 where cd: "c = Le_Constraint rel1 lc rc" "d = Le_Constraint rel2 ld rd" by (cases c; cases d, auto) have 1: "rel_of rel1 (lc\v\) rc" using assms cd by auto have 2: "rel_of rel2 (ld\v\) rd" using assms cd by auto from 1 have le1: "lc\v\ \ rc" by (cases rel1, auto) from 2 have le2: "ld\v\ \ rd" by (cases rel2, auto) from 1 2 le1 le2 have "rel_of (rel1 + rel2) ((lc\v\) + (ld\v\)) (rc + rd)" apply (cases rel1; cases rel2; simp add: add_mono) by (metis add.commute le_less_trans order.strict_iff_order plus_less)+ thus ?thesis by (auto simp: cd valuate_add) qed lemma satisfies_sumlist_le_constraints: assumes "\ c. c \ set (cs :: 'a :: lrv le_constraint list) \ v \\<^sub>l\<^sub>e c" shows "v \\<^sub>l\<^sub>e sum_list cs" using assms by (induct cs, auto intro: satisfies_zero_le_constraint satisfies_sum_le_constraints) lemma sum_list_lec: "sum_list ls = Le_Constraint (sum_list (map lec_rel ls)) (sum_list (map lec_poly ls)) (sum_list (map lec_const ls))" proof (induct ls) case Nil show ?case by (auto simp: zero_le_constraint_def Leq_Rel_0) next case (Cons l ls) show ?case by (cases l, auto simp: Cons) qed lemma sum_list_Leq_Rel: "((\x\C. lec_rel (f x)) = Leq_Rel) \ (\ x \ set C. lec_rel (f x) = Leq_Rel)" proof (induct C) case (Cons c C) show ?case proof (cases "lec_rel (f c)") case Leq_Rel show ?thesis using Cons by (simp add: Leq_Rel Leq_Rel_0) qed simp qed (simp add: Leq_Rel_0) subsection \Farkas' Lemma on Layer 4\ text \On layer 4 the algorithm works on a state containing a tableau, atoms (or bounds), an assignment and a satisfiability flag. Only non-strict inequalities appear at this level. In order to even state a variant of Farkas' Lemma on layer 4, we need conversions from atoms to non-strict constraints and then further to linear inequalities of type @{type le_constraint}. The latter conversion is a partial operation, since non-strict constraints of type @{type ns_constraint} permit greater-or-equal constraints, whereas @{type le_constraint} allows only less-or-equal.\ text \The advantage of first going via @{type ns_constraint} is that this type permits a multiplication with arbitrary rational numbers (the direction of the inequality must be flipped when multiplying by a negative number, which is not possible with @{type le_constraint}).\ instantiation ns_constraint :: (scaleRat) scaleRat begin fun scaleRat_ns_constraint :: "rat \ 'a ns_constraint \ 'a ns_constraint" where "scaleRat_ns_constraint r (LEQ_ns p c) = (if (r < 0) then GEQ_ns (r *R p) (r *R c) else LEQ_ns (r *R p) (r *R c))" | "scaleRat_ns_constraint r (GEQ_ns p c) = (if (r > 0) then GEQ_ns (r *R p) (r *R c) else LEQ_ns (r *R p) (r *R c))" instance .. end lemma sat_scale_rat_ns: assumes "v \\<^sub>n\<^sub>s ns" shows "v \\<^sub>n\<^sub>s (f *R ns)" proof - have "f < 0 | f = 0 | f > 0" by auto then show ?thesis using assms by (cases ns, auto simp: valuate_scaleRat scaleRat_leq1 scaleRat_leq2) qed lemma scaleRat_scaleRat_ns_constraint: assumes "a \ 0 \ b \ 0" shows "a *R (b *R (c :: 'a :: lrv ns_constraint)) = (a * b) *R c" proof - have "b > 0 \ b < 0 \ b = 0" by linarith moreover have "a > 0 \ a < 0 \ a = 0" by linarith ultimately show ?thesis using assms by (elim disjE; cases c, auto simp add: not_le not_less mult_neg_pos mult_neg_neg mult_nonpos_nonneg mult_nonpos_nonpos mult_nonneg_nonpos mult_pos_neg) qed fun lec_of_nsc where "lec_of_nsc (LEQ_ns p c) = (Leqc p c)" fun is_leq_ns where "is_leq_ns (LEQ_ns p c) = True" | "is_leq_ns (GEQ_ns p c) = False" lemma lec_of_nsc: assumes "is_leq_ns c" shows "(v \\<^sub>l\<^sub>e lec_of_nsc c) \ (v \\<^sub>n\<^sub>s c)" using assms by (cases c, auto) fun nsc_of_atom where "nsc_of_atom (Leq var b) = LEQ_ns (lp_monom 1 var) b" | "nsc_of_atom (Geq var b) = GEQ_ns (lp_monom 1 var) b" lemma nsc_of_atom: "v \\<^sub>n\<^sub>s nsc_of_atom a \ v \\<^sub>a a" by (cases a, auto) text \We say that $C$ is a list of Farkas coefficients \emph{for a given tableau $t$ and atom set $as$}, if it is a list of pairs $(r,a)$ such that $a \in as$, $r$ is non-zero, $r \cdot a$ is a `less-than-or-equal'-constraint, and the linear combination of inequalities must result in an inequality of the form $p \leq c$, where $c < 0$ and $t \models p = 0$.\ definition farkas_coefficients_atoms_tableau where "farkas_coefficients_atoms_tableau (as :: 'a :: lrv atom set) t C = (\ p c. (\(r,a) \ set C. a \ as \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0) \ (\(r,a) \ C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c \ c < 0 \ (\ v :: 'a valuation. v \\<^sub>t t \(p\v\ = 0)))" text \We first prove that if the check-function detects a conflict, then Farkas coefficients do exist for the tableau and atom set for which the conflict is detected.\ definition bound_atoms :: "('i, 'a) state \ 'a atom set" ("\\<^sub>A") where "bound_atoms s = (\(v,x). Geq v x) ` (set_of_map (\\<^sub>l s)) \ (\(v,x). Leq v x) ` (set_of_map (\\<^sub>u s))" context PivotUpdateMinVars begin lemma farkas_check: assumes check: "check s' = s" and U: "\ s" "\ \ s'" and inv: "\ s'" "\ (\ s')" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'" "\ s'" and index: "index_valid as s'" shows "\ C. farkas_coefficients_atoms_tableau (snd ` as) (\ s') C" proof - let ?Q = "\ s f p c C. set C \ \\<^sub>A s \ distinct C \ (\a \ set C. is_leq_ns (f (atom_var a) *R nsc_of_atom a) \ f (atom_var a) \ 0) \ (\a \ C. lec_of_nsc (f (atom_var a) *R nsc_of_atom a)) = Leqc p c \ c < 0 \ (\ v :: 'a valuation. v \\<^sub>t \ s \(p\v\ = 0))" let ?P = "\ s. \ s \ (\ f p c C. ?Q s f p c C)" have "?P (check s')" proof (induct rule: check_induct''[OF inv, of ?P]) case (3 s x\<^sub>i dir I) have dir: "dir = Positive \ dir = Negative" by fact let ?eq = "(eq_for_lvar (\ s) x\<^sub>i)" define X\<^sub>j where "X\<^sub>j = rvars_eq ?eq" define XL\<^sub>j where "XL\<^sub>j = Abstract_Linear_Poly.vars_list (rhs ?eq)" have [simp]: "set XL\<^sub>j = X\<^sub>j" unfolding XL\<^sub>j_def X\<^sub>j_def using set_vars_list by blast have XL\<^sub>j_distinct: "distinct XL\<^sub>j" unfolding XL\<^sub>j_def using distinct_vars_list by simp define A where "A = coeff (rhs ?eq)" have bounds_id: "\\<^sub>A (set_unsat I s) = \\<^sub>A s" "\\<^sub>u (set_unsat I s) = \\<^sub>u s" "\\<^sub>l (set_unsat I s) = \\<^sub>l s" by (auto simp: boundsl_def boundsu_def bound_atoms_def) have t_id: "\ (set_unsat I s) = \ s" by simp have u_id: "\ (set_unsat I s) = True" by simp let ?p = "rhs ?eq - lp_monom 1 x\<^sub>i" have p_eval_zero: "?p \ v \ = 0" if "v \\<^sub>t \ s" for v :: "'a valuation" proof - have eqT: "?eq \ set (\ s)" by (simp add: 3(7) eq_for_lvar local.min_lvar_not_in_bounds_lvars) have "v \\<^sub>e ?eq" using that eqT satisfies_tableau_def by blast also have "?eq = (lhs ?eq, rhs ?eq)" by (cases ?eq, auto) also have "lhs ?eq = x\<^sub>i" by (simp add: 3(7) eq_for_lvar local.min_lvar_not_in_bounds_lvars) finally have "v \\<^sub>e (x\<^sub>i, rhs ?eq)" . then show ?thesis by (auto simp: satisfies_eq_iff valuate_minus) qed have Xj_rvars: "X\<^sub>j \ rvars (\ s)" unfolding X\<^sub>j_def using 3 min_lvar_not_in_bounds_lvars rvars_of_lvar_rvars by blast have xi_lvars: "x\<^sub>i \ lvars (\ s)" using 3 min_lvar_not_in_bounds_lvars rvars_of_lvar_rvars by blast have "lvars (\ s) \ rvars (\ s) = {}" using 3 normalized_tableau_def by auto with xi_lvars Xj_rvars have xi_Xj: "x\<^sub>i \ X\<^sub>j" by blast have rhs_eval_xi: "(rhs (eq_for_lvar (\ s) x\<^sub>i)) \\\ s\\ = \\ s\ x\<^sub>i" proof - have *: "(rhs eq) \ v \ = v (lhs eq)" if "v \\<^sub>e eq" for v :: "'a valuation" and eq using satisfies_eq_def that by metis moreover have "\\ s\ \\<^sub>e eq_for_lvar (\ s) x\<^sub>i" using 3 satisfies_tableau_def eq_for_lvar curr_val_satisfies_no_lhs_def xi_lvars by blast ultimately show ?thesis using eq_for_lvar xi_lvars by simp qed let ?\\<^sub>l = "Direction.LB dir" let ?\\<^sub>u = "Direction.UB dir" let ?lt = "Direction.lt dir" let ?le = "Simplex.le ?lt" let ?Geq = "Direction.GE dir" let ?Leq = "Direction.LE dir" have 0: "(if A x < 0 then ?\\<^sub>l s x = Some (\\ s\ x) else ?\\<^sub>u s x = Some (\\ s\ x)) \ A x \ 0" if x: "x \ X\<^sub>j" for x proof - have "Some (\\ s\ x) = (?\\<^sub>l s x)" if "A x < 0" proof - have cmp: "\ \\<^sub>l\<^sub>b ?lt (\\ s\ x) (?\\<^sub>l s x)" using x that dir min_rvar_incdec_eq_None[OF 3(9)] unfolding X\<^sub>j_def A_def by auto then obtain c where c: "?\\<^sub>l s x = Some c" by (cases "?\\<^sub>l s x", auto simp: bound_compare_defs) also have "c = \\ s\ x" proof - have "x \ rvars (\ s)" using that x Xj_rvars by blast then have "x \ (- lvars (\ s))" using 3 unfolding normalized_tableau_def by auto moreover have "\x\(- lvars (\ s)). in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" using 3 unfolding curr_val_satisfies_no_lhs_def by (simp add: satisfies_bounds_set.simps) ultimately have "in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" by blast moreover have "?le (\\ s\ x) c" using cmp c dir unfolding bound_compare_defs by auto ultimately show ?thesis using c dir by (auto simp del: Simplex.bounds_lg) qed then show ?thesis using c by simp qed moreover have "Some (\\ s\ x) = (?\\<^sub>u s x)" if "0 < A x" proof - have cmp: "\ \\<^sub>u\<^sub>b ?lt (\\ s\ x) (?\\<^sub>u s x)" using x that min_rvar_incdec_eq_None[OF 3(9)] unfolding X\<^sub>j_def A_def by auto then obtain c where c: "?\\<^sub>u s x = Some c" by (cases "?\\<^sub>u s x", auto simp: bound_compare_defs) also have "c = \\ s\ x" proof - have "x \ rvars (\ s)" using that x Xj_rvars by blast then have "x \ (- lvars (\ s))" using 3 unfolding normalized_tableau_def by auto moreover have "\x\(- lvars (\ s)). in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" using 3 unfolding curr_val_satisfies_no_lhs_def by (simp add: satisfies_bounds_set.simps) ultimately have "in_bounds x \\ s\ (\\<^sub>l s, \\<^sub>u s)" by blast moreover have "?le c (\\ s\ x)" using cmp c dir unfolding bound_compare_defs by auto ultimately show ?thesis using c dir by (auto simp del: Simplex.bounds_lg) qed then show ?thesis using c by simp qed moreover have "A x \ 0" using that coeff_zero unfolding A_def X\<^sub>j_def by auto ultimately show ?thesis using that by auto qed have l_Ba: "l \ \\<^sub>A s" if "l \ {?Geq x\<^sub>i (the (?\\<^sub>l s x\<^sub>i))}" for l proof - from that have l: "l = ?Geq x\<^sub>i (the (?\\<^sub>l s x\<^sub>i))" by simp from 3(8) obtain c where bl': "?\\<^sub>l s x\<^sub>i = Some c" by (cases "?\\<^sub>l s x\<^sub>i", auto simp: bound_compare_defs) hence bl: "(x\<^sub>i, c) \ set_of_map (?\\<^sub>l s)" unfolding set_of_map_def by auto show "l \ \\<^sub>A s" unfolding l bound_atoms_def using bl bl' dir by auto qed let ?negA = "filter (\ x. A x < 0) XL\<^sub>j" let ?posA = "filter (\ x. \ A x < 0) XL\<^sub>j" define neg where "neg = (if dir = Positive then (\ x :: rat. x) else uminus)" define negP where "negP = (if dir = Positive then (\ x :: linear_poly. x) else uminus)" define nega where "nega = (if dir = Positive then (\ x :: 'a. x) else uminus)" from dir have dirn: "dir = Positive \ neg = (\ x. x) \ negP = (\ x. x) \ nega = (\ x. x) \ dir = Negative \ neg = uminus \ negP = uminus \ nega = uminus" unfolding neg_def negP_def nega_def by auto define C where "C = map (\x. ?Geq x (the (?\\<^sub>l s x))) ?negA @ map (\ x. ?Leq x (the (?\\<^sub>u s x))) ?posA @ [?Geq x\<^sub>i (the (?\\<^sub>l s x\<^sub>i))]" define f where "f = (\x. if x = x\<^sub>i then neg (-1) else neg (A x))" define c where "c = (\x\C. lec_const (lec_of_nsc (f (atom_var x) *R nsc_of_atom x)))" let ?q = "negP ?p" show ?case unfolding bounds_id t_id u_id proof (intro exI impI conjI allI) show "v \\<^sub>t \ s \ ?q \ v \ = 0" for v :: "'a valuation" using dirn p_eval_zero[of v] by (auto simp: valuate_minus) show "set C \ \\<^sub>A s" unfolding C_def set_append set_map set_filter list.simps using 0 l_Ba dir by (intro Un_least subsetI) (force simp: bound_atoms_def set_of_map_def)+ show is_leq: "\a\set C. is_leq_ns (f (atom_var a) *R nsc_of_atom a) \ f (atom_var a) \ 0" using dirn xi_Xj 0 unfolding C_def f_def by (elim disjE, auto) show "(\a \ C. lec_of_nsc (f (atom_var a) *R nsc_of_atom a)) = Leqc ?q c" unfolding sum_list_lec le_constraint.simps map_map o_def proof (intro conjI) define scale_poly :: "'a atom \ linear_poly" where "scale_poly = (\x. lec_poly (lec_of_nsc (f (atom_var x) *R nsc_of_atom x)))" have "(\x\C. scale_poly x) = (\x\?negA. scale_poly (?Geq x (the (?\\<^sub>l s x)))) + (\x\?posA. scale_poly (?Leq x (the (?\\<^sub>u s x)))) - negP (lp_monom 1 x\<^sub>i)" unfolding C_def using dirn by (auto simp add: comp_def scale_poly_def f_def) also have "(\x\?negA. scale_poly (?Geq x (the (?\\<^sub>l s x)))) = (\x\ ?negA. negP (A x *R lp_monom 1 x))" unfolding scale_poly_def f_def using dirn xi_Xj by (subst map_cong) auto also have "(\x\?posA. scale_poly (?Leq x (the (?\\<^sub>u s x)))) = (\x\ ?posA. negP (A x *R lp_monom 1 x))" unfolding scale_poly_def f_def using dirn xi_Xj by (subst map_cong) auto also have "(\x\ ?negA. negP (A x *R lp_monom 1 x)) + (\x\ ?posA. negP (A x *R lp_monom 1 x)) = negP (rhs (eq_for_lvar (\ s) x\<^sub>i))" using dirn XL\<^sub>j_distinct coeff_zero by (elim disjE; intro poly_eqI, auto intro!: poly_eqI simp add: coeff_sum_list A_def X\<^sub>j_def uminus_sum_list_map[unfolded o_def, symmetric]) finally show "(\x\C. lec_poly (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))) = ?q" unfolding scale_poly_def using dirn by auto show "(\x\C. lec_rel (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))) = Leq_Rel" unfolding sum_list_Leq_Rel proof fix c assume c: "c \ set C" show "lec_rel (lec_of_nsc (f (atom_var c) *R nsc_of_atom c)) = Leq_Rel" using is_leq[rule_format, OF c] by (cases "f (atom_var c) *R nsc_of_atom c", auto) qed qed (simp add: c_def) show "c < 0" proof - define scale_const_f :: "'a atom \ 'a" where "scale_const_f x = lec_const (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))" for x obtain d where bl': "?\\<^sub>l s x\<^sub>i = Some d" using 3 by (cases "?\\<^sub>l s x\<^sub>i", auto simp: bound_compare_defs) have "c = (\x\map (\x. ?Geq x (the (?\\<^sub>l s x))) ?negA. scale_const_f x) + (\x\ map (\x. ?Leq x (the (?\\<^sub>u s x))) ?posA. scale_const_f x) - nega d" unfolding c_def C_def f_def scale_const_f_def using dirn rhs_eval_xi bl' by auto also have "(\x\map (\x. ?Geq x (the (?\\<^sub>l s x))) ?negA. scale_const_f x) = (\x\ ?negA. nega (A x *R the (?\\<^sub>l s x)))" using xi_Xj dirn by (subst map_cong) (auto simp add: f_def scale_const_f_def) also have "\ = (\x\?negA. nega (A x *R \\ s\ x))" using 0 by (subst map_cong) auto also have "(\x\map (\x. ?Leq x (the (?\\<^sub>u s x))) ?posA. scale_const_f x) = (\x\ ?posA. nega (A x *R the (?\\<^sub>u s x)))" using xi_Xj dirn by (subst map_cong) (auto simp add: f_def scale_const_f_def) also have "\ = (\x\ ?posA. nega (A x *R \\ s\ x))" using 0 by (subst map_cong) auto also have "(\x\?negA. nega (A x *R \\ s\ x)) + (\x\?posA. nega (A x *R \\ s\ x)) = (\x\?negA @ ?posA. nega (A x *R \\ s\ x))" by auto also have "\ = (\x\ X\<^sub>j. nega (A x *R \\ s\ x))" using XL\<^sub>j_distinct by (subst sum_list_distinct_conv_sum_set) (auto intro!: sum.cong) also have "\ = nega (\x\ X\<^sub>j. (A x *R \\ s\ x))" using dirn by (auto simp: sum_negf) also have "(\x\ X\<^sub>j. (A x *R \\ s\ x)) = ((rhs ?eq) \\\ s\\)" unfolding A_def X\<^sub>j_def by (subst linear_poly_sum) (auto simp add: sum_negf) also have "\ = \\ s\ x\<^sub>i" using rhs_eval_xi by blast also have "nega (\\ s\ x\<^sub>i) - nega d < 0" proof - have "?lt (\\ s\ x\<^sub>i) d" using dirn 3(2-) bl' by (elim disjE, auto simp: bound_compare_defs) thus ?thesis using dirn unfolding minus_lt[symmetric] by auto qed finally show ?thesis . qed show "distinct C" unfolding C_def using XL\<^sub>j_distinct xi_Xj dirn by (auto simp add: inj_on_def distinct_map) qed qed (insert U, blast+) then obtain f p c C where Qs: "?Q s f p c C" using U unfolding check by blast from index[folded check_tableau_index_valid[OF U(2) inv(3,4,2,1)]] check have index: "index_valid as s" by auto from check_tableau_equiv[OF U(2) inv(3,4,2,1), unfolded check] have id: "v \\<^sub>t \ s = v \\<^sub>t \ s'" for v :: "'a valuation" by auto let ?C = "map (\ a. (f (atom_var a), a)) C" have "set C \ \\<^sub>A s" using Qs by blast also have "\ \ snd ` as" using index unfolding bound_atoms_def index_valid_def set_of_map_def boundsl_def boundsu_def o_def by force finally have sub: "snd ` set ?C \ snd ` as" by force show ?thesis unfolding farkas_coefficients_atoms_tableau_def by (intro exI[of _ p] exI[of _ c] exI[of _ ?C] conjI, insert Qs[unfolded id] sub, (force simp: o_def)+) qed end text \Next, we show that a conflict found by the assert-bound function also gives rise to Farkas coefficients.\ context Update begin lemma farkas_assert_bound: assumes inv: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" "\ s" and index: "index_valid as s" and U: "\ (assert_bound ia s)" shows "\ C. farkas_coefficients_atoms_tableau (snd ` (insert ia as)) (\ s) C" proof - obtain i a where ia[simp]: "ia = (i,a)" by force let ?A = "snd ` insert ia as" have "\ x c d. Leq x c \ ?A \ Geq x d \ ?A \ c < d" proof (cases a) case (Geq x d) let ?s = "update\\ (Direction.UBI_upd (Direction (\x y. y < x) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>l_update Geq Leq (\))) i x d s" have id: "\ ?s = \ s" by auto have norm: "\ (\ ?s)" using inv by auto have val: "\ ?s" using inv(4) unfolding tableau_valuated_def by simp have idd: "x \ lvars (\ ?s) \ \ (update x d ?s) = \ ?s" by (rule update_unsat_id[OF norm val]) from U[unfolded ia Geq] inv(1) id idd have "\\<^sub>l\<^sub>b (\x y. y < x) d (\\<^sub>u s x)" by (auto split: if_splits simp: Let_def) then obtain c where Bu: "\\<^sub>u s x = Some c" and lt: "c < d" by (cases "\\<^sub>u s x", auto simp: bound_compare_defs) from Bu obtain j where "Mapping.lookup (\\<^sub>i\<^sub>u s) x = Some (j,c)" unfolding boundsu_def by auto with index[unfolded index_valid_def] have "(j, Leq x c) \ as" by auto hence xc: "Leq x c \ ?A" by force have xd: "Geq x d \ ?A" unfolding ia Geq by force from xc xd lt show ?thesis by auto next case (Leq x c) let ?s = "update\\ (Direction.UBI_upd (Direction (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>u_update Leq Geq (\))) i x c s" have id: "\ ?s = \ s" by auto have norm: "\ (\ ?s)" using inv by auto have val: "\ ?s" using inv(4) unfolding tableau_valuated_def by simp have idd: "x \ lvars (\ ?s) \ \ (update x c ?s) = \ ?s" by (rule update_unsat_id[OF norm val]) from U[unfolded ia Leq] inv(1) id idd have "\\<^sub>l\<^sub>b (<) c (\\<^sub>l s x)" by (auto split: if_splits simp: Let_def) then obtain d where Bl: "\\<^sub>l s x = Some d" and lt: "c < d" by (cases "\\<^sub>l s x", auto simp: bound_compare_defs) from Bl obtain j where "Mapping.lookup (\\<^sub>i\<^sub>l s) x = Some (j,d)" unfolding boundsl_def by auto with index[unfolded index_valid_def] have "(j, Geq x d) \ as" by auto hence xd: "Geq x d \ ?A" by force have xc: "Leq x c \ ?A" unfolding ia Leq by force from xc xd lt show ?thesis by auto qed then obtain x c d where c: "Leq x c \ ?A" and d: "Geq x d \ ?A" and cd: "c < d" by blast show ?thesis unfolding farkas_coefficients_atoms_tableau_def proof (intro exI conjI allI) let ?C = "[(-1, Geq x d), (1,Leq x c)]" show "\(r,a)\set ?C. a \ ?A \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0" using c d by auto show "c - d < 0" using cd using minus_lt by auto qed (auto simp: valuate_zero) qed end text \Moreover, we prove that all other steps of the simplex algorithm on layer~4, such as pivoting, asserting bounds without conflict, etc., preserve Farkas coefficients.\ lemma farkas_coefficients_atoms_tableau_mono: assumes "as \ bs" shows "farkas_coefficients_atoms_tableau as t C \ farkas_coefficients_atoms_tableau bs t C" using assms unfolding farkas_coefficients_atoms_tableau_def by force locale AssertAllState''' = AssertAllState'' init ass_bnd chk + Update update + PivotUpdateMinVars eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update for init and ass_bnd :: "'i \ 'a :: lrv atom \ _" and chk :: "('i, 'a) state \ ('i, 'a) state" and update :: "nat \ 'a :: lrv \ ('i, 'a) state \ ('i, 'a) state" and 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" + assumes ass_bnd: "ass_bnd = Update.assert_bound update" and chk: "chk = PivotUpdateMinVars.check eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update" context AssertAllState''' begin lemma farkas_assert_bound_loop: assumes "\ (assert_bound_loop as (init t))" and norm: "\ t" shows "\ C. farkas_coefficients_atoms_tableau (snd ` set as) t C" proof - let ?P = "\ as s. \ s \ (\ C. farkas_coefficients_atoms_tableau (snd ` as) (\ s) C)" let ?s = "assert_bound_loop as (init t)" have "\ \ (init t)" by (rule init_unsat_flag) have "\ (assert_bound_loop as (init t)) = t \ (\ (assert_bound_loop as (init t)) \ (\ C. farkas_coefficients_atoms_tableau (snd ` set as) (\ (init t)) C))" proof (rule AssertAllState''Induct[OF norm], unfold ass_bnd, goal_cases) case 1 have "\ \ (init t)" by (rule init_unsat_flag) moreover have "\ (init t) = t" by (rule init_tableau_id) ultimately show ?case by auto next case (2 as bs s) hence "snd ` as \ snd ` bs" by auto from farkas_coefficients_atoms_tableau_mono[OF this] 2(2) show ?case by auto next case (3 s a ats) let ?s = "assert_bound a s" have tab: "\ ?s = \ s" unfolding ass_bnd by (rule assert_bound_nolhs_tableau_id, insert 3, auto) have t: "t = \ s" using 3 by simp show ?case unfolding t tab proof (intro conjI impI refl) assume "\ ?s" from farkas_assert_bound[OF 3(1,3-6,8) this] show "\ C. farkas_coefficients_atoms_tableau (snd ` insert a (set ats)) (\ (init (\ s))) C" unfolding t[symmetric] init_tableau_id . qed qed thus ?thesis unfolding init_tableau_id using assms by blast qed text \Now we get to the main result for layer~4: If the main algorithm returns unsat, then there are Farkas coefficients for the tableau and atom set that were given as input for this layer.\ lemma farkas_assert_all_state: assumes U: "\ (assert_all_state t as)" and norm: "\ t" shows "\ C. farkas_coefficients_atoms_tableau (snd ` set as) t C" proof - let ?s = "assert_bound_loop as (init t)" show ?thesis proof (cases "\ (assert_bound_loop as (init t))") case True from farkas_assert_bound_loop[OF this norm] show ?thesis by auto next case False from AssertAllState''_tableau_id[OF norm] have T: "\ ?s = t" unfolding init_tableau_id . from U have U: "\ (check ?s)" unfolding chk[symmetric] by simp show ?thesis proof (rule farkas_check[OF refl U False, unfolded T, OF _ norm]) from AssertAllState''_precond[OF norm, unfolded Let_def] False show "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s" "\ ?s" "\ ?s" by blast+ from AssertAllState''_index_valid[OF norm] show "index_valid (set as) ?s" . qed qed qed subsection \Farkas' Lemma on Layer 3\ text \There is only a small difference between layers 3 and 4, namely that there is no simplex algorithm (@{const assert_all_state}) on layer 3, but just a tableau and atoms.\ text \Hence, one task is to link the unsatisfiability flag on layer 4 with unsatisfiability of the original tableau and atoms (layer 3). This can be done via the existing soundness results of the simplex algorithm. Moreover, we give an easy proof that the existence of Farkas coefficients for a tableau and set of atoms implies unsatisfiability.\ end lemma farkas_coefficients_atoms_tableau_unsat: assumes "farkas_coefficients_atoms_tableau as t C" shows "\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" proof assume "\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" then obtain v where *: "v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" by auto then obtain p c where isleq: "(\(r,a) \ set C. a \ as \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0)" and leq: "(\(r,a) \ C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c" and cltz: "c < 0" and p0: "p\v\ = 0" using assms farkas_coefficients_atoms_tableau_def by blast have fa: "\(r,a) \ set C. v \\<^sub>a a" using * isleq leq satisfies_atom_set_def by force { fix r a assume a: "(r,a) \ set C" from a fa have va: "v \\<^sub>a a" unfolding satisfies_atom_set_def by auto hence v: "v \\<^sub>n\<^sub>s (r *R nsc_of_atom a)" by (auto simp: nsc_of_atom sat_scale_rat_ns) from a isleq have "is_leq_ns (r *R nsc_of_atom a)" by auto from lec_of_nsc[OF this] v have "v \\<^sub>l\<^sub>e lec_of_nsc (r *R nsc_of_atom a)" by blast } note v = this have "v \\<^sub>l\<^sub>e Leqc p c" unfolding leq[symmetric] by (rule satisfies_sumlist_le_constraints, insert v, auto) then have "0 \ c" using p0 by auto then show False using cltz by auto qed text \Next is the main result for layer~3: a tableau and a finite set of atoms are unsatisfiable if and only if there is a list of Farkas coefficients for the set of atoms and the tableau.\ lemma farkas_coefficients_atoms_tableau: assumes norm: "\ t" and fin: "finite as" shows "(\ C. farkas_coefficients_atoms_tableau as t C) \ (\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as)" proof from finite_list[OF fin] obtain bs where as: "as = set bs" by auto assume unsat: "\ v. v \\<^sub>t t \ v \\<^sub>a\<^sub>s as" let ?as = "map (\ x. ((),x)) bs" interpret AssertAllState''' init_state assert_bound_code check_code update_code eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code by (unfold_locales, auto simp: assert_bound_code_def check_code_def) let ?call = "assert_all t ?as" have id: "snd ` set ?as = as" unfolding as by force from assert_all_sat[OF norm, of ?as, unfolded id] unsat obtain I where "?call = Inl I" by (cases ?call, auto) from this[unfolded assert_all_def Let_def] have "\ (assert_all_state_code t ?as)" by (auto split: if_splits simp: assert_all_state_code_def) from farkas_assert_all_state[OF this[unfolded assert_all_state_code_def] norm, unfolded id] show "\ C. farkas_coefficients_atoms_tableau as t C" . qed (insert farkas_coefficients_atoms_tableau_unsat, auto) subsection \Farkas' Lemma on Layer 2\ text \The main difference between layers 2 and 3 is the introduction of slack-variables in layer 3 via the preprocess-function. Our task here is to show that Farkas coefficients at layer 3 (where slack-variables are used) can be converted into Farkas coefficients for layer 2 (before the preprocessing).\ text \We also need to adapt the previos notion of Farkas coefficients, which was used in @{const farkas_coefficients_atoms_tableau}, for layer~2. At layer 3, Farkas coefficients are the coefficients in a linear combination of atoms that evaluates to an inequality of the form $p \leq c$, where $p$ is a linear polynomial, $c < 0$, and $t \models p = 0$ holds. At layer 2, the atoms are replaced by non-strict constraints where the left-hand side is a polynomial in the original variables, but the corresponding linear combination (with Farkas coefficients) evaluates directly to the inequality $0 \leq c$, with $c < 0$. The implication $t \models p = 0$ is no longer possible in this layer, since there is no tableau $t$, nor is it needed, since $p$ is $0$. Thus, the statement defining Farkas coefficients must be changed accordingly.\ definition farkas_coefficients_ns where "farkas_coefficients_ns ns C = (\ c. (\(r, n) \ set C. n \ ns \ is_leq_ns (r *R n) \ r \ 0) \ (\(r, n) \ C. lec_of_nsc (r *R n)) = Leqc 0 c \ c < 0)" text \The easy part is to prove that Farkas coefficients imply unsatisfiability.\ lemma farkas_coefficients_ns_unsat: assumes "farkas_coefficients_ns ns C" shows "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" proof assume "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" then obtain v where *: "v \\<^sub>n\<^sub>s\<^sub>s ns" by auto obtain c where isleq: "(\(a,n) \ set C. n \ ns \ is_leq_ns (a *R n) \ a \ 0)" and leq: "(\(a,n) \ C. lec_of_nsc (a *R n)) = Leqc 0 c" and cltz: "c < 0" using assms farkas_coefficients_ns_def by blast { fix a n assume n: "(a,n) \ set C" from n * isleq have "v \\<^sub>n\<^sub>s n" by auto hence v: "v \\<^sub>n\<^sub>s (a *R n)" by (rule sat_scale_rat_ns) from n isleq have "is_leq_ns (a *R n)" by auto from lec_of_nsc[OF this] v have "v \\<^sub>l\<^sub>e lec_of_nsc (a *R n)" by blast } note v = this have "v \\<^sub>l\<^sub>e Leqc 0 c" unfolding leq[symmetric] by (rule satisfies_sumlist_le_constraints, insert v, auto) then show False using cltz by (metis leD satisfiable_le_constraint.simps valuate_zero rel_of.simps(1)) qed text \In order to eliminate the need for a tableau, we require the notion of an arbitrary substitution on polynomials, where all variables can be replaced at once. The existing simplex formalization provides only a function to replace one variable at a time.\ definition subst_poly :: "(var \ linear_poly) \ linear_poly \ linear_poly" where "subst_poly \ p = (\ x \ vars p. coeff p x *R \ x)" lemma subst_poly_0[simp]: "subst_poly \ 0 = 0" unfolding subst_poly_def by simp lemma valuate_subst_poly: "(subst_poly \ p) \ v \ = (p \ (\ x. ((\ x) \ v \)) \)" by (subst (2) linear_poly_sum, unfold subst_poly_def valuate_sum valuate_scaleRat, simp) lemma subst_poly_add: "subst_poly \ (p + q) = subst_poly \ p + subst_poly \ q" by (rule linear_poly_eqI, unfold valuate_add valuate_subst_poly, simp) fun subst_poly_lec :: "(var \ linear_poly) \ 'a le_constraint \ 'a le_constraint" where "subst_poly_lec \ (Le_Constraint rel p c) = Le_Constraint rel (subst_poly \ p) c" lemma subst_poly_lec_0[simp]: "subst_poly_lec \ 0 = 0" unfolding zero_le_constraint_def by simp lemma subst_poly_lec_add: "subst_poly_lec \ (c1 + c2) = subst_poly_lec \ c1 + subst_poly_lec \ c2" by (cases c1; cases c2, auto simp: subst_poly_add) lemma subst_poly_lec_sum_list: "subst_poly_lec \ (sum_list ps) = sum_list (map (subst_poly_lec \) ps)" by (induct ps, auto simp: subst_poly_lec_add) lemma subst_poly_lp_monom[simp]: "subst_poly \ (lp_monom r x) = r *R \ x" unfolding subst_poly_def by (simp add: vars_lp_monom) lemma subst_poly_scaleRat: "subst_poly \ (r *R p) = r *R (subst_poly \ p)" by (rule linear_poly_eqI, unfold valuate_scaleRat valuate_subst_poly, simp) text \We need several auxiliary properties of the preprocess-function which are not present in the simplex formalization.\ lemma Tableau_is_monom_preprocess': assumes "(x, p) \ set (Tableau (preprocess' cs start))" shows "\ is_monom p" using assms by(induction cs start rule: preprocess'.induct) (auto simp add: Let_def split: if_splits option.splits) lemma preprocess'_atoms_to_constraints': assumes "preprocess' cs start = S" shows "set (Atoms S) \ {(i,qdelta_constraint_to_atom c v) | i c v. (i,c) \ set cs \ (\ is_monom (poly c) \ Poly_Mapping S (poly c) = Some v)}" unfolding assms(1)[symmetric] by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits, force+) lemma monom_of_atom_coeff: assumes "is_monom (poly ns)" "a = qdelta_constraint_to_atom ns v" shows "(monom_coeff (poly ns)) *R nsc_of_atom a = ns" using assms is_monom_monom_coeff_not_zero by(cases a; cases ns) (auto split: atom.split ns_constraint.split simp add: monom_poly_assemble field_simps) text \The next lemma provides the functionality that is required to convert an atom back to a non-strict constraint, i.e., it is a kind of inverse of the preprocess-function.\ lemma preprocess'_atoms_to_constraints: assumes S: "preprocess' cs start = S" and start: "start = start_fresh_variable cs" and ns: "ns = (case a of Leq v c \ LEQ_ns q c | Geq v c \ GEQ_ns q c)" and "a \ snd ` set (Atoms S)" shows "(atom_var a \ fst ` set (Tableau S) \ (\ r. r \ 0 \ r *R nsc_of_atom a \ snd ` set cs)) \ ((atom_var a, q) \ set (Tableau S) \ ns \ snd ` set cs)" proof - let ?S = "preprocess' cs start" from assms(4) obtain i where ia: "(i,a) \ set (Atoms S)" by auto with preprocess'_atoms_to_constraints'[OF assms(1)] obtain c v where a: "a = qdelta_constraint_to_atom c v" and c: "(i,c) \ set cs" and nmonom: "\ is_monom (poly c) \ Poly_Mapping S (poly c) = Some v" by blast hence c': "c \ snd ` set cs" by force let ?p = "poly c" show ?thesis proof (cases "is_monom ?p") case True hence av: "atom_var a = monom_var ?p" unfolding a by (cases c, auto) from Tableau_is_monom_preprocess'[of _ ?p cs start] True have "(x, ?p) \ set (Tableau ?S)" for x by blast { assume "(atom_var a, q) \ set (Tableau S)" hence "(monom_var ?p, q) \ set (Tableau S)" unfolding av by simp hence "monom_var ?p \ lvars (Tableau S)" unfolding lvars_def by force from lvars_tableau_ge_start[rule_format, OF this[folded S]] have "monom_var ?p \ start" unfolding S . moreover have "monom_var ?p \ vars_constraints (map snd cs)" using True c by (auto intro!: bexI[of _ "(i,c)"] simp: monom_var_in_vars) ultimately have False using start_fresh_variable_fresh[of cs, folded start] by force } moreover from monom_of_atom_coeff[OF True a] True have "\r. r \ 0 \ r *R nsc_of_atom a = c" by (intro exI[of _ "monom_coeff ?p"], auto, cases a, auto) ultimately show ?thesis using c' by auto next case False hence av: "atom_var a = v" unfolding a by (cases c, auto) from nmonom[OF False] have "Poly_Mapping S ?p = Some v" . from preprocess'_Tableau_Poly_Mapping_Some[OF this[folded S]] have tab: "(atom_var a, ?p) \ set (Tableau (preprocess' cs start))" unfolding av by simp hence "atom_var a \ fst ` set (Tableau S)" unfolding S by force moreover { assume "(atom_var a, q) \ set (Tableau S)" from tab this have qp: "q = ?p" unfolding S using lvars_distinct[of cs start, unfolded S lhs_def] by (simp add: case_prod_beta' eq_key_imp_eq_value) have "ns = c" unfolding ns qp using av a False by (cases c, auto) hence "ns \ snd ` set cs" using c' by blast } ultimately show ?thesis by blast qed qed text \Next follows the major technical lemma of this part, namely that Farkas coefficients on layer~3 for preprocessed constraints can be converted into Farkas coefficients on layer~2.\ lemma farkas_coefficients_preprocess': assumes pp: "preprocess' cs (start_fresh_variable cs) = S" and ft: "farkas_coefficients_atoms_tableau (snd ` set (Atoms S)) (Tableau S) C" shows "\ C. farkas_coefficients_ns (snd ` set cs) C" proof - note ft[unfolded farkas_coefficients_atoms_tableau_def] obtain p c where 0: "\ (r,a) \ set C. a \ snd ` set (Atoms S) \ is_leq_ns (r *R nsc_of_atom a) \ r \ 0" "(\(r,a)\C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c" "c < 0" "\v :: QDelta valuation. v \\<^sub>t Tableau S \ p \ v \ = 0" using ft unfolding farkas_coefficients_atoms_tableau_def by blast note 0 = 0(1)[rule_format, of "(a, b)" for a b, unfolded split] 0(2-) let ?T = "Tableau S" define \ :: "var \ linear_poly" where "\ = (\ x. case map_of ?T x of Some p \ p | None \ lp_monom 1 x)" let ?P = "(\r a s ns. ns \ (snd ` set cs) \ is_leq_ns (s *R ns) \ s \ 0 \ subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (s *R ns))" have "\s ns. ?P r a s ns" if ra: "(r,a) \ set C" for r a proof - have a: "a \ snd ` set (Atoms S)" using ra 0 by force from 0 ra have is_leq: "is_leq_ns (r *R nsc_of_atom a)" and r0: "r \ 0" by auto let ?x = "atom_var a" show ?thesis proof (cases "map_of ?T ?x") case (Some q) hence \: "\ ?x = q" unfolding \_def by auto from Some have xqT: "(?x, q) \ set ?T" by (rule map_of_SomeD) define ns where "ns = (case a of Leq v c \ LEQ_ns q c | Geq v c \ GEQ_ns q c)" from preprocess'_atoms_to_constraints[OF pp refl ns_def a] xqT have ns_mem: "ns \ snd ` set cs" by blast have id: "subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (r *R ns)" using is_leq \ by (cases a, auto simp: ns_def subst_poly_scaleRat) from id is_leq \ have is_leq: "is_leq_ns (r *R ns)" by (cases a, auto simp: ns_def) show ?thesis by (intro exI[of _ r] exI[of _ ns] conjI ns_mem id is_leq conjI r0) next case None hence "?x \ fst ` set ?T" by (meson map_of_eq_None_iff) from preprocess'_atoms_to_constraints[OF pp refl refl a] this obtain rr where rr: "rr *R nsc_of_atom a \ (snd ` set cs)" and rr0: "rr \ 0" by blast from None have \: "\ ?x = lp_monom 1 ?x" unfolding \_def by simp define ns where "ns = rr *R nsc_of_atom a" define s where "s = r / rr" from rr0 r0 have s0: "s \ 0" unfolding s_def by auto from is_leq \ have "subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (r *R nsc_of_atom a)" by (cases a, auto simp: subst_poly_scaleRat) moreover have "r *R nsc_of_atom a = s *R ns" unfolding ns_def s_def scaleRat_scaleRat_ns_constraint[OF rr0] using rr0 by simp ultimately have "subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (s *R ns)" "is_leq_ns (s *R ns)" using is_leq by auto then show ?thesis unfolding ns_def using rr s0 by blast qed qed hence "\ ra. \ s ns. (fst ra, snd ra) \ set C \ ?P (fst ra) (snd ra) s ns" by blast from choice[OF this] obtain s where s: "\ ra. \ ns. (fst ra, snd ra) \ set C \ ?P (fst ra) (snd ra) (s ra) ns" by blast from choice[OF this] obtain ns where ns: "\ r a. (r,a) \ set C \ ?P r a (s (r,a)) (ns (r,a))" by force define NC where "NC = map (\(r,a). (s (r,a), ns (r,a))) C" have "(\(s, ns)\map (\(r,a). (s (r,a), ns (r,a))) C'. lec_of_nsc (s *R ns)) = (\(r, a)\C'. subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)))" if "set C' \ set C" for C' using that proof (induction C') case Nil then show ?case by simp next case (Cons a C') have "(\x\a # C'. lec_of_nsc (s x *R ns x)) = lec_of_nsc (s a *R ns a) + (\x\C'. lec_of_nsc (s x *R ns x))" by simp also have "(\x\C'. lec_of_nsc (s x *R ns x)) = (\(r, a)\C'. subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a)))" using Cons by (auto simp add: case_prod_beta' comp_def) also have "lec_of_nsc (s a *R ns a) = subst_poly_lec \ (lec_of_nsc (fst a *R nsc_of_atom (snd a)))" proof - have "a \ set C" using Cons by simp then show ?thesis using ns by auto qed finally show ?case by (auto simp add: case_prod_beta' comp_def) qed also have "(\(r, a)\C. subst_poly_lec \ (lec_of_nsc (r *R nsc_of_atom a))) = subst_poly_lec \ (\(r, a)\C. (lec_of_nsc (r *R nsc_of_atom a)))" by (auto simp add: subst_poly_lec_sum_list case_prod_beta' comp_def) also have "(\(r, a)\C. (lec_of_nsc (r *R nsc_of_atom a))) = Leqc p c" using 0 by blast also have "subst_poly_lec \ (Leqc p c) = Leqc (subst_poly \ p) c" by simp also have "subst_poly \ p = 0" proof (rule all_valuate_zero) fix v :: "QDelta valuation" have "(subst_poly \ p) \ v \ = (p \ \x. ((\ x) \ v \) \)" by (rule valuate_subst_poly) also have "\ = 0" proof (rule 0(4)) have "(\ a) \ v \ = (q \ \x. ((\ x) \ v \) \)" if "(a, q) \ set (Tableau S)" for a q proof - have "distinct (map fst ?T)" using normalized_tableau_preprocess' assms unfolding normalized_tableau_def lhs_def by (auto simp add: case_prod_beta') then have 0: "\ a = q" unfolding \_def using that by auto have "q \ v \ = (q \ \x. ((\ x) \ v \) \)" proof - have "vars q \ rvars ?T" unfolding rvars_def using that by force moreover have "(\ x) \ v \ = v x" if "x \ rvars ?T" for x proof - have "x \ lvars (Tableau S)" using that normalized_tableau_preprocess' assms unfolding normalized_tableau_def by blast then have "x \ fst ` set (Tableau S)" unfolding lvars_def by force then have "map_of ?T x = None" using map_of_eq_None_iff by metis then have "\ x = lp_monom 1 x" unfolding \_def by auto also have "(lp_monom 1 x) \ v \ = v x" by auto finally show ?thesis . qed ultimately show ?thesis by (auto intro!: valuate_depend) qed then show ?thesis using 0 by blast qed then show "(\x. ((\ x) \ v \)) \\<^sub>t ?T" using 0 by (auto simp add: satisfies_tableau_def satisfies_eq_def) qed finally show "(subst_poly \ p) \ v \ = 0" . qed finally have "(\(s, n)\NC. lec_of_nsc (s *R n)) = Le_Constraint Leq_Rel 0 c" unfolding NC_def by blast moreover have "ns (r,a) \ snd ` set cs" "is_leq_ns (s (r, a) *R ns (r, a))" "s (r, a) \ 0" if "(r, a) \ set C" for r a using ns that by force+ ultimately have "farkas_coefficients_ns (snd ` set cs) NC" unfolding farkas_coefficients_ns_def NC_def using 0 by force then show ?thesis by blast qed lemma preprocess'_unsat_indexD: "i \ set (UnsatIndices (preprocess' ns j)) \ \ c. poly c = 0 \ \ zero_satisfies c \ (i,c) \ set ns" by (induct ns j rule: preprocess'.induct, auto simp: Let_def split: if_splits option.splits) lemma preprocess'_unsat_index_farkas_coefficients_ns: assumes "i \ set (UnsatIndices (preprocess' ns j))" shows "\ C. farkas_coefficients_ns (snd ` set ns) C" proof - from preprocess'_unsat_indexD[OF assms] obtain c where contr: "poly c = 0" "\ zero_satisfies c" and mem: "(i,c) \ set ns" by auto from mem have mem: "c \ snd ` set ns" by force let ?c = "ns_constraint_const c" define r where "r = (case c of LEQ_ns _ _ \ 1 | _ \ (-1 :: rat))" define d where "d = (case c of LEQ_ns _ _ \ ?c | _ \ - ?c)" have [simp]: "(- x < 0) = (0 < x)" for x :: QDelta using uminus_less_lrv[of _ 0] by simp show ?thesis unfolding farkas_coefficients_ns_def by (intro exI[of _ "[(r,c)]"] exI[of _ d], insert mem contr, cases "c", auto simp: r_def d_def) qed text \The combination of the previous results easily provides the main result of this section: a finite set of non-strict constraints on layer~2 is unsatisfiable if and only if there are Farkas coefficients. Again, here we use results from the simplex formalization, namely soundness of the preprocess-function.\ lemma farkas_coefficients_ns: assumes "finite (ns :: QDelta ns_constraint set)" shows "(\ C. farkas_coefficients_ns ns C) \ (\ v. v \\<^sub>n\<^sub>s\<^sub>s ns)" proof assume "\ C. farkas_coefficients_ns ns C" from farkas_coefficients_ns_unsat this show "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" by blast next assume unsat: "\ v. v \\<^sub>n\<^sub>s\<^sub>s ns" from finite_list[OF assms] obtain nsl where ns: "ns = set nsl" by auto let ?cs = "map (Pair ()) nsl" obtain I t ias where part1: "preprocess_part_1 ?cs = (t,ias,I)" by (cases "preprocess_part_1 ?cs", auto) let ?as = "snd ` set ias" let ?s = "start_fresh_variable ?cs" have fin: "finite ?as" by auto have id: "ias = Atoms (preprocess' ?cs ?s)" "t = Tableau (preprocess' ?cs ?s)" "I = UnsatIndices (preprocess' ?cs ?s)" using part1 unfolding preprocess_part_1_def Let_def by auto have norm: "\ t" using normalized_tableau_preprocess'[of ?cs] unfolding id . { fix v assume "v \\<^sub>a\<^sub>s ?as" "v \\<^sub>t t" from preprocess'_sat[OF this[unfolded id], folded id] unsat[unfolded ns] have "set I \ {}" by auto then obtain i where "i \ set I" using all_not_in_conv by blast from preprocess'_unsat_index_farkas_coefficients_ns[OF this[unfolded id]] have "\C. farkas_coefficients_ns (snd ` set ?cs) C" by simp } with farkas_coefficients_atoms_tableau[OF norm fin] obtain C where "farkas_coefficients_atoms_tableau ?as t C \ (\C. farkas_coefficients_ns (snd ` set ?cs) C)" by blast from farkas_coefficients_preprocess'[of ?cs, OF refl] this have "\ C. farkas_coefficients_ns (snd ` set ?cs) C" using part1 unfolding preprocess_part_1_def Let_def by auto also have "snd ` set ?cs = ns" unfolding ns by force finally show "\ C. farkas_coefficients_ns ns C" . qed subsection \Farkas' Lemma on Layer 1\ text \The main difference of layers 1 and 2 is the restriction to non-strict constraints via delta-rationals. Since we now work with another constraint type, @{type constraint}, we again need translations into linear inequalities of type @{type le_constraint}. Moreover, we also need to define scaling of constraints where flipping the comparison sign may be required.\ fun is_le :: "constraint \ bool" where "is_le (LT _ _) = True" | "is_le (LEQ _ _) = True" | "is_le _ = False" fun lec_of_constraint where "lec_of_constraint (LEQ p c) = (Le_Constraint Leq_Rel p c)" | "lec_of_constraint (LT p c) = (Le_Constraint Lt_Rel p c)" lemma lec_of_constraint: assumes "is_le c" shows "(v \\<^sub>l\<^sub>e (lec_of_constraint c)) \ (v \\<^sub>c c)" using assms by (cases c, auto) instantiation constraint :: scaleRat begin fun scaleRat_constraint :: "rat \ constraint \ constraint" where "scaleRat_constraint r cc = (if r = 0 then LEQ 0 0 else (case cc of LEQ p c \ (if (r < 0) then GEQ (r *R p) (r *R c) else LEQ (r *R p) (r *R c)) | LT p c \ (if (r < 0) then GT (r *R p) (r *R c) else LT (r *R p) (r *R c)) | GEQ p c \ (if (r > 0) then GEQ (r *R p) (r *R c) else LEQ (r *R p) (r *R c)) | GT p c \ (if (r > 0) then GT (r *R p) (r *R c) else LT (r *R p) (r *R c)) - | LTPP p q \ - (if (r < 0) then GT (r *R (p - q)) 0 else LT (r *R (p - q)) 0) - | LEQPP p q \ - (if (r < 0) then GEQ (r *R (p - q)) 0 else LEQ (r *R (p - q)) 0) - | GTPP p q \ - (if (r > 0) then GT (r *R (p - q)) 0 else LT (r *R (p - q)) 0) - | GEQPP p q \ - (if (r > 0) then GEQ (r *R (p - q)) 0 else LEQ (r *R (p - q)) 0) - | EQPP p q \ LEQ (r *R (p - q)) 0 \ \We do not keep equality, since the aim is + | EQ p c \ LEQ (r *R p) (r *R c) \ \We do not keep equality, since the aim is to convert the scaled constraints into inequalities, which will then be summed up.\ - | EQ p c \ LEQ (r *R p) (r *R c) ))" instance .. end lemma sat_scale_rat: assumes "(v :: rat valuation) \\<^sub>c c" shows "v \\<^sub>c (r *R c)" proof - have "r < 0 \ r = 0 \ r > 0" by auto then show ?thesis using assms by (cases c, auto simp: right_diff_distrib valuate_minus valuate_scaleRat scaleRat_leq1 scaleRat_leq2 valuate_zero) qed text \In the following definition of Farkas coefficients (for layer 1), the main difference to @{const farkas_coefficients_ns} is that the linear combination evaluates either to a strict inequality where the constant must be non-positive, or to a non-strict inequality where the constant must be negative.\ definition farkas_coefficients where "farkas_coefficients cs C = (\ d rel. (\ (r,c) \ set C. c \ cs \ is_le (r *R c) \ r \ 0) \ (\ (r,c) \ C. lec_of_constraint (r *R c)) = Le_Constraint rel 0 d \ (rel = Leq_Rel \ d < 0 \ rel = Lt_Rel \ d \ 0))" text \Again, the existence Farkas coefficients immediately implies unsatisfiability.\ lemma farkas_coefficients_unsat: assumes "farkas_coefficients cs C" shows "\ v. v \\<^sub>c\<^sub>s cs" proof assume "\ v. v \\<^sub>c\<^sub>s cs" then obtain v where *: "v \\<^sub>c\<^sub>s cs" by auto obtain d rel where isleq: "(\(r,c) \ set C. c \ cs \ is_le (r *R c) \ r \ 0)" and leq: "(\ (r,c) \ C. lec_of_constraint (r *R c)) = Le_Constraint rel 0 d" and choice: "rel = Lt_Rel \ d \ 0 \ rel = Leq_Rel \ d < 0" using assms farkas_coefficients_def by blast { fix r c assume c: "(r,c) \ set C" from c * isleq have "v \\<^sub>c c" by auto hence v: "v \\<^sub>c (r *R c)" by (rule sat_scale_rat) from c isleq have "is_le (r *R c)" by auto from lec_of_constraint[OF this] v have "v \\<^sub>l\<^sub>e lec_of_constraint (r *R c)" by blast } note v = this have "v \\<^sub>l\<^sub>e Le_Constraint rel 0 d" unfolding leq[symmetric] by (rule satisfies_sumlist_le_constraints, insert v, auto) then show False using choice by (cases rel, auto simp: valuate_zero) qed text \Now follows the difficult implication. The major part is proving that the translation @{const constraint_to_qdelta_constraint} preserves the existence of Farkas coefficients via pointwise compatibility of the sum. Here, compatibility links a strict or non-strict inequality from the input constraint to a translated non-strict inequality over delta-rationals.\ fun compatible_cs where "compatible_cs (Le_Constraint Leq_Rel p c) (Le_Constraint Leq_Rel q d) = (q = p \ d = QDelta c 0)" | "compatible_cs (Le_Constraint Lt_Rel p c) (Le_Constraint Leq_Rel q d) = (q = p \ qdfst d = c)" | "compatible_cs _ _ = False" lemma compatible_cs_0_0: "compatible_cs 0 0" by code_simp lemma compatible_cs_plus: "compatible_cs c1 d1 \ compatible_cs c2 d2 \ compatible_cs (c1 + c2) (d1 + d2)" by (cases c1; cases d1; cases c2; cases d2; cases "lec_rel c1"; cases "lec_rel d1"; cases "lec_rel c2"; cases "lec_rel d2"; auto simp: plus_QDelta_def) lemma unsat_farkas_coefficients: assumes "\ v. v \\<^sub>c\<^sub>s cs" and fin: "finite cs" shows "\ C. farkas_coefficients cs C" proof - from finite_list[OF fin] obtain csl where cs: "cs = set csl" by blast let ?csl = "map (Pair ()) csl" let ?ns = "(snd ` set (to_ns ?csl))" let ?nsl = "concat (map constraint_to_qdelta_constraint csl)" have id: "snd ` set ?csl = cs" unfolding cs by force have id2: "?ns = set ?nsl" unfolding to_ns_def set_concat by force from SolveExec'Default.to_ns_sat[of ?csl, unfolded id] assms have unsat: "\ v. \v\ \\<^sub>n\<^sub>s\<^sub>s ?ns" by metis have fin: "finite ?ns" by auto have "\ v. v \\<^sub>n\<^sub>s\<^sub>s ?ns" proof assume "\ v. v \\<^sub>n\<^sub>s\<^sub>s ?ns" then obtain v where model: "v \\<^sub>n\<^sub>s\<^sub>s ?ns" by blast let ?v = "Mapping.Mapping (\ x. Some (v x))" have "v = \?v\" by (intro ext, auto simp: map2fun_def Mapping.lookup.abs_eq) from model this unsat show False by metis qed from farkas_coefficients_ns[OF fin] this id2 obtain N where farkas: "farkas_coefficients_ns (set ?nsl) N" by metis from this[unfolded farkas_coefficients_ns_def] obtain d where is_leq: "\ a n. (a,n) \ set N \ n \ set ?nsl \ is_leq_ns (a *R n) \ a \ 0" and sum: "(\(a,n)\N. lec_of_nsc (a *R n)) = Le_Constraint Leq_Rel 0 d" and d0: "d < 0" by blast let ?prop = "\ NN C. (\ (a,c) \ set C. c \ cs \ is_le (a *R c) \ a \ 0) \ compatible_cs (\ (a,c) \ C. lec_of_constraint (a *R c)) (\(a,n)\NN. lec_of_nsc (a *R n))" have "set NN \ set N \ \ C. ?prop NN C" for NN proof (induct NN) case Nil have "?prop Nil Nil" by (simp add: compatible_cs_0_0) thus ?case by blast next case (Cons an NN) obtain a n where an: "an = (a,n)" by force from Cons an obtain C where IH: "?prop NN C" and n: "(a,n) \ set N" by auto have compat_CN: "compatible_cs (\(f, c)\C. lec_of_constraint (f *R c)) (\(a,n)\NN. lec_of_nsc (a *R n))" using IH by blast from n is_leq obtain c where c: "c \ cs" and nc: "n \ set (constraint_to_qdelta_constraint c)" unfolding cs by force from is_leq[OF n] have is_leq: "is_leq_ns (a *R n) \ a \ 0" by blast have is_less: "is_le (a *R c)" and a0: "a \ 0" and compat_cn: "compatible_cs (lec_of_constraint (a *R c)) (lec_of_nsc (a *R n))" by (atomize(full), cases c, insert is_leq nc, auto simp: QDelta_0_0 scaleRat_QDelta_def qdsnd_0 qdfst_0) let ?C = "Cons (a, c) C" let ?N = "Cons (a, n) NN" have "?prop ?N ?C" unfolding an proof (intro conjI) show "\ (a,c) \ set ?C. c \ cs \ is_le (a *R c) \ a \ 0" using IH is_less a0 c by auto show "compatible_cs (\(a, c)\?C. lec_of_constraint (a *R c)) (\(a,n)\?N. lec_of_nsc (a *R n))" using compatible_cs_plus[OF compat_cn compat_CN] by simp qed thus ?case unfolding an by blast qed from this[OF subset_refl, unfolded sum] obtain C where is_less: "(\(a, c)\set C. c \ cs \ is_le (a *R c) \ a \ 0)" and compat: "compatible_cs (\(f, c)\C. lec_of_constraint (f *R c)) (Le_Constraint Leq_Rel 0 d)" (is "compatible_cs ?sum _") by blast obtain rel p e where "?sum = Le_Constraint rel p e" by (cases ?sum) with compat have sum: "?sum = Le_Constraint rel 0 e" by (cases rel, auto) have e: "(rel = Leq_Rel \ e < 0 \ rel = Lt_Rel \ e \ 0)" using compat[unfolded sum] d0 by (cases rel, auto simp: less_QDelta_def qdfst_0 qdsnd_0) show ?thesis unfolding farkas_coefficients_def by (intro exI conjI, rule is_less, rule sum, insert e, auto) qed text \Finally we can prove on layer 1 that a finite set of constraints is unsatisfiable if and only if there are Farkas coefficients.\ lemma farkas_coefficients: assumes "finite cs" shows "(\ C. farkas_coefficients cs C) \ (\ v. v \\<^sub>c\<^sub>s cs)" using farkas_coefficients_unsat unsat_farkas_coefficients[OF _ assms] by blast section \Corollaries from the Literature\ text \In this section, we convert the previous variations of Farkas' Lemma into more well-known forms of this result. Moreover, instead of referring to the various constraint types of the simplex formalization, we now speak solely about constraints of type @{type le_constraint}.\ subsection \Farkas' Lemma on Delta-Rationals\ text \We start with Lemma~2 of \<^cite>\"Bromberger2017"\, a variant of Farkas' Lemma for delta-rationals. To be more precise, it states that a set of non-strict inequalities over delta-rationals is unsatisfiable if and only if there is a linear combination of the inequalities that results in a trivial unsatisfiable constraint $0 < const$ for some negative constant $const$. We can easily prove this statement via the lemma @{thm [source] farkas_coefficients_ns} and some conversions between the different constraint types.\ lemma Farkas'_Lemma_Delta_Rationals: fixes cs :: "QDelta le_constraint set" assumes only_non_strict: "lec_rel ` cs \ {Leq_Rel}" and fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const. (\ (r, c) \ set C. r > 0 \ c \ cs) \ (\ (r,c) \ C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const \ const < 0)" (is "?lhs = ?rhs") proof - { fix c assume "c \ cs" with only_non_strict have "lec_rel c = Leq_Rel" by auto then have "\ p const. c = Leqc p const" by (cases c, auto) } note leqc = this let ?to_ns = "\ c. LEQ_ns (lec_poly c) (lec_const c)" let ?ns = "?to_ns ` cs" from fin have fin: "finite ?ns" by auto have "v \\<^sub>n\<^sub>s\<^sub>s ?ns \ (\ c \ cs. v \\<^sub>l\<^sub>e c)" for v using leqc by fastforce hence "?lhs = (\ v. v \\<^sub>n\<^sub>s\<^sub>s ?ns)" by simp also have "\ = (\C. farkas_coefficients_ns ?ns C)" unfolding farkas_coefficients_ns[OF fin] .. also have "\ = ?rhs" proof assume "\ C. farkas_coefficients_ns ?ns C" then obtain C const where is_leq: "\ (s, n) \ set C. n \ ?ns \ is_leq_ns (s *R n) \ s \ 0" and sum: "(\(s, n)\C. lec_of_nsc (s *R n)) = Leqc 0 const" and c0: "const < 0" unfolding farkas_coefficients_ns_def by blast let ?C = "map (\ (s,n). (s,lec_of_nsc n)) C" show ?rhs proof (intro exI[of _ ?C] exI[of _ const] conjI c0, unfold sum[symmetric] map_map o_def set_map, intro ballI, clarify) { fix s n assume sn: "(s, n) \ set C" with is_leq have n_ns: "n \ ?ns" and is_leq: "is_leq_ns (s *R n)" "s \ 0" by force+ from n_ns obtain c where c: "c \ cs" and n: "n = LEQ_ns (lec_poly c) (lec_const c)" by auto with leqc[OF c] obtain p d where cs: "Leqc p d \ cs" and n: "n = LEQ_ns p d" by auto from is_leq[unfolded n] have s0: "s > 0" by (auto split: if_splits) let ?n = "lec_of_nsc n" from cs n have mem: "?n \ cs" by auto show "0 < s \ ?n \ cs" using s0 mem by blast have "Leqc (s *R lec_poly ?n) (s *R lec_const ?n) = lec_of_nsc (s *R n)" unfolding n using s0 by simp } note id = this show "(\x\C. case case x of (s, n) \ (s, lec_of_nsc n) of (r, c) \ Leqc (r *R lec_poly c) (r *R lec_const c)) = (\(s, n)\C. lec_of_nsc (s *R n))" (is "sum_list (map ?f C) = sum_list (map ?g C)") proof (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl]) fix pair assume mem: "pair \ set C" then obtain s n where pair: "pair = (s,n)" by force show "?f pair = ?g pair" unfolding pair split using id[OF mem[unfolded pair]] . qed qed next assume ?rhs then obtain C const where C: "\ r c. (r,c) \ set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const" and const: "const < 0" by blast let ?C = "map (\ (r,c). (r, ?to_ns c)) C" show "\ C. farkas_coefficients_ns ?ns C" unfolding farkas_coefficients_ns_def proof (intro exI[of _ ?C] exI[of _ const] conjI const, unfold sum[symmetric]) show "\(s, n)\set ?C. n \ ?ns \ is_leq_ns (s *R n) \ s \ 0" using C by fastforce show "(\(s, n)\?C. lec_of_nsc (s *R n)) = (\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c))" unfolding map_map o_def by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, force) qed qed finally show ?thesis . qed subsection \Motzkin's Transposition Theorem or the Kuhn-Fourier Theorem\ text \Next, we prove a generalization of Farkas' Lemma that permits arbitrary combinations of strict and non-strict inequalities: Motzkin's Transposition Theorem which is also known as the Kuhn--Fourier Theorem. The proof is mainly based on the lemma @{thm [source] farkas_coefficients}, again requiring conversions between constraint types.\ theorem Motzkin's_transposition_theorem: fixes cs :: "rat le_constraint set" assumes fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const rel. (\ (r, c) \ set C. r > 0 \ c \ cs) \ (\ (r,c) \ C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const \ (rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0))" (is "?lhs = ?rhs") proof - let ?to_cs = "\ c. (case lec_rel c of Leq_Rel \ LEQ | _ \ LT) (lec_poly c) (lec_const c)" have to_cs: "v \\<^sub>c ?to_cs c \ v \\<^sub>l\<^sub>e c" for v c by (cases c, cases "lec_rel c", auto) let ?cs = "?to_cs ` cs" from fin have fin: "finite ?cs" by auto have "v \\<^sub>c\<^sub>s ?cs \ (\ c \ cs. v \\<^sub>l\<^sub>e c)" for v using to_cs by auto hence "?lhs = (\ v. v \\<^sub>c\<^sub>s ?cs)" by simp also have "\ = (\C. farkas_coefficients ?cs C)" unfolding farkas_coefficients[OF fin] .. also have "\ = ?rhs" proof assume "\ C. farkas_coefficients ?cs C" then obtain C const rel where is_leq: "\ (s, n) \ set C. n \ ?cs \ is_le (s *R n) \ s \ 0" and sum: "(\(s, n)\C. lec_of_constraint (s *R n)) = Le_Constraint rel 0 const" and c0: "(rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0)" unfolding farkas_coefficients_def by blast let ?C = "map (\ (s,n). (s,lec_of_constraint n)) C" show ?rhs proof (intro exI[of _ ?C] exI[of _ const] exI[of _ rel] conjI c0, unfold map_map o_def set_map sum[symmetric], intro ballI, clarify) { fix s n assume sn: "(s, n) \ set C" with is_leq have n_ns: "n \ ?cs" and is_leq: "is_le (s *R n)" and s0: "s \ 0" by force+ from n_ns obtain c where c: "c \ cs" and n: "n = ?to_cs c" by auto from is_leq[unfolded n] have "s \ 0" by (cases "lec_rel c", auto split: if_splits) with s0 have s0: "s > 0" by auto let ?c = "lec_of_constraint n" from c n have mem: "?c \ cs" by (cases c, cases "lec_rel c", auto) show "0 < s \ ?c \ cs" using s0 mem by blast have "lec_of_constraint (s *R n) = Le_Constraint (lec_rel ?c) (s *R lec_poly ?c) (s *R lec_const ?c)" unfolding n using s0 by (cases c, cases "lec_rel c", auto) } note id = this show "(\x\C. case case x of (s, n) \ (s, lec_of_constraint n) of (r, c) \ Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = (\(s, n)\C. lec_of_constraint (s *R n))" (is "sum_list (map ?f C) = sum_list (map ?g C)") proof (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl]) fix pair assume mem: "pair \ set C" obtain r c where pair: "pair = (r,c)" by force show "?f pair = ?g pair" unfolding pair split id[OF mem[unfolded pair]] .. qed qed next assume ?rhs then obtain C const rel where C: "\ r c. (r,c) \ set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const" and const: "rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0" by blast let ?C = "map (\ (r,c). (r, ?to_cs c)) C" show "\ C. farkas_coefficients ?cs C" unfolding farkas_coefficients_def proof (intro exI[of _ ?C] exI[of _ const] exI[of _ rel] conjI const, unfold sum[symmetric]) show "\(s, n)\set ?C. n \ ?cs \ is_le (s *R n) \ s \ 0" using C by (fastforce split: le_rel.splits) show "(\(s, n)\?C. lec_of_constraint (s *R n)) = (\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c))" unfolding map_map o_def by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, fastforce split: le_rel.splits) qed qed finally show ?thesis . qed subsection \Farkas' Lemma\ text \Finally we derive the commonly used form of Farkas' Lemma, which easily follows from @{thm [source] Motzkin's_transposition_theorem}. It only permits non-strict inequalities and, as a result, the sum of inequalities will always be non-strict.\ lemma Farkas'_Lemma: fixes cs :: "rat le_constraint set" assumes only_non_strict: "lec_rel ` cs \ {Leq_Rel}" and fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const. (\ (r, c) \ set C. r > 0 \ c \ cs) \ (\ (r,c) \ C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const \ const < 0)" (is "_ = ?rhs") proof - { fix c assume "c \ cs" with only_non_strict have "lec_rel c = Leq_Rel" by auto then have "\ p const. c = Leqc p const" by (cases c, auto) } note leqc = this let ?lhs = "\C const rel. (\(r, c)\set C. 0 < r \ c \ cs) \ (\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const \ (rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0)" show ?thesis unfolding Motzkin's_transposition_theorem[OF fin] proof assume ?rhs then obtain C const where C: "\ r c. (r, c)\set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const" and const: "const < 0" by blast show ?lhs proof (intro exI[of _ C] exI[of _ const] exI[of _ Leq_Rel] conjI) show "\ (r,c) \ set C. 0 < r \ c \ cs" using C by force show "(\(r, c)\ C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const" unfolding sum[symmetric] by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, force dest!: leqc) qed (insert const, auto) next assume ?lhs then obtain C const rel where C: "\ r c. (r, c)\set C \ 0 < r \ c \ cs" and sum: "(\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint rel 0 const" and const: "rel = Leq_Rel \ const < 0 \ rel = Lt_Rel \ const \ 0" by blast have id: "(\(r, c)\C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) = (\(r, c)\C. Leqc (r *R lec_poly c) (r *R lec_const c))" (is "_ = ?sum") by (rule arg_cong[of _ _ sum_list], rule map_cong, auto dest!: C leqc) have "lec_rel ?sum = Leq_Rel" unfolding sum_list_lec by (auto simp add: sum_list_Leq_Rel o_def) with sum[unfolded id] have rel: "rel = Leq_Rel" by auto with const have const: "const < 0" by auto show ?rhs by (intro exI[of _ C] exI[of _ const] conjI const, insert sum id C rel, force+) qed qed text \We also present slightly modified versions\ lemma sum_list_map_filter_sum: fixes f :: "'a \ 'b :: comm_monoid_add" shows "sum_list (map f (filter g xs)) + sum_list (map f (filter (Not o g) xs)) = sum_list (map f xs)" by (induct xs, auto simp: ac_simps) text \A version where every constraint obtains exactly one coefficient and where 0 coefficients are allowed.\ lemma Farkas'_Lemma_set_sum: fixes cs :: "rat le_constraint set" assumes only_non_strict: "lec_rel ` cs \ {Leq_Rel}" and fin: "finite cs" shows "(\ v. \ c \ cs. v \\<^sub>l\<^sub>e c) \ (\ C const. (\ c \ cs. C c \ 0) \ (\ c \ cs. Leqc ((C c) *R lec_poly c) ((C c) *R lec_const c)) = Leqc 0 const \ const < 0)" unfolding Farkas'_Lemma[OF only_non_strict fin] proof ((standard; elim exE conjE), goal_cases) case (2 C const) from finite_distinct_list[OF fin] obtain csl where csl: "set csl = cs" and dist: "distinct csl" by auto let ?list = "filter (\ c. C c \ 0) csl" let ?C = "map (\ c. (C c, c)) ?list" show ?case proof (intro exI[of _ ?C] exI[of _ const] conjI) have "(\(r, c)\?C. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = (\(r, c)\map (\c. (C c, c)) csl. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c))" unfolding map_map by (rule sum_list_map_filter, auto simp: zero_le_constraint_def) also have "\ = Le_Constraint Leq_Rel 0 const" unfolding 2(2)[symmetric] csl[symmetric] unfolding sum.distinct_set_conv_list[OF dist] map_map o_def split .. finally show "(\(r, c)\?C. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint Leq_Rel 0 const" by auto show "const < 0" by fact show "\(r, c)\set ?C. 0 < r \ c \ cs" using 2(1) unfolding set_map set_filter csl by auto qed next case (1 C const) define CC where "CC = (\ c. sum_list (map fst (filter (\ rc. snd rc = c) C)))" show "(\ C const. (\ c \ cs. C c \ 0) \ (\ c \ cs. Leqc ((C c) *R lec_poly c) ((C c) *R lec_const c)) = Leqc 0 const \ const < 0)" proof (intro exI[of _ CC] exI[of _ const] conjI) show "\c\cs. 0 \ CC c" unfolding CC_def using 1(1) by (force intro!: sum_list_nonneg) show "const < 0" by fact from 1 have snd: "snd ` set C \ cs" by auto show "(\c\cs. Le_Constraint Leq_Rel (CC c *R lec_poly c) (CC c *R lec_const c)) = Le_Constraint Leq_Rel 0 const" unfolding 1(2)[symmetric] using fin snd unfolding CC_def proof (induct cs arbitrary: C rule: finite_induct) case empty hence C: "C = []" by auto thus ?case by simp next case *: (insert c cs C) let ?D = "filter (Not \ (\rc. snd rc = c)) C" from * have "snd ` set ?D \ cs" by auto note IH = *(3)[OF this] have id: "(\a\ ?D. case a of (r, c) \ Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = (\(r, c)\?D. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c))" by (induct C, force+) show ?case unfolding sum.insert[OF *(1,2)] unfolding sum_list_map_filter_sum[of _ "\ rc. snd rc = c" C, symmetric] proof (rule arg_cong2[of _ _ _ _ "(+)"], goal_cases) case 2 show ?case unfolding IH[symmetric] by (rule sum.cong, insert *(2,1), auto intro!: arg_cong[of _ _ "\ xs. sum_list (map _ xs)"], (induct C, auto)+) next case 1 show ?case proof (rule sym, induct C) case (Cons rc C) thus ?case by (cases "rc", cases "snd rc = c", auto simp: field_simps scaleRat_left_distrib) qed (auto simp: zero_le_constraint_def) qed qed qed qed text \A version with indexed constraints, i.e., in particular where constraints may occur several times.\ lemma Farkas'_Lemma_indexed: fixes c :: "nat \ rat le_constraint" assumes only_non_strict: "lec_rel ` c ` Is \ {Leq_Rel}" and fin: "finite Is" shows "(\ v. \ i \ Is. v \\<^sub>l\<^sub>e c i) \ (\ C const. (\ i \ Is. C i \ 0) \ (\ i \ Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const \ const < 0)" proof - let ?C = "c ` Is" have fin: "finite ?C" using fin by auto have "(\ v. \ i \ Is. v \\<^sub>l\<^sub>e c i) = (\ v. \ cc \ ?C. v \\<^sub>l\<^sub>e cc)" by force also have "\ = (\ C const. (\ i \ Is. C i \ 0) \ (\ i \ Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const \ const < 0)" (is "?l = ?r") proof assume ?r then obtain C const where r: "(\ i \ Is. C i \ 0)" and eq: "(\ i \ Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const" and "const < 0" by auto from finite_distinct_list[OF \finite Is\] obtain Isl where isl: "set Isl = Is" and dist: "distinct Isl" by auto let ?CC = "filter (\ rc. fst rc \ 0) (map (\ i. (C i, c i)) Isl)" show ?l unfolding Farkas'_Lemma[OF only_non_strict fin] proof (intro exI[of _ ?CC] exI[of _ const] conjI) show "const < 0" by fact show "\ (r, ca) \ set ?CC. 0 < r \ ca \ ?C" using r(1) isl by auto show "(\(r, c)\?CC. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint Leq_Rel 0 const" unfolding eq[symmetric] by (subst sum_list_map_filter, force simp: zero_le_constraint_def, unfold map_map o_def, subst sum_list_distinct_conv_sum_set[OF dist], rule sum.cong, auto simp: isl) qed next assume ?l from this[unfolded Farkas'_Lemma_set_sum[OF only_non_strict fin]] obtain C const where nonneg: "(\c\ ?C. 0 \ C c)" and sum: "(\c\ ?C. Le_Constraint Leq_Rel (C c *R lec_poly c) (C c *R lec_const c)) = Le_Constraint Leq_Rel 0 const" and const: "const < 0" by blast define I where "I = (\ i. (C (c i) / rat_of_nat (card (Is \ { j. c i = c j}))))" show ?r proof (intro exI[of _ I] exI[of _ const] conjI const) show "\i \ Is. 0 \ I i" using nonneg unfolding I_def by auto show "(\ i \ Is. Le_Constraint Leq_Rel (I i *R lec_poly (c i)) (I i *R lec_const (c i))) = Le_Constraint Leq_Rel 0 const" unfolding sum[symmetric] unfolding sum.image_gen[OF \finite Is\, of _ c] proof (rule sum.cong[OF refl], goal_cases) case (1 cc) define II where "II = (Is \ {j. cc = c j})" from 1 have "II \ {}" unfolding II_def by auto moreover have finII: "finite II" using \finite Is\ unfolding II_def by auto ultimately have card: "card II \ 0" by auto let ?C = "\ II. rat_of_nat (card II)" define ii where "ii = C cc / rat_of_nat (card II)" have "(\i\{x \ Is. c x = cc}. Le_Constraint Leq_Rel (I i *R lec_poly (c i)) (I i *R lec_const (c i))) = (\ i\ II. Le_Constraint Leq_Rel (ii *R lec_poly cc) (ii *R lec_const cc))" unfolding I_def ii_def II_def by (rule sum.cong, auto) also have "\ = Le_Constraint Leq_Rel ((?C II * ii) *R lec_poly cc) ((?C II * ii) *R lec_const cc)" using finII by (induct II rule: finite_induct, auto simp: zero_le_constraint_def field_simps scaleRat_left_distrib) also have "?C II * ii = C cc" unfolding ii_def using card by auto finally show ?case . qed qed qed finally show ?thesis . qed end diff --git a/thys/Linear_Programming/LP_Preliminaries.thy b/thys/Linear_Programming/LP_Preliminaries.thy --- a/thys/Linear_Programming/LP_Preliminaries.thy +++ b/thys/Linear_Programming/LP_Preliminaries.thy @@ -1,1200 +1,1185 @@ theory LP_Preliminaries imports More_Jordan_Normal_Forms Matrix_LinPoly Jordan_Normal_Form.Matrix_Impl Farkas.Simplex_for_Reals "HOL-Library.Mapping" begin (* Component wise greater equal constraints for vector b starting at index \ [x\<^sub>i\<^sub>n\<^sub>d\<^sub>e\<^sub>x, x\<^sub>i\<^sub>n\<^sub>d\<^sub>e\<^sub>x\<^sub>+\<^sub>1,\,x\<^sub>i\<^sub>n\<^sub>d\<^sub>e\<^sub>x\<^sub>+\<^sub>n] \ [b\<^sub>0, b\<^sub>1,\,b\<^sub>n] \ *) fun vars_from_index_geq_vec where "vars_from_index_geq_vec index b = [GEQ (lp_monom 1 (i+index)) (b$i). i \ [0.. {0..i. GEQ (lp_monom 1 (i+start)) (b$i))" "dim_vec b"] by auto lemma vars_from_index_geq_sat: assumes "\x\ \\<^sub>c\<^sub>s set (vars_from_index_geq_vec start b)" assumes "i < dim_vec b" shows "\x\ (i+start) \ b$i" proof - have e_e:"GEQ (lp_monom 1 (i+start)) (b$i) \ set (vars_from_index_geq_vec start b)" using constraints_set_vars_geq_vec_def[of start b] using assms(2) by auto then have "\x\ \\<^sub>c GEQ (lp_monom 1 (i+start)) (b$i)" using assms(1) by blast then have "(lp_monom 1 (i+start)) \\x\\ \ (b$i)" using satisfies_constraint.simps(4)[of "\x\" "lp_monom 1 (i + start)" "b$i"] by simp then show ?thesis by simp qed (* Matrix A less equal vector b (A \ b): a1 b1 c1 d1 \ X \ b_1, a2 b2 c2 d2 \ X \ b_2, ... *) fun mat_x_leq_vec where "mat_x_leq_vec A b = [LEQ (matrix_to_lpolies A!i) (b$i) . i <- [0..x\ \\<^sub>c\<^sub>s set (mat_x_leq_vec A b)" assumes "i < dim_vec b" shows "((matrix_to_lpolies A)!i) \\x\\ \ b$i" proof - have e_e: "LEQ ((matrix_to_lpolies A)!i) (b$i) \ set (mat_x_leq_vec A b)" by (simp add: assms(2)) then have "\x\ \\<^sub>c LEQ ((matrix_to_lpolies A)!i) (b$i)" using assms(1) by blast then show ?thesis using satisfies_constraint.simps by auto qed (* Matrix A less equal vector b (A = b): a1 b1 c1 d1 \ X = b_1, a2 b2 c2 d2 \ X = b_2, ... *) fun x_mat_eq_vec where "x_mat_eq_vec b A = [EQ (matrix_to_lpolies A!i) (b$i) . i <- [0..\<^sub>c\<^sub>s set (x_mat_eq_vec b A)" assumes "i < dim_vec b" shows "((matrix_to_lpolies A)!i) \ x \ = b$i" proof - have e_e: "EQ ((matrix_to_lpolies A)!i) (b$i) \ set (x_mat_eq_vec b A)" by (simp add: assms(2)) then have "x \\<^sub>c EQ ((matrix_to_lpolies A)!i) (b$i)" using assms(1) by blast then show ?thesis using satisfies_constraint.simps by auto qed section \ Get different matrices into same space, without interference \ (* Given matrix A and B create: A 0 0 B *) fun two_block_non_interfering where "two_block_non_interfering A B = (let z1 = 0\<^sub>m (dim_row A) (dim_col B); z2 = 0\<^sub>m (dim_row B) (dim_col A) in four_block_mat A z1 z2 B)" lemma split_two_block_non_interfering: assumes "split_block (two_block_non_interfering A B) (dim_row A) (dim_col A) = (Q1, Q2, Q3, Q4)" shows "Q1 = A" "Q4 = B" using split_four_block_dual_fst_lst[of A _ _ B Q1 Q2 Q3 Q4] assms by auto lemma two_block_non_interfering_dims: "dim_row (two_block_non_interfering A B) = dim_row A + dim_row B" "dim_col (two_block_non_interfering A B) = dim_col A + dim_col B" by (simp)+ lemma two_block_non_interfering_zeros_are_0: assumes "i < dim_row A" and "j \ dim_col A" and "j < dim_col (two_block_non_interfering A B)" shows "(two_block_non_interfering A B)$$(i,j) = 0" "(two_block_non_interfering A B)$$(i,j) = 0" using four_block_mat_def assms two_block_non_interfering_dims[of A B] by auto lemma two_block_non_interfering_row_comp1: assumes "i v (0\<^sub>v (dim_col B))" using assms by auto lemma two_block_non_interfering_row_comp2: assumes "i dim_row A" shows "row (two_block_non_interfering A B) i = (0\<^sub>v (dim_col A)) @\<^sub>v row B (i - dim_row A)" using assms by (auto) lemma first_vec_two_block_non_inter_is_first_vec: assumes "dim_col A + dim_col B = dim_vec v" assumes "dim_row A = n" shows "vec_first (two_block_non_interfering A B *\<^sub>v v) n = A *\<^sub>v (vec_first v (dim_col A))" proof fix i assume a: "i < dim_vec (A *\<^sub>v vec_first v (dim_col A))" let ?tb = "two_block_non_interfering A B" have i_n: "i < n" using a assms(2) by auto have "vec_first (?tb *\<^sub>v v) n $ i = vec_first (vec (dim_row ?tb) (\ i. row ?tb i \ v)) n $ i" unfolding mult_mat_vec_def by simp also have "... = (vec n (\ i. row ?tb i \ v)) $ i" unfolding vec_first_def using trans_less_add1 by (metis a assms(2) dim_mult_mat_vec index_vec two_block_non_interfering_dims(1)) also have "... = row ?tb i \ v" by (simp add: i_n) also have "... = (row A i @\<^sub>v 0\<^sub>v (dim_col B)) \ v" using assms(2) i_n two_block_non_interfering_row_comp1 by fastforce also have "... = row A i \ vec_first v (dim_vec (row A i)) + 0\<^sub>v (dim_col B) \ vec_last v (dim_vec (0\<^sub>v (dim_col B)))" using append_split_vec_distrib_scalar_prod[of "row A i" "0\<^sub>v (dim_col B)" v] assms(1) by auto then have "vec_first (two_block_non_interfering A B *\<^sub>v v) n $ i = row A i \ vec_first v (dim_vec (row A i))" using calculation by auto then show "vec_first (two_block_non_interfering A B *\<^sub>v v) n $ i = (A *\<^sub>v vec_first v (dim_col A)) $ i" by (simp add: assms(2) i_n) next have "dim_vec (A *\<^sub>v v) = dim_row A" using dim_vec_def dim_mult_mat_vec[of A v] by auto then have "dim_vec (vec_first (two_block_non_interfering A B *\<^sub>v v) n) = n" by auto then show "dim_vec (vec_first (two_block_non_interfering A B *\<^sub>v v) n) = dim_vec (A *\<^sub>v vec_first v (dim_col A))" by (simp add: assms(2)) qed lemma last_vec_two_block_non_inter_is_last_vec: assumes "dim_col A + dim_col B = dim_vec v" assumes "dim_row B = n" shows "vec_last ((two_block_non_interfering A B) *\<^sub>v v) n = B *\<^sub>v (vec_last v (dim_col B))" proof fix i assume a: "i < dim_vec (B *\<^sub>v vec_last v (dim_col B))" let ?tb = "two_block_non_interfering A B" let ?vl = "(vec (dim_row ?tb) (\ i. row ?tb i \ v))" have i_n: "i < n" using assms(2) using a by auto have in3: "(dim_row ?tb) - n + i \ dim_row A" by (simp add: assms(2)) have in3': "(dim_row ?tb) - n + i < dim_row ?tb" by (simp add: assms(2) i_n two_block_non_interfering_dims(1)) have "dim_row A + n = dim_row (two_block_non_interfering A B)" by (simp add: assms(2) two_block_non_interfering_dims(1)) then have dim_a: "dim_row A = dim_row (two_block_non_interfering A B) - n" by (metis (no_types) diff_add_inverse2) have "vec_last (?tb *\<^sub>v v) n $ i = vec_last (vec (dim_row ?tb) (\ i. row ?tb i \ v)) n $ i" unfolding mult_mat_vec_def by auto also have "... = ?vl $ (dim_vec ?vl - n + i)" unfolding vec_last_def using i_n index_vec by blast also have "... = row ?tb ((dim_row ?tb) - n + i) \ v" unfolding index_vec by (simp add: assms(2) i_n two_block_non_interfering_dims(1)) also have "... = row B i \ vec_last v (dim_vec (row B i))" proof - have "dim_vec (0\<^sub>v (dim_col A) @\<^sub>v row B i) = dim_vec v" by (simp add: \dim_col A + dim_col B = dim_vec v\) then show ?thesis using dim_a assms(1) in3' two_block_non_interfering_row_comp2 append_split_vec_distrib_scalar_prod[of "0\<^sub>v (dim_col A)" "row B i" v] by (metis add.commute add.right_neutral diff_add_inverse in3 index_zero_vec(2) scalar_prod_left_zero vec_first_carrier) qed also have "... = row B i \ vec_last v (dim_col B)" by simp thus "vec_last (two_block_non_interfering A B *\<^sub>v v) n $ i = (B *\<^sub>v vec_last v (dim_col B)) $ i" using assms(2) calculation i_n by auto qed (simp add: assms(2)) lemma two_block_non_interfering_mult_decomposition: assumes "dim_col A + dim_col B = dim_vec v" shows "two_block_non_interfering A B *\<^sub>v v = A *\<^sub>v vec_first v (dim_col A) @\<^sub>v B *\<^sub>v vec_last v (dim_col B)" proof - let ?tb = "two_block_non_interfering A B" from first_vec_two_block_non_inter_is_first_vec[of A B v "dim_row A", OF assms] have "vec_first (?tb *\<^sub>v v) (dim_row A) = A *\<^sub>v vec_first v (dim_col A)" by blast moreover from last_vec_two_block_non_inter_is_last_vec[of A B v "dim_row B", OF assms] have "vec_last (?tb *\<^sub>v v) (dim_row B) = B *\<^sub>v vec_last v (dim_col B)" by blast ultimately show ?thesis using vec_first_last_append[of "?tb *\<^sub>v v" "(dim_row A)" "(dim_row B)"] dim_mult_mat_vec[of "?tb" v] two_block_non_interfering_dims(1)[of A B] by (metis carrier_vec_dim_vec) qed (* A \ b A = c *) fun mat_leqb_eqc where "mat_leqb_eqc A b c = (let lst = matrix_to_lpolies (two_block_non_interfering A A\<^sup>T) in [LEQ (lst!i) (b$i) . i <- [0..vc)$i) . i <- [dim_vec b ..< dim_vec (b@\<^sub>vc)]])" lemma mat_leqb_eqc_for_LEQ: assumes "i < dim_vec b" assumes "i < dim_row A" shows "(mat_leqb_eqc A b c)!i = LEQ ((matrix_to_lpolies A)!i) (b$i)" proof - define lst where lst: "lst = (mat_leqb_eqc A b c)" define l where l: "l = matrix_to_lpolies (two_block_non_interfering A A\<^sup>T)" have ileqA: "i < dim_row A" using assms by auto have "l!i = vec_to_lpoly ((row A i)@\<^sub>v 0\<^sub>v (dim_row A))" unfolding l using two_block_non_interfering_row_comp1[of i A "A\<^sup>T", OF ileqA] by (metis ileqA lpoly_of_v_equals_v_append0 matrix_to_lp_vec_to_lpoly_row trans_less_add1 two_block_non_interfering_dims(1)) then have leq: "l!i = (matrix_to_lpolies A)!i" using lpoly_of_v_equals_v_append0[of "row A i" "(dim_row A)"] l by (simp add: ileqA) have *: "lst = [LEQ (l!i) (b$i) . i <- [0..vc)$i) . i <- [dim_vec b ..< dim_vec (b@\<^sub>vc)] ]" unfolding l lst by (metis mat_leqb_eqc.simps) have "([LEQ (l!i) (b$i). i <- [0..vc)$i). i <- [dim_vec b ..< dim_vec (b@\<^sub>vc)]]) ! i = [LEQ (l!i) (b$i). i <- [0.. i" and "i < dim_vec (b@\<^sub>vc)" assumes "dim_row A = dim_vec b" and "dim_col A \ dim_vec c" shows "(mat_leqb_eqc A b c)!i = EQ (vec_to_lpoly (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T (i-dim_vec b))) (c$(i-dim_vec b))" proof - define lst where lst: "lst = (mat_leqb_eqc A b c)" define l where l: "l = matrix_to_lpolies (two_block_non_interfering A A\<^sup>T)" have i_s: "i < dim_row (two_block_non_interfering A A\<^sup>T)" using assms by (simp add: assms(2) assms(3) two_block_non_interfering_dims(1)) have l':"l!i = vec_to_lpoly ((0\<^sub>v (dim_col A)) @\<^sub>v (row A\<^sup>T (i - dim_vec b)))" using l two_block_non_interfering_row_comp2[of i A "A\<^sup>T", OF i_s] assms(1) assms(3) i_s matrix_to_lp_vec_to_lpoly_row by presburger have "([LEQ (l!i) (b$i) . i <- [0..vc)$i) . i <- [dim_vec b ..< dim_vec (b@\<^sub>vc)]])!i = [EQ (l!i) ((b@\<^sub>vc)$i) . i <- [dim_vec b..< dim_vec (b@\<^sub>vc)]] ! (i - dim_vec b)" by (simp add: assms(1) leD nth_append) also have "... = EQ (l!i) ((b@\<^sub>vc)$i)" using assms(1) assms(2) by auto also have "... = EQ (l!i) (c$(i-dim_vec b))" using assms(1) assms(2) by auto then show ?thesis using mat_leqb_eqc.simps by (metis (full_types) calculation l l') qed lemma mat_leqb_eqc_satisfies1: assumes "x \\<^sub>c\<^sub>s set (mat_leqb_eqc A b c)" assumes "i < dim_vec b" and "i < dim_row A" shows "(matrix_to_lpolies A!i) \x\ \ b$i" proof - have e_e: "LEQ (matrix_to_lpolies A ! i) (b$i) \ set (mat_leqb_eqc A b c)" using mat_leqb_eqc_for_LEQ[of i b A c, OF assms(2) assms(3)] nth_mem[of i "matrix_to_lpolies A"] mat_leqb_eqc.simps by (metis (no_types, lifting) assms(2) diff_zero in_set_conv_nth length_append length_map length_upt trans_less_add1) then have "x \\<^sub>c LEQ ((matrix_to_lpolies A)!i) (b$i)" using assms by blast then show ?thesis using satisfies_constraint.simps by auto qed lemma mat_leqb_eqc_satisfies2: assumes "x \\<^sub>c\<^sub>s set (mat_leqb_eqc A b c)" assumes "dim_vec b \ i" and "i < dim_vec (b@\<^sub>vc)" and "dim_row A = dim_vec b" and "dim_vec c \ dim_col A" shows "(matrix_to_lpolies (two_block_non_interfering A A\<^sup>T) ! i) \x\ = (b @\<^sub>v c) $ i" proof - have e_e: "EQ (vec_to_lpoly (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T (i - dim_vec b))) (c $ (i - dim_vec b)) \ set (mat_leqb_eqc A b c)" using assms(2) mat_leqb_eqc.simps[of A b c] nth_mem[of i "(mat_leqb_eqc A b c)"] using mat_leqb_eqc_for_EQ[of b i c A, OF assms(2) assms(3) assms(4) assms(5)] by (metis (mono_tags, lifting) add_diff_cancel_left' assms(3) diff_zero index_append_vec(2) length_append length_map length_upt) hence sateq: "x \\<^sub>c EQ (vec_to_lpoly (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T (i - dim_vec b))) (c $ (i - dim_vec b))" using assms(1) by blast have *: "i < dim_row (two_block_non_interfering A A\<^sup>T)" by (metis assms(3) assms(4) assms(5) dual_order.order_iff_strict dual_order.strict_trans index_append_vec(2) index_transpose_mat(2) nat_add_left_cancel_less two_block_non_interfering_dims(1)) have **: "dim_row A \ i" by (simp add: assms(2) assms(4)) then have "x \\<^sub>c EQ ((matrix_to_lpolies (two_block_non_interfering A A\<^sup>T))!i) ((b@\<^sub>vc)$i)" using two_block_non_interfering_row_comp2[of i A "A\<^sup>T", OF * **] by (metis "*" sateq assms(3) assms(4) index_append_vec(1) index_append_vec(2) leD matrix_to_lp_vec_to_lpoly_row) then show ?thesis using satisfies_constraint.simps(5) by simp qed lemma mat_leqb_eqc_simplex_satisfies2: assumes "simplex (mat_leqb_eqc A b c) = Sat x" assumes "dim_vec b \ i" and "i < dim_vec (b@\<^sub>vc)" and "dim_row A = dim_vec b" and "dim_vec c \ dim_col A" shows "(matrix_to_lpolies (two_block_non_interfering A A\<^sup>T) ! i) \\x\\ = (b @\<^sub>v c) $ i" using mat_leqb_eqc_satisfies2 assms(1) assms(2) assms(3) assms(4) assms(5) simplex(3) by blast fun index_geq_n where "index_geq_n i n = GEQ (lp_monom 1 i) n" lemma index_geq_n_simplex: assumes "\x\ \\<^sub>c (index_geq_n i n)" shows "\x\ i \ n" using assms by simp (* In the variables x_i to x_i+(length v) we synthesise a vector that is pointwise greater than v *) fun from_index_geq0_vector where "from_index_geq0_vector i v = [GEQ (lp_monom 1 (i+j)) (v$j) . j <-[0..\<^sub>c\<^sub>s set (from_index_geq0_vector i v)" "j < dim_vec v" shows "x (i + j) \ v$j" proof - have "GEQ (lp_monom 1 (i+j)) (v$j)\ set (from_index_geq0_vector i v)" by (simp add: assms(2)) moreover have "x \\<^sub>c GEQ (lp_monom 1 (i+j)) (v$j)" using calculation(1) assms by force ultimately show ?thesis by simp qed lemma from_index_geq0_vector_simplex2: assumes "\x\ \\<^sub>c\<^sub>s set (from_index_geq0_vector i v)" assumes "i \ j" and "j < (dim_vec v) + i" shows "\x\ j \ v$(j - i)" by (metis assms(1) assms(2) assms(3) from_index_geq_vector_simplex le_add_diff_inverse less_diff_conv2) (* [c1, ... cm, 01, ... 0n] * X \ [01, ... 0m, b1,...,bn] * X *) -fun x_times_c_geq_y_times_b where - "x_times_c_geq_y_times_b c b = GEQPP (vec_to_lpoly (c @\<^sub>v 0\<^sub>v (dim_vec b))) - (vec_to_lpoly (0\<^sub>v (dim_vec c) @\<^sub>v b))" +definition x_times_c_geq_y_times_b where + "x_times_c_geq_y_times_b c b = GEQ ( + vec_to_lpoly (c @\<^sub>v 0\<^sub>v (dim_vec b)) - vec_to_lpoly (0\<^sub>v (dim_vec c) @\<^sub>v b)) 0" lemma x_times_c_geq_y_times_b_correct: assumes "simplex [x_times_c_geq_y_times_b c b] = Sat x" shows "((vec_to_lpoly (c @\<^sub>v 0\<^sub>v (dim_vec b))) \ \x\ \) \ ((vec_to_lpoly (0\<^sub>v (dim_vec c) @\<^sub>v b)) \ \x\ \)" - using assms simplex(3) by fastforce + using simplex(3)[OF assms] unfolding x_times_c_geq_y_times_b_def + by (simp add: valuate_minus) (* Splitting an assignment into two vectors *) (* The first [0...(i-1)] elements and [i...j] elements *) definition split_i_j_x where "split_i_j_x i j x = (vec i \x\, vec (j - i) (\y. \x\ (y+i)))" abbreviation split_n_m_x where "split_n_m_x n m x \ split_i_j_x n (n+m) x" lemma split_vec_dims: assumes "split_i_j_x i j x = (a ,b)" shows "dim_vec a = i" "dim_vec b = (j - i)" using assms(1) unfolding split_i_j_x_def by auto+ lemma split_n_m_x_abbrev_dims: assumes "split_n_m_x n m x = (a, b)" shows "dim_vec a = n" "dim_vec b = m" using split_vec_dims using assms apply blast using assms split_vec_dims(2) by fastforce lemma split_access_fst_1: assumes "k < i" assumes "split_i_j_x i j x = (a, b)" shows "a $ k = \x\ k" by (metis Pair_inject assms(1) assms(2) index_vec split_i_j_x_def) lemma split_access_snd_1: assumes "i \ k" and "k < j" assumes "split_i_j_x i j x = (a, b)" shows "b $ (k - i) = \x\ k" proof - have "vec (j - i) (\n. \x\ (n + i)) = b" by (metis (no_types) assms(3) prod.sel(2) split_i_j_x_def) then show ?thesis using assms(1) assms(2) by fastforce qed lemma split_access_fst_2: assumes "(x, y) = split_i_j_x i j Z" assumes "k < dim_vec x" shows "x$k = \Z\ k" by (metis assms(1) assms(2) split_access_fst_1 split_vec_dims(1)) lemma split_access_snd_2: assumes "(x, y) = split_i_j_x i j Z" assumes "k < dim_vec y" shows "y$k = \Z\ (k+dim_vec x)" using assms split_i_j_x_def[of i j Z] by auto lemma from_index_geq0_vector_split_snd: assumes "\X\ \\<^sub>c\<^sub>s set (from_index_geq0_vector d v)" assumes "(x, y) = split_n_m_x d m X" shows "\i. i < dim_vec v \ i < m \ y$i \ v$i" using assms unfolding split_i_j_x_def using from_index_geq_vector_simplex[of d v "\X\" _] index_vec by (simp add: add.commute) lemma split_coeff_vec_index_sum: assumes "(x,y) = split_i_j_x (dim_vec (lpoly_to_vec v)) l X" shows "(\i = 0..X\ i) = (\i = 0..X\", symmetric] have "(\i = 0..X\ i) = (\i = 0.. x = (v \\X\\)" proof - from valuate_with_dim_poly[of v "\X\", symmetric] have 1: "(v \\X\\) = (\i = 0..X\ i)" by simp have "(\i = 0..X\ i) = (\i = 0.. x" unfolding scalar_prod_def by blast finally show ?thesis by (metis (no_types, lifting) "1" dim_poly_dim_vec_equiv lin_poly_to_vec_coeff_access split_vec_dims(1) sum.ivl_cong assms) qed definition mat_times_vec_leq ("[_*\<^sub>v_]\_" [1000,1000,100]) where "[A *\<^sub>v x]\b \ (\i < dim_vec b. (A *\<^sub>v x)$i \ b$i) \ (dim_row A = dim_vec b) \ (dim_col A = dim_vec x)" definition vec_times_mat_eq ("[_\<^sub>v*_]=_" [1000,1000,100]) where "[y \<^sub>v* A]=c \ (\i < dim_vec c. (A\<^sup>T *\<^sub>v y)$i = c$i) \ (dim_col A\<^sup>T = dim_vec y) \ (dim_row A\<^sup>T = dim_vec c)" definition vec_times_mat_leq ("[_\<^sub>v*_]\_" [1000,1000,100]) where "[y \<^sub>v* A]\c \ (\i < dim_vec c. (A\<^sup>T *\<^sub>v y)$i \ c$i) \ (dim_col A\<^sup>T = dim_vec y) \ (dim_row A\<^sup>T = dim_vec c)" lemma mat_times_vec_leqI[intro]: assumes "dim_row A = dim_vec b" assumes "dim_col A = dim_vec x" assumes "\i. i < dim_vec b \ (A *\<^sub>v x)$i \ b$i" shows "[A *\<^sub>v x]\b" unfolding mat_times_vec_leq_def using assms by auto lemma mat_times_vec_leqD[dest]: assumes "[A *\<^sub>v x]\b" shows "dim_row A = dim_vec b" "dim_col A = dim_vec x" "\i. i < dim_vec b \ (A *\<^sub>v x)$i \ b$i" using assms mat_times_vec_leq_def by blast+ lemma vec_times_mat_eqD[dest]: assumes "[y \<^sub>v* A]=c" shows "(\i < dim_vec c. (A\<^sup>T *\<^sub>v y)$i = c$i)" "(dim_col A\<^sup>T = dim_vec y)" "(dim_row A\<^sup>T = dim_vec c)" using assms vec_times_mat_eq_def by blast+ lemma vec_times_mat_leqD[dest]: assumes "[y \<^sub>v* A]\c" shows "(\i < dim_vec c. (A\<^sup>T *\<^sub>v y)$i \ c$i)" "(dim_col A\<^sup>T = dim_vec y)" "(dim_row A\<^sup>T = dim_vec c)" using assms vec_times_mat_leq_def by blast+ lemma mat_times_vec_eqI[intro]: assumes "dim_col A\<^sup>T = dim_vec x" assumes "dim_row A\<^sup>T = dim_vec c" assumes "\i. i < dim_vec c \ (A\<^sup>T *\<^sub>v x)$i = c$i" shows "[x \<^sub>v* A]=c" unfolding vec_times_mat_eq_def using assms by blast lemma mat_leqb_eqc_split_correct1: assumes "dim_vec b = dim_row A" assumes "\X\ \\<^sub>c\<^sub>s set (mat_leqb_eqc A b c)" assumes "(x,y) = split_i_j_x (dim_col A) l X" shows "[A *\<^sub>v x]\b" proof (standard, goal_cases) case 1 then show ?case using assms(1)[symmetric] . case 2 then show ?case using assms(3) unfolding split_i_j_x_def using split_vec_dims[of 0 "dim_col A" X x y] by auto case (3 i) with mat_leqb_eqc_satisfies1[of A b c "\X\" i] have m: "(matrix_to_lpolies A ! i) \ \X\ \ \ b $ i" using assms(1) assms(2) by linarith have leq: "dim_poly (vec_to_lpoly (row A i)) \ dim_col A" using vec_to_poly_dim_less[of "row A i"] by simp have i: "i < dim_row A" using "3" assms(1) by linarith from two_block_non_interfering_row_comp1[of i A "A\<^sup>T"] have "row (two_block_non_interfering A A\<^sup>T) i = row A i @\<^sub>v 0\<^sub>v (dim_col A\<^sup>T)" using "3" assms(1) by linarith have "(vec_to_lpoly (row A i @\<^sub>v 0\<^sub>v (dim_col A\<^sup>T))) \\X\\ = ((vec_to_lpoly (row A i)) \\X\\)" using lpoly_of_v_equals_v_append0 by auto also have "... = (\a = 0..X\ a)" using valuate_with_dim_poly[of "vec_to_lpoly (row A i)" "\X\"] by blast also have "... = (\a = 0..X\ a)" using split_coeff_vec_index_sum[of x y] sum_dim_vec_equals_sum_dim_poly[of "row A i" "\X\"] by auto also have "... = row A i \ x" unfolding scalar_prod_def using \dim_col A = dim_vec x\ i assms(3) using matrix_to_lpolies_coeff_access[of i A] matrix_to_lp_vec_to_lpoly_row[of i A] split_access_fst_1[of _ "(dim_col A)" l X x y] by fastforce finally show ?case using m i lpoly_of_v_equals_v_append0 by auto qed lemma mat_leqb_eqc_split_simplex_correct1: assumes "dim_vec b = dim_row A" assumes "simplex (mat_leqb_eqc A b c) = Sat X" assumes "(x,y) = split_i_j_x (dim_col A) l X" shows "[A *\<^sub>v x]\b" using mat_leqb_eqc_split_correct1[of b A c X x y] assms(1) assms(2) assms(3) simplex(3) by blast lemma sat_mono: assumes "set A \ set B" shows "\X\ \\<^sub>c\<^sub>s set B \ \X\ \\<^sub>c\<^sub>s set A" using assms(1) assms by blast lemma mat_leqb_eqc_split_subset_correct1: assumes "dim_vec b = dim_row A" assumes "set (mat_leqb_eqc A b c) \ set S" assumes "simplex S = Sat X" assumes "(x,y) = split_i_j_x (dim_col A) l X" shows "[A *\<^sub>v x]\b" using sat_mono assms(1) assms(2) assms(3) assms(4) mat_leqb_eqc_split_correct1 simplex(3) by blast lemma mat_leqb_eqc_split_correct2: assumes "dim_vec c = dim_row A\<^sup>T" assumes "dim_vec b = dim_col A\<^sup>T" assumes "\X\ \\<^sub>c\<^sub>s set (mat_leqb_eqc A b c)" assumes "(x, y) = split_n_m_x (dim_row A\<^sup>T) (dim_col A\<^sup>T) X" shows "[y \<^sub>v* A]=c" proof (standard, goal_cases) case 1 then show ?case using assms split_n_m_x_abbrev_dims(2)[OF assms(4)[symmetric]] by linarith case 2 then show ?case using assms(1)[symmetric] . case (3 i) define lst where lst: "lst = matrix_to_lpolies (two_block_non_interfering A A\<^sup>T)" define db where db: "db = dim_vec b" define dc where dc: "dc = dim_vec c" have cA: "dim_vec c \ dim_col A" by (simp add: assms(1)) have dbi_dim: "db+i < dim_vec (b @\<^sub>v c)" by (simp add: "3" db) have *: "dim_vec b \ db+i" by (simp add: db) have "([LEQ (lst!i) (b$i) . i <- [0..vc)$i) . i <- [dim_vec b ..< dim_vec (b@\<^sub>vc)]]) ! (db + i) = EQ (lst!(db+i)) ((b@\<^sub>vc)$(db+i))" using mat_leqb_eqc_for_EQ[of b "db+i" c A] nth_append[of "[LEQ (lst!i) (b$i) . i <- [0..vc)$i) . i <- [dim_vec b ..< dim_vec (b@\<^sub>vc)]]"] by (simp add: "3" db) have rowA: "dim_vec b = dim_row A" using assms index_transpose_mat(3)[of A] by linarith have "\X\ \\<^sub>c EQ (lst!(db+i)) (c$i)" proof - have "db + i - dim_vec b = i" using db diff_add_inverse by blast then have "(lst ! (db + i)) \ \X\ \ = c $ i" by (metis dbi_dim rowA * cA assms(3) index_append_vec(1) index_append_vec(2) leD lst mat_leqb_eqc_satisfies2) then show ?thesis using satisfies_constraint.simps(5)[of "\X\" "(lst ! (db + i))" "(c $ i)"] by simp qed then have sat: "(lst!(db+i)) \\X\\ = c$i" by simp define V where V: "V = vec (db+dc) (\i. \X\ i)" have vdim: "dim_vec V = dim_vec (b@\<^sub>vc)" using V db dc by simp have *: "db + i < dim_row (two_block_non_interfering A A\<^sup>T)" by (metis dbi_dim assms(1) index_append_vec(2) rowA two_block_non_interfering_dims(1)) have **: "dim_row A \ db + i" by (simp add: assms(2) db) from two_block_non_interfering_row_comp2[of "db+i" A "A\<^sup>T", OF * **] have eql: "row (two_block_non_interfering A A\<^sup>T) (db + i) = 0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T i" by (simp add: assms(2) db) with matrix_to_lp_vec_to_lpoly_row[of i "A\<^sup>T"] have eqv: "lst!(db+i) = vec_to_lpoly (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T i)" using "*" lst matrix_to_lp_vec_to_lpoly_row by presburger then have "\jj\db+dc. Abstract_Linear_Poly.coeff (lst!(db+i)) j = 0" by (metis (mono_tags, lifting) eqv index_transpose_mat(3) index_zero_vec(2) leD add.commute assms(1) assms(2) coeff_nonzero_dim_vec_non_zero(2) index_append_vec(2) index_row(2) index_transpose_mat(2) db dc) moreover have "vars (lst!(db+i)) \ {dim_col A..\X\\ = (\j\{dim_col A..X\ j)" using eval_poly_with_sum_superset[of "{dim_col A..X\"] by blast also have "... = (\j\{dim_col A..j\{dim_col A..v (dim_col A) @\<^sub>v row A\<^sup>T i)$j * V$j)" proof - have "\j\{dim_col A..v (dim_col A) @\<^sub>v row A\<^sup>T i)$j" by (metis \V \ vec (db + dc) \X\\ vdim assms(1) assms(2) index_transpose_mat(2) atLeastLessThan_iff dim_vec eql eqv index_append_vec(2) index_row(2) vec_to_lin_poly_coeff_access semiring_normalization_rules(24) two_block_non_interfering_dims(2)) then show ?thesis by (metis (mono_tags, lifting) sum.cong) qed also have "... = (\j\{0..v (dim_col A) @\<^sub>v row A\<^sup>T i)$j * V$j) + (\j\{dim_col A..v (dim_col A) @\<^sub>v row A\<^sup>T i)$j * V$j)" by (metis (no_types, lifting) add_cancel_left_left atLeastLessThan_iff mult_eq_0_iff class_semiring.add.finprod_all1 index_append_vec(1) index_zero_vec(1) index_zero_vec(2) trans_less_add1) also have "... = (\j\{0..v (dim_col A) @\<^sub>v row A\<^sup>T i)$j * V$j)" by (metis (no_types, lifting) add.commute assms(1) dc index_transpose_mat(2) le_add1 le_add_same_cancel1 sum.atLeastLessThan_concat) also have "... = (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T i) \ V" unfolding scalar_prod_def by (simp add: V) also have "... = 0\<^sub>v (dim_col A) \ vec_first V (dim_vec (0\<^sub>v (dim_col A))) + row A\<^sup>T i \ vec_last V (dim_vec (row A\<^sup>T i))" using append_split_vec_distrib_scalar_prod[of "0\<^sub>v (dim_col A)" "row A\<^sup>T i" V] by (metis (no_types, lifting) \dim_vec V = dim_vec (b @\<^sub>v c)\ add.commute assms(1) assms(2) index_append_vec(2) index_row(2) index_transpose_mat(2) index_transpose_mat(3) index_zero_vec(2)) also have "0\<^sub>v (dim_col A) \ vec_first V (dim_vec (0\<^sub>v (dim_col A))) + row A\<^sup>T i \ vec_last V (dim_vec (row A\<^sup>T i)) = (row A\<^sup>T i) \ y" proof - have "vec_last V (dim_vec (row A\<^sup>T i)) = y" proof (standard, goal_cases) case (1 i) then show ?case proof - have f1: "dim_col A\<^sup>T = db" by (simp add: assms(2) db) then have "\v va. vec db (\n. \X\ (n + dc)) = v \ (x, y) \ (va, v)" by (metis Pair_inject add_diff_cancel_left' assms(1) assms(4) dc split_i_j_x_def) then show ?case unfolding V vec_last_def using split_access_fst_1[of "(dim_row A\<^sup>T)" i "(dim_col A\<^sup>T)" X x y] by (metis "1" add.commute add_diff_cancel_left' add_less_cancel_left dim_vec f1 index_row(2) index_vec) qed next case 2 then show ?case using \dim_col A\<^sup>T = dim_vec y\ by auto qed then show ?thesis by (simp add: assms(1)) qed then show ?case unfolding mult_mat_vec_def by (metis "3" assms(1) calculation index_vec sat) qed lemma mat_leqb_eqc_split_simplex_correct2: assumes "dim_vec c = dim_row A\<^sup>T" assumes "dim_vec b = dim_col A\<^sup>T" assumes "simplex (mat_leqb_eqc A b c) = Sat X" assumes "(x, y) = split_n_m_x (dim_row A\<^sup>T) (dim_col A\<^sup>T) X" shows "[y \<^sub>v* A]=c" using assms(1) assms(2) assms(3) assms(4) mat_leqb_eqc_split_correct2 simplex(3) by blast lemma mat_leqb_eqc_correct: assumes "dim_vec c = dim_row A\<^sup>T" and "dim_vec b = dim_col A\<^sup>T" assumes "simplex (mat_leqb_eqc A b c) = Sat X" assumes "(x, y) = split_n_m_x (dim_row A\<^sup>T) (dim_col A\<^sup>T) X" shows "[y \<^sub>v* A]=c" "[A *\<^sub>v x]\b" using mat_leqb_eqc_split_simplex_correct1[of b A c X x y] using assms(1) assms(2) assms(3) assms(4) mat_leqb_eqc_split_simplex_correct2 apply blast using mat_leqb_eqc_split_correct2[of b A c X x y] by (metis (no_types) Matrix.transpose_transpose assms(2) assms(3) assms(4) index_transpose_mat(3) mat_leqb_eqc_split_simplex_correct1[of b A c X x y]) lemma eval_lpoly_eq_dot_prod_split1: assumes "(x, y) = split_n_m_x (dim_vec c) (dim_vec b) X" shows"(vec_to_lpoly c) \\X\\ = c \ x" proof - have *: "(vec_to_lpoly c) \\X\\ = (\i\vars (vec_to_lpoly c). Abstract_Linear_Poly.coeff (vec_to_lpoly c) i * \X\ i)" using linear_poly_sum sum.cong eval_poly_with_sum by auto also have "... = (\i\{0..X\ i)" using vars_subset_dim_vec_to_lpoly_dim[of c] linear_poly_sum[of "vec_to_lpoly c" "\X\"] eval_poly_with_sum_superset[of "{0..X\"] by auto also have "... = (\i\{0.. x" unfolding scalar_prod_def using split_vec_dims(1)[of "dim_vec c" "(dim_vec c) + (dim_vec b)" X x y] assms by auto finally show ?thesis . qed lemma eval_lpoly_eq_dot_prod_split2: assumes "(x, y) = split_n_m_x (dim_vec b) (dim_vec c) X" shows"(vec_to_lpoly (0\<^sub>v (dim_vec b) @\<^sub>v c)) \\X\\ = c \ y" proof - let ?p = "(vec_to_lpoly ((0\<^sub>v (dim_vec b) @\<^sub>v c)))" let ?v0 = "(0\<^sub>v (dim_vec b) @\<^sub>v c)" have *: "\i\X\\ = (\i\vars ?p. Abstract_Linear_Poly.coeff ?p i * \X\ i)" using eval_poly_with_sum by blast also have "... = (\i\{0..X\ i)" using eval_poly_with_sum_superset[of "{0..X\"] calculation vars_subset_dim_vec_to_lpoly_dim[of ?v0] by force also have "... = (\i\{0..X\ i) + (\i\{(dim_vec b)..X\ i)" by (simp add: sum.atLeastLessThan_concat) also have "... = (\i\{(dim_vec b)..X\ i)" using * by simp also have "... = (\i\{(dim_vec b)..X\ i)" using vec_to_lin_poly_coeff_access by auto also have "... = (\i\{0..X\ (i+dim_vec b))" using index_zero_vec(2)[of "dim_vec b"] index_append_vec(2)[of "0\<^sub>v (dim_vec b)" c] ** * sum.shift_bounds_nat_ivl[of "(\i. ?v0$i * \X\ i)" 0 "dim_vec b" "dim_vec c"] by (simp add: add.commute) also have "... = (\i\{0..X\ (i+dim_vec b))" by auto also have "... = (\i\{0.. y" by (metis assms scalar_prod_def split_n_m_x_abbrev_dims(2)) finally show ?thesis . qed lemma x_times_c_geq_y_times_b_split_dotP: assumes "\X\ \\<^sub>c x_times_c_geq_y_times_b c b" assumes "(x, y) = split_n_m_x (dim_vec c) (dim_vec b) X" shows "c \ x \ b \ y" using assms lpoly_of_v_equals_v_append0 eval_lpoly_eq_dot_prod_split2[of x y c b X] - eval_lpoly_eq_dot_prod_split1[of x y c b X] by auto + eval_lpoly_eq_dot_prod_split1[of x y c b X] + by (auto simp: x_times_c_geq_y_times_b_def valuate_minus) lemma mult_right_leq: fixes A :: "('a::{comm_semiring_1,ordered_semiring}) mat" assumes "dim_vec y = dim_vec b" and "\i < dim_vec y. y$i \ 0" and "[A *\<^sub>v x]\ b" shows "(A *\<^sub>v x) \ y \ b \ y" proof - have "(\nv x) $ n * y $ n) \ (\nv* A]=c" shows "(A\<^sup>T *\<^sub>v y) \ x = c \ x" unfolding scalar_prod_def using atLeastLessThan_iff[of _ 0 "dim_vec x"] vec_times_mat_eq_def[of y A c] sum.cong[of _ _ "\i. (A\<^sup>T *\<^sub>v y) $ i * x $ i" "\i. c $ i * x $ i"] by (metis (mono_tags, lifting) assms(1) assms(2)) lemma soundness_mat_x_leq: assumes "dim_row A = dim_vec b" assumes "simplex (mat_x_leq_vec A b) = Sat X" shows "\x. [A *\<^sub>v x]\b" proof define x where x: "x = fst (split_n_m_x (dim_col A) (dim_row A) X)" have *: "dim_vec x = dim_col A" by (simp add: split_i_j_x_def x) have "\iv x) $ i \ b $ i" proof (standard, standard) fix i assume "i < dim_vec b" have "row A i \ x \ b$i" using mat_x_leq_vec_sol[of A b X i] by (metis \i < dim_vec b\ assms(1) assms(2) eval_lpoly_eq_dot_prod_split1 fst_conv index_row(2) matrix_to_lp_vec_to_lpoly_row simplex(3) split_i_j_x_def x) then show "(A *\<^sub>v x) $ i \ b $ i" by (simp add: \i < dim_vec b\ assms(1)) qed then show "[A *\<^sub>v x]\b" using mat_times_vec_leqI[of A b x, OF assms(1) *[symmetric]] by auto qed lemma completeness_mat_x_leq: assumes "\x. [A *\<^sub>v x]\b" shows "\X. simplex (mat_x_leq_vec A b) = Sat X" proof (rule ccontr) assume 1: "\X. simplex (mat_x_leq_vec A b) = Inr X" have *: "\v. v \\<^sub>c\<^sub>s set (mat_x_leq_vec A b)" using simplex(1)[of "mat_x_leq_vec A b"] using "1" sum.exhaust_sel by blast then have "dim_vec b = dim_row A" using assms mat_times_vec_leqD(1)[of A _ b] by auto then obtain x where x: "[A *\<^sub>v x]\b" using assms by blast have x_A: "dim_vec x = dim_col A" using x by auto define v where v: "v = (\i. (if i < dim_vec x then x$i else 0))" have v_d: "\i < dim_vec x. x$i = v i" by (simp add: v) have "v \\<^sub>c\<^sub>s set (mat_x_leq_vec A b)" proof fix c assume "c \ set (mat_x_leq_vec A b)" then obtain i where i: "c = LEQ (matrix_to_lpolies A!i) (b$i)" "i < dim_vec b" by auto let ?p = "matrix_to_lpolies A!i" have 2: "?p\ v \ = (row A i) \ x" using matrix_to_lpolies_lambda_valuate_scalarP[of i A x] v by (metis \dim_vec b = dim_row A\ i(2) x_A) also have "... \ b$i" by (metis i(2) index_mult_mat_vec mat_times_vec_leq_def x) finally show "v \\<^sub>c c" using i(1) satisfies_constraint.simps(3)[of v "(matrix_to_lpolies A ! i)" "b $ i"] 2 \row A i \ x \ b $ i\ by simp qed then show False using * by auto qed lemma soundness_mat_x_eq_vec: assumes "dim_row A\<^sup>T = dim_vec c" assumes "simplex (x_mat_eq_vec c A\<^sup>T) = Sat X" shows "\x. [x \<^sub>v* A]=c" proof define x where x: "x = fst (split_n_m_x (dim_col A\<^sup>T) (dim_row A\<^sup>T) X)" have "dim_vec x = dim_col A\<^sup>T" unfolding split_i_j_x_def using split_vec_dims(1)[of "(dim_col A\<^sup>T)" _ X x] fst_conv[of x] by (simp add: split_i_j_x_def x) have "\i < dim_vec c. (A\<^sup>T *\<^sub>v x)$i = c$i" proof (standard, standard) fix i assume a: "i < dim_vec c" have *: "\X\ \\<^sub>c\<^sub>s set (x_mat_eq_vec c A\<^sup>T)" using assms(2) simplex(3) by blast then have "row A\<^sup>T i \ x = c$i" using x_mat_eq_vec_sol[of c "A\<^sup>T" "\X\" i, OF * a] eval_lpoly_eq_dot_prod_split1 fstI by (metis a assms(1) index_row(2) matrix_to_lpolies_vec_of_row split_i_j_x_def x) then show "(A\<^sup>T *\<^sub>v x) $ i = c $ i" unfolding mult_mat_vec_def using a assms(1) by auto qed then show "[x \<^sub>v* A]=c" using mat_times_vec_eqI[of A x c, OF \dim_vec x = dim_col A\<^sup>T\[symmetric] assms(1)] by auto qed lemma completeness_mat_x_eq_vec: assumes "\x. [x \<^sub>v* A]=c" shows "\X. simplex (x_mat_eq_vec c A\<^sup>T) = Sat X" proof (rule ccontr) assume 1: "\X. simplex (x_mat_eq_vec c A\<^sup>T) = Inr X" then have *: "\v. v \\<^sub>c\<^sub>s set (x_mat_eq_vec c A\<^sup>T)" using simplex(1)[of "x_mat_eq_vec c A\<^sup>T"] using sum.exhaust_sel 1 by blast then have "dim_vec c = dim_col A" using assms by (metis index_transpose_mat(2) vec_times_mat_eqD(3)) obtain x where " [x \<^sub>v* A]=c" using assms by auto then have "dim_vec x = dim_col A\<^sup>T" using assms by (metis \[x \<^sub>v* A]=c\ vec_times_mat_eq_def) define v where v: "v = (\i. (if i < dim_vec x then x$i else 0))" have v_d: "\i < dim_vec x. x$i = v i" by (simp add: v) have "v \\<^sub>c\<^sub>s set (x_mat_eq_vec c A\<^sup>T)" proof fix a assume "a \ set (x_mat_eq_vec c A\<^sup>T)" then obtain i where i: "a = EQ (matrix_to_lpolies A\<^sup>T!i) (c$i)" "i < dim_vec c" by (metis (no_types, lifting) add_cancel_right_left diff_zero in_set_conv_nth length_map length_upt nth_map_upt x_mat_eq_vec.simps) let ?p = "matrix_to_lpolies A\<^sup>T!i" have 2: "?p\ v \ = (row A\<^sup>T i) \ x" using matrix_to_lpolies_lambda_valuate_scalarP[of i "A\<^sup>T" x] v by (metis \dim_vec c = dim_col A\ \dim_vec x = dim_col A\<^sup>T\ i(2) index_transpose_mat(2)) also have "... = c$i" by (metis \[x \<^sub>v* A]=c\ \dim_vec c = dim_col A\ i(2) index_mult_mat_vec index_transpose_mat(2) vec_times_mat_eqD(1)) finally show "v \\<^sub>c a" using i(1) satisfies_constraint.simps(5)[of v "(matrix_to_lpolies A\<^sup>T ! i)" "(c $ i)"] by simp qed then show False using "*" by blast qed lemma soundness_mat_leqb_eqc1: assumes "dim_row A = dim_vec b" assumes "simplex (mat_leqb_eqc A b c) = Sat X" shows "\x. [A *\<^sub>v x]\b" proof define x where x: "x = fst (split_n_m_x (dim_col A) (dim_row A) X)" have *: "dim_vec x = dim_col A" by (simp add: split_i_j_x_def x) have "\iv x) $ i \ b $ i" proof (standard, standard) fix i assume "i < dim_vec b" have "row A i \ x \ b$i" using mat_x_leq_vec_sol[of A b X i] by (metis \i < dim_vec b\ assms(1) assms(2) fst_conv split_i_j_x_def x index_mult_mat_vec mat_leqb_eqc_split_simplex_correct1 mat_times_vec_leqD(3)) then show "(A *\<^sub>v x) $ i \ b $ i" by (simp add: \i < dim_vec b\ assms(1)) qed then show "[A *\<^sub>v x]\b" using mat_times_vec_leqI[of A b x, OF assms(1) *[symmetric]] by auto qed lemma soundness_mat_leqb_eqc2: assumes "dim_row A\<^sup>T = dim_vec c" assumes "dim_col A\<^sup>T = dim_vec b" assumes "simplex (mat_leqb_eqc A b c) = Sat X" shows "\y. [y \<^sub>v* A]=c" proof (standard, intro mat_times_vec_eqI) define y where x: "y = snd (split_n_m_x (dim_col A) (dim_row A) X)" have *: "dim_vec y = dim_row A" by (simp add: split_i_j_x_def x) show "dim_col A\<^sup>T = dim_vec y" by (simp add: "*") show "dim_row A\<^sup>T = dim_vec c" using assms(1) by blast show "\i. i < dim_vec c \ (A\<^sup>T *\<^sub>v y) $ i = c $ i" proof - fix i assume a: "i < dim_vec c" have "[y \<^sub>v* A]=c" using mat_leqb_eqc_split_correct2[of c A b _ _ y, OF assms(1)[symmetric] assms(2)[symmetric]] by (metis Matrix.transpose_transpose assms(3) index_transpose_mat(2) simplex(3) snd_conv split_i_j_x_def x) then show "(A\<^sup>T *\<^sub>v y) $ i = c $ i" by (metis a vec_times_mat_eq_def) qed qed lemma completeness_mat_leqb_eqc: assumes "\x. [A *\<^sub>v x]\b" and "\y. [y \<^sub>v* A]=c" shows "\X. simplex (mat_leqb_eqc A b c) = Sat X" proof (rule ccontr) assume 1: "\X. simplex (mat_leqb_eqc A b c) = Sat X" have *: "\v. v \\<^sub>c\<^sub>s set (mat_leqb_eqc A b c)" using simplex(1)[of "mat_leqb_eqc A b c"] using "1" sum.exhaust_sel by blast then have "dim_vec b = dim_row A" using assms mat_times_vec_leqD(1)[of A _ b] by presburger then obtain x y where x: "[A *\<^sub>v x]\b" "[y \<^sub>v* A]=c" using assms by blast have x_A: "dim_vec x = dim_col A" using x by auto have yr: "dim_vec y = dim_row A" using vec_times_mat_eq_def x(2) by force define v where v: "v = (\i. (if i < dim_vec (x@\<^sub>vy) then (x@\<^sub>vy)$i else 0))" have v_d: "\i < dim_vec (x@\<^sub>vy). (x@\<^sub>vy)$i = v i" by (simp add: v) have i_in: "\i \ {0..< dim_vec y}. y$i = v (i+dim_vec x)" by (simp add: v) have "v \\<^sub>c\<^sub>s set (mat_leqb_eqc A b c)" proof fix e assume asm: "e \ set (mat_leqb_eqc A b c)" define lst where lst: "lst = matrix_to_lpolies (two_block_non_interfering A A\<^sup>T)" let ?L = "[LEQ (lst!i) (b$i) . i <- [0..vc)$i) . i <- [dim_vec b ..< dim_vec (b@\<^sub>vc)]]" have L: "mat_leqb_eqc A b c = ?L" by (metis (full_types) lst mat_leqb_eqc.simps) then obtain i where i: "e = ?L!i" "i \{0..vc)" using i(2) by auto consider (leqb) "i \ {0.. {dim_vec b..\<^sub>c e" proof (cases) case leqb have il: "i < dim_vec b" using atLeastLessThan_iff leqb by blast have iA: "i < dim_row A" using \dim_vec b = dim_row A\ \i < dim_vec b\ by linarith then have *: "e = LEQ (lst!i) (b$i)" by (simp add: i(1) nth_append il) then have "... = LEQ ((matrix_to_lpolies A)!i) (b$i)" using mat_leqb_eqc_for_LEQ[of i b A c, OF il \i < dim_row A\] L i(1) by simp then have eqmp: "lst!i = ((matrix_to_lpolies A)!i)" by blast have sset: "vars (lst!i) \ {0..i < dim_row A\ eqmp index_row(2) vars_subset_dim_vec_to_lpoly_dim x_A) have **: "((lst!i) \ v \) = ((vec_to_lpoly (row A i)) \ v \)" by (simp add: \i < dim_row A\ eqmp) also have "... = (\j\vars(lst!i). Abstract_Linear_Poly.coeff (lst!i) j * v j)" using ** eval_poly_with_sum by auto also have "... = (\j\{0..j\{0..j\{0.. x" unfolding scalar_prod_def by (simp) also have "... \ b$i" by (metis \i < dim_vec b\ index_mult_mat_vec mat_times_vec_leq_def x(1)) finally show ?thesis by (simp add: "*") next case eqc have igeq: "i \ dim_vec b" using atLeastLessThan_iff eqc by blast have *: "i < length ?L" using atLeastLessThan_iff eqc by blast have "e =?L!i" using L i(1) by presburger have "?L!i \ set [EQ (lst!i) ((b@\<^sub>vc)$i). i <- [dim_vec b..< dim_vec (b@\<^sub>vc)]]" using in_second_append_list length_map by (metis (no_types, lifting) igeq * length_upt minus_nat.diff_0) then have "?L!i = [EQ (lst!i) ((b@\<^sub>vc)$i). i <- [dim_vec b..< dim_vec (b@\<^sub>vc)]]!(i-dim_vec b)" by (metis (no_types, lifting) \dim_vec b \ i\ diff_zero leD length_map length_upt nth_append) then have "?L!i = EQ (lst!i) ((b@\<^sub>vc)$i)" using add_diff_inverse_nat diff_less_mono by (metis (no_types, lifting) \dim_vec b \ i\ * ldimbc leD nth_map_upt) then have e: "e = EQ (lst!i) ((b@\<^sub>vc)$i)" using i(1) by blast with mat_leqb_eqc_for_EQ[of b i c A, OF igeq] have lsta: "(lst!i) = (vec_to_lpoly (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T (i - dim_vec b)))" by (metis (no_types, lifting) \dim_vec b = dim_row A\ * ldimbc assms(2) igeq index_append_vec(2) lst matrix_to_lpolies_vec_of_row vec_times_mat_eq_def two_block_non_interfering_dims(1) two_block_non_interfering_row_comp2 ) let ?p = "(vec_to_lpoly (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T (i - dim_vec b)))" have "dim_poly ?p \ dim_col A + dim_row A" using dim_poly_of_append_vec[of "0\<^sub>v (dim_col A)" "row A\<^sup>T (i - dim_vec b)"] index_zero_vec(2)[of "dim_col A"] by (metis \dim_vec (0\<^sub>v (dim_col A)) = dim_col A\ index_row(2) index_transpose_mat(3)) have "\i < dim_col A. Abstract_Linear_Poly.coeff ?p i = 0" using vec_coeff_append1[of _ "0\<^sub>v (dim_col A)" "row A\<^sup>T (i - dim_vec b)"] by (metis atLeastLessThan_iff index_zero_vec(1) index_zero_vec(2) zero_le) then have "dim_vec (0\<^sub>v (dim_col A) @\<^sub>v row A\<^sup>T (i - dim_vec b)) = dim_col A + dim_row A" by (metis index_append_vec(2) index_row(2) index_transpose_mat(3) index_zero_vec(2)) then have allcr: "\j\{0..T (i - dim_vec b))$j" by (metis add_diff_cancel_right' atLeastLessThan_iff diff_add_inverse index_zero_vec(2) le_add_same_cancel2 less_diff_conv vec_coeff_append2) have vs: "vars ?p \ {dim_col A.. v \ = (\j\vars ?p. Abstract_Linear_Poly.coeff ?p j * v j)" using eval_poly_with_sum by blast also have "... = (\j\{dim_col A..j\{0..j. Abstract_Linear_Poly.coeff ?p j * v j" 0 "dim_col A" "dim_row A"] by (metis (no_types, lifting) add.commute add_cancel_left_left) also have "... = (\j\{0..j\{0..T (i - dim_vec b))$j * y$j)" using allcr by (metis (no_types, lifting) sum.cong) also have "... = (row A\<^sup>T (i - dim_vec b)) \ y" by (metis \dim_vec y = dim_row A\ scalar_prod_def) also have "... = (b@\<^sub>vc)$i" using vec_times_mat_eqD[OF x(2)] * igeq by auto finally show ?thesis using e lsta satisfies_constraint.simps(5)[of _ "(lst ! i)" "((b @\<^sub>v c) $ i)"] by simp qed qed then show False using * by blast qed lemma sound_and_compltete_mat_leqb_eqc [iff]: assumes "dim_row A\<^sup>T = dim_vec c" assumes "dim_col A\<^sup>T = dim_vec b" shows "(\x. [A *\<^sub>v x]\b) \ (\y. [y \<^sub>v* A]=c) \ (\X. simplex (mat_leqb_eqc A b c) = Sat X)" by (metis assms(1) assms(2) completeness_mat_leqb_eqc index_transpose_mat(3) soundness_mat_leqb_eqc1 soundness_mat_leqb_eqc2) section \ Translate Inequalities to Matrix Form \ (* We (obviously) cannot use strict inequalities hence we use the option type *) fun nonstrict_constr where "nonstrict_constr (LEQ p r) = True" | "nonstrict_constr (GEQ p r) = True" | "nonstrict_constr (EQ p r) = True" | - "nonstrict_constr (LEQPP p q) = True" | - "nonstrict_constr (GEQPP p q) = True" | - "nonstrict_constr (EQPP p q) = True" | "nonstrict_constr _ = False" abbreviation "nonstrict_constrs cs \ (\a \ set cs. nonstrict_constr a)" fun transf_constraint where "transf_constraint (LEQ p r) = [LEQ p r]" | "transf_constraint (GEQ p r) = [LEQ (-p) (-r)]" | "transf_constraint (EQ p r) = [LEQ p r, LEQ (-p) (-r)]" | - "transf_constraint (LEQPP p q) = [LEQ (p - q) 0]" | - "transf_constraint (GEQPP p q) = [LEQ (-(p - q)) 0]" | - "transf_constraint (EQPP p q) = [LEQ (p - q) 0, LEQ (-(p - q)) 0]" | "transf_constraint _ = []" fun transf_constraints where "transf_constraints [] = []" | "transf_constraints (x#xs) = transf_constraint x @ (transf_constraints xs)" lemma trans_constraint_creats_LEQ_only: assumes "transf_constraint x \ []" shows "(\x \ set (transf_constraint x). \a b. x = LEQ a b)" using assms by (cases x, auto+) lemma trans_constraints_creats_LEQ_only: assumes "transf_constraints xs \ []" assumes "x \ set (transf_constraints xs)" shows "\p r. LEQ p r = x" using assms apply(induction xs) using trans_constraint_creats_LEQ_only apply(auto) apply fastforce apply (metis in_set_simps(3) trans_constraint_creats_LEQ_only) by fastforce lemma non_strict_constr_no_LT: assumes "nonstrict_constrs cs" shows "\x \ set cs. \(\a b. LT a b = x)" - using assms nonstrict_constr.simps(7) by blast + using assms nonstrict_constr.simps(4) by blast lemma non_strict_constr_no_GT: assumes "nonstrict_constrs cs" shows "\x \ set cs. \(\a b. GT a b = x)" - using assms nonstrict_constr.simps(8) by blast + using assms nonstrict_constr.simps(5) by blast -lemma non_strict_constr_no_LTPP: - assumes "nonstrict_constrs cs" - shows "\x \ set cs. \(\a b. LTPP a b = x)" - using assms nonstrict_constr.simps(9) by blast - -lemma non_strict_constr_no_GTPP: - assumes "nonstrict_constrs cs" - shows "\x \ set cs. \(\a b. GTPP a b = x)" - using assms nonstrict_constr.simps(10) by blast lemma non_strict_consts_cond: assumes "\x. x \ set cs \ \(\a b. LT a b = x)" assumes "\x. x \ set cs \ \(\a b. GT a b = x)" - assumes "\x. x \ set cs \ \(\a b. LTPP a b = x)" - assumes "\x. x \ set cs \ \(\a b. GTPP a b = x)" shows "nonstrict_constrs cs" - by (metis assms(1) assms(2) assms(3) assms(4) nonstrict_constr.elims(3)) + by (metis assms(1-2) nonstrict_constr.elims(3)) lemma sat_constr_sat_transf_constrs: assumes "v \\<^sub>c cs" shows "v \\<^sub>c\<^sub>s set (transf_constraint cs)" using assms by (cases cs) (simp add: valuate_uminus valuate_minus)+ lemma sat_constrs_sat_transf_constrs: assumes "v \\<^sub>c\<^sub>s set cs" shows "v \\<^sub>c\<^sub>s set (transf_constraints cs)" using assms by(induction cs, simp) (metis UnE list.set_intros(1) list.set_intros(2) sat_constr_sat_transf_constrs set_append transf_constraints.simps(2)) lemma sat_transf_constrs_sat_constr: assumes "nonstrict_constr cs" assumes "v \\<^sub>c\<^sub>s set (transf_constraint cs)" shows "v \\<^sub>c cs" using assms by (cases cs) (simp add: valuate_uminus valuate_minus)+ lemma sat_transf_constrs_sat_constrs: assumes "nonstrict_constrs cs" assumes "v \\<^sub>c\<^sub>s set (transf_constraints cs)" shows "v \\<^sub>c\<^sub>s set cs" using assms by (induction cs, auto) (simp add: sat_transf_constrs_sat_constr) end \ No newline at end of file diff --git a/thys/Simplex/Simplex.thy b/thys/Simplex/Simplex.thy --- a/thys/Simplex/Simplex.thy +++ b/thys/Simplex/Simplex.thy @@ -1,8345 +1,8203 @@ (* Authors: F. Maric, M. Spasic, R. Thiemann *) section \The Simplex Algorithm\ theory Simplex imports Linear_Poly_Maps QDelta Rel_Chain Simplex_Algebra "HOL-Library.Multiset" "HOL-Library.RBT_Mapping" "HOL-Library.Code_Target_Numeral" begin -text\Linear constraints are of the form \p \ c\ or \p\<^sub>1 \ p\<^sub>2\, where \p\, \p\<^sub>1\, and \p\<^sub>2\, are -linear polynomials, \c\ is a rational constant and \\ \ +text\Linear constraints are of the form \p \ c\, where \p\ is +a homogenenous linear polynomial, \c\ is a rational constant and \\ \ {<, >, \, \, =}\. Their abstract syntax is given by the \constraint\ type, and semantics is given by the relation \\\<^sub>c\, defined straightforwardly by primitive recursion over the \constraint\ type. A set of constraints is satisfied, denoted by \\\<^sub>c\<^sub>s\, if all constraints are. There is also an indexed version \\\<^sub>i\<^sub>c\<^sub>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 "\\<^sub>c" 100) where "v \\<^sub>c (LT l r) \ (l\v\) < r *R 1" | "v \\<^sub>c GT l r \ (l\v\) > r *R 1" | "v \\<^sub>c LEQ l r \ (l\v\) \ r *R 1" | "v \\<^sub>c GEQ l r \ (l\v\) \ r *R 1" | "v \\<^sub>c EQ l r \ (l\v\) = r *R 1" -| "v \\<^sub>c LTPP l1 l2 \ (l1\v\) < (l2\v\)" -| "v \\<^sub>c GTPP l1 l2 \ (l1\v\) > (l2\v\)" -| "v \\<^sub>c LEQPP l1 l2 \ (l1\v\) \ (l2\v\)" -| "v \\<^sub>c GEQPP l1 l2 \ (l1\v\) \ (l2\v\)" -| "v \\<^sub>c EQPP l1 l2 \ (l1\v\) = (l2\v\)" abbreviation satisfies_constraints :: "rat valuation \ constraint set \ bool" (infixl "\\<^sub>c\<^sub>s" 100) where "v \\<^sub>c\<^sub>s cs \ \ c \ cs. v \\<^sub>c c" lemma unsat_mono: assumes "\ (\ v. v \\<^sub>c\<^sub>s cs)" and "cs \ ds" shows "\ (\ v. v \\<^sub>c\<^sub>s ds)" using assms by auto fun i_satisfies_cs (infixl "\\<^sub>i\<^sub>c\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>c\<^sub>s cs \ v \\<^sub>c\<^sub>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) \\<^sub>i\<^sub>c\<^sub>s set ics)) \ (distinct_indices ics \ (\ J. J \ I \ (\ v. (J,v) \\<^sub>i\<^sub>c\<^sub>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 = \ \Decide if the given list of constraints is satisfiable. Return either an unsat core, or a satisfying valuation.\ fixes solve :: "'i i_constraint list \ 'i list + rat valuation" \ \If the status @{const Sat} is returned, then returned valuation satisfies all constraints.\ assumes simplex_sat: "solve cs = Sat v \ v \\<^sub>c\<^sub>s flat (set cs)" \ \If the status @{const Unsat} is returned, then constraints are unsatisfiable, i.e., an unsatisfiable core is returned.\ 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\ \\<^sub>c\<^sub>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 \\\<^sub>\\ 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 \\\<^sub>n\<^sub>s\ and \\\<^sub>n\<^sub>s\<^sub>s\. The indexed variant is \\\<^sub>i\<^sub>n\<^sub>s\<^sub>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 "\\<^sub>n\<^sub>s" 100) where "v \\<^sub>n\<^sub>s LEQ_ns l r \ l\v\ \ r" | "v \\<^sub>n\<^sub>s GEQ_ns l r \ l\v\ \ r" abbreviation satisfies_ns_constraints :: "'a::lrv valuation \ 'a ns_constraint set \ bool" (infixl "\\<^sub>n\<^sub>s\<^sub>s " 100) where "v \\<^sub>n\<^sub>s\<^sub>s cs \ \ c \ cs. v \\<^sub>n\<^sub>s c" fun i_satisfies_ns_constraints :: "'i set \ 'a::lrv valuation \ ('i,'a) i_ns_constraint set \ bool" (infixl "\\<^sub>i\<^sub>n\<^sub>s\<^sub>s " 100) where "(I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs \ v \\<^sub>n\<^sub>s\<^sub>s restrict_to I cs" lemma i_satisfies_ns_constraints_mono: "(I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs \ J \ I \ (J,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs)) \ (distinct_indices_ns cs \ (\ J \ I. \ v. (J,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s cs)))" text\Specification of reduction of constraints to non-strict form is given by:\ locale To_ns = \ \Convert a constraint to an equisatisfiable non-strict constraint list. The conversion must work for arbitrary subsets of constraints -- selected by some index set I -- in order to carry over unsat-cores and in order to support incremental simplex solving.\ fixes to_ns :: "'i i_constraint list \ ('i,'a::lrv) i_ns_constraint list" \ \Convert the valuation that satisfies all non-strict constraints to the valuation that satisfies all initial constraints.\ 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'\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set (to_ns cs) \ (I,\from_ns v' (flat_list (to_ns cs))\) \\<^sub>i\<^sub>c\<^sub>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'\ \\<^sub>n\<^sub>s\<^sub>s flat (set (to_ns cs)) \ \from_ns v' (flat_list (to_ns cs))\ \\<^sub>c\<^sub>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\ \\<^sub>n\<^sub>s\<^sub>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\<^sub>i \ b\<^sub>i\ where \x\<^sub>i\ is a variable and \b\<^sub>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\<^sub>1 + x\<^sub>2 \ b\<^sub>1, x\<^sub>1 + x\<^sub>2 \ b\<^sub>2, x\<^sub>2 \ b\<^sub>3]\ can be transformed to the tableau \[x\<^sub>3 = x\<^sub>1 + x\<^sub>2]\ and atoms \[x\<^sub>3 \ b\<^sub>1, x\<^sub>3 \ b\<^sub>2, x\<^sub>2 \ b\<^sub>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 "\\<^sub>e" 100) where "v \\<^sub>e eq \ v (lhs eq) = (rhs eq)\v\" lemma satisfies_eq_iff: "v \\<^sub>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 "\\<^sub>t" 100) where "v \\<^sub>t t \ \ e \ set t. v \\<^sub>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' \\<^sub>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' \\<^sub>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' \\<^sub>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 \\<^sub>t t1 \ v \\<^sub>t t2" shows "mset t1 = mset t2" proof- { fix t1 t2 assume "lvars t1 = lvars t2" "rvars t1 = rvars t2" "\ t1" "\ v::'a::lrv valuation. v \\<^sub>t t1 \ v \\<^sub>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' \\<^sub>t t1" "\ x \ rvars t1. v x = v' x" using recalc_tableau_lvars[of t1] \\ t1\ by auto have "v' \\<^sub>t t2" using \v' \\<^sub>t t1\ \\ v. v \\<^sub>t t1 \ v \\<^sub>t t2\ by simp have "b \v'\ = b' \v'\" using \(a, b) \ set t1\ \v' \\<^sub>t t1\ using \(a, b') \ set t2\ \v' \\<^sub>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) qed text\Elementary atoms are represented by the type \'a atom\ and semantics for atoms and sets of atoms is denoted by \\\<^sub>a\ and \\\<^sub>a\<^sub>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 "\\<^sub>a" 100) where "v \\<^sub>a Leq x c \ v x \ c" | "v \\<^sub>a Geq x c \ v x \ c" definition satisfies_atom_set :: "'a::linorder valuation \ 'a atom set \ bool" (infixl "\\<^sub>a\<^sub>s" 100) where "v \\<^sub>a\<^sub>s as \ \ a \ as. v \\<^sub>a a" definition satisfies_atom' :: "'a::linorder valuation \ 'a atom \ bool" (infixl "\\<^sub>a\<^sub>e" 100) where "v \\<^sub>a\<^sub>e a \ v (atom_var a) = atom_const a" lemma satisfies_atom'_stronger: "v \\<^sub>a\<^sub>e a \ v \\<^sub>a a" by (cases a, auto simp: satisfies_atom'_def) abbreviation satisfies_atom_set' :: "'a::linorder valuation \ 'a atom set \ bool" (infixl "\\<^sub>a\<^sub>e\<^sub>s" 100) where "v \\<^sub>a\<^sub>e\<^sub>s as \ \ a \ as. v \\<^sub>a\<^sub>e a" lemma satisfies_atom_set'_stronger: "v \\<^sub>a\<^sub>e\<^sub>s as \ v \\<^sub>a\<^sub>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 "\\<^sub>i\<^sub>a\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>a\<^sub>s as \ v \\<^sub>a\<^sub>s restrict_to I as" fun i_satisfies_atom_set' :: "'i set \ 'a::linorder valuation \ ('i,'a) i_atom set \ bool" (infixl "\\<^sub>i\<^sub>a\<^sub>e\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as \ v \\<^sub>a\<^sub>e\<^sub>s restrict_to I as" lemma i_satisfies_atom_set'_stronger: "Iv \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as \ Iv \\<^sub>i\<^sub>a\<^sub>s as" using satisfies_atom_set'_stronger[of _ "snd Iv"] by (cases Iv, auto) lemma satisifies_atom_restrict_to_Cons: "v \\<^sub>a\<^sub>s restrict_to I (set as) \ (i \ I \ v \\<^sub>a a) \ v \\<^sub>a\<^sub>s restrict_to I (set ((i,a) # as))" unfolding satisfies_atom_set_def by auto lemma satisfies_tableau_Cons: "v \\<^sub>t t \ v \\<^sub>e e \ v \\<^sub>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 \ \The returned tableau is always normalized.\ preprocess_tableau_normalized: "preprocess cs = (t,as,trans_v,U) \ \ t" and \ \Tableau and atoms are equisatisfiable with starting non-strict constraints.\ i_preprocess_sat: "\ v. preprocess cs = (t,as,trans_v,U) \ I \ set U = {} \ (I,\v\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t \ (I,\trans_v v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" and preprocess_unsat: "preprocess cs = (t, as,trans_v,U) \ (I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs \ \ v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>t t" and \ \distinct indices on ns-constraints ensures distinct indices in atoms\ preprocess_distinct: "preprocess cs = (t, as,trans_v, U) \ distinct_indices_ns (set cs) \ distinct_indices_atoms (set as)" and \ \unsat indices\ preprocess_unsat_indices: "preprocess cs = (t, as,trans_v, U) \ i \ set U \ \ (\ v. ({i},v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs)" and \ \preprocessing cannot introduce new indices\ 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\ \\<^sub>a\<^sub>s flat (set as) \ \v\ \\<^sub>t t \ \trans_v v\ \\<^sub>n\<^sub>s\<^sub>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 \\<^sub>t t \ (I,v) \\<^sub>i\<^sub>a\<^sub>s as)) \ (distinct_indices_atoms as \ (\ J \ I. \ v. v \\<^sub>t t \ (J,v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as)))" lemma minimal_unsat_core_tabl_atomsD: assumes "minimal_unsat_core_tabl_atoms I t as" shows "I \ fst ` as" "\ (\ v. v \\<^sub>t t \ (I,v) \\<^sub>i\<^sub>a\<^sub>s as)" "distinct_indices_atoms as \ J \ I \ \ v. v \\<^sub>t t \ (J,v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>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\ \\<^sub>t t \ \v\ \\<^sub>a\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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 \\<^sub>t t" "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s set as" by auto from i_satisfies_atom_set'_stronger[OF model(2)] have model': "(J, v) \\<^sub>i\<^sub>a\<^sub>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\ \\<^sub>t t" "(J, \w\) \\<^sub>i\<^sub>a\<^sub>s set as" by auto from i_preprocess_sat[OF prep _ this(2,1)] \J \ I\ inter have "(J, \trans_v w\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" by auto then show "\ w. (J, w) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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\ \\<^sub>n\<^sub>s\<^sub>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\<^sub>i \ x\<^sub>i \ u\<^sub>i\, where \l\<^sub>i\ and \u\<^sub>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 \\<^sub>u\<^sub>b b \ case b of None \ False | Some b' \ c \ b'\ \c \\<^sub>u\<^sub>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 ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ False | Some b' \ le lt b' c" definition gtub ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ False | Some b' \ lt b' c" definition leub ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ True | Some b' \ le lt c b'" definition ltub ("\\<^sub>u\<^sub>b") where "\\<^sub>u\<^sub>b lt c b \ case b of None \ True | Some b' \ lt c b'" definition lelb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ False | Some b' \ le lt c b'" definition ltlb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ False | Some b' \ lt c b'" definition gelb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ True | Some b' \ le lt b' c" definition gtlb ("\\<^sub>l\<^sub>b") where "\\<^sub>l\<^sub>b lt c b \ case b of None \ True | Some b' \ lt b' c" definition ge_ubound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>u\<^sub>b" 100) where "c \\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition gt_ubound :: "'a::linorder \ 'a option \ bool" (infixl ">\<^sub>u\<^sub>b" 100) where "c >\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition le_ubound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>u\<^sub>b" 100) where "c \\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition lt_ubound :: "'a::linorder \ 'a option \ bool" (infixl "<\<^sub>u\<^sub>b" 100) where "c <\<^sub>u\<^sub>b b = \\<^sub>u\<^sub>b (<) c b" definition le_lbound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>l\<^sub>b" 100) where "c \\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>b (<) c b" definition lt_lbound :: "'a::linorder \ 'a option \ bool" (infixl "<\<^sub>l\<^sub>b" 100) where "c <\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>b (<) c b" definition ge_lbound :: "'a::linorder \ 'a option \ bool" (infixl "\\<^sub>l\<^sub>b" 100) where "c \\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>b (<) c b" definition gt_lbound :: "'a::linorder \ 'a option \ bool" (infixl ">\<^sub>l\<^sub>b" 100) where "c >\<^sub>l\<^sub>b b = \\<^sub>l\<^sub>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]: "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" "\\<^sub>l\<^sub>b (>) a b = \\<^sub>u\<^sub>b (<) a b" "\\<^sub>u\<^sub>b (>) a b = \\<^sub>l\<^sub>b (<) a b" by (case_tac[!] b) (auto simp add: bound_compare'_defs) (* Auxiliary lemmas about bound comparison *) lemma [simp]: "\ c \\<^sub>u\<^sub>b None " "\ c \\<^sub>l\<^sub>b None" by (auto simp add: bound_compare_defs) lemma neg_bounds_compare: "(\ (c \\<^sub>u\<^sub>b b)) \ c <\<^sub>u\<^sub>b b" "(\ (c \\<^sub>u\<^sub>b b)) \ c >\<^sub>u\<^sub>b b" "(\ (c >\<^sub>u\<^sub>b b)) \ c \\<^sub>u\<^sub>b b" "(\ (c <\<^sub>u\<^sub>b b)) \ c \\<^sub>u\<^sub>b b" "(\ (c \\<^sub>l\<^sub>b b)) \ c >\<^sub>l\<^sub>b b" "(\ (c \\<^sub>l\<^sub>b b)) \ c <\<^sub>l\<^sub>b b" "(\ (c <\<^sub>l\<^sub>b b)) \ c \\<^sub>l\<^sub>b b" "(\ (c >\<^sub>l\<^sub>b b)) \ c \\<^sub>l\<^sub>b b" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma bounds_compare_contradictory [simp]: "\c \\<^sub>u\<^sub>b b; c <\<^sub>u\<^sub>b b\ \ False" "\c \\<^sub>u\<^sub>b b; c >\<^sub>u\<^sub>b b\ \ False" "\c >\<^sub>u\<^sub>b b; c \\<^sub>u\<^sub>b b\ \ False" "\c <\<^sub>u\<^sub>b b; c \\<^sub>u\<^sub>b b\ \ False" "\c \\<^sub>l\<^sub>b b; c >\<^sub>l\<^sub>b b\ \ False" "\c \\<^sub>l\<^sub>b b; c <\<^sub>l\<^sub>b b\ \ False" "\c <\<^sub>l\<^sub>b b; c \\<^sub>l\<^sub>b b\ \ False" "\c >\<^sub>l\<^sub>b b; c \\<^sub>l\<^sub>b b\ \ False" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma compare_strict_nonstrict: "x <\<^sub>u\<^sub>b b \ x \\<^sub>u\<^sub>b b" "x >\<^sub>u\<^sub>b b \ x \\<^sub>u\<^sub>b b" "x <\<^sub>l\<^sub>b b \ x \\<^sub>l\<^sub>b b" "x >\<^sub>l\<^sub>b b \ x \\<^sub>l\<^sub>b b" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma [simp]: "\ x \ c; c <\<^sub>u\<^sub>b b \ \ x <\<^sub>u\<^sub>b b" "\ x < c; c \\<^sub>u\<^sub>b b \ \ x <\<^sub>u\<^sub>b b" "\ x \ c; c \\<^sub>u\<^sub>b b \ \ x \\<^sub>u\<^sub>b b" "\ x \ c; c >\<^sub>l\<^sub>b b \ \ x >\<^sub>l\<^sub>b b" "\ x > c; c \\<^sub>l\<^sub>b b \ \ x >\<^sub>l\<^sub>b b" "\ x \ c; c \\<^sub>l\<^sub>b b \ \ x \\<^sub>l\<^sub>b b" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma bounds_lg [simp]: "\ c >\<^sub>u\<^sub>b b; x \\<^sub>u\<^sub>b b\ \ x < c" "\ c \\<^sub>u\<^sub>b b; x <\<^sub>u\<^sub>b b\ \ x < c" "\ c \\<^sub>u\<^sub>b b; x \\<^sub>u\<^sub>b b\ \ x \ c" "\ c <\<^sub>l\<^sub>b b; x \\<^sub>l\<^sub>b b\ \ x > c" "\ c \\<^sub>l\<^sub>b b; x >\<^sub>l\<^sub>b b\ \ x > c" "\ c \\<^sub>l\<^sub>b b; x \\<^sub>l\<^sub>b b\ \ x \ c" by (case_tac[!] b) (auto simp add: bound_compare_defs) lemma bounds_compare_Some [simp]: "x \\<^sub>u\<^sub>b Some c \ x \ c" "x \\<^sub>u\<^sub>b Some c \ x \ c" "x <\<^sub>u\<^sub>b Some c \ x < c" "x >\<^sub>u\<^sub>b Some c \ x > c" "x \\<^sub>l\<^sub>b Some c \ x \ c" "x \\<^sub>l\<^sub>b Some c \ x \ c" "x >\<^sub>l\<^sub>b Some c \ x > c" "x <\<^sub>l\<^sub>b Some c \ x < c" by (auto simp add: bound_compare_defs) fun in_bounds where "in_bounds x v (lb, ub) = (v x \\<^sub>l\<^sub>b lb x \ v x \\<^sub>u\<^sub>b ub x)" fun satisfies_bounds :: "'a::linorder valuation \ 'a bounds \ 'a bounds \ bool" (infixl "\\<^sub>b" 100) where "v \\<^sub>b b \ (\ x. in_bounds x v b)" declare satisfies_bounds.simps [simp del] lemma satisfies_bounds_iff: "v \\<^sub>b (lb, ub) \ (\ x. v x \\<^sub>l\<^sub>b lb x \ v x \\<^sub>u\<^sub>b ub x)" by (auto simp add: satisfies_bounds.simps) lemma not_in_bounds: "\ (in_bounds x v (lb, ub)) = (v x <\<^sub>l\<^sub>b lb x \ v x >\<^sub>u\<^sub>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 \\<^sub>a\<^sub>s as \ v \\<^sub>b (lb, ub))" declare atoms_equiv_bounds.simps [simp del] lemma atoms_equiv_bounds_simps: "as \ (lb, ub) \ \ v. v \\<^sub>a\<^sub>s as \ v \\<^sub>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 \\\<^sub>l\ and \\\<^sub>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") (\\<^sub>i\<^sub>l: "('i,'a) bounds_index") (\\<^sub>i\<^sub>u: "('i,'a) bounds_index") (\: "(var, 'a) mapping") (\: bool) (\\<^sub>c: "'i list option") definition indexl :: "('i,'a) state \ 'i bound_index" ("\\<^sub>l") where "\\<^sub>l s = (fst o the) o look (\\<^sub>i\<^sub>l s)" definition boundsl :: "('i,'a) state \ 'a bounds" ("\\<^sub>l") where "\\<^sub>l s = map_option snd o look (\\<^sub>i\<^sub>l s)" definition indexu :: "('i,'a) state \ 'i bound_index" ("\\<^sub>u") where "\\<^sub>u s = (fst o the) o look (\\<^sub>i\<^sub>u s)" definition boundsu :: "('i,'a) state \ 'a bounds" ("\\<^sub>u") where "\\<^sub>u s = map_option snd o look (\\<^sub>i\<^sub>u s)" abbreviation BoundsIndicesMap ("\\<^sub>i") where "\\<^sub>i s \ (\\<^sub>i\<^sub>l s, \\<^sub>i\<^sub>u s)" abbreviation Bounds :: "('i,'a) state \ 'a bounds \ 'a bounds" ("\") where "\ s \ (\\<^sub>l s, \\<^sub>u s)" abbreviation Indices :: "('i,'a) state \ 'i bound_index \ 'i bound_index" ("\") where "\ s \ (\\<^sub>l s, \\<^sub>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 "\\<^sub>i\<^sub>b" 100) where "(I,v) \\<^sub>i\<^sub>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 "\\<^sub>i\<^sub>b\<^sub>e" 100) where "(I,v) \\<^sub>i\<^sub>b\<^sub>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 "\\<^sub>i" 100) where "as \\<^sub>i bi \ (\ I v. (I,v) \\<^sub>i\<^sub>a\<^sub>s as \ (I,v) \\<^sub>i\<^sub>b bi)" declare atoms_imply_bounds_index.simps[simp del] lemma i_satisfies_atom_set_mono: "as \ as' \ v \\<^sub>i\<^sub>a\<^sub>s as' \ v \\<^sub>i\<^sub>a\<^sub>s as" by (cases v, auto simp: satisfies_atom_set_def) lemma atoms_imply_bounds_index_mono: "as \ as' \ as \\<^sub>i bi \ as' \\<^sub>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 "\\<^sub>s" 100) where "v \\<^sub>s s \ v \\<^sub>b \ s \ v \\<^sub>t \ s" definition curr_val_satisfies_state :: "('i,'a::lrv) state \ bool" ("\") where "\ s \ \\ s\ \\<^sub>s s" fun satisfies_state_index :: "'i set \ 'a::lrv valuation \ ('i,'a) state \ bool" (infixl "\\<^sub>i\<^sub>s" 100) where "(I,v) \\<^sub>i\<^sub>s s \ (v \\<^sub>t \ s \ (I,v) \\<^sub>i\<^sub>b \\ s)" declare satisfies_state_index.simps[simp del] fun satisfies_state_index' :: "'i set \ 'a::lrv valuation \ ('i,'a) state \ bool" (infixl "\\<^sub>i\<^sub>s\<^sub>e" 100) where "(I,v) \\<^sub>i\<^sub>s\<^sub>e s \ (v \\<^sub>t \ s \ (I,v) \\<^sub>i\<^sub>b\<^sub>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 (\\<^sub>i\<^sub>l s) x = Some (i,b) \ look (\\<^sub>i\<^sub>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 (\\<^sub>i\<^sub>l s) x = Some (i,b) \ look (\\<^sub>i\<^sub>u s) x = Some (i,b)) \ (look (\\<^sub>i\<^sub>l s) x' = Some (i,b') \ look (\\<^sub>i\<^sub>u s) x' = Some (i,b')) \ (x = x' \ b = b')))" lemma distinct_indices_stateD: assumes "distinct_indices_state s" shows "look (\\<^sub>i\<^sub>l s) x = Some (i,b) \ look (\\<^sub>i\<^sub>u s) x = Some (i,b) \ look (\\<^sub>i\<^sub>l s) x' = Some (i,b') \ look (\\<^sub>i\<^sub>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 (\\<^sub>c s)) \ indices_state s \ (\ (\ v. (set (the (\\<^sub>c s)),v) \\<^sub>i\<^sub>s s)))" definition subsets_sat_core :: "('i,'a::lrv) state \ bool" where "subsets_sat_core s = ((\ I. I \ set (the (\\<^sub>c s)) \ (\ v. (I,v) \\<^sub>i\<^sub>s\<^sub>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 \\<^sub>t t \ (I, v) \\<^sub>i\<^sub>a\<^sub>s as)" by blast with i_satisfies_atom_set_mono[OF sub] show "(\v. v \\<^sub>t t \ (I, v) \\<^sub>i\<^sub>a\<^sub>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 \\<^sub>t t" "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as" by blast have "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>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 \\<^sub>a\<^sub>e b" by force thus "v \\<^sub>a\<^sub>e a" using id unfolding satisfies_atom'_def by auto qed with v show "\v. v \\<^sub>t t \ (J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s bs" by blast qed lemma state_satisfies_index: assumes "v \\<^sub>s s" shows "(I,v) \\<^sub>i\<^sub>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 \\<^sub>t \ s" and bnd: "v x \\<^sub>l\<^sub>b \\<^sub>l s x" "v x \\<^sub>u\<^sub>b \\<^sub>u s x" by auto show "v \\<^sub>t \ s" by fact show "\\<^sub>l s x = Some c \ \\<^sub>l s x \ I \ c \ v x" using bnd(1) by auto show "\\<^sub>u s x = Some c \ \\<^sub>u s x \ I \ v x \ c" using bnd(2) by auto qed lemma unsat_state_core_unsat: "unsat_state_core s \ \ (\ v. v \\<^sub>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 (\\<^sub>i\<^sub>l s) x = Some (i,b) \ ((i, Geq x b) \ as)) \ (look (\\<^sub>i\<^sub>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 \\\<^sub>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 \\<^sub>i\<^sub>u_update where "\\<^sub>i\<^sub>u_update up (State T BIL BIU V U UC) = State T BIL (up BIU) V U UC" fun \\<^sub>i\<^sub>l_update where "\\<^sub>i\<^sub>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]: "\\<^sub>i\<^sub>u (\\<^sub>i\<^sub>u_update up s) = up (\\<^sub>i\<^sub>u s)" "\\<^sub>i\<^sub>l (\\<^sub>i\<^sub>u_update up s) = \\<^sub>i\<^sub>l s" "\ (\\<^sub>i\<^sub>u_update up s) = \ s" "\ (\\<^sub>i\<^sub>u_update up s) = \ s" "\ (\\<^sub>i\<^sub>u_update up s) = \ s" "\\<^sub>c (\\<^sub>i\<^sub>u_update up s) = \\<^sub>c s" "\\<^sub>i\<^sub>l (\\<^sub>i\<^sub>l_update up s) = up (\\<^sub>i\<^sub>l s)" "\\<^sub>i\<^sub>u (\\<^sub>i\<^sub>l_update up s) = \\<^sub>i\<^sub>u s" "\ (\\<^sub>i\<^sub>l_update up s) = \ s" "\ (\\<^sub>i\<^sub>l_update up s) = \ s" "\ (\\<^sub>i\<^sub>l_update up s) = \ s" "\\<^sub>c (\\<^sub>i\<^sub>l_update up s) = \\<^sub>c s" "\ (\_update V s) = V" "\\<^sub>i\<^sub>l (\_update V s) = \\<^sub>i\<^sub>l s" "\\<^sub>i\<^sub>u (\_update V s) = \\<^sub>i\<^sub>u s" "\ (\_update V s) = \ s" "\ (\_update V s) = \ s" "\\<^sub>c (\_update V s) = \\<^sub>c s" "\ (\_update T s) = T" "\\<^sub>i\<^sub>l (\_update T s) = \\<^sub>i\<^sub>l s" "\\<^sub>i\<^sub>u (\_update T s) = \\<^sub>i\<^sub>u s" "\ (\_update T s) = \ s" "\ (\_update T s) = \ s" "\\<^sub>c (\_update T s) = \\<^sub>c s" by (atomize(full), cases s, auto) declare \\<^sub>i\<^sub>u_update.simps[simp del] \\<^sub>i\<^sub>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]: "\\<^sub>i\<^sub>l (set_unsat I s) = \\<^sub>i\<^sub>l s" "\\<^sub>i\<^sub>u (set_unsat I s) = \\<^sub>i\<^sub>u s" "\ (set_unsat I s) = \ s" "\ (set_unsat I s) = \ s" "\ (set_unsat I s) = True" "\\<^sub>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 (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>u_update Leq Geq (\)" definition Negative where [simp]: "Negative \ Direction (>) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>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 \ \The final and the initial tableau are equivalent.\ assert_all_state_tableau_equiv: "\ t \ assert_all_state t as = s' \ (v::'a valuation) \\<^sub>t t \ v \\<^sub>t \ s'" and \ \If @{term \} is not raised, then the valuation in the final state satisfies its tableau and its bounds (that are, in this case, equivalent to the set of all asserted bounds).\ 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 \ \If @{term \} is raised, then there is no valuation satisfying the tableau and the bounds in the final state (that are, in this case, equivalent to a subset of asserted atoms).\ 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 \\<^sub>i \\ s'" and \ \The set of indices is taken from the constraints\ 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 (\\<^sub>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) \\<^sub>t t = v \\<^sub>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 \ {} \\<^sub>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 \ \Tableau in the initial state for @{text t} is @{text t}:\ init_tableau_id: "\ (init t) = t" and \ \Since unsatisfiability is not detected, @{text \} flag must not be set:\ init_unsat_flag: "\ \ (init t)" and \ \The current valuation must satisfy the tableau:\ init_satisfies_tableau: "\\ (init t)\ \\<^sub>t t" and \ \In an inital state no atoms are yet asserted so the bounds must be empty:\ init_bounds: "\\<^sub>i\<^sub>l (init t) = Mapping.empty" "\\<^sub>i\<^sub>u (init t) = Mapping.empty" and \ \All tableau vars are valuated:\ init_tableau_valuated: "\ (init t)" begin lemma init_satisfies_bounds: "\\ (init t)\ \\<^sub>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: "{} \\<^sub>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 \ \Tableau remains equivalent to the previous one and normalized and valuated.\ assert_tableau: "\\ \ s; \ s; \ (\ s); \ s\ \ let s' = assert a s in ((v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ s') \ \ (\ s') \ \ s'" and \ \If the @{term \} flag is not raised, then the current valuation is updated so that it satisfies the current tableau and the current bounds.\ assert_sat: "\\ \ s; \ s; \ (\ s); \ s\ \ \ \ (assert a s) \ \ (assert a s)" and \ \The set of asserted atoms remains equivalent to the bounds in the state.\ assert_atoms_equiv_bounds: "\\ \ s; \ s; \ (\ s); \ s\ \ flat ats \ \ s \ flat (ats \ {a}) \ \ (assert a s)" and \ \There is a subset of asserted atoms which remains index-equivalent to the bounds in the state.\ assert_atoms_imply_bounds_index: "\\ \ s; \ s; \ (\ s); \ s\ \ ats \\<^sub>i \\ s \ insert a ats \\<^sub>i \\ (assert a s)" and \ \If the @{term \} flag is raised, then there is no valuation that satisfies both the current tableau and the current bounds.\ 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) \\<^sub>t \ s \ v \\<^sub>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 \\<^sub>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' \\<^sub>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 \\<^sub>t t = v \\<^sub>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 \\<^sub>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" ("_ \\<^sub>b _ \/ _") translations "v \\<^sub>b b \ S" == "CONST satisfies_bounds_set v b S" lemma satisfies_bounds_set_iff: "v \\<^sub>b (lb, ub) \ S \ (\ x \ S. v x \\<^sub>l\<^sub>b lb x \ v x \\<^sub>u\<^sub>b ub x)" by (simp add: satisfies_bounds_set.simps) definition curr_val_satisfies_no_lhs ("\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s") where "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s \ \\ s\ \\<^sub>t (\ s) \ (\\ s\ \\<^sub>b (\ s) \ (- lvars (\ s)))" lemma satisfies_satisfies_no_lhs: "\ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\<^sub>l s x = None \ \\<^sub>u s x = None then True else the (\\<^sub>l s x) \ the (\\<^sub>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 \\<^sub>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) \\<^sub>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; \\<^sub>u s x\<^sub>i = Some c\ \ c \\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i" unfolding bounds_consistent_def by (cases "\\<^sub>l s x\<^sub>i", auto simp add: bound_compare_defs split: if_splits) (erule_tac x="x\<^sub>i" in allE, auto) lemma bounds_consistent_leq_ub: "\\ s; \\<^sub>l s x\<^sub>i = Some c\ \ c \\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i" unfolding bounds_consistent_def by (cases "\\<^sub>u s x\<^sub>i", auto simp add: bound_compare_defs split: if_splits) (erule_tac x="x\<^sub>i" in allE, auto) lemma bounds_consistent_gt_ub: "\\ s; c <\<^sub>l\<^sub>b \\<^sub>l s x \ \ \ c >\<^sub>u\<^sub>b \\<^sub>u s x" unfolding bounds_consistent_def by (case_tac[!] "\\<^sub>l s x", case_tac[!] "\\<^sub>u s x") (auto simp add: bound_compare_defs, erule_tac x="x" in allE, simp) lemma bounds_consistent_lt_lb: "\\ s; c >\<^sub>u\<^sub>b \\<^sub>u s x \ \ \ c <\<^sub>l\<^sub>b \\<^sub>l s x" unfolding bounds_consistent_def by (case_tac[!] "\\<^sub>l s x", case_tac[!] "\\<^sub>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 \ \The tableau remains unchanged and valuated.\ assert_bound_tableau: "\\ \ s; \ s; \ (\ s); \ s\ \ assert_bound a s = s' \ \ s' = \ s \ \ s'" and \ \If the @{term \} flag is not set, all but the lhs variables in the tableau remain within their bounds, the new valuation satisfies the tableau, and bounds do not overlap.\ assert_bound_sat: "\\ \ s; \ s; \ (\ s); \ s\ \ assert_bound a s = s' \ \ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ s'" and \ \The set of asserted atoms remains equivalent to the bounds in the state.\ 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 \\<^sub>i \\ s \ insert a ats \\<^sub>i \\ (assert_bound a s)" and \ \@{term \} flag is raised, only if the bounds became inconsistent:\ 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; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ \ (assert_bound a s) = \ s" assumes assert_bound_nolhs_sat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ \ \ (assert_bound a s) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (assert_bound a s) \ \ (assert_bound a s)" assumes assert_bound_nolhs_atoms_equiv_bounds: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ flat ats \ \ s \ flat (ats \ {a}) \ \ (assert_bound a s)" assumes assert_bound_nolhs_atoms_imply_bounds_index: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ ats \\<^sub>i \\ s \ insert a ats \\<^sub>i \\ (assert_bound a s)" assumes assert_bound_nolhs_unsat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ (\ s); \ s; \ s\ \ \ (assert_bound a s)" assumes assert_bound_nolhs_index_valid: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\, \\ s\, \\ (\ s)\, and \\ s\ hold.\ locale Check = fixes check::"('i,'a::lrv) state \ ('i,'a) state" assumes \ \If @{text check} is called from an inconsistent state, the state is unchanged.\ check_unsat_id: "\ s \ check s = s" and \ \The tableau remains equivalent to the previous one, normalized and valuated, the state stays consistent.\ check_tableau: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ let s' = check s in ((v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ s') \ \ (\ s') \ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ s'" and \ \The bounds remain unchanged.\ check_bounds_id: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \\<^sub>i (check s) = \\<^sub>i s" and \ \If @{term \} flag is not raised, the current valuation @{text "\"} satisfies both the tableau and the bounds and if it is raised, there is no valuation that satisfies them.\ check_sat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ \ (check s) \ \ (check s)" and check_unsat: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ (check s) \ minimal_unsat_state_core (check s)" begin lemma check_tableau_equiv: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ (v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ (check s)" using check_tableau by (simp add: Let_def) lemma check_tableau_index_valid: assumes "\ \ s" " \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ (\ (check s))" using check_tableau by (simp add: Let_def) lemma check_bounds_consistent: assumes "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" "\ (\ s)" "\ s" shows "\ (check s)" using check_bounds_id[OF assms] assms(3) unfolding Let_def bounds_consistent_def boundsl_def boundsu_def by (metis Pair_inject) lemma check_tableau_valuated: "\\ \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ (\ s); \ s\ \ \ (check s)" using check_tableau by (simp add: Let_def) lemma check_indices_state: assumes "\ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\<^sub>t \ s = v \\<^sub>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 \\<^sub>t \ s = v \\<^sub>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 \\<^sub>i \\ s" from assert_bound_atoms_imply_bounds_index[OF this, of a] have as: "insert a ats \\<^sub>i \\ (assert_bound a s)" by auto show "insert a ats \\<^sub>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 "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s \ \ s" shows "let s' = assert_bound_loop ats s in \ (\ s') \ \ s' \ (\ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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\ \\<^sub>t \ s; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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)) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\<^sub>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' \\<^sub>i \\ ?s" by auto from AssertAllState''_precond[of t ats', OF assms, unfolded Let_def] False have *: "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\<^sub>t t = v \\<^sub>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 \\<^sub>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 \ \Tableau, bounds, and the unsatisfiability flag are preserved.\ update_id: "\\ (\ s); \ s; x \ lvars (\ s)\ \ let s' = update x c s in \ s' = \ s \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>c s" and \ \Tableau remains valuated.\ update_tableau_valuated: "\\ (\ s); \ s; x \ lvars (\ s)\ \ \ (update x v s)" and \ \The given variable @{text "x"} in the updated valuation is set to the given value @{text "v"} while all other variables (except those on the lhs of the tableau) are unchanged.\ 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 \ \Updated valuation continues to satisfy the tableau.\ update_satisfies_tableau: "\\ (\ s); \ s; x \ lvars (\ s)\ \ \\ s\ \\<^sub>t \ s \ \\ (update x c s)\ \\<^sub>t \ s" begin lemma update_bounds_id: assumes "\ (\ s)" "\ s" "x \ lvars (\ s)" shows "\\<^sub>i (update x c s) = \\<^sub>i s" "\\ (update x c s) = \\ s" "\\<^sub>l (update x c s) = \\<^sub>l s" "\\<^sub>u (update x c s) = \\<^sub>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)\ \ \\<^sub>c (update x c s) = \\<^sub>c s" using update_id by (auto simp add: Let_def) definition assert_bound' where [simp]: "assert_bound' dir i x c s \ (if (\\<^sub>u\<^sub>b (lt dir)) c (UB dir s x) then s else let s' = update\\ (UBI_upd dir) i x c s in if (\\<^sub>l\<^sub>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 "\\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \ P s" assumes "\\ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \\<^sub>l\<^sub>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); \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x))\ \ P (update x c (update\\ (UBI_upd dir) i x c s))" assumes "\\ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)); x \ lvars (\ s)\ \ P (update\\ (UBI_upd dir) i x c s)" assumes "\\ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>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 "\\<^sub>u\<^sub>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 "\\<^sub>l\<^sub>b (lt dir) c ((LB dir) s x)") case True then show ?thesis using \\ \\<^sub>u\<^sub>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 \\ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \\ \\<^sub>l\<^sub>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 \\ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \\ \\<^sub>l\<^sub>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 \\ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)\ \\ \\<^sub>l\<^sub>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; \\<^sub>u\<^sub>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; \ \\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x); \\<^sub>l\<^sub>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); \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>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); \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>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; \ (\\<^sub>u\<^sub>b (lt dir) c ((UB dir) s x)); \ (\\<^sub>l\<^sub>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' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s" assumes "\ s. P s = P' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>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: "\\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" and dir: "dir = Positive \ dir = Negative" assumes v: "v \\<^sub>b \ (update\\ (UBI_upd dir) i x c s)" shows "v \\<^sub>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" "\ \\<^sub>u\<^sub>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 \\<^sub>a\<^sub>s (ats \ {LE dir x c}) = v \\<^sub>b \ ?s" proof assume "v \\<^sub>a\<^sub>s (ats \ {LE dir x c})" then have "v \\<^sub>a\<^sub>s ats" "le (lt dir) (v x) c" using \dir = Positive \ dir = Negative\ unfolding satisfies_atom_set_def by auto show "v \\<^sub>b \ ?s" unfolding satisfies_bounds.simps proof fix x' show "in_bounds x' v (\ ?s)" using \v \\<^sub>a\<^sub>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 \\<^sub>b \ ?s" then have "v \\<^sub>b \ s" using \\ \\<^sub>u\<^sub>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 "\\<^sub>u s x"] using neg_bounds_compare(5)[of c "\\<^sub>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 \\<^sub>a\<^sub>s (ats \ {LE dir x c})" unfolding satisfies_atom_set_def proof fix a assume "a \ ats \ {LE dir x c}" then show "v \\<^sub>a a" proof assume "a \ {LE dir x c}" then show ?thesis using \v \\<^sub>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 \\<^sub>b \ s\ unfolding atoms_equiv_bounds.simps satisfies_atom_set_def by auto qed qed qed qed lemma bounds_updates: "\\<^sub>l (\\<^sub>i\<^sub>u_update u s) = \\<^sub>l s" "\\<^sub>u (\\<^sub>i\<^sub>l_update u s) = \\<^sub>u s" "\\<^sub>u (\\<^sub>i\<^sub>u_update (upd x (i,c)) s) = (\\<^sub>u s) (x \ c)" "\\<^sub>l (\\<^sub>i\<^sub>l_update (upd x (i,c)) s) = (\\<^sub>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" \ \@{text rhs_eq_val} computes the value of the rhs of @{text e} in @{text "\v\(x := c)"}.\ assumes rhs_eq_val: "\v\ \\<^sub>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 \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>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\ \\<^sub>t \ s \ \e \ set (\ s). \\ (update x c s)\ \\<^sub>e e" proof fix e assume "e \ set (\ s)" "\\ s\ \\<^sub>t \ s" then have "\\ s\ \\<^sub>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)\ \\<^sub>e e" using \\\ s\ \\<^sub>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\ \\<^sub>t \ s \ \\ (update x c s)\ \\<^sub>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 \ \Valuation, bounds, and the unsatisfiability flag are not changed.\ pivot_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ let s' = pivot x\<^sub>i x\<^sub>j s in \ s' = \ s \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>c s" and \ \The tableau remains equivalent to the previous one and normalized.\ pivot_tableau: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ let s' = pivot x\<^sub>i x\<^sub>j s in ((v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ s') \ \ (\ s') " and \ \@{text "x\<^sub>i"} and @{text "x\<^sub>j"} are swapped, while the other variables do not change sides.\ pivot_vars': "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ let s' = pivot x\<^sub>i x\<^sub>j s in rvars(\ s') = rvars(\ s)-{x\<^sub>j}\{x\<^sub>i} \ lvars(\ s') = lvars(\ s)-{x\<^sub>i}\{x\<^sub>j}" begin lemma pivot_bounds_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>i (pivot x\<^sub>i x\<^sub>j s) = \\<^sub>i s" using pivot_id by (simp add: Let_def) lemma pivot_bounds_id': assumes "\ (\ s)" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\\ (pivot x\<^sub>i x\<^sub>j s) = \\ s" "\ (pivot x\<^sub>i x\<^sub>j s) = \ s" "\ (pivot x\<^sub>i x\<^sub>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\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot x\<^sub>i x\<^sub>j s) = \ s" using pivot_id by (simp add: Let_def) lemma pivot_unsat_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot x\<^sub>i x\<^sub>j s) = \ s" using pivot_id by (simp add: Let_def) lemma pivot_unsat_core_id: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>c (pivot x\<^sub>i x\<^sub>j s) = \\<^sub>c s" using pivot_id by (simp add: Let_def) lemma pivot_tableau_equiv: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ (v::'a valuation) \\<^sub>t \ s = v \\<^sub>t \ (pivot x\<^sub>i x\<^sub>j s)" using pivot_tableau by (simp add: Let_def) lemma pivot_tableau_normalized: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (\ (pivot x\<^sub>i x\<^sub>j s))" using pivot_tableau by (simp add: Let_def) lemma pivot_rvars: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ rvars (\ (pivot x\<^sub>i x\<^sub>j s)) = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" using pivot_vars' by (simp add: Let_def) lemma pivot_lvars: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ lvars (\ (pivot x\<^sub>i x\<^sub>j s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" using pivot_vars' by (simp add: Let_def) lemma pivot_vars: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ tvars (\ (pivot x\<^sub>i x\<^sub>j s)) = tvars (\ s) " using pivot_lvars[of s x\<^sub>i x\<^sub>j] pivot_rvars[of s x\<^sub>i x\<^sub>j] using rvars_of_lvar_rvars[of x\<^sub>i "\ s"] by auto lemma pivot_tableau_valuated: "\\ (\ s); x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i; \ s\ \ \ (pivot x\<^sub>i x\<^sub>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\<^sub>i x\<^sub>j c s = update x\<^sub>i c (pivot x\<^sub>i x\<^sub>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\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \ s" assumes pivotandupdate_unsat_core_id: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>c (pivot_and_update x\<^sub>i x\<^sub>j c s) = \\<^sub>c s" assumes pivotandupdate_bounds_id: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\<^sub>i (pivot_and_update x\<^sub>i x\<^sub>j c s) = \\<^sub>i s" assumes pivotandupdate_tableau_normalized: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (\ (pivot_and_update x\<^sub>i x\<^sub>j c s))" assumes pivotandupdate_tableau_equiv: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ (v::'a valuation) \\<^sub>t \ s \ v \\<^sub>t \ (pivot_and_update x\<^sub>i x\<^sub>j c s)" assumes pivotandupdate_satisfies_tableau: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \\ s\ \\<^sub>t \ s \ \\ (pivot_and_update x\<^sub>i x\<^sub>j c s)\ \\<^sub>t \ s" assumes pivotandupdate_rvars: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ rvars (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" assumes pivotandupdate_lvars: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ lvars (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" assumes pivotandupdate_valuation_nonlhs: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ x \ lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j} \ look (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) x = (if x = x\<^sub>i then Some c else look (\ s) x)" assumes pivotandupdate_tableau_valuated: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ \ (pivot_and_update x\<^sub>i x\<^sub>j c s)" begin lemma pivotandupdate_bounds_id': assumes "\ (\ s)" "\ s" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\\ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \\ s" "\ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \ s" "\ (pivot_and_update x\<^sub>i x\<^sub>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\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i\ \ look (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) x\<^sub>i = Some c" using pivotandupdate_valuation_nonlhs[of s x\<^sub>i x\<^sub>j x\<^sub>i c] using rvars_of_lvar_rvars by (auto simp add: normalized_tableau_def) lemma pivotandupdate_valuation_other_nolhs: "\\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i; x \ lvars (\ s); x \ x\<^sub>j\ \ look (\ (pivot_and_update x\<^sub>i x\<^sub>j c s)) x = look (\ s) x" using pivotandupdate_valuation_nonlhs[of s x\<^sub>i x\<^sub>j x c] by auto lemma pivotandupdate_nolhs: "\ \ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \\<^sub>l s x\<^sub>i = Some c \ \\<^sub>u s x\<^sub>i = Some c\ \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (pivot_and_update x\<^sub>i x\<^sub>j c s)" using pivotandupdate_satisfies_tableau[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_tableau_equiv[of s x\<^sub>i x\<^sub>j _ c] using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_valuation_other_nolhs[of s x\<^sub>i x\<^sub>j _ c] using pivotandupdate_lvars[of s x\<^sub>i x\<^sub>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\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\ (pivot_and_update x\<^sub>i x\<^sub>j c s) = \ s" using assms pivotandupdate_bounds_id'[of s x\<^sub>i x\<^sub>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\<^sub>i x\<^sub>j c s \ update x\<^sub>i c (pivot x\<^sub>i x\<^sub>j s)" lemma pivot_update_precond: assumes "\ (\ s)" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" shows "\ (\ (pivot x\<^sub>i x\<^sub>j s))" "x\<^sub>i \ lvars (\ (pivot x\<^sub>i x\<^sub>j s))" proof- from assms have "x\<^sub>i \ x\<^sub>j" using rvars_of_lvar_rvars[of x\<^sub>i "\ s"] by (auto simp add: normalized_tableau_def) then show "\ (\ (pivot x\<^sub>i x\<^sub>j s))" "x\<^sub>i \ lvars (\ (pivot x\<^sub>i x\<^sub>j s))" using assms using pivot_tableau_normalized[of s x\<^sub>i x\<^sub>j] using pivot_lvars[of s x\<^sub>i x\<^sub>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 \\<^sub>u\<^sub>b \\<^sub>u s x then s else let s' = s \ \\<^sub>u := (\\<^sub>u s) (x := Some c) \ in if c <\<^sub>l\<^sub>b \\<^sub>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 \\<^sub>t t) = (w \\<^sub>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) \\<^sub>i\<^sub>s\<^sub>e s" and ivalid: "index_valid as s" and dist: "distinct_indices_atoms as" shows "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as" "v \\<^sub>t \ s" unfolding i_satisfies_atom_set'.simps proof (intro ballI) from model[unfolded satisfies_state_index'.simps] have model: "v \\<^sub>t \ s" "(J, v) \\<^sub>i\<^sub>b\<^sub>e \\ s" by auto show "v \\<^sub>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 (\\<^sub>i\<^sub>l s) x = Some (i, c) \ look (\\<^sub>i\<^sub>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 \\<^sub>a\<^sub>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\ \\<^sub>t t \ \v\ \\<^sub>a\<^sub>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 (\\<^sub>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 \\<^sub>t t" "(set I, v) \\<^sub>i\<^sub>a\<^sub>s set as" from unsat have no_model: "\ ((set I, v) \\<^sub>i\<^sub>s ?s)" by auto from assert_all_state_unsat_atoms_equiv_bounds[OF D refl U] have equiv: "set as \\<^sub>i \\ ?s" by auto from assert_all_state_tableau_equiv[OF D refl, of v] model have model_t: "v \\<^sub>t \ ?s" by auto have model_as': "(set I, v) \\<^sub>i\<^sub>a\<^sub>s set as" using model(2) by (auto simp: satisfies_atom_set_def) with equiv model_t have "(set I, v) \\<^sub>i\<^sub>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) \\<^sub>i\<^sub>s\<^sub>e ?s" by blast have "(J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s set as" "v \\<^sub>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 \\<^sub>t t \ (J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>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: "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 (\\<^sub>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 (\\<^sub>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 (\\<^sub>c s)) \ indices_state s \ \ (\v. (v \\<^sub>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 (\\<^sub>c s)) \ (\ v. v \\<^sub>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' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u undefined \\<^sub>l \\<^sub>u Geq Leq s' = ?P' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l undefined \\<^sub>u \\<^sub>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 (\\<^sub>c s')). y x"] ext arg_cong[of _ _ Not], unfold id, auto) show "?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u undefined \\<^sub>l \\<^sub>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 (\\<^sub>c s')). y x"] ext; intro arg_cong[of _ _ Ex] ext), auto) then show "?P s' = ?P' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l undefined \\<^sub>u \\<^sub>l Leq Geq s'" unfolding id' . next fix c::'a and x::nat and dir assume "\\<^sub>l\<^sub>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 simp del: Simplex.bounds_lg) } 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 (\\<^sub>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 \\<^sub>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\ \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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)" "\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" "\ \\<^sub>l\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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)\ \\<^sub>t \ (assert_bound ia s)" proof- let ?P = "\ lt UBI LBI UB LB UBI_upd UI LI LE GE s. \\ s\ \\<^sub>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')\ \\<^sub>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)\ \\<^sub>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). \\<^sub>l\<^sub>b lt (\\ s\ x) (LB s x) \ \\<^sub>u\<^sub>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' (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>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] \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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)" "\ \\<^sub>l\<^sub>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 "\\<^sub>l\<^sub>b (lt dir) (\\ ?s'\ y) (LB dir ?s' y) \ \\<^sub>u\<^sub>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 \\ \\<^sub>l\<^sub>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' (<) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>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)" "\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" "\ \\<^sub>l\<^sub>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 "\ \\<^sub>u\<^sub>b (lt dir) c (UB dir s x)" "\ \\<^sub>l\<^sub>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 "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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" "\\<^sub>u\<^sub>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" "\ (\\<^sub>u\<^sub>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" "\ (\\<^sub>u\<^sub>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" "\ (\\<^sub>u\<^sub>b (lt dir) c (UB dir s x))" "x \ lvars (\ s)" then have "\ (\ ?s)" "\ ?s" "x \ lvars (\ ?s)" using \\ (\ s)\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ \\ s\ by (auto simp: tableau_valuated_def) from update_bounds_id[OF this, of c] have "\\<^sub>i ?s' = \\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = P (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" "\ s" and ats: "ats \\<^sub>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 \\<^sub>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) \\<^sub>i\<^sub>a\<^sub>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 \\<^sub>i \\ s' \ Q bs (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" "bs \\<^sub>i \\ s' \ Q bs (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>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' (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s'" and xx: "\ s'. ?P s' = ?P' (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>u Geq Leq s'" using equiv by blast+ from ats equiv[of ats s] have Q_ats: "Q ats (<) \\<^sub>i\<^sub>u \\<^sub>i\<^sub>l \\<^sub>u \\<^sub>l \\<^sub>i\<^sub>u_update \\<^sub>u \\<^sub>l Leq Geq s" "Q ats (>) \\<^sub>i\<^sub>l \\<^sub>i\<^sub>u \\<^sub>l \\<^sub>u \\<^sub>i\<^sub>l_update \\<^sub>l \\<^sub>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) \\<^sub>i\<^sub>a\<^sub>s A" from i_satisfies_atom_set_mono[OF ats_sub this] have "(I, v) \\<^sub>i\<^sub>a\<^sub>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 \\<^sub>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) 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 \\<^sub>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 \ \Lhs var of @{text eq} and @{text x\<^sub>j} are swapped, while the other variables do not change sides.\ vars_pivot_eq:" \x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ let eq' = pivot_eq eq x\<^sub>j in lhs eq' = x\<^sub>j \ rvars_eq eq' = {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" and \ \Pivoting keeps the equation equisatisfiable.\ equiv_pivot_eq: "\x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ (v::'a::lrv valuation) \\<^sub>e pivot_eq eq x\<^sub>j \ v \\<^sub>e eq" begin lemma lhs_pivot_eq: "\x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ lhs (pivot_eq eq x\<^sub>j) = x\<^sub>j" using vars_pivot_eq by (simp add: Let_def) lemma rvars_pivot_eq: "\x\<^sub>j \ rvars_eq eq; lhs eq \ rvars_eq eq \ \ rvars_eq (pivot_eq eq x\<^sub>j) = {lhs eq} \ (rvars_eq eq - {x\<^sub>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 \ \Effect of @{text "subst_var x\<^sub>j lp' lp"} on @{text lp} variables.\ vars_subst_var': "(vars lp - {x\<^sub>j}) - vars lp' \s vars (subst_var x\<^sub>j lp' lp) \s (vars lp - {x\<^sub>j}) \ vars lp'"and subst_no_effect: "x\<^sub>j \ vars lp \ subst_var x\<^sub>j lp' lp = lp" and subst_with_effect: "x\<^sub>j \ vars lp \ x \ vars lp' - vars lp \ x \ vars (subst_var x\<^sub>j lp' lp)" and \ \Effect of @{text "subst_var x\<^sub>j lp' lp"} on @{text lp} value.\ equiv_subst_var: "(v::'a :: lrv valuation) x\<^sub>j = lp' \v\ \ lp \v\ = (subst_var x\<^sub>j lp' lp) \v\" begin lemma vars_subst_var: "vars (subst_var x\<^sub>j lp' lp) \ (vars lp - {x\<^sub>j}) \ vars lp'" using vars_subst_var' by simp lemma vars_subst_var_supset: "vars (subst_var x\<^sub>j lp' lp) \ (vars lp - {x\<^sub>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\<^sub>j lp eq) \ (rvars_eq eq - {x\<^sub>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\<^sub>j lp eq) \ (rvars_eq eq) - {x\<^sub>j} - (vars lp)" unfolding subst_var_eq_def by (simp add: vars_subst_var_supset) lemma equiv_subst_var_eq: assumes "(v::'a valuation) \\<^sub>e (x\<^sub>j, lp')" shows "v \\<^sub>e eq \ v \\<^sub>e subst_var_eq x\<^sub>j lp' eq" using assms unfolding subst_var_eq_def unfolding satisfies_eq_def using equiv_subst_var[of v x\<^sub>j lp' "rhs eq"] by auto end locale Pivot' = EqForLVar + PivotEq + SubstVar begin definition pivot_tableau' :: "var \ var \ tableau \ tableau" where "pivot_tableau' x\<^sub>i x\<^sub>j t \ let x\<^sub>i_idx = eq_idx_for_lvar t x\<^sub>i; eq = t ! x\<^sub>i_idx; eq' = pivot_eq eq x\<^sub>j in map (\ idx. if idx = x\<^sub>i_idx then eq' else subst_var_eq x\<^sub>j (rhs eq') (t ! idx) ) [0.. var \ ('i,'a::lrv) state \ ('i,'a) state" where "pivot' x\<^sub>i x\<^sub>j s \ \_update (pivot_tableau' x\<^sub>i x\<^sub>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\<^sub>i x\<^sub>j t \ let eq = eq_for_lvar t x\<^sub>i; eq' = pivot_eq eq x\<^sub>j in map (\ e. if lhs e = lhs eq then eq' else subst_var_eq x\<^sub>j (rhs eq') e) t" definition pivot :: "var \ var \ ('i,'a::lrv) state \ ('i,'a) state" where "pivot x\<^sub>i x\<^sub>j s \ \_update (pivot_tableau x\<^sub>i x\<^sub>j (\ s)) s" lemma pivot_tableau'pivot_tableau: assumes "\ t" "x\<^sub>i \ lvars t" shows "pivot_tableau' x\<^sub>i x\<^sub>j t = pivot_tableau x\<^sub>i x\<^sub>j t" proof- let ?f = "\idx. if idx = eq_idx_for_lvar t x\<^sub>i then pivot_eq (t ! eq_idx_for_lvar t x\<^sub>i) x\<^sub>j else subst_var_eq x\<^sub>j (rhs (pivot_eq (t ! eq_idx_for_lvar t x\<^sub>i) x\<^sub>j)) (t ! idx)" let ?f' = "\e. if lhs e = lhs (eq_for_lvar t x\<^sub>i) then pivot_eq (eq_for_lvar t x\<^sub>i) x\<^sub>j else subst_var_eq x\<^sub>j (rhs (pivot_eq (eq_for_lvar t x\<^sub>i) x\<^sub>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\<^sub>i \ set t" "eq_idx_for_lvar t x\<^sub>i < length t" using eq_for_lvar[of x\<^sub>i t] \x\<^sub>i \ lvars t\ eq_idx_for_lvar[of x\<^sub>i t] by (auto simp add: eq_for_lvar_def) ultimately have "lhs (t ! i) = lhs (t ! eq_idx_for_lvar t x\<^sub>i) \ t ! i = t ! (eq_idx_for_lvar t x\<^sub>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\<^sub>i) \ i = eq_idx_for_lvar t x\<^sub>i" using \i < length t\ \eq_idx_for_lvar t x\<^sub>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\<^sub>i x\<^sub>j t = pivot_tableau x\<^sub>i x\<^sub>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\<^sub>i \ lvars (\ s)" shows "pivot' x\<^sub>i x\<^sub>j s = pivot x\<^sub>i x\<^sub>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\<^sub>i x\<^sub>j and v::"'a valuation" assume "\ (\ s)" "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" show "let s' = pivot x\<^sub>i x\<^sub>j s in \ s' = \ s \ \\<^sub>i s' = \\<^sub>i s \ \ s' = \ s \ \\<^sub>c s' = \\<^sub>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\<^sub>i" let ?eq = "?t ! ?idx" let ?eq' = "pivot_eq ?eq x\<^sub>j" have "?idx < length ?t" "lhs (?t ! ?idx) = x\<^sub>i" using \x\<^sub>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\<^sub>j \ rvars_eq ?eq" using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ unfolding eq_for_lvar_def by simp then have "x\<^sub>j \ rvars ?t" using \?idx < length ?t\ using in_set_conv_nth[of ?eq ?t] by (auto simp add: rvars_def) then have "x\<^sub>j \ lvars ?t" using \\ ?t\ unfolding normalized_tableau_def by auto have "x\<^sub>i \ rvars ?t" using \x\<^sub>i \ lvars ?t\ \\ ?t\ unfolding normalized_tableau_def rvars_def by auto then have "x\<^sub>i \ rvars_eq ?eq" unfolding rvars_def using \?idx < length ?t\ using in_set_conv_nth[of ?eq ?t] by auto have "x\<^sub>i \ x\<^sub>j" using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ rvars_eq ?eq\ by auto have "?eq' = (x\<^sub>j, rhs ?eq')" using lhs_pivot_eq[of x\<^sub>j ?eq] using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ \lhs (?t ! ?idx) = x\<^sub>i\ \x\<^sub>i \ rvars_eq ?eq\ by (auto simp add: eq_for_lvar_def) (cases "?eq'", simp)+ let ?I1 = "[0..?idx < length ?t\ by (rule interval_3split) then have map_lhs_pivot: "map lhs (\ (pivot' x\<^sub>i x\<^sub>j s)) = map (\idx. lhs (?t ! idx)) ?I1 @ [x\<^sub>j] @ map (\idx. lhs (?t ! idx)) ?I2" using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ \lhs (?t ! ?idx) = x\<^sub>i\ \x\<^sub>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\<^sub>i x\<^sub>j s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" proof- have "lvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = {x\<^sub>j} \ (\idx. lhs (?t ! idx)) ` ({0..?idx < length ?t\ \?eq' = (x\<^sub>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\<^sub>j} \ (((\idx. lhs (?t ! idx)) ` {0..?idx < length ?t\ \distinct (map lhs ?t)\ by (auto simp add: distinct_conv_nth) also have "... = {x\<^sub>j} \ (set (map lhs ?t) - {x\<^sub>i})" using \lhs (?t ! ?idx) = x\<^sub>i\ by (auto simp add: in_set_conv_nth rev_image_eqI) (auto simp add: image_def) finally show "lvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" by (simp add: lvars_def) qed moreover have rvars_pivot: "rvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" proof- have "rvars_eq ?eq' = {x\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>j})" using rvars_pivot_eq[of x\<^sub>j ?eq] using \lhs (?t ! ?idx) = x\<^sub>i\ using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ rvars_eq ?eq\ by simp let ?S1 = "rvars_eq ?eq'" let ?S2 = "\idx\({0..j (rhs ?eq') (?t ! idx))" have "rvars (\ (pivot' x\<^sub>i x\<^sub>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\<^sub>i} \ (rvars ?t - {x\<^sub>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\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>j})\ by (force simp add: in_set_conv_nth) moreover have "?S2 \ ?rhs" proof- have "?S2 \ (\idx\{0..j}) \ rvars_eq ?eq')" apply (rule UN_mono) using rvars_eq_subst_var_eq by auto also have "... \ rvars_eq ?eq' \ (\idx\{0..j})" by auto also have "... = rvars_eq ?eq' \ (rvars ?t - {x\<^sub>j})" unfolding rvars_def by (force simp add: in_set_conv_nth) finally show ?thesis using \rvars_eq ?eq' = {x\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>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..j}) - rvars_eq ?eq'" have "x \ ?S2'" using False \x \ ?rhs\ using \rvars_eq ?eq' = {x\<^sub>i} \ (rvars_eq ?eq - {x\<^sub>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\<^sub>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\<^sub>i x\<^sub>j s in rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i} \ lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" using pivot'pivot[where ?'i = 'i] using \\ (\ s)\ \x\<^sub>i \ lvars (\ s)\ by (simp add: Let_def) have "\ (\ (pivot' x\<^sub>i x\<^sub>j s))" unfolding normalized_tableau_def proof have "lvars (\ (pivot' x\<^sub>i x\<^sub>j s)) \ rvars (\ (pivot' x\<^sub>i x\<^sub>j s)) = {}" (is ?g1) using \\ (\ s)\ unfolding normalized_tableau_def using lvars_pivot rvars_pivot using \x\<^sub>i \ x\<^sub>j\ by auto moreover have "0 \ rhs ` set (\ (pivot' x\<^sub>i x\<^sub>j s))" (is ?g2) proof let ?eq = "eq_for_lvar (\ s) x\<^sub>i" from eq_for_lvar[OF \x\<^sub>i \ lvars (\ s)\] have "?eq \ set (\ s)" and var: "lhs ?eq = x\<^sub>i" by auto have "lhs ?eq \ rvars_eq ?eq" using \\ (\ s)\ \?eq \ set (\ s)\ using \x\<^sub>i \ rvars_eq (\ s ! eq_idx_for_lvar (\ s) x\<^sub>i)\ eq_for_lvar_def var by auto from vars_pivot_eq[OF \x\<^sub>j \ rvars_eq ?eq\ this] have vars_pivot: "lhs (pivot_eq ?eq x\<^sub>j) = x\<^sub>j" "rvars_eq (pivot_eq ?eq x\<^sub>j) = {lhs (eq_for_lvar (\ s) x\<^sub>i)} \ (rvars_eq (eq_for_lvar (\ s) x\<^sub>i) - {x\<^sub>j})" unfolding Let_def by auto from vars_pivot(2) have rhs_pivot0: "rhs (pivot_eq ?eq x\<^sub>j) \ 0" using vars_zero by auto assume "0 \ rhs ` set (\ (pivot' x\<^sub>i x\<^sub>j s))" from this[unfolded pivot'pivot[OF \\ (\ s)\ \x\<^sub>i \ lvars (\ s)\] pivot_def] have "0 \ rhs ` set (pivot_tableau x\<^sub>i x\<^sub>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\<^sub>i" and rvars_eq: "rvars_eq (subst_var_eq x\<^sub>j (rhs (pivot_eq ?eq x\<^sub>j)) e) = {}" by (auto simp: vars_zero) from rvars_eq[unfolded subst_var_eq_def] have empty: "vars (subst_var x\<^sub>j (rhs (pivot_eq ?eq x\<^sub>j)) (rhs e)) = {}" by auto show False proof (cases "x\<^sub>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\<^sub>i \ vars (rhs (pivot_eq ?eq x\<^sub>j)) - rvars_eq e" unfolding vars_pivot(2) var using \\ (\ s)\[unfolded normalized_tableau_def] \x\<^sub>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\<^sub>i x\<^sub>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\<^sub>j \ lvars ?t\ unfolding lvars_def by auto qed moreover have "v \\<^sub>t ?t = v \\<^sub>t \ (pivot' x\<^sub>i x\<^sub>j s)" unfolding satisfies_tableau_def proof assume "\e\set (?t). v \\<^sub>e e" show "\e\set (\ (pivot' x\<^sub>i x\<^sub>j s)). v \\<^sub>e e" proof- have "v \\<^sub>e ?eq'" using \x\<^sub>i \ rvars_eq ?eq\ using \?idx < length ?t\ \\e\set (?t). v \\<^sub>e e\ using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ lvars ?t\ by (simp add: equiv_pivot_eq eq_idx_for_lvar) moreover { fix idx assume "idx < length ?t" "idx \ ?idx" have "v \\<^sub>e subst_var_eq x\<^sub>j (rhs ?eq') (?t ! idx)" using \?eq' = (x\<^sub>j, rhs ?eq')\ using \v \\<^sub>e ?eq'\ \idx < length ?t\ \\e\set (?t). v \\<^sub>e e\ using equiv_subst_var_eq[of v x\<^sub>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\<^sub>i x\<^sub>j s)). v \\<^sub>e e" then have "v \\<^sub>e ?eq'" "\ idx. \idx < length ?t; idx \ ?idx \ \ v \\<^sub>e subst_var_eq x\<^sub>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 \\<^sub>e e" proof- { fix idx assume "idx < length ?t" have "v \\<^sub>e (?t ! idx)" proof (cases "idx = ?idx") case True then show ?thesis using \v \\<^sub>e ?eq'\ using \x\<^sub>j \ rvars_eq ?eq\ \x\<^sub>i \ lvars ?t\ \x\<^sub>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 \\<^sub>e subst_var_eq x\<^sub>j (rhs ?eq') (?t ! idx)\ using \v \\<^sub>e ?eq'\ \?eq' = (x\<^sub>j, rhs ?eq')\ using equiv_subst_var_eq[of v x\<^sub>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\<^sub>i x\<^sub>j s in v \\<^sub>t \ s = v \\<^sub>t \ s' \ \ (\ s')" using pivot'pivot[where ?'i = 'i] using \\ (\ s)\ \x\<^sub>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\<^sub>i\ out of its bounds, a rhs variable \x\<^sub>j\ is sought which allows pivoting with \x\<^sub>i\ and updating \x\<^sub>i\ to its violated bound. If \x\<^sub>i\ is under its lower bound it must be increased, and if \x\<^sub>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\<^sub>i\ is above its upper bound is symmetric (avoiding symmetries is discussed in Section \ref{sec:symmetries}). If there is no such \x\<^sub>j\, unsatisfiability is detected and \check\ fails. The procedure is recursively repeated, until it either succeeds or fails. To ensure termination, variables \x\<^sub>i\ and \x\<^sub>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\<^sub>i \ x\<^sub>i\lvars (\ s) \ \in_bounds x\<^sub>i \\ s\ (\ s) \ (\x\lvars (\ s). x < x\<^sub>i \ in_bounds x \\ s\ (\ s))" begin lemma min_lvar_not_in_bounds_None': "min_lvar_not_in_bounds s = None \ (\\ s\ \\<^sub>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\<^sub>i \ x\<^sub>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\<^sub>i \ \ in_bounds x\<^sub>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\<^sub>i \ (\ x \ lvars (\ s). x < x\<^sub>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 \ \\<^sub>u\<^sub>b (lt dir) (\\ s\ x) (UB dir s x)) \ (coeff (rhs eq) x < 0 \ \\<^sub>l\<^sub>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\<^sub>j \ x\<^sub>j \ rvars_eq eq" assumes min_rvar_incdec_eq_Some_incdec: "min_rvar_incdec_eq dir s eq = Inr x\<^sub>j \ reasable_var dir x\<^sub>j eq s" assumes min_rvar_incdec_eq_Some_min: "min_rvar_incdec_eq dir s eq = Inr x\<^sub>j \ (\ x \ rvars_eq eq. x < x\<^sub>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) \\<^sub>i\<^sub>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 "\ (\\<^sub>u\<^sub>b (lt dir) (\\ s\ x) (UB dir s x))" using \\ x \ rvars_eq eq. \ reasable_var dir x eq s\ by auto then have "\\<^sub>u\<^sub>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 "\\<^sub>u\<^sub>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 simp del: Simplex.bounds_lg) next fix x assume x: "x \ rvars_eq eq" "0 > coeff (rhs eq) x" "\\ s\ x - v x \ 0" then have "\ (\\<^sub>l\<^sub>b (lt dir) (\\ s\ x) (LB dir s x))" using \\ x \ rvars_eq eq. \ reasable_var dir x eq s\ by auto then have "\\<^sub>l\<^sub>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 "\\<^sub>l\<^sub>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 simp del: Simplex.bounds_lg) 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 simp del: Simplex.bounds_lg) using valuate_nonpos[of "rhs eq" "\x. \\ s\ x - v x"] apply (force simp del: Simplex.bounds_lg) 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] simp del: Simplex.bounds_lg) 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\<^sub>i \ min_rvar_incdec_eq dir s (eq_for_lvar (\ s) x\<^sub>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\<^sub>i s \ let l\<^sub>i = the (LB dir s x\<^sub>i); x\<^sub>j' = min_rvar_incdec dir s x\<^sub>i in case x\<^sub>j' of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s" lemma check'_cases: assumes "\ I. \min_rvar_incdec dir s x\<^sub>i = Inl I; check' dir x\<^sub>i s = set_unsat I s\ \ P (set_unsat I s)" assumes "\ x\<^sub>j l\<^sub>i. \min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j; l\<^sub>i = the (LB dir s x\<^sub>i); check' dir x\<^sub>i s = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s\ \ P (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)" shows "P (check' dir x\<^sub>i s)" using assms unfolding check'_def by (cases "min_rvar_incdec dir s x\<^sub>i", auto) partial_function (tailrec) check where "check s = (if \ s then s else let x\<^sub>i' = min_lvar_not_in_bounds s in case x\<^sub>i' of None \ s | Some x\<^sub>i \ let dir = if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative in check (check' dir x\<^sub>i s))" declare check.simps[code] inductive check_dom where step: "\\x\<^sub>i. \\ \ s; Some x\<^sub>i = min_lvar_not_in_bounds s; \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ \ check_dom (check' Positive x\<^sub>i s); \x\<^sub>i. \\ \ s; Some x\<^sub>i = min_lvar_not_in_bounds s; \ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ \ check_dom (check' Negative x\<^sub>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\<^sub>i' = min_lvar_not_in_bounds s in case x\<^sub>i' of None \ s | Some x\<^sub>i \ if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then check (check_inc x\<^sub>i s) else check (check_dec x\<^sub>i s)" } @{text[display] "check_inc x\<^sub>i s \ let l\<^sub>i = the (\\<^sub>l s x\<^sub>i); x\<^sub>j' = min_rvar_inc s x\<^sub>i in case x\<^sub>j' of None \ s \ \ := True \ | Some x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>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\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative)" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j" "l\<^sub>i = the (LB dir s x\<^sub>i)" "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" " \ s" shows "\ (\ (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)) \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s) \ \ (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s) \ \ (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s)" proof- have "\\<^sub>l s x\<^sub>i = Some l\<^sub>i \ \\<^sub>u s x\<^sub>i = Some l\<^sub>i" using \l\<^sub>i = the (LB dir s x\<^sub>i)\ \dir = (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative)\ using \min_lvar_not_in_bounds s = Some x\<^sub>i\ min_lvar_not_in_bounds_Some[of s x\<^sub>i] using \\ s\ by (case_tac[!] "\\<^sub>l s x\<^sub>i", case_tac[!] "\\<^sub>u s x\<^sub>i") (auto simp add: bounds_consistent_def bound_compare_defs) then show ?thesis using assms using pivotandupdate_tableau_normalized[of s x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_nolhs[of s x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_bounds_consistent[of s x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_tableau_valuated[of s x\<^sub>i x\<^sub>j l\<^sub>i] by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars) qed (* -------------------------------------------------------------------------- *) (* Termination *) (* -------------------------------------------------------------------------- *) abbreviation gt_state' where "gt_state' dir s s' x\<^sub>i x\<^sub>j l\<^sub>i \ min_lvar_not_in_bounds s = Some x\<^sub>i \ l\<^sub>i = the (LB dir s x\<^sub>i) \ min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j \ s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s" definition gt_state :: "('i,'a) state \ ('i,'a) state \ bool" (infixl "\\<^sub>x" 100) where "s \\<^sub>x s' \ \ x\<^sub>i x\<^sub>j l\<^sub>i. let dir = if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative in gt_state' dir s s' x\<^sub>i x\<^sub>j l\<^sub>i" abbreviation succ :: "('i,'a) state \ ('i,'a) state \ bool" (infixl "\" 100) where "s \ s' \ \ (\ s) \ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s \ \ s \ s \\<^sub>x s' \ \\<^sub>i s' = \\<^sub>i s \ \\<^sub>c s' = \\<^sub>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 "\\<^sup>+" 100) where "s \\<^sup>+ s' \ (s, s') \ succ_rel\<^sup>+" abbreviation succ_rel_rtrancl :: "('i,'a) state \ ('i,'a) state \ bool" (infixl "\\<^sup>*" 100) where "s \\<^sup>* s' \ (s, s') \ succ_rel\<^sup>*" lemma succ_vars: assumes "s \ s'" obtains x\<^sub>i x\<^sub>j where "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" "x\<^sub>j \ rvars (\ s)" "lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" proof- from assms obtain x\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" unfolding gt_state_def by (auto split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" "lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j] using pivotandupdate_lvars[of s x\<^sub>i x\<^sub>j] by auto moreover have "x\<^sub>j \ rvars (\ s)" using \x\<^sub>i \ lvars (\ s)\ using \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ using eq_for_lvar[of x\<^sub>i "\ s"] unfolding rvars_def by auto ultimately have "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" "x\<^sub>j \ rvars (\ s)" "lvars (\ s') = lvars (\ s) - {x\<^sub>i} \ {x\<^sub>j}" "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>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'" "\\<^sub>i s = \\<^sub>i s'" "(v::'a valuation) \\<^sub>t (\ s) \ v \\<^sub>t (\ s')" proof- from assms obtain x\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" unfolding gt_state_def by (auto split: if_splits) then show "\ (\ s')" "\ s'" "\ s'" "\\<^sub>i s = \\<^sub>i s'" "(v::'a valuation) \\<^sub>t (\ s) \ v \\<^sub>t (\ s')" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using pivotandupdate_tableau_normalized[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_bounds_consistent[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_bounds_id[of s x\<^sub>i x\<^sub>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\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>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\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using \x \ rvars (\ s)\ \x \ rvars (\ s')\ using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] using pivotandupdate_valuation_other_nolhs[of s x\<^sub>i x\<^sub>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\<^sub>i x\<^sub>j c where *: "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" unfolding gt_state_def by (auto split: if_splits) then have "x\<^sub>i = xr" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>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\<^sub>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 "\ \\<^sub>l\<^sub>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\<^sub>i x\<^sub>j c where"\ (\ s) \ \ s \ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "gt_state' (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative) s s' x\<^sub>i x\<^sub>j c" by (auto simp add: gt_state_def Let_def) then have "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "s' = pivot_and_update x\<^sub>i x\<^sub>j c s" and *: "(\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i \ min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j) \ (\ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j)" by (auto split: if_splits) then have "xr = x\<^sub>j" "xs = x\<^sub>i" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>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 "\ (\\<^sub>l\<^sub>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 "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ xs) (LB dir s xs)" then have "\\<^sub>l\<^sub>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 "\ (\\<^sub>u\<^sub>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\<^sub>j\ \xs = x\<^sub>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\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ s')" and dir: "dir = Positive \ dir = Negative" shows "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (LB dir s x\<^sub>i)" "\\ s'\ x\<^sub>i = the (\\<^sub>l s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (\\<^sub>u s x\<^sub>i)" proof- from assms(1) obtain x\<^sub>i' x\<^sub>j c where"\ (\ s) \ \ s \ \ s \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "gt_state' (if \\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' then Positive else Negative) s s' x\<^sub>i' x\<^sub>j c" by (auto simp add: gt_state_def Let_def) then have "\ (\ s)" "\ s" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i'" "s' = pivot_and_update x\<^sub>i' x\<^sub>j c s" and *: "(\\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' \ c = the (\\<^sub>l s x\<^sub>i') \ min_rvar_incdec Positive s x\<^sub>i' = Inr x\<^sub>j) \ (\ \\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' \ c = the (\\<^sub>u s x\<^sub>i') \ min_rvar_incdec Negative s x\<^sub>i' = Inr x\<^sub>j)" by (auto split: if_splits) then have "x\<^sub>i = x\<^sub>i'" "x\<^sub>i' \ lvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i')" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i'] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i'" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i'" x\<^sub>j] using \x\<^sub>i \ lvars (\ s)\ \x\<^sub>i \ rvars (\ s')\ using pivotandupdate_rvars by (auto simp add: normalized_tableau_def) show "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (LB dir s x\<^sub>i)" proof assume "\ \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i)" then have "\\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i)" using dir by (cases "LB dir s x\<^sub>i") (auto simp add: bound_compare_defs) moreover then have "\ \\<^sub>u\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (UB dir s x\<^sub>i)" using \\ s\ dir using bounds_consistent_gt_ub bounds_consistent_lt_lb by (force simp add: bound_compare''_defs) ultimately show "\\ s'\ x\<^sub>i = the (LB dir s x\<^sub>i)" using * \x\<^sub>i = x\<^sub>i'\ \s' = pivot_and_update x\<^sub>i' x\<^sub>j c s\ using \\ (\ s)\ \\ s\ \x\<^sub>i' \ lvars (\ s)\ \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i')\ using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] dir by (case_tac[!] "\\<^sub>l s x\<^sub>i'", case_tac[!] "\\<^sub>u s x\<^sub>i'") (auto simp add: bound_compare_defs map2fun_def) qed have "\ \\ s\ x\<^sub>i' <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i' \ \\ s\ x\<^sub>i' >\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i'" using \min_lvar_not_in_bounds s = Some x\<^sub>i'\ using min_lvar_not_in_bounds_Some[of s x\<^sub>i'] using not_in_bounds[of x\<^sub>i' "\\ s\" "\\<^sub>l s" "\\<^sub>u s"] by auto then show "\\ s'\ x\<^sub>i = the (\\<^sub>l s x\<^sub>i) \ \\ s'\ x\<^sub>i = the (\\<^sub>u s x\<^sub>i)" using \\ (\ s)\ \\ s\ \x\<^sub>i' \ lvars (\ s)\ \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i')\ using \s' = pivot_and_update x\<^sub>i' x\<^sub>j c s\ \x\<^sub>i = x\<^sub>i'\ using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j c] using * by (case_tac[!] "\\<^sub>l s x\<^sub>i'", case_tac[!] "\\<^sub>u s x\<^sub>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 (\\<^sub>l s x) \ \\ s'\ x = the (\\<^sub>u s x)" proof- from assms obtain x\<^sub>i x\<^sub>j b where "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>i)" "s' = pivot_and_update x\<^sub>i x\<^sub>j b s" unfolding gt_state_def by (auto simp add: Let_def split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" "x\<^sub>j \ rvars (\ s)" "x\<^sub>j \ lvars (\ s)" "x\<^sub>i \ x\<^sub>j" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using rvars_of_lvar_rvars \\ (\ s)\ by (auto simp add: normalized_tableau_def) then have "rvars (\ s') = rvars (\ s) - {x\<^sub>j} \ {x\<^sub>i}" "x \ rvars (\ s) \ x = x\<^sub>i" "x \ x\<^sub>j" "x \ x\<^sub>i \ x \ lvars (\ s)" using \x \ rvars (\ s')\ using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j] using \\ (\ s)\ \\ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>j b s\ by (auto simp add: normalized_tableau_def) then show ?thesis using pivotandupdate_valuation_xi[of s x\<^sub>i x\<^sub>j b] using pivotandupdate_valuation_other_nolhs[of s x\<^sub>i x\<^sub>j x b] using \x\<^sub>i \ lvars (\ s)\ \x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ using \\ (\ s)\ \\ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>j b s\ \b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>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\<^sub>i x\<^sub>j b where "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>i)" "s' = pivot_and_update x\<^sub>i x\<^sub>j b s" unfolding gt_state_def by (auto simp add: Let_def split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ s)" "x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)" "x\<^sub>j \ rvars (\ s)" "x\<^sub>j \ lvars (\ s)" "x\<^sub>i \ x\<^sub>j" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>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\<^sub>i x\<^sub>j x b] using \\ (\ s)\ \\ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>j b s\ using \x \ tvars (\ s')\ using pivotandupdate_rvars[of s x\<^sub>i x\<^sub>j] using pivotandupdate_lvars[of s x\<^sub>i x\<^sub>j] by (auto simp add: map2fun_def) qed lemma succ_valuation_satisfies: assumes "s \ s'" "\\ s\ \\<^sub>t \ s" shows "\\ s'\ \\<^sub>t \ s'" proof- from \s \ s'\ obtain x\<^sub>i x\<^sub>j b where "\ (\ s)" "\ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "min_rvar_incdec Positive s x\<^sub>i = Inr x\<^sub>j \ min_rvar_incdec Negative s x\<^sub>i = Inr x\<^sub>j" "b = the (\\<^sub>l s x\<^sub>i) \ b = the (\\<^sub>u s x\<^sub>i)" "s' = pivot_and_update x\<^sub>i x\<^sub>j b s" unfolding gt_state_def by (auto simp add: Let_def split: if_splits) then have "x\<^sub>i \ lvars (\ s)" "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" using min_lvar_not_in_bounds_lvars[of s x\<^sub>i] using min_rvar_incdec_eq_Some_rvars[of Positive s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using min_rvar_incdec_eq_Some_rvars[of Negative s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] \\ (\ s)\ by (auto simp add: normalized_tableau_def) then show ?thesis using pivotandupdate_satisfies_tableau[of s x\<^sub>i x\<^sub>j b] using pivotandupdate_tableau_equiv[of s x\<^sub>i x\<^sub>j ] using \\ (\ s)\ \\ s\ \\\ s\ \\<^sub>t \ s\ \s' = pivot_and_update x\<^sub>i x\<^sub>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 "\\<^sub>i (l ! i) = \\<^sub>i (l ! j)" using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "\\<^sub>i (l ! i) = \\<^sub>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) \\<^sub>t \ (l ! i) \ v \\<^sub>t \ (l ! j)" using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "v \\<^sub>t \ (l ! i) = v \\<^sub>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) \\<^sub>t \ (l ! i) \ v \\<^sub>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 (\\<^sub>l (l ! i) x) \ \\ (l ! j)\ x = the (\\<^sub>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\<^sub>i x\<^sub>j where vars: "x\<^sub>i \ lvars (\ (l ! (j - 1)))" "x\<^sub>j \ rvars (\ (l ! (j - 1)))" "rvars (\ (l ! j)) = rvars (\ (l ! (j - 1))) - {x\<^sub>j} \ {x\<^sub>i}" using \l ! (j - 1) \ (l ! j)\ by (rule succ_vars) simp then have bounds: "\\<^sub>l (l ! (j - 1)) = \\<^sub>l (l ! i)" "\\<^sub>l (l ! j) = \\<^sub>l (l ! i)" "\\<^sub>u (l ! (j - 1)) = \\<^sub>u (l ! i)" "\\<^sub>u (l ! j) = \\<^sub>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\<^sub>j \ x \ rvars (\ (l ! (j - 1))) \ x = x\<^sub>i" using vars by auto then show "\\ (l ! j)\ x = \\ (l ! i)\ x \ \\ (l ! j)\ x = the (\\<^sub>l (l ! i) x) \ \\ (l ! j)\ x = the (\\<^sub>u (l ! i) x)" proof assume "x \ x\<^sub>j \ x \ rvars (\ (l ! (j - 1)))" then show ?thesis using jj \x \ rvars (\ (l ! j))\ ji using bounds by force next assume "x = x\<^sub>i" then show ?thesis using succ_set_on_bound(2)[of "l ! (j - 1)" "l ! j" x\<^sub>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)\ \\<^sub>t \ (l ! i) \ \\ (l ! j)\ \\<^sub>t \ (l ! j)" using assms proof (rule succ_chain_induct) fix i assume "l ! i \ (l ! (i + 1))" then show "\\ (l ! i)\ \\<^sub>t \ (l ! i) \ \\ (l ! (i + 1))\ \\<^sub>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 \\<^sub>t t = v \\<^sub>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. mset t = mset ?t}" proof fix x assume "x \ ?Al L" have "mset x = mset ?t" apply (rule tableau_perm) using \?t \ ?Al L\ \x \ ?Al L\ by auto then show "x \ {t. mset t = mset ?t}" by simp qed moreover have "finite {t. mset t = mset ?t}" by (fact mset_eq_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 \\<^sub>t t = v \\<^sub>t t0)}" (is "finite ?A") proof- let ?Al = "\ L. {t. lvars t = L \ rvars t = V - L \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>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 \\<^sup>* s'})" proof- have "{s'. s \\<^sup>* s'} = {s'. s \\<^sup>+ s'} \ {s}" by (auto simp add: rtrancl_eq_or_trancl) moreover have "finite (\ ` {s'. s \\<^sup>+ s'})" (is "finite ?A") proof- let ?T = "{t. tvars t = tvars (\ s) \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t(\ s))}" have "?A \ ?T" proof fix t assume "t \ ?A" then obtain s' where "s \\<^sup>+ 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 \\<^sup>+ s'\ using succ_inv(1)[of _ s'] by (auto dest: tranclD2) moreover have "\v::'a valuation. v \\<^sub>t \ s' = v \\<^sub>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 \\<^sub>t t = v \\<^sub>t t0) \ v \\<^sub>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 \\<^sub>t t = v \\<^sub>t t0)}" let ?Vt = "\ t. {v::'a valuation. v \\<^sub>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 \\<^sub>t t = v \\<^sub>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 \\<^sub>t t" "v2 \\<^sub>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 \\<^sub>t t\ \v2 \\<^sub>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 \\<^sup>* s'})" proof- have "{s'. s \\<^sup>* s'} = {s'. s \\<^sup>+ s'} \ {s}" by (auto simp add: rtrancl_eq_or_trancl) moreover have "finite (\ ` {s'. s \\<^sup>+ s'})" (is "finite ?A") proof- let ?P = "\ v. check_valuation v (\\ s\) (\ x. the (\\<^sub>l s x)) (\ x. the (\\<^sub>u s x)) (\ s) (tvars (\ s))" let ?P' = "\ v::(var, 'a) mapping. \ t. tvars t = tvars (\ s) \ \ t \ (\ v::'a valuation. v \\<^sub>t t = v \\<^sub>t \ s) \ \v\ \\<^sub>t t \ (\ x \ rvars t. \v\ x = \\ s\ x \ \v\ x = the (\\<^sub>l s x) \ \v\ x = the (\\<^sub>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 (\\<^sub>l s x)" "\ x. the (\\<^sub>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 \\<^sup>+ 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\ \\<^sub>t \ s" using \s \\<^sup>+ 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 \\<^sup>+ s'\ using succ_inv(1)[of _ s'] by (auto dest: tranclD2) moreover have "\v::'a valuation. v \\<^sub>t \ s' = v \\<^sub>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'\ \\<^sub>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\ \\<^sub>t \ s\ by simp moreover have "\x\rvars (\ s'). \\ s'\ x = \\ s\ x \ \\ s'\ x = the (\\<^sub>l s x) \ \\ s'\ x = the (\\<^sub>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 "\\<^sub>i ` {s'. s \\<^sup>* s'} = {\\<^sub>i s}" proof - have "s \\<^sup>* s' \ \\<^sub>i s' = \\<^sub>i s" for s' by (induct s s' rule: rtrancl.induct, auto) then show ?thesis by blast qed lemma accessible_unsat_core: shows "\\<^sub>c ` {s'. s \\<^sup>* s'} = {\\<^sub>c s}" proof - have "s \\<^sup>* s' \ \\<^sub>c s' = \\<^sub>c s" for s' by (induct s s' rule: rtrancl.induct, auto) then show ?thesis by blast qed lemma state_eqI: "\\<^sub>i\<^sub>l s = \\<^sub>i\<^sub>l s' \ \\<^sub>i\<^sub>u s = \\<^sub>i\<^sub>u s' \ \ s = \ s' \ \ s = \ s' \ \ s = \ s' \ \\<^sub>c s = \\<^sub>c s' \ s = s'" by (cases s, cases s', auto) lemma finite_accessible_states: shows "finite {s'. s \\<^sup>* s'}" (is "finite ?A") proof- let ?V = "\ ` ?A" let ?T = "\ ` ?A" let ?P = "?V \ ?T \ {\\<^sub>i s} \ {True, False} \ {\\<^sub>c s}" have "finite ?P" using finite_accessible_valuations finite_accessible_tableaus by auto moreover let ?f = "\ s. (\ s, \ s, \\<^sub>i s, \ s, \\<^sub>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\<^sup>+" proof assume "s \\<^sup>+ 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\<^sub>i x\<^sub>j :: var where "x\<^sub>i \ lvars (\ s)" "x\<^sub>i \ rvars (\ (l ! 1))" by (rule succ_vars) auto then have "x\<^sub>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 "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?sl" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?sp" using \l ! sl \ (l ! (sl + 1))\ using \l' ! sp \ (l' ! (sp+ 1))\ by simp_all have "\\<^sub>i ?sp = \\<^sub>i ?sl" proof- have "\\<^sub>i (l' ! sp) = \\<^sub>i (l' ! (length l' - 1))" using \sp + 1 < length l'\ \succ_chain l'\ using succ_chain_bounds_id by auto then have "\\<^sub>i (last l') = \\<^sub>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\ \\<^sub>e ?eq" using \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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\ \\<^sub>e ?eq" proof- have "\\ (l ! sl)\ \\<^sub>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 \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 <\<^sub>l\<^sub>b \\<^sub>l ?sl xr then Positive else Negative)" then have "\\<^sub>l\<^sub>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 "\ \\<^sub>l\<^sub>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 \\\<^sub>l\<^sub>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 <\<^sub>l\<^sub>b \\<^sub>l ?sp xs then Positive else Negative)" then have "\\<^sub>l\<^sub>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 "\ \\<^sub>l\<^sub>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 \\\<^sub>l\<^sub>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 \ \\<^sub>u\<^sub>b (lt dir2) (\?bp\ x) (UB dir2 ?sp x)) \ (coeff (rhs ?eq) x < 0 \ \\<^sub>l\<^sub>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 "\ \\<^sub>u\<^sub>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 "\\<^sub>u\<^sub>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 "\ \\<^sub>l\<^sub>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 "\\<^sub>l\<^sub>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 \\\<^sub>l\<^sub>b (lt dir2) (\?bp\ xs) (LB dir2 ?sp xs)\ using \\\<^sub>i ?sp = \\<^sub>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 "\\<^sub>u\<^sub>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)\ \\\<^sub>i ?sp = \\<^sub>i ?sl\ dir2 apply (auto simp add: bound_compare''_defs) using bounds_lg(3)[of "\?bp\ x" "\\<^sub>u (l ! sl) x" "\?bl\ x"] using bounds_lg(6)[of "\?bp\ x" "\\<^sub>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 simp del: Simplex.bounds_lg) next assume "coeff (rhs ?eq) x < 0" "\?bp\ x - \?bl\ x \ 0" then have "\\<^sub>l\<^sub>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)\ \\\<^sub>i ?sp = \\<^sub>i ?sl\ dir2 apply (auto simp add: bound_compare''_defs) using bounds_lg(3)[of "\?bp\ x" "\\<^sub>u (l ! sl) x" "\?bl\ x"] using bounds_lg(6)[of "\?bp\ x" "\\<^sub>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 simp del: Simplex.bounds_lg) 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 \\ \\<^sub>l\<^sub>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 \\\<^sub>l\<^sub>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 \\ \\<^sub>l\<^sub>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 \\\<^sub>i ?sp = \\<^sub>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 simp del: Simplex.bounds_lg) 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 simp del: Simplex.bounds_lg) using valuate_nonpos[of "rhs ?eq" "\ x. \?bp\ x - \?bl\ x"] apply (force simp del: Simplex.bounds_lg) 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 simp del: Simplex.bounds_lg) 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 simp del: Simplex.bounds_lg) qed ultimately have False using diff_satified dir2 by (auto split: if_splits simp del: Simplex.bounds_lg) } 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\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative)" and *: "\ s'. \s \ s'; \ s'; \ (\ s'); \ s'; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ check_dom s'" and "\ s" "\ (\ s)" "\ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ \ s" "min_lvar_not_in_bounds s = Some x\<^sub>i" "\\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i)" shows "check_dom (case min_rvar_incdec dir s x\<^sub>i of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j (the (LB dir s x\<^sub>i)) s)" proof (cases "min_rvar_incdec dir s x\<^sub>i") case Inl then show ?thesis using check_unsat_terminates by simp next case (Inr x\<^sub>j) then have xj: "x\<^sub>j \ rvars_of_lvar (\ s) x\<^sub>i" using min_rvar_incdec_eq_Some_rvars[of _ s "eq_for_lvar (\ s) x\<^sub>i" x\<^sub>j] using dir by simp let ?s' = "pivot_and_update x\<^sub>i x\<^sub>j (the (LB dir s x\<^sub>i)) s" have "check_dom ?s'" proof (rule * ) show **: "\ ?s'" "\ (\ ?s')" "\ ?s'" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s'" using \min_lvar_not_in_bounds s = Some x\<^sub>i\ Inr using \\ s\ \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ dir using pivotandupdate_check_precond by auto have xi: "x\<^sub>i \ lvars (\ s)" using assms(8) min_lvar_not_in_bounds_lvars by blast show "s \ ?s'" unfolding gt_state_def using \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ \\ s\ using \min_lvar_not_in_bounds s = Some x\<^sub>i\ \\\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "s\<^sub>0 \\<^sup>* s" shows "check_dom s" using assms proof (induct s rule: wf_induct[of "{(y, x). s\<^sub>0 \\<^sup>* x \ x \ y}"]) show "wf {(y, x). s\<^sub>0 \\<^sup>* x \ x \ y}" proof (rule finite_acyclic_wf) let ?A = "{(s', s). s\<^sub>0 \\<^sup>* s \ s \ s'}" let ?B = "{s. s\<^sub>0 \\<^sup>* 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\<^sub>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\<^sub>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\<^sub>0 \\<^sup>* x \ x \ y} \ \ s' \ \ (\ s') \ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ s\<^sub>0 \\<^sup>* s' \ check_dom s'" "\ s" "\ (\ s)" "\ s" " \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "s\<^sub>0 \\<^sup>* s" then have *: "\ s'. \s \ s'; \ s'; \ (\ s'); \ s'; \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ check_dom s'" using rtrancl_into_trancl1[of s\<^sub>0 s succ_rel] using trancl_into_rtrancl[of s\<^sub>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\<^sub>i assume "\ \ s" "Some x\<^sub>i = min_lvar_not_in_bounds s" "\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i" have "\\<^sub>l s x\<^sub>i = LB Positive s x\<^sub>i" by simp show "check_dom (case min_rvar_incdec Positive s x\<^sub>i of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j (the (\\<^sub>l s x\<^sub>i)) s)" apply (subst \\\<^sub>l s x\<^sub>i = LB Positive s x\<^sub>i\) apply (rule check_sat_terminates'_aux[of Positive s x\<^sub>i]) using \\ s\ \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ * using \\ \ s\ \Some x\<^sub>i = min_lvar_not_in_bounds s\ \\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ by (simp_all add: bound_compare''_defs) next fix x\<^sub>i assume "\ \ s" "Some x\<^sub>i = min_lvar_not_in_bounds s" "\ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i" then have "\\ s\ x\<^sub>i >\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i" using min_lvar_not_in_bounds_Some[of s x\<^sub>i] using neg_bounds_compare(7) neg_bounds_compare(2) by auto have "\\<^sub>u s x\<^sub>i = LB Negative s x\<^sub>i" by simp show "check_dom (case min_rvar_incdec Negative s x\<^sub>i of Inl I \ set_unsat I s | Inr x\<^sub>j \ pivot_and_update x\<^sub>i x\<^sub>j (the (\\<^sub>u s x\<^sub>i)) s)" apply (subst \\\<^sub>u s x\<^sub>i = LB Negative s x\<^sub>i\) apply (rule check_sat_terminates'_aux) using \\ s\ \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ * using \\ \ s\ \Some x\<^sub>i = min_lvar_not_in_bounds s\ \\\ s\ x\<^sub>i >\<^sub>u\<^sub>b \\<^sub>u s x\<^sub>i\ \\ \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i\ by (simp_all add: bound_compare''_defs) qed qed lemma check_sat_terminates: assumes "\ s" "\ (\ s)" "\ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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\<^sub>i dir I. \dir = Positive \ dir = Negative; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inl I\ \ P (set_unsat I s)" assumes "\ x\<^sub>i x\<^sub>j l\<^sub>i dir. \dir = (if \\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i then Positive else Negative); \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inr x\<^sub>j; l\<^sub>i = the (LB dir s x\<^sub>i); check' dir x\<^sub>i s = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s\ \ P (check (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s))" assumes "\ (\ s)" "\ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ using check.simps[of s] by simp next case (Some x\<^sub>i) let ?dir = "if (\\ s\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s x\<^sub>i) then (Positive :: ('i,'a)Direction) else Negative" let ?s' = "check' ?dir x\<^sub>i s" have "\\<^sub>l\<^sub>b (lt ?dir) (\\ s\ x\<^sub>i) (LB ?dir s x\<^sub>i)" using \min_lvar_not_in_bounds s = Some x\<^sub>i\ using min_lvar_not_in_bounds_Some[of s x\<^sub>i] using not_in_bounds[of x\<^sub>i "\\ s\" "\\<^sub>l s" "\\<^sub>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\<^sub>i\ \\\<^sub>l\<^sub>b (lt ?dir) (\\ s\ x\<^sub>i) (LB ?dir s x\<^sub>i)\ using assms(3)[of ?dir x\<^sub>i] using assms(4)[of ?dir x\<^sub>i] using check.simps[of "set_unsat (_ :: 'i list) s"] using \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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\<^sub>i\ using check.simps[of s] using \\ (\ s)\ \\ s\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s\ by auto qed qed lemma check_induct: fixes s :: "('i,'a) state" assumes *: "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" assumes **: "\ s. \ s \ P s s" "\ s. \\ \ s; min_lvar_not_in_bounds s = None\ \ P s s" "\ s x\<^sub>i dir I. \dir = Positive \ dir = Negative; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inl I\ \ P s (set_unsat I s)" assumes step': "\ s x\<^sub>i x\<^sub>j l\<^sub>i. \\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i)\ \ P s (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>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\<^sub>i x\<^sub>j l\<^sub>i dir let ?dir = "if \\ s'\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s' x\<^sub>i then Positive else Negative" let ?s' = "check' dir x\<^sub>i s'" assume "\ \ s'" "min_lvar_not_in_bounds s' = Some x\<^sub>i" "min_rvar_incdec dir s' x\<^sub>i = Inr x\<^sub>j" "l\<^sub>i = the (LB dir s' x\<^sub>i)" "?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'" "dir = ?dir" moreover then have "\ ?s'" "\ (\ ?s')" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s'" "\ ?s'" using \\ s'\ \\ (\ s')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'\ \\ s'\ using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ using pivotandupdate_check_precond[of dir s' x\<^sub>i x\<^sub>j l\<^sub>i] by auto ultimately have "P (check' dir x\<^sub>i s') (check (check' dir x\<^sub>i s'))" using step(2)[of x\<^sub>i] step(4)[of x\<^sub>i] \\ (\ s')\ \\ s'\ by auto then show "P s' (check (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'))" using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ \\ (\ s')\ \\ s'\ using \min_lvar_not_in_bounds s' = Some x\<^sub>i\ \min_rvar_incdec dir s' x\<^sub>i = Inr x\<^sub>j\ using step'[of s' x\<^sub>i x\<^sub>j l\<^sub>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')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'\ \\ s'\ **) qed qed lemma check_induct': fixes s :: "('i,'a) state" assumes "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" assumes "\ s x\<^sub>i dir I. \dir = Positive \ dir = Negative; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>i = Inl I; P s\ \ P (set_unsat I s)" assumes "\ s x\<^sub>i x\<^sub>j l\<^sub>i. \\ (\ s); \ s; x\<^sub>i \ lvars (\ s); x\<^sub>j \ rvars_eq (eq_for_lvar (\ s) x\<^sub>i); P s\ \ P (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>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)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" assumes **: "\ s \ P s" "\ s. \\ s; \ (\ s); \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ \ s; min_lvar_not_in_bounds s = None\ \ P s" "\ s x\<^sub>i dir I. \dir = Positive \ dir = Negative; \ s; \ (\ s); \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s; \ s; \ \ s; min_lvar_not_in_bounds s = Some x\<^sub>i; \\<^sub>l\<^sub>b (lt dir) (\\ s\ x\<^sub>i) (LB dir s x\<^sub>i); min_rvar_incdec dir s x\<^sub>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\<^sub>i x\<^sub>j l\<^sub>i dir let ?dir = "if \\ s'\ x\<^sub>i <\<^sub>l\<^sub>b \\<^sub>l s' x\<^sub>i then Positive else Negative" let ?s' = "check' dir x\<^sub>i s'" assume "\ \ s'" "min_lvar_not_in_bounds s' = Some x\<^sub>i" "min_rvar_incdec dir s' x\<^sub>i = Inr x\<^sub>j" "l\<^sub>i = the (LB dir s' x\<^sub>i)" "?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'" "dir = ?dir" moreover then have "\ ?s'" "\ (\ ?s')" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s'" "\ ?s'" "\ \ ?s'" using \\ s'\ \\ (\ s')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'\ \\ s'\ using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ using pivotandupdate_check_precond[of dir s' x\<^sub>i x\<^sub>j l\<^sub>i] using pivotandupdate_unsat_id[of s' x\<^sub>i x\<^sub>j l\<^sub>i] by (auto simp add: min_lvar_not_in_bounds_lvars min_rvar_incdec_eq_Some_rvars) ultimately have "P (check (check' dir x\<^sub>i s'))" using step(2)[of x\<^sub>i] step(4)[of x\<^sub>i] \\ (\ s')\ \\ s'\ by auto then show "P (check (pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'))" using \?s' = pivot_and_update x\<^sub>i x\<^sub>j l\<^sub>i s'\ by simp qed (simp_all add: \\ s'\ \\ (\ s')\ \\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 lemma bounds_consistent_set_unsat[simp]: "\ (set_unsat I s) = \ s" unfolding bounds_consistent_def boundsl_def boundsu_def set_unsat_simps by simp lemma curr_val_satisfies_no_lhs_set_unsat[simp]: "(\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (set_unsat I s)) = (\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s)" unfolding curr_val_satisfies_no_lhs_def boundsl_def boundsu_def set_unsat_simps by auto 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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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\<^sub>i dir I assume nolhs: "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'" and min_rvar: "min_rvar_incdec dir s' x\<^sub>i = Inl I" and sat: "\ \ s'" and min_lvar: "min_lvar_not_in_bounds s' = Some x\<^sub>i" and dir: "dir = Positive \ dir = Negative" and lt: "\\<^sub>l\<^sub>b (lt dir) (\\ s'\ x\<^sub>i) (LB dir s' x\<^sub>i)" and norm: "\ (\ s')" and valuated: "\ s'" let ?eq = "eq_for_lvar (\ s') x\<^sub>i" have unsat_core: "set (the (\\<^sub>c (set_unsat I s'))) = set I" by auto obtain l\<^sub>i where LB_Some: "LB dir s' x\<^sub>i = Some l\<^sub>i" and lt: "lt dir (\\ s'\ x\<^sub>i) l\<^sub>i" using lt by (cases "LB dir s' x\<^sub>i") (auto simp add: bound_compare_defs) from LB_Some dir obtain i where LBI: "look (LBI dir s') x\<^sub>i = Some (i,l\<^sub>i)" and LI: "LI dir s' x\<^sub>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\<^sub>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\<^sub>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\<^sub>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\<^sub>i \ lvars (\ s')" using min_lvar by (simp add: min_lvar_not_in_bounds_lvars) then have **: "?eq \ set (\ s')" "lhs ?eq = x\<^sub>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'\ \\<^sub>t \ s'" and b: "\\ s'\ \\<^sub>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 (\\<^sub>i\<^sub>l s') x = Some (i,c) \ look (\\<^sub>i\<^sub>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: "\ (\\<^sub>l\<^sub>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 simp del: Simplex.bounds_lg) 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: "\ (\\<^sub>u\<^sub>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 simp del: Simplex.bounds_lg) 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 (\\<^sub>i\<^sub>l s') x = Some (i,c) \ look (\\<^sub>i\<^sub>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'\) \\<^sub>i\<^sub>s\<^sub>e s'" unfolding satisfies_state_index'.simps proof (intro conjI) show "\\ s'\ \\<^sub>t \ s'" by fact show "(?R1 \ ?R2, \\ s'\) \\<^sub>i\<^sub>b\<^sub>e \\ s'" unfolding satisfies_bounds_index'.simps proof (intro conjI impI allI) fix x c assume c: "\\<^sub>l s' x = Some c" and i: "\\<^sub>l s' x \ ?R1 \ ?R2" from c have ci: "look (\\<^sub>i\<^sub>l s') x = Some (\\<^sub>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: "\\<^sub>u s' x = Some c" and i: "\\<^sub>u s' x \ ?R1 \ ?R2" from c have ci: "look (\\<^sub>i\<^sub>u s') x = Some (\\<^sub>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 (\\<^sub>c (set_unsat I s'))) = set I" "\ x. x \\<^sub>i\<^sub>s\<^sub>e set_unsat I s' \ x \\<^sub>i\<^sub>s\<^sub>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) \\<^sub>i\<^sub>s\<^sub>e s'" proof (cases "J \ ?R1 \ ?R2") case True with R1R2 have "(J, \\ s'\) \\<^sub>i\<^sub>s\<^sub>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\<^sub>i - \\ s'\ x\<^sub>i" have "\\ s'\ x\<^sub>i < l\<^sub>i \ 0 < l\<^sub>i - \\ s'\ x\<^sub>i" "l\<^sub>i < \\ s'\ x\<^sub>i \ l\<^sub>i - \\ s'\ x\<^sub>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 simp del: Simplex.bounds_lg) 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 \\<^sub>t \ s'" unfolding v_def using rev.update_satisfies_tableau[OF norm valuated yl] \\\ s'\ \\<^sub>t \ s'\ by auto with **(1) have "v \\<^sub>e ?eq" unfolding satisfies_tableau_def by auto from this[unfolded satisfies_eq_def id] have v_xi: "v x\<^sub>i = (rhs ?eq \ v \)" . from \\\ s'\ \\<^sub>t \ s'\ **(1) have "\\ s'\ \\<^sub>e ?eq" unfolding satisfies_tableau_def by auto hence V_xi: "\\ s'\ x\<^sub>i = (rhs ?eq \ \\ s'\ \)" unfolding satisfies_eq_def id . have "v x\<^sub>i = \\ s'\ x\<^sub>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\<^sub>i" unfolding up_def diff_def scaleRat_scaleRat using cy0 by simp finally have v_xi_l: "v x\<^sub>i = l\<^sub>i" . { assume both: "\\<^sub>u s' y \ ?R1 \ ?R2" "\\<^sub>u s' y \ None" "\\<^sub>l s' y \ ?R1 \ ?R2" "\\<^sub>l s' y \ None" and diff: "\\<^sub>l s' y \ \\<^sub>u s' y" from both(1) dir obtain xu cu where looku: "look (\\<^sub>i\<^sub>l s') xu = Some (\\<^sub>u s' y, cu) \ look (\\<^sub>i\<^sub>u s') xu = Some (\\<^sub>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 \ \\<^sub>u s' y = LI dir s' xu' \ coeff (rhs ?eq) xu' > 0 \ \\<^sub>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 \ \\<^sub>u s' y = LI dir s' xu \ coeff (rhs ?eq) xu > 0 \ \\<^sub>u s' y = UI dir s' xu" by auto { assume "xu \ y" with dist[OF looku, of y] have "look (\\<^sub>i\<^sub>u s') y = None" by (cases "look (\\<^sub>i\<^sub>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 (\\<^sub>i\<^sub>l s') xl = Some (\\<^sub>l s' y, cl) \ look (\\<^sub>i\<^sub>u s') xl = Some (\\<^sub>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 \ \\<^sub>l s' y = LI dir s' xl' \ coeff (rhs ?eq) xl' > 0 \ \\<^sub>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 \ \\<^sub>l s' y = LI dir s' xl \ coeff (rhs ?eq) xl > 0 \ \\<^sub>l s' y = UI dir s' xl" by auto { assume "xl \ y" with dist[OF lookl, of y] have "look (\\<^sub>i\<^sub>l s') y = None" by (cases "look (\\<^sub>i\<^sub>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) \\<^sub>i\<^sub>b\<^sub>e \\ s'" unfolding satisfies_bounds_index'.simps proof (intro conjI allI impI) fix x c assume x: "\\<^sub>l s' x = Some c" "\\<^sub>l s' x \ J" with k have not_k: "\\<^sub>l s' x \ k" by auto from x have ci: "look (\\<^sub>i\<^sub>l s') x = Some (\\<^sub>l s' x, c)" unfolding boundsl_def indexl_def by auto show "v x = c" proof (cases "\\<^sub>l s' x = i") case False hence iR12: "\\<^sub>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: "\\<^sub>u s' x = k" by auto with ci Iu x(2) k sub False True have both: "\\<^sub>u s' y \ ?R1 \ ?R2" "\\<^sub>l s' y \ ?R1 \ ?R2" and diff: "\\<^sub>l s' y \ \\<^sub>u s' y" unfolding setI LI by auto have "\\<^sub>l s' y \ None" using x True by simp from both_y_False[OF both(1) _ both(2) this diff] have "\\<^sub>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 = \\<^sub>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\<^sub>i l\<^sub>i] have xxi: "x = x\<^sub>i" and c: "c = l\<^sub>i" by auto have vxi: "v x = l\<^sub>i" unfolding xxi v_xi_l .. thus ?thesis unfolding c by simp qed next fix x c assume x: "\\<^sub>u s' x = Some c" "\\<^sub>u s' x \ J" with k have not_k: "\\<^sub>u s' x \ k" by auto from x have ci: "look (\\<^sub>i\<^sub>u s') x = Some (\\<^sub>u s' x, c)" unfolding boundsu_def indexu_def by auto show "v x = c" proof (cases "\\<^sub>u s' x = i") case False hence iR12: "\\<^sub>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: "\\<^sub>l s' x = k" by auto with ci Iu x(2) k sub False True have both: "\\<^sub>u s' y \ ?R1 \ ?R2" "\\<^sub>l s' y \ ?R1 \ ?R2" and diff: "\\<^sub>l s' y \ \\<^sub>u s' y" unfolding setI LI by auto have "\\<^sub>u s' y \ None" using x True by simp from both_y_False[OF both(1) this both(2) _ diff] have "\\<^sub>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 = \\<^sub>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\<^sub>i l\<^sub>i] have xxi: "x = x\<^sub>i" and c: "c = l\<^sub>i" by auto have vxi: "v x = l\<^sub>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) \\<^sub>i\<^sub>s set_unsat I s'" then have Iv: "(set I, v) \\<^sub>i\<^sub>s s'" unfolding satisfies_state_index.simps by (auto simp: indexl_def indexu_def boundsl_def boundsu_def) from Iv have vt: "v \\<^sub>t \ s'" and Iv: "(set I, v) \\<^sub>i\<^sub>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\<^sub>i (v x\<^sub>i)" . have "\\ s'\ \\<^sub>e ?eq" using nolhs \?eq \ set (\ s')\ unfolding curr_val_satisfies_no_lhs_def by (simp add: satisfies_tableau_def) then have "\\ s'\ x\<^sub>i = (rhs ?eq) \ \\ s'\ \" using \lhs ?eq = x\<^sub>i\ by (simp add: satisfies_eq_def) moreover have "v \\<^sub>e ?eq" using vt \?eq \ set (\ s')\ by (simp add: satisfies_state_def satisfies_tableau_def) then have "v x\<^sub>i = (rhs ?eq) \ v \" using \lhs ?eq = x\<^sub>i\ by (simp add: satisfies_eq_def) moreover have "\\<^sub>l\<^sub>b (lt dir) (v x\<^sub>i) (LB dir s' x\<^sub>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)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" then have "v \\<^sub>t \ s = v \\<^sub>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)" by (rule check_induct', simp_all add: * pivotandupdate_tableau_normalized pivotandupdate_bounds_consistent) moreover have "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s (check s)" by (rule check_induct'', simp_all add: *) 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 \\<^sub>t \ s = v \\<^sub>t \ s' \ \ (\ s') \ \ s' \ \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s' \ \ s'" by (simp add: Let_def) next fix s :: "('i,'a) state" assume *: "\ s" "\ (\ s)" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ s" from * show "\\<^sub>i (check s) = \\<^sub>i s" by (rule check_induct, simp_all add: pivotandupdate_bounds_id) next fix s :: "('i,'a) state" assume *: "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>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 (\\\<^sub>l\ vs \\\<^sub>u\) and their updating functions, comparisons with bounds (e.g., \\\<^sub>u\<^sub>b\ vs \\\<^sub>l\<^sub>b\ or \<\<^sub>l\<^sub>b\ vs \>\<^sub>u\<^sub>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, \\\<^sub>l\ for lower and \\\<^sub>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\ (*-------------------------------------------------------------------------- *) (* -------------------------------------------------------------------------- *) (* Init init_state *) (* -------------------------------------------------------------------------- *) 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\ \\<^sub>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) (* -------------------------------------------------------------------------- *) (* MinVars min_lvar_not_in_bounds min_rvar_incdec_eq *) (* -------------------------------------------------------------------------- *) 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\<^sub>i show "min_lvar_not_in_bounds s = Some x\<^sub>i \ x\<^sub>i \ lvars (\ s) \ \ in_bounds x\<^sub>i \\ s\ (\ s) \ (\x\lvars (\ s). x < x\<^sub>i \ in_bounds x \\ s\ (\ s))" unfolding min_lvar_not_in_bounds_def lvars_def using min_satisfying_Some by blast+ qed \ \all variables in vs have either a positive or a negative coefficient, so no equal-zero test required.\ 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\<^sub>j \ Inr x\<^sub>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 "\ (\\<^sub>l\<^sub>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 "\ (\\<^sub>u\<^sub>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\<^sub>j assume "min_rvar_incdec_eq dir s eq = Inr x\<^sub>j" from this[unfolded min_rvar_incdec_eq_def Let_def] have "?min = Some x\<^sub>j" by (cases ?min, auto) then show "x\<^sub>j \ rvars_eq eq" "reasable_var dir x\<^sub>j eq s" "(\ x' \ rvars_eq eq. x' < x\<^sub>j \ \ reasable_var dir x' eq s)" using min_satisfying_Some set_vars_list by blast+ qed (* -------------------------------------------------------------------------- *) (* EqForLVar eq_idx_for_lvar *) (* -------------------------------------------------------------------------- *) 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 (* -------------------------------------------------------------------------- *) (* PivotEq pivot_eq *) (* -------------------------------------------------------------------------- *) 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 \\<^sub>e e = v \\<^sub>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\<^sub>j assume "x\<^sub>j \ rvars_eq eq" "lhs eq \ rvars_eq eq" have "lhs (pivot_eq eq x\<^sub>j) = x\<^sub>j" unfolding pivot_eq_def by (simp add: Let_def) moreover have "rvars_eq (pivot_eq eq x\<^sub>j) = {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" proof show "rvars_eq (pivot_eq eq x\<^sub>j) \ {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" proof fix x assume "x \ rvars_eq (pivot_eq eq x\<^sub>j)" have *: "coeff (rhs (pivot_eq eq x\<^sub>j)) x\<^sub>j = 0" using \x\<^sub>j \ rvars_eq eq\ \lhs eq \ rvars_eq eq\ using coeff_Var2[of "lhs eq" x\<^sub>j] by (auto simp add: Let_def pivot_eq_def) have "coeff (rhs eq) x\<^sub>j \ 0" using \x\<^sub>j \ rvars_eq eq\ using coeff_zero by (cases eq) (auto simp add:) then show "x \ {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" using pivot_eq_rvars[of x eq x\<^sub>j] using \x \ rvars_eq (pivot_eq eq x\<^sub>j)\ \x\<^sub>j \ rvars_eq eq\ \lhs eq \ rvars_eq eq\ using coeff_zero * by auto qed show "{lhs eq} \ (rvars_eq eq - {x\<^sub>j}) \ rvars_eq (pivot_eq eq x\<^sub>j)" proof fix x assume "x \ {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" have *: "coeff (rhs eq) (lhs eq) = 0" using coeff_zero using \lhs eq \ rvars_eq eq\ by auto have **: "coeff (rhs eq) x\<^sub>j \ 0" using \x\<^sub>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\<^sub>j) (lhs eq) = 0" using \x\<^sub>j \ rvars_eq eq\ \lhs eq \ rvars_eq eq\ using coeff_Var2[of x\<^sub>j "lhs eq"] by auto then have "coeff (rhs (pivot_eq eq x\<^sub>j)) x \ 0" using \x \ {lhs eq} \ (rvars_eq eq - {x\<^sub>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\<^sub>j)" by (simp add: coeff_zero) qed qed ultimately show "let eq' = pivot_eq eq x\<^sub>j in lhs eq' = x\<^sub>j \ rvars_eq eq' = {lhs eq} \ (rvars_eq eq - {x\<^sub>j})" by (simp add: Let_def) next fix v eq x\<^sub>j assume "x\<^sub>j \ rvars_eq eq" then show "v \\<^sub>e pivot_eq eq x\<^sub>j = v \\<^sub>e eq" using pivot_eq_satisfies_eq by blast qed (* -------------------------------------------------------------------------- *) (* SubstVar subst_var *) (* -------------------------------------------------------------------------- *) 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\<^sub>j lp' lp have *: "\x. \x \ vars (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j); x \ vars lp'\ \ x \ vars lp" proof- fix x assume "x \ vars (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j)" then have "coeff (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>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\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j) x \ 0\ using \coeff lp' x = 0\ by (cases "x = x\<^sub>j") (auto simp add: coeff_Var2) qed qed have "vars (subst_var x\<^sub>j lp' lp) \ (vars lp - {x\<^sub>j}) \ vars lp'" unfolding subst_var_def using coeff_zero[of "lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j" x\<^sub>j] using coeff_zero[of lp' x\<^sub>j] using * by auto moreover have "\x. \x \ vars (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j); x \ vars lp; x \ vars lp'\ \ x = x\<^sub>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\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j)" then have "coeff (lp + coeff lp x\<^sub>j *R lp' - coeff lp x\<^sub>j *R Var x\<^sub>j) x = 0" using coeff_zero by force then show "x = x\<^sub>j" using \coeff lp x \ 0\ \coeff lp' x = 0\ by (cases "x = x\<^sub>j") (auto simp add: coeff_Var2) qed then have "vars lp - {x\<^sub>j} - vars lp' \ vars (subst_var x\<^sub>j lp' lp)" by (auto simp add: subst_var_def) ultimately show "vars lp - {x\<^sub>j} - vars lp' \s vars (subst_var x\<^sub>j lp' lp) \s vars lp - {x\<^sub>j} \ vars lp'" by simp next fix v x\<^sub>j lp' lp show "v x\<^sub>j = lp' \ v \ \ lp \ v \ = (subst_var x\<^sub>j lp' lp) \ v \" unfolding subst_var_def using valuate_minus[of "lp + coeff lp x\<^sub>j *R lp'" "coeff lp x\<^sub>j *R Var x\<^sub>j" v] using valuate_add[of lp "coeff lp x\<^sub>j *R lp'" v] using valuate_scaleRat[of "coeff lp x\<^sub>j" lp' v] valuate_scaleRat[of "coeff lp x\<^sub>j" "Var x\<^sub>j" v] using valuate_Var[of x\<^sub>j v] by auto next fix x\<^sub>j lp lp' assume "x\<^sub>j \ vars lp" hence 0: "coeff lp x\<^sub>j = 0" using coeff_zero by blast show "subst_var x\<^sub>j lp' lp = lp" unfolding subst_var_def 0 by simp next fix x\<^sub>j lp x lp' assume "x\<^sub>j \ vars lp" "x \ vars lp' - vars lp" hence x: "x \ x\<^sub>j" and 0: "coeff lp x = 0" and no0: "coeff lp x\<^sub>j \ 0" "coeff lp' x \ 0" using coeff_zero by blast+ from x have 00: "coeff (Var x\<^sub>j) x = 0" using coeff_Var2 by auto show "x \ vars (subst_var x\<^sub>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) (* -------------------------------------------------------------------------- *) (* Update update *) (* -------------------------------------------------------------------------- *) definition rhs_eq_val where "rhs_eq_val v x\<^sub>i c e \ let x\<^sub>j = lhs e; a\<^sub>i\<^sub>j = coeff (rhs e) x\<^sub>i in \v\ x\<^sub>j + a\<^sub>i\<^sub>j *R (c - \v\ x\<^sub>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\ \\<^sub>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) (* -------------------------------------------------------------------------- *) (* Preprocess preprocess *) (* -------------------------------------------------------------------------- *) 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 \\<^sub>n\<^sub>s c" by (cases c, auto simp: valuate_zero) lemma not_zero_satisfies: "poly c = 0 \ \ zero_satisfies c \ \ v \\<^sub>n\<^sub>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 \\<^sub>n\<^sub>s (normalize_ns_constraint c) \ v \\<^sub>n\<^sub>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 \\<^sub>i\<^sub>n\<^sub>s\<^sub>s (map_prod id normalize_ns_constraint ` cs) \ Iv \\<^sub>i\<^sub>n\<^sub>s\<^sub>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 \\<^sub>a qdelta_constraint_to_atom h var" "v \\<^sub>e linear_poly_to_eq (poly h) var" "\ is_monom (poly h)" shows "v \\<^sub>n\<^sub>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 \\<^sub>a qdelta_constraint_to_atom h var \ v \\<^sub>n\<^sub>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) \\<^sub>i\<^sub>a\<^sub>s set (Atoms (preprocess' s start))" "v \\<^sub>t Tableau (preprocess' s start)" "I \ set (UnsatIndices (preprocess' s start)) = {}" shows "(I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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 \\<^sub>a\<^sub>s flat (set (Atoms (preprocess' s start)))" "v \\<^sub>t Tableau (preprocess' s start)" "set (UnsatIndices (preprocess' s start)) = {}" shows "v \\<^sub>n\<^sub>s\<^sub>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 \\<^sub>n\<^sub>s c \ v2 \\<^sub>n\<^sub>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 \\<^sub>t t" "\ var \ tvars t. v1 var = v2 var" shows "v2 \\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set s" "vars_constraints (flat_list s) \ V" "\var \ V. var < start" shows "\v'. (\var \ V. v var = v' var) \ v' \\<^sub>a\<^sub>s restrict_to I (set (Atoms (preprocess' s start))) \ v' \\<^sub>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' \\<^sub>a\<^sub>s restrict_to I (set (Atoms (preprocess' t start)))" and v'_t: "v' \\<^sub>t Tableau (preprocess' t start)" and vars_h: "vars_constraints [h] \ V" by auto from Cons(2)[unfolded hh] have i: "i \ I \ v \\<^sub>n\<^sub>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 \\<^sub>n\<^sub>s h \ v' \\<^sub>n\<^sub>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' \\<^sub>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' \\<^sub>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' \\<^sub>a\<^sub>s ?atoms" unfolding satisfies_atom_set_def proof fix a assume "a \ ?atoms" then have "v' \\<^sub>a a" using \v' \\<^sub>a\<^sub>s ?atoms\ hh by (force simp add: satisfies_atom_set_def) then show "?v' \\<^sub>a a" using \a \ ?atoms\ atom_var_first[of a t start] using Cons(3) Cons(4) by (cases a) auto qed show "?v' \\<^sub>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' \\<^sub>t Tableau (preprocess' t start)" using \v' \\<^sub>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' \\<^sub>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' \\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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'" \ \new tableau is normalized\ "\w\ \\<^sub>t t' \ \tv w\ \\<^sub>t t" \ \solution of new tableau is translated to solution of old tableau\ "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" \ \solution translation also works for bounds\ "v \\<^sub>t t \ v \\<^sub>t t'" \ \solution of old tableau is also solution for new tableau\ 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 \\<^sub>t t \ v \\<^sub>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\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \w\) \\<^sub>i\<^sub>a\<^sub>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\ \\<^sub>a a = \w\ \\<^sub>a a" unfolding tv by (cases a, auto) qed then show "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" by blast assume w: "\w\ \\<^sub>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\ \\<^sub>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\ \\<^sub>e eq" unfolding tv eq satisfies_eq_iff using xy xq by (auto intro!: valuate_depend) } moreover have "\tv w\ \\<^sub>e (x,p)" unfolding satisfies_eq_iff tv using xp by (auto intro!: valuate_depend) ultimately show "\tv w\ \\<^sub>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) \\<^sub>t t' \ v \\<^sub>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 \\<^sub>t t' \ v \\<^sub>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) \\<^sub>t t' \ \tv w\ \\<^sub>t t" "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" "v \\<^sub>t t \ (v :: 'a valuation) \\<^sub>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) \\<^sub>t t' \ \tv w\ \\<^sub>t t" "(I, \w\) \\<^sub>i\<^sub>a\<^sub>s set as \ (I, \tv w\) \\<^sub>i\<^sub>a\<^sub>s set as" "v \\<^sub>t t \ (v :: 'a valuation) \\<^sub>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\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t \ (I, \trans_v v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" "i \ set ui \ \v. ({i}, v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" "\ v. (I,v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs \ \v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>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 \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set ncs \ Iv \\<^sub>i\<^sub>n\<^sub>s\<^sub>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\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t1 \ I \ set ui = {} \ (I,\v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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\) \\<^sub>i\<^sub>a\<^sub>s set as \ \v\ \\<^sub>t t \ (I,\trans_v v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" then obtain w where "(I,w) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set cs" by blast hence "(I,w) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set ncs" unfolding ncs . from preprocess'_unsat[OF this _ start_fresh_variable_fresh, of ncs] have "\v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>t t1" using part1 unfolding preprocess_part_1_def Let_def by auto then show "\v'. (I,v') \\<^sub>i\<^sub>a\<^sub>s set as \ v' \\<^sub>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 (* -------------------------------------------------------------------------- *) (* To_ns to_ns from_ns *) (* -------------------------------------------------------------------------- *) 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'\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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)\) \\<^sub>i\<^sub>c\<^sub>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)" (* index-filtered list *) let ?f_list = "flat_list (to_ns ics)" let ?flist = "?listf ?f_list" (* full 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'\ \\<^sub>n\<^sub>s\<^sub>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\ \\<^sub>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) \\<^sub>i\<^sub>c\<^sub>s set cs" let ?v = "\var. QDelta.QDelta (v var) 0" have "(I,?v) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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) \\<^sub>i\<^sub>c\<^sub>s set cs" and c: "i \ I \ v \\<^sub>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 \\<^sub>c c" by force from c n jn have "?v \\<^sub>n\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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 \\<^sub>n\<^sub>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) \\<^sub>i\<^sub>n\<^sub>s\<^sub>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\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set (to_ns cs)" by auto from sat[OF this] show " \v. (J, v) \\<^sub>i\<^sub>c\<^sub>s set cs" by blast qed qed (* cleanup *) 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 (* -------------------------------------------------------------------------- *) (* Main soundness lemma and executability *) (* -------------------------------------------------------------------------- *) 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) \\<^sub>i\<^sub>c\<^sub>s set cs) \ (distinct_indices cs \ (\ J \ set I. (\ v. (J, v) \\<^sub>i\<^sub>c\<^sub>s set cs)))" \ \minimal unsat core\ "simplex_index cs = Sat v \ \v\ \\<^sub>c\<^sub>s (snd ` set cs)" \ \satisfying assingment\ 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) \\<^sub>i\<^sub>c\<^sub>s set cs) \ (distinct_indices cs \ (\J\set I. \v. (J, v) \\<^sub>i\<^sub>c\<^sub>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\ \\<^sub>c\<^sub>s (snd ` set cs)" . qed text \Simplex version where indices will be created\ definition simplex where "simplex cs = simplex_index (zip [0.. \ (\ v. v \\<^sub>c\<^sub>s set cs)" \ \unsat of original constraints\ "simplex cs = Unsat I \ set I \ {0.. \ (\ v. v \\<^sub>c\<^sub>s {cs ! i | i. i \ set I}) \ (\J\set I. \v. v \\<^sub>c\<^sub>s {cs ! i |i. i \ J})" \ \minimal unsat core\ "simplex cs = Sat v \ \v\ \\<^sub>c\<^sub>s set cs" \ \satisfying assignment\ proof (unfold simplex_def) let ?cs = "zip [0.. {0 ..< length cs}" and core: "\v. v \\<^sub>c\<^sub>s (snd ` (set ?cs \ set I \ UNIV))" "(distinct_indices (zip [0.. (\ J \ set I. \v. v \\<^sub>c\<^sub>s (snd ` (set ?cs \ J \ UNIV))))" by (auto simp flip: set_map) note core(2) also have "distinct_indices (zip [0.. J \ set I. \v. v \\<^sub>c\<^sub>s (snd ` (set ?cs \ J \ UNIV))) = (\ J \ set I. \v. v \\<^sub>c\<^sub>s { cs ! i | i. i \ J})" using index by (intro all_cong1 imp_cong ex_cong1 arg_cong[of _ _ "\ x. _ \\<^sub>c\<^sub>s x"] refl, force simp: set_zip) finally have core': "(\J\set I. \v. v \\<^sub>c\<^sub>s {cs ! i |i. i \ J}) " . note unsat = unsat_mono[OF core(1)] show "\ (\ v. v \\<^sub>c\<^sub>s set cs)" by (rule unsat, auto simp: set_zip) show "set I \ {0.. \ (\ v. v \\<^sub>c\<^sub>s {cs ! i | i. i \ set I}) \ (\J\set I. \v. v \\<^sub>c\<^sub>s {cs ! i |i. i \ J})" by (intro conjI index core', rule unsat, auto simp: set_zip) next assume "simplex_index (zip [0..v\ \\<^sub>c\<^sub>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] +lemma "case simplex [LT (lp_monom 2 1 - lp_monom 3 2 + lp_monom 3 0) 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)), + (2, GT (lp_monom 2 1 - lp_monom 1 2) 0), + (3, EQ (lp_monom 1 1 - lp_monom 2 2) 0), (4, GT (lp_monom 2 2) 5), (5, GT (lp_monom 3 0) 7)] of Sat _ \ False | Unsat I \ set I = {1,3,4}" \ \Constraints 1,3,4 are unsat core\ by eval end \ No newline at end of file diff --git a/thys/Simplex/Simplex_Incremental.thy b/thys/Simplex/Simplex_Incremental.thy --- a/thys/Simplex/Simplex_Incremental.thy +++ b/thys/Simplex/Simplex_Incremental.thy @@ -1,1034 +1,1034 @@ (* Author: R. Thiemann *) section \The Incremental Simplex Algorithm\ text \In this theory we specify operations which permit to incrementally add constraints. To this end, first an indexed list of potential constraints is used to construct the initial state, and then one can activate indices, extract solutions or unsat cores, do backtracking, etc.\ theory Simplex_Incremental imports Simplex begin subsection \Lowest Layer: Fixed Tableau and Incremental Atoms\ text \Interface\ locale Incremental_Atom_Ops = fixes init_s :: "tableau \ 's" and assert_s :: "('i,'a :: lrv) i_atom \ 's \ 'i list + 's" and check_s :: "'s \ 's \ ('i list option)" and solution_s :: "'s \ (var, 'a) mapping" and checkpoint_s :: "'s \ 'c" and backtrack_s :: "'c \ 's \ 's" and precond_s :: "tableau \ bool" and weak_invariant_s :: "tableau \ ('i,'a) i_atom set \ 's \ bool" and invariant_s :: "tableau \ ('i,'a) i_atom set \ 's \ bool" and checked_s :: "tableau \ ('i,'a) i_atom set \ 's \ bool" assumes assert_s_ok: "invariant_s t as s \ assert_s a s = Inr s' \ invariant_s t (insert a as) s'" and assert_s_unsat: "invariant_s t as s \ assert_s a s = Unsat I \ minimal_unsat_core_tabl_atoms (set I) t (insert a as)" and check_s_ok: "invariant_s t as s \ check_s s = (s', None) \ checked_s t as s'" and check_s_unsat: "invariant_s t as s \ check_s s = (s',Some I) \ weak_invariant_s t as s' \ minimal_unsat_core_tabl_atoms (set I) t as" and init_s: "precond_s t \ checked_s t {} (init_s t)" and solution_s: "checked_s t as s \ solution_s s = v \ \v\ \\<^sub>t t \ \v\ \\<^sub>a\<^sub>s Simplex.flat as" and backtrack_s: "checked_s t as s \ checkpoint_s s = c \ weak_invariant_s t bs s' \ backtrack_s c s' = s'' \ as \ bs \ invariant_s t as s''" and weak_invariant_s: "invariant_s t as s \ weak_invariant_s t as s" and checked_invariant_s: "checked_s t as s \ invariant_s t as s" begin fun assert_all_s :: "('i,'a) i_atom list \ 's \ 'i list + 's" where "assert_all_s [] s = Inr s" | "assert_all_s (a # as) s = (case assert_s a s of Unsat I \ Unsat I | Inr s' \ assert_all_s as s')" lemma assert_all_s_ok: "invariant_s t as s \ assert_all_s bs s = Inr s' \ invariant_s t (set bs \ as) s'" proof (induct bs arbitrary: s as) case (Cons b bs s as) from Cons(3) obtain s'' where ass: "assert_s b s = Inr s''" and rec: "assert_all_s bs s'' = Inr s'" by (auto split: sum.splits) from Cons(1)[OF assert_s_ok[OF Cons(2) ass] rec] show ?case by auto qed auto lemma assert_all_s_unsat: "invariant_s t as s \ assert_all_s bs s = Unsat I \ minimal_unsat_core_tabl_atoms (set I) t (as \ set bs)" proof (induct bs arbitrary: s as) case (Cons b bs s as) show ?case proof (cases "assert_s b s") case unsat: (Inl J) with Cons have J: "J = I" by auto from assert_s_unsat[OF Cons(2) unsat] J have min: "minimal_unsat_core_tabl_atoms (set I) t (insert b as)" by auto show ?thesis by (rule minimal_unsat_core_tabl_atoms_mono[OF _ min], auto) next case (Inr s') from Cons(1)[OF assert_s_ok[OF Cons(2) Inr]] Cons(3) Inr show ?thesis by auto qed qed simp end text \Implementation of the interface via the Simplex operations init, check, and assert-bound.\ locale Incremental_State_Ops_Simplex = AssertBoundNoLhs assert_bound + Init init + Check check for assert_bound :: "('i,'a::lrv) i_atom \ ('i,'a) state \ ('i,'a) state" and init :: "tableau \ ('i,'a) state" and check :: "('i,'a) state \ ('i,'a) state" begin definition weak_invariant_s where "weak_invariant_s t (as :: ('i,'a)i_atom set) s = (\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s \ \ (\ s) \ \ s \ \ s \ (\ v :: (var \ 'a). v \\<^sub>t \ s \ v \\<^sub>t t) \ index_valid as s \ Simplex.flat as \ \ s \ as \\<^sub>i \\ s)" definition invariant_s where "invariant_s t (as :: ('i,'a)i_atom set) s = (weak_invariant_s t as s \ \ \ s)" definition checked_s where "checked_s t as s = (invariant_s t as s \ \ s)" definition assert_s where "assert_s a s = (let s' = assert_bound a s in if \ s' then Inl (the (\\<^sub>c s')) else Inr s')" definition check_s where "check_s s = (let s' = check s in if \ s' then (s', Some (the (\\<^sub>c s'))) else (s', None))" definition checkpoint_s where "checkpoint_s s = \\<^sub>i s" fun backtrack_s :: "_ \ ('i, 'a) state \ ('i, 'a) state" where "backtrack_s (bl, bu) (State t bl_old bu_old v u uc) = State t bl bu v False None" lemmas invariant_defs = weak_invariant_s_def invariant_s_def checked_s_def lemma invariant_sD: assumes "invariant_s t as s" shows "\ \ s" "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" "\ s" "Simplex.flat as \ \ s" "as \\<^sub>i \\ s" "index_valid as s" "(\ v :: (var \ 'a). v \\<^sub>t \ s \ v \\<^sub>t t)" using assms unfolding invariant_defs by auto lemma weak_invariant_sD: assumes "weak_invariant_s t as s" shows "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\ (\ s)" "\ s" "\ s" "Simplex.flat as \ \ s" "as \\<^sub>i \\ s" "index_valid as s" "(\ v :: (var \ 'a). v \\<^sub>t \ s \ v \\<^sub>t t)" using assms unfolding invariant_defs by auto lemma minimal_unsat_state_core_translation: assumes unsat: "minimal_unsat_state_core (s :: ('i,'a::lrv)state)" and tabl: "\(v :: 'a valuation). v \\<^sub>t \ s = v \\<^sub>t t" and index: "index_valid as s" and imp: "as \\<^sub>i \\ s" and I: "I = the (\\<^sub>c s)" shows "minimal_unsat_core_tabl_atoms (set I) t as" unfolding minimal_unsat_core_tabl_atoms_def proof (intro conjI impI notI allI; (elim exE conjE)?) from unsat[unfolded minimal_unsat_state_core_def] have unsat: "unsat_state_core s" and minimal: "distinct_indices_state s \ subsets_sat_core s" by auto from unsat[unfolded unsat_state_core_def I[symmetric]] have Is: "set I \ indices_state s" and unsat: "(\v. (set I, v) \\<^sub>i\<^sub>s s)" by auto from Is index show "set I \ fst ` as" using index_valid_indices_state by blast { fix v assume t: "v \\<^sub>t t" and as: "(set I, v) \\<^sub>i\<^sub>a\<^sub>s as" from t tabl have t: "v \\<^sub>t \ s" by auto then have "(set I, v) \\<^sub>i\<^sub>s s" using as imp using atoms_imply_bounds_index.simps satisfies_state_index.simps by blast with unsat show False by blast } { fix J assume dist: "distinct_indices_atoms as" and J: "J \ set I" from J Is have J': "J \ indices_state s" by auto from dist index have "distinct_indices_state s" by (metis index_valid_distinct_indices) with minimal have "subsets_sat_core s" . from this[unfolded subsets_sat_core_def I[symmetric], rule_format, OF J] obtain v where "(J, v) \\<^sub>i\<^sub>s\<^sub>e s" by blast from satisfying_state_valuation_to_atom_tabl[OF J' this index dist] tabl show "\v. v \\<^sub>t t \ (J, v) \\<^sub>i\<^sub>a\<^sub>e\<^sub>s as" by blast } qed sublocale Incremental_Atom_Ops init assert_s check_s \ checkpoint_s backtrack_s \ weak_invariant_s invariant_s checked_s proof (unfold_locales, goal_cases) case (1 t as s a s') (* assert ok *) from 1(2)[unfolded assert_s_def Let_def] have U: "\ \ (assert_bound a s)" and s': "s' = assert_bound a s" by (auto split: if_splits) note * = invariant_sD[OF 1(1)] from assert_bound_nolhs_tableau_id[OF *(1-5)] have T: "\ s' = \ s" unfolding s' by auto from *(3,9) have "\ (\ s')" "\ v :: var \ 'a. v \\<^sub>t \ s' = v \\<^sub>t t" unfolding T by blast+ moreover from assert_bound_nolhs_sat[OF *(1-5) U] have " \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'" "\ s'" unfolding s' by auto moreover from assert_bound_nolhs_atoms_equiv_bounds[OF *(1-6), of a] have "Simplex.flat (insert a as) \ \ s'" unfolding s' by auto moreover from assert_bound_nolhs_atoms_imply_bounds_index[OF *(1-5,7)] have "insert a as \\<^sub>i \\ s'" unfolding s' . moreover from assert_bound_nolhs_tableau_valuated[OF *(1-5)] have "\ s'" unfolding s' . moreover from assert_bound_nolhs_index_valid[OF *(1-5,8)] have "index_valid (insert a as) s'" unfolding s' by auto moreover from U s' have "\ \ s'" by auto ultimately show ?case unfolding invariant_defs by auto next case (2 t as s a I) (* assert unsat *) from 2(2)[unfolded assert_s_def Let_def] obtain s' where s': "s' = assert_bound a s" and U: "\ (assert_bound a s)" and I: "I = the (\\<^sub>c s')" by (auto split: if_splits) note * = invariant_sD[OF 2(1)] from assert_bound_nolhs_tableau_id[OF *(1-5)] have T: "\ s' = \ s" unfolding s' by auto from *(3,9) have tabl: "\ v :: var \ 'a. v \\<^sub>t \ s' = v \\<^sub>t t" unfolding T by blast+ from assert_bound_nolhs_unsat[OF *(1-5,8) U] s' have unsat: "minimal_unsat_state_core s'" by auto from assert_bound_nolhs_index_valid[OF *(1-5,8)] have index: "index_valid (insert a as) s'" unfolding s' by auto from assert_bound_nolhs_atoms_imply_bounds_index[OF *(1-5,7)] have imp: "insert a as \\<^sub>i \\ s'" unfolding s' . from minimal_unsat_state_core_translation[OF unsat tabl index imp I] show ?case . next case (3 t as s s') (* check ok *) from 3(2)[unfolded check_s_def Let_def] have U: "\ \ (check s)" and s': "s' = check s" by (auto split: if_splits) note * = invariant_sD[OF 3(1)] note ** = *(1,2,5,3,4) from check_tableau_equiv[OF **] *(9) have "\v :: _ \ 'a. v \\<^sub>t \ s' = v \\<^sub>t t" unfolding s' by auto moreover from check_tableau_index_valid[OF **] *(8) have "index_valid as s'" unfolding s' by auto moreover from check_tableau_normalized[OF **] have "\ (\ s')" unfolding s' . moreover from check_tableau_valuated[OF **] have "\ s'" unfolding s' . moreover from check_sat[OF ** U] have "\ s'" unfolding s'. moreover from satisfies_satisfies_no_lhs[OF this] satisfies_consistent[of s'] this have " \\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'" "\ s'" by blast+ moreover from check_bounds_id[OF **] *(6) have "Simplex.flat as \ \ s'" unfolding s' by (auto simp: boundsu_def boundsl_def) moreover from check_bounds_id[OF **] *(7) have "as \\<^sub>i \\ s'" unfolding s' by (auto simp: boundsu_def boundsl_def indexu_def indexl_def) moreover from U have "\ \ s'" unfolding s' . ultimately show ?case unfolding invariant_defs by auto next case (4 t as s s' I) (* check Unsat *) from 4(2)[unfolded check_s_def Let_def] have s': "s' = check s" and U: "\ (check s)" and I: "I = the (\\<^sub>c s')" by (auto split: if_splits) note * = invariant_sD[OF 4(1)] note ** = *(1,2,5,3,4) from check_unsat[OF ** U] have unsat: "minimal_unsat_state_core s'" unfolding s' by auto from check_tableau_equiv[OF **] *(9) have tabl: "\v :: _ \ 'a. v \\<^sub>t \ s' = v \\<^sub>t t" unfolding s' by auto from check_tableau_index_valid[OF **] *(8) have index: "index_valid as s'" unfolding s' by auto from check_bounds_id[OF **] *(7) have imp: "as \\<^sub>i \\ s'" unfolding s' by (auto simp: boundsu_def boundsl_def indexu_def indexl_def) from check_bounds_id[OF **] *(6) have bequiv: "Simplex.flat as \ \ s'" unfolding s' by (auto simp: boundsu_def boundsl_def) have "weak_invariant_s t as s'" unfolding invariant_defs using check_tableau_normalized[OF **] check_tableau_valuated[OF **] check_tableau[OF **] unfolding s'[symmetric] by (intro conjI index imp tabl bequiv, auto) with minimal_unsat_state_core_translation[OF unsat tabl index imp I] show ?case by auto next case *: (5 t) (* init *) show ?case unfolding invariant_defs using init_tableau_normalized[OF *] init_index_valid[of _ t] init_atoms_imply_bounds_index[of t] init_satisfies[of t] init_atoms_equiv_bounds[of t] init_tableau_id[of t] init_unsat_flag[of t] init_tableau_valuated[of t] satisfies_consistent[of "init t"] satisfies_satisfies_no_lhs[of "init t"] by auto next case (6 t as s v) (* solution *) then show ?case unfolding invariant_defs by (meson atoms_equiv_bounds.simps curr_val_satisfies_state_def satisfies_state_def) next case (7 t as s c bs s' s'') (* checkpoint and backtrack *) from 7(1)[unfolded checked_s_def] have inv_s: "invariant_s t as s" and s: "\ s" by auto from 7(2) have c: "c = \\<^sub>i s" unfolding checkpoint_s_def by auto have s'': "\ s'' = \ s'" "\ s'' = \ s'" "\\<^sub>i s'' = \\<^sub>i s" "\ s'' = False" "\\<^sub>c s'' = None" unfolding 7(4)[symmetric] c by (atomize(full), cases s', auto) then have BI: "\ s'' = \ s" "\ s'' = \ s" by (auto simp: boundsu_def boundsl_def indexu_def indexl_def) note * = invariant_sD[OF inv_s] note ** = weak_invariant_sD[OF 7(3)] have "\ \ s''" unfolding s'' by auto moreover from **(2) have "\ (\ s'')" unfolding s'' . moreover from **(3) have "\ s''" unfolding tableau_valuated_def s'' . moreover from **(8) have "\v :: _ \ 'a. v \\<^sub>t \ s'' = v \\<^sub>t t" unfolding s'' . moreover from *(6) have "Simplex.flat as \ \ s''" unfolding BI . moreover from *(7) have "as \\<^sub>i \\ s''" unfolding BI . moreover from *(8) have "index_valid as s''" unfolding index_valid_def using s'' by auto moreover from **(3) have "\ s''" unfolding tableau_valuated_def s'' . moreover from satisfies_consistent[of s] s have "\ s''" unfolding bounds_consistent_def using BI by auto moreover from 7(5) *(6) **(5) have vB: "v \\<^sub>b \ s' \ v \\<^sub>b \ s''" for v unfolding atoms_equiv_bounds.simps satisfies_atom_set_def BI by force from **(1) have t: "\\ s'\ \\<^sub>t \ s'" and b: "\\ s'\ \\<^sub>b \ s' \ - lvars (\ s')" unfolding curr_val_satisfies_no_lhs_def by auto let ?v = "\ x. if x \ lvars (\ s') then case \\<^sub>l s' x of None \ the (\\<^sub>u s' x) | Some b \ b else \\ s'\ x" have "?v \\<^sub>b \ s'" unfolding satisfies_bounds.simps proof (intro allI) fix x :: var show "in_bounds x ?v (\ s')" proof (cases "x \ lvars (\ s')") case True with **(4)[unfolded bounds_consistent_def, rule_format, of x] show ?thesis by (cases "\\<^sub>l s' x"; cases "\\<^sub>u s' x", auto simp: bound_compare_defs) next case False with b show ?thesis unfolding satisfies_bounds_set.simps by auto qed qed from vB[OF this] have v: "?v \\<^sub>b \ s''" . have "\\ s'\ \\<^sub>b \ s'' \ - lvars (\ s')" unfolding satisfies_bounds_set.simps proof clarify fix x assume "x \ lvars (\ s')" with v[unfolded satisfies_bounds.simps, rule_format, of x] show "in_bounds x \\ s'\ (\ s'')" by auto qed with t have "\\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s''" unfolding curr_val_satisfies_no_lhs_def s'' by auto ultimately show ?case unfolding invariant_defs by blast qed (auto simp: invariant_defs) end subsection \Intermediate Layer: Incremental Non-Strict Constraints\ text \Interface\ locale Incremental_NS_Constraint_Ops = fixes init_nsc :: "('i,'a :: lrv) i_ns_constraint list \ 's" and assert_nsc :: "'i \ 's \ 'i list + 's" and check_nsc :: "'s \ 's \ ('i list option)" and solution_nsc :: "'s \ (var, 'a) mapping" and checkpoint_nsc :: "'s \ 'c" and backtrack_nsc :: "'c \ 's \ 's" and weak_invariant_nsc :: "('i,'a) i_ns_constraint list \ 'i set \ 's \ bool" and invariant_nsc :: "('i,'a) i_ns_constraint list \ 'i set \ 's \ bool" and checked_nsc :: "('i,'a) i_ns_constraint list \ 'i set \ 's \ bool" assumes assert_nsc_ok: "invariant_nsc nsc J s \ assert_nsc j s = Inr s' \ invariant_nsc nsc (insert j J) s'" and assert_nsc_unsat: "invariant_nsc nsc J s \ assert_nsc j s = Unsat I \ set I \ insert j J \ minimal_unsat_core_ns (set I) (set nsc)" and check_nsc_ok: "invariant_nsc nsc J s \ check_nsc s = (s', None) \ checked_nsc nsc J s'" and check_nsc_unsat: "invariant_nsc nsc J s \ check_nsc s = (s',Some I) \ set I \ J \ weak_invariant_nsc nsc J s' \ minimal_unsat_core_ns (set I) (set nsc)" and init_nsc: "checked_nsc nsc {} (init_nsc nsc)" and solution_nsc: "checked_nsc nsc J s \ solution_nsc s = v \ (J, \v\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set nsc" and backtrack_nsc: "checked_nsc nsc J s \ checkpoint_nsc s = c \ weak_invariant_nsc nsc K s' \ backtrack_nsc c s' = s'' \ J \ K \ invariant_nsc nsc J s''" and weak_invariant_nsc: "invariant_nsc nsc J s \ weak_invariant_nsc nsc J s" and checked_invariant_nsc: "checked_nsc nsc J s \ invariant_nsc nsc J s" text \Implementation via the Simplex operation preprocess and the incremental operations for atoms.\ fun create_map :: "('i \ 'a)list \ ('i, ('i \ 'a) list)mapping" where "create_map [] = Mapping.empty" | "create_map ((i,a) # xs) = (let m = create_map xs in case Mapping.lookup m i of None \ Mapping.update i [(i,a)] m | Some ias \ Mapping.update i ((i,a) # ias) m)" definition list_map_to_fun :: "('i, ('i \ 'a) list)mapping \ 'i \ ('i \ 'a) list" where "list_map_to_fun m i = (case Mapping.lookup m i of None \ [] | Some ias \ ias)" lemma list_map_to_fun_create_map: "set (list_map_to_fun (create_map ias) i) = set ias \ {i} \ UNIV" proof (induct ias) case Nil show ?case unfolding list_map_to_fun_def by auto next case (Cons ja ias) obtain j a where ja: "ja = (j,a)" by force show ?case proof (cases "j = i") case False then have id: "list_map_to_fun (create_map (ja # ias)) i = list_map_to_fun (create_map ias) i" unfolding ja list_map_to_fun_def by (auto simp: Let_def split: option.splits) show ?thesis unfolding id Cons unfolding ja using False by auto next case True with ja have ja: "ja = (i,a)" by auto have id: "list_map_to_fun (create_map (ja # ias)) i = ja # list_map_to_fun (create_map ias) i" unfolding ja list_map_to_fun_def by (auto simp: Let_def split: option.splits) show ?thesis unfolding id using Cons unfolding ja by auto qed qed fun prod_wrap :: "('c \ 's \ 's \ ('i list option)) \ 'c \ 's \ ('c \ 's) \ ('i list option)" where "prod_wrap f (asi,s) = (case f asi s of (s', info) \ ((asi,s'), info))" lemma prod_wrap_def': "prod_wrap f (asi,s) = map_prod (Pair asi) id (f asi s)" unfolding prod_wrap.simps by (auto split: prod.splits) locale Incremental_Atom_Ops_For_NS_Constraint_Ops = Incremental_Atom_Ops init_s assert_s check_s solution_s checkpoint_s backtrack_s \ weak_invariant_s invariant_s checked_s + Preprocess preprocess for init_s :: "tableau \ 's" and assert_s :: "('i :: linorder,'a :: lrv) i_atom \ 's \ 'i list + 's" and check_s :: "'s \ 's \ 'i list option" and solution_s :: "'s \ (var, 'a) mapping" and checkpoint_s :: "'s \ 'c" and backtrack_s :: "'c \ 's \ 's" and weak_invariant_s :: "tableau \ ('i,'a) i_atom set \ 's \ bool" and invariant_s :: "tableau \ ('i,'a) i_atom set \ 's \ bool" and checked_s :: "tableau \ ('i,'a) i_atom set \ 's \ bool" and preprocess :: "('i,'a) i_ns_constraint list \ tableau \ ('i,'a) i_atom list \ ((var,'a)mapping \ (var,'a)mapping) \ 'i list" begin definition check_nsc where "check_nsc = prod_wrap (\ asitv. check_s)" definition assert_nsc where "assert_nsc i = (\ ((asi,tv,ui),s). if i \ set ui then Unsat [i] else case assert_all_s (list_map_to_fun asi i) s of Unsat I \ Unsat I | Inr s' \ Inr ((asi,tv,ui),s'))" fun checkpoint_nsc where "checkpoint_nsc (asi_tv_ui,s) = checkpoint_s s" fun backtrack_nsc where "backtrack_nsc c (asi_tv_ui,s) = (asi_tv_ui, backtrack_s c s)" fun solution_nsc where "solution_nsc ((asi,tv,ui),s) = tv (solution_s s)" definition "init_nsc nsc = (case preprocess nsc of (t,as,trans_v,ui) \ ((create_map as, trans_v, remdups ui), init_s t))" fun invariant_as_asi where "invariant_as_asi as asi tc tc' ui ui' = (tc = tc' \ set ui = set ui' \ (\ i. set (list_map_to_fun asi i) = (as \ ({i} \ UNIV))))" fun weak_invariant_nsc where "weak_invariant_nsc nsc J ((asi,tv,ui),s) = (case preprocess nsc of (t,as,tv',ui') \ invariant_as_asi (set as) asi tv tv' ui ui' \ weak_invariant_s t (set as \ (J \ UNIV)) s \ J \ set ui = {})" fun invariant_nsc where "invariant_nsc nsc J ((asi,tv,ui),s) = (case preprocess nsc of (t,as,tv',ui') \ invariant_as_asi (set as) asi tv tv' ui ui' \ invariant_s t (set as \ (J \ UNIV)) s \ J \ set ui = {})" fun checked_nsc where "checked_nsc nsc J ((asi,tv,ui),s) = (case preprocess nsc of (t,as,tv',ui') \ invariant_as_asi (set as) asi tv tv' ui ui' \ checked_s t (set as \ (J \ UNIV)) s \ J \ set ui = {})" lemma i_satisfies_atom_set_inter_right: "((I, v) \\<^sub>i\<^sub>a\<^sub>s (ats \ (J \ UNIV))) \ ((I \ J, v) \\<^sub>i\<^sub>a\<^sub>s ats)" unfolding i_satisfies_atom_set.simps by (rule arg_cong[of _ _ "\ x. v \\<^sub>a\<^sub>s x"], auto) lemma ns_constraints_ops: "Incremental_NS_Constraint_Ops init_nsc assert_nsc check_nsc solution_nsc checkpoint_nsc backtrack_nsc weak_invariant_nsc invariant_nsc checked_nsc" proof (unfold_locales, goal_cases) case (1 nsc J S j S') (* assert ok *) obtain asi tv s ui where S: "S = ((asi,tv,ui),s)" by (cases S, auto) obtain t as tv' ui' where prep[simp]: "preprocess nsc = (t, as, tv', ui')" by (cases "preprocess nsc") note pre = 1[unfolded S assert_nsc_def] from pre(2) obtain s' where ok: "assert_all_s (list_map_to_fun asi j) s = Inr s'" and S': "S' = ((asi,tv,ui),s')" and j: "j \ set ui" by (auto split: sum.splits if_splits) from pre(1)[simplified] have inv: "invariant_s t (set as \ J \ UNIV) s" and asi: "set (list_map_to_fun asi j) = set as \ {j} \ UNIV" "invariant_as_asi (set as) asi tv tv' ui ui'" "J \ set ui = {}" by auto from assert_all_s_ok[OF inv ok, unfolded asi] asi(2-) j show ?case unfolding invariant_nsc.simps S' prep split by (metis Int_insert_left Sigma_Un_distrib1 inf_sup_distrib1 insert_is_Un) next case (2 nsc J S j I) (* assert unsat *) obtain asi s tv ui where S: "S = ((asi,tv,ui),s)" by (cases S, auto) obtain t as tv' ui' where prep[simp]: "preprocess nsc = (t, as, tv', ui')" by (cases "preprocess nsc") note pre = 2[unfolded S assert_nsc_def split] show ?case proof (cases "j \ set ui") case False with pre(2) have unsat: "assert_all_s (list_map_to_fun asi j) s = Unsat I" by (auto split: sum.splits) from pre(1) have inv: "invariant_s t (set as \ J \ UNIV) s" and asi: "set (list_map_to_fun asi j) = set as \ {j} \ UNIV" by auto from assert_all_s_unsat[OF inv unsat, unfolded asi] have "minimal_unsat_core_tabl_atoms (set I) t (set as \ J \ UNIV \ set as \ {j} \ UNIV)" . also have "set as \ J \ UNIV \ set as \ {j} \ UNIV = set as \ insert j J \ UNIV" by blast finally have unsat: "minimal_unsat_core_tabl_atoms (set I) t (set as \ insert j J \ UNIV)" . hence I: "set I \ insert j J" unfolding minimal_unsat_core_tabl_atoms_def by force with False pre have empty: "set I \ set ui' = {}" by auto have "minimal_unsat_core_tabl_atoms (set I) t (set as)" by (rule minimal_unsat_core_tabl_atoms_mono[OF _ unsat], auto) from preprocess_minimal_unsat_core[OF prep this empty] have "minimal_unsat_core_ns (set I) (set nsc)" . then show ?thesis using I by blast next case True with pre(2) have I: "I = [j]" by auto from pre(1)[unfolded invariant_nsc.simps prep split invariant_as_asi.simps] have "set ui = set ui'" by simp with True have j: "j \ set ui'" by auto from preprocess_unsat_index[OF prep j] show ?thesis unfolding I by auto qed next case (3 nsc J S S') (* check ok *) then show ?case using check_s_ok unfolding check_nsc_def by (cases S, auto split: prod.splits, blast) next case (4 nsc J S S' I) (* check unsat *) obtain asi s tv ui where S: "S = ((asi,tv,ui),s)" by (cases S, auto) obtain t as tv' ui' where prep[simp]: "preprocess nsc = (t, as, tv', ui')" by (cases "preprocess nsc") from 4(2)[unfolded S check_nsc_def, simplified] obtain s' where unsat: "check_s s = (s', Some I)" and S': "S' = ((asi, tv, ui), s')" by (cases "check_s s", auto) note pre = 4[unfolded S check_nsc_def unsat, simplified] from pre have inv: "invariant_s t (set as \ J \ UNIV) s" by auto from check_s_unsat[OF inv unsat] have weak: "weak_invariant_s t (set as \ J \ UNIV) s'" and unsat: "minimal_unsat_core_tabl_atoms (set I) t (set as \ J \ UNIV)" by auto hence I: "set I \ J" unfolding minimal_unsat_core_tabl_atoms_def by force with pre have empty: "set I \ set ui' = {}" by auto have "minimal_unsat_core_tabl_atoms (set I) t (set as)" by (rule minimal_unsat_core_tabl_atoms_mono[OF _ unsat], auto) from preprocess_minimal_unsat_core[OF prep this empty] have "minimal_unsat_core_ns (set I) (set nsc)" . then show ?case using I weak unfolding S' using pre by auto next case (5 nsc) (* init *) obtain t as tv' ui' where prep[simp]: "preprocess nsc = (t, as, tv', ui')" by (cases "preprocess nsc") show ?case unfolding init_nsc_def using init_s preprocess_tableau_normalized[OF prep] by (auto simp: list_map_to_fun_create_map) next case (6 nsc J S v) (* solution *) obtain asi s tv ui where S: "S = ((asi,tv,ui),s)" by (cases S, auto) obtain t as tv' ui' where prep[simp]: "preprocess nsc = (t, as, tv', ui')" by (cases "preprocess nsc") have "(J,\solution_s s\) \\<^sub>i\<^sub>a\<^sub>s set as" "\solution_s s\ \\<^sub>t t" using 6 S solution_s[of t _ s] by auto from i_preprocess_sat[OF prep _ this] show ?case using 6 S by auto next case (7 nsc J S c K S' S'') (* backtrack *) obtain t as tvp uip where prep[simp]: "preprocess nsc = (t, as, tvp, uip)" by (cases "preprocess nsc") obtain asi s tv ui where S: "S = ((asi,tv,ui),s)" by (cases S, auto) obtain asi' s' tv' ui' where S': "S' = ((asi',tv',ui'),s')" by (cases S', auto) obtain asi'' s'' tv'' ui'' where S'': "S'' = ((asi'',tv'',ui''),s'')" by (cases S'', auto) from backtrack_s[of t _ s c _ s' s''] show ?case using 7 S S' S'' by auto next case (8 nsc J S) then show ?case using weak_invariant_s by (cases S, auto) next case (9 nsc J S) then show ?case using checked_invariant_s by (cases S, auto) qed end subsection \Highest Layer: Incremental Constraints\ text \Interface\ locale Incremental_Simplex_Ops = fixes init_cs :: "'i i_constraint list \ 's" and assert_cs :: "'i \ 's \ 'i list + 's" and check_cs :: "'s \ 's \ 'i list option" and solution_cs :: "'s \ rat valuation" and checkpoint_cs :: "'s \ 'c" and backtrack_cs :: "'c \ 's \ 's" and weak_invariant_cs :: "'i i_constraint list \ 'i set \ 's \ bool" and invariant_cs :: "'i i_constraint list \ 'i set \ 's \ bool" and checked_cs :: "'i i_constraint list \ 'i set \ 's \ bool" assumes assert_cs_ok: "invariant_cs cs J s \ assert_cs j s = Inr s' \ invariant_cs cs (insert j J) s'" and assert_cs_unsat: "invariant_cs cs J s \ assert_cs j s = Unsat I \ set I \ insert j J \ minimal_unsat_core (set I) cs" and check_cs_ok: "invariant_cs cs J s \ check_cs s = (s', None) \ checked_cs cs J s'" and check_cs_unsat: "invariant_cs cs J s \ check_cs s = (s',Some I) \ weak_invariant_cs cs J s' \ set I \ J \ minimal_unsat_core (set I) cs" and init_cs: "checked_cs cs {} (init_cs cs)" and solution_cs: "checked_cs cs J s \ solution_cs s = v \ (J, v) \\<^sub>i\<^sub>c\<^sub>s set cs" and backtrack_cs: "checked_cs cs J s \ checkpoint_cs s = c \ weak_invariant_cs cs K s' \ backtrack_cs c s' = s'' \ J \ K \ invariant_cs cs J s''" and weak_invariant_cs: "invariant_cs cs J s \ weak_invariant_cs cs J s" and checked_invariant_cs: "checked_cs cs J s \ invariant_cs cs J s" text \Implementation via the Simplex-operation To-Ns and the Incremental Operations for Non-Strict Constraints\ locale Incremental_NS_Constraint_Ops_To_Ns_For_Incremental_Simplex = Incremental_NS_Constraint_Ops init_nsc assert_nsc check_nsc solution_nsc checkpoint_nsc backtrack_nsc weak_invariant_nsc invariant_nsc checked_nsc + To_ns to_ns from_ns for init_nsc :: "('i,'a :: lrv) i_ns_constraint list \ 's" and assert_nsc :: "'i \ 's \ 'i list + 's" and check_nsc :: "'s \ 's \ 'i list option" and solution_nsc :: "'s \ (var, 'a) mapping" and checkpoint_nsc :: "'s \ 'c" and backtrack_nsc :: "'c \ 's \ 's" and weak_invariant_nsc :: "('i,'a) i_ns_constraint list \ 'i set \ 's \ bool" and invariant_nsc :: "('i,'a) i_ns_constraint list \ 'i set \ 's \ bool" and checked_nsc :: "('i,'a) i_ns_constraint list \ 'i set \ 's \ bool" and to_ns :: "'i i_constraint list \ ('i,'a) i_ns_constraint list" and from_ns :: "(var, 'a) mapping \ 'a ns_constraint list \ (var, rat) mapping" begin fun assert_cs where "assert_cs i (cs,s) = (case assert_nsc i s of Unsat I \ Unsat I | Inr s' \ Inr (cs, s'))" definition "init_cs cs = (let tons_cs = to_ns cs in (map snd (tons_cs), init_nsc tons_cs))" definition "check_cs s = prod_wrap (\ cs. check_nsc) s" fun checkpoint_cs where "checkpoint_cs (cs,s) = (checkpoint_nsc s)" fun backtrack_cs where "backtrack_cs c (cs,s) = (cs, backtrack_nsc c s)" fun solution_cs where "solution_cs (cs,s) = (\from_ns (solution_nsc s) cs\)" fun weak_invariant_cs where "weak_invariant_cs cs J (ds,s) = (ds = map snd (to_ns cs) \ weak_invariant_nsc (to_ns cs) J s)" fun invariant_cs where "invariant_cs cs J (ds,s) = (ds = map snd (to_ns cs) \ invariant_nsc (to_ns cs) J s)" fun checked_cs where "checked_cs cs J (ds,s) = (ds = map snd (to_ns cs) \ checked_nsc (to_ns cs) J s)" sublocale Incremental_Simplex_Ops init_cs assert_cs check_cs solution_cs checkpoint_cs backtrack_cs weak_invariant_cs invariant_cs checked_cs proof (unfold_locales, goal_cases) case (1 cs J S j S') (* assert ok *) then obtain s where S: "S = (map snd (to_ns cs),s)" by (cases S, auto) note pre = 1[unfolded S assert_cs.simps] from pre(2) obtain s' where ok: "assert_nsc j s = Inr s'" and S': "S' = (map snd (to_ns cs),s')" by (auto split: sum.splits) from pre(1) have inv: "invariant_nsc (to_ns cs) J s" by simp from assert_nsc_ok[OF inv ok] show ?case unfolding invariant_cs.simps S' split by auto next case (2 cs J S j I) (* assert unsat *) then obtain s where S: "S = (map snd (to_ns cs), s)" by (cases S, auto) note pre = 2[unfolded S assert_cs.simps] from pre(2) have unsat: "assert_nsc j s = Unsat I" by (auto split: sum.splits) from pre(1) have inv: "invariant_nsc (to_ns cs) J s" by auto from assert_nsc_unsat[OF inv unsat] have "set I \ insert j J" "minimal_unsat_core_ns (set I) (set (to_ns cs))" by auto from to_ns_unsat[OF this(2)] this(1) show ?case by blast next case (3 cs J S S') (* check ok *) then show ?case using check_nsc_ok unfolding check_cs_def by (cases S, auto split: prod.splits) next case (4 cs J S S' I) (* check unsat *) then obtain s where S: "S = (map snd (to_ns cs),s)" by (cases S, auto) note pre = 4[unfolded S check_cs_def] from pre(2) obtain s' where unsat: "check_nsc s = (s',Some I)" and S': "S' = (map snd (to_ns cs),s')" by (auto split: prod.splits) from pre(1) have inv: "invariant_nsc (to_ns cs) J s" by auto from check_nsc_unsat[OF inv unsat] have "set I \ J" "weak_invariant_nsc (to_ns cs) J s'" "minimal_unsat_core_ns (set I) (set (to_ns cs))" unfolding minimal_unsat_core_ns_def by auto from to_ns_unsat[OF this(3)] this(1,2) show ?case unfolding S' using S by auto next case (5 cs) (* init *) show ?case unfolding init_cs_def Let_def using init_nsc by auto next case (6 cs J S v) (* solution *) then obtain s where S: "S = (map snd (to_ns cs),s)" by (cases S, auto) obtain w where w: "solution_nsc s = w" by auto note pre = 6[unfolded S solution_cs.simps w Let_def] from pre have inv: "checked_nsc (to_ns cs) J s" and v: "v = \from_ns w (map snd (to_ns cs))\" by auto from solution_nsc[OF inv w] have w: "(J, \w\) \\<^sub>i\<^sub>n\<^sub>s\<^sub>s set (to_ns cs)" . from i_to_ns_sat[OF w] show ?case unfolding v . next case (7 cs J S c K S' S'') (* backtrack *) then show ?case using backtrack_nsc[of "to_ns cs" J] by (cases S, cases S', cases S'', auto) next case (8 cs J S) then show ?case using weak_invariant_nsc by (cases S, auto) next case (9 cs J S) then show ?case using checked_invariant_nsc by (cases S, auto) qed end subsection \Concrete Implementation\ subsubsection \Connecting all the locales\ global_interpretation Incremental_State_Ops_Simplex_Default: Incremental_State_Ops_Simplex assert_bound_code init_state check_code defines assert_s = Incremental_State_Ops_Simplex_Default.assert_s and check_s = Incremental_State_Ops_Simplex_Default.check_s and backtrack_s = Incremental_State_Ops_Simplex_Default.backtrack_s and checkpoint_s = Incremental_State_Ops_Simplex_Default.checkpoint_s and weak_invariant_s = Incremental_State_Ops_Simplex_Default.weak_invariant_s and invariant_s = Incremental_State_Ops_Simplex_Default.invariant_s and checked_s = Incremental_State_Ops_Simplex_Default.checked_s and assert_all_s = Incremental_State_Ops_Simplex_Default.assert_all_s .. (* RT: I don't know why the following two lemmas and the declare are required, and not done automatically via the global-interpretation *) lemma Incremental_State_Ops_Simplex_Default_assert_all_s[simp]: "Incremental_State_Ops_Simplex_Default.assert_all_s = assert_all_s" by (metis assert_all_s_def assert_s_def) lemmas assert_all_s_code = Incremental_State_Ops_Simplex_Default.assert_all_s.simps[unfolded Incremental_State_Ops_Simplex_Default_assert_all_s] declare assert_all_s_code[code] global_interpretation Incremental_Atom_Ops_For_NS_Constraint_Ops_Default: Incremental_Atom_Ops_For_NS_Constraint_Ops init_state assert_s check_s \ checkpoint_s backtrack_s weak_invariant_s invariant_s checked_s preprocess defines init_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.init_nsc and check_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.check_nsc and assert_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.assert_nsc and checkpoint_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.checkpoint_nsc and solution_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.solution_nsc and backtrack_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.backtrack_nsc and invariant_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.invariant_nsc and weak_invariant_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.weak_invariant_nsc and checked_nsc = Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.checked_nsc .. type_synonym 'i simplex_state' = "QDelta ns_constraint list \ (('i, ('i \ QDelta atom) list) mapping \ ((var,QDelta)mapping \ (var,QDelta)mapping) \ 'i list) \ ('i, QDelta) state" global_interpretation Incremental_Simplex: Incremental_NS_Constraint_Ops_To_Ns_For_Incremental_Simplex init_nsc assert_nsc check_nsc solution_nsc checkpoint_nsc backtrack_nsc weak_invariant_nsc invariant_nsc checked_nsc to_ns from_ns defines init_simplex' = Incremental_Simplex.init_cs and assert_simplex' = Incremental_Simplex.assert_cs and check_simplex' = Incremental_Simplex.check_cs and backtrack_simplex' = Incremental_Simplex.backtrack_cs and checkpoint_simplex' = Incremental_Simplex.checkpoint_cs and solution_simplex' = Incremental_Simplex.solution_cs and weak_invariant_simplex' = Incremental_Simplex.weak_invariant_cs and invariant_simplex' = Incremental_Simplex.invariant_cs and checked_simplex' = Incremental_Simplex.checked_cs proof - interpret Incremental_NS_Constraint_Ops init_nsc assert_nsc check_nsc solution_nsc checkpoint_nsc backtrack_nsc weak_invariant_nsc invariant_nsc checked_nsc using Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.ns_constraints_ops . show "Incremental_NS_Constraint_Ops_To_Ns_For_Incremental_Simplex init_nsc assert_nsc check_nsc solution_nsc checkpoint_nsc backtrack_nsc weak_invariant_nsc invariant_nsc checked_nsc to_ns from_ns" .. qed subsubsection \An implementation which encapsulates the state\ text \In principle, we now already have a complete implementation of the incremental simplex algorithm with @{const init_simplex'}, @{const assert_simplex'}, etc. However, this implementation results in code where the interal type @{typ "'i simplex_state'"} becomes visible. Therefore, we now define all operations on a new type which encapsulates the internal construction.\ datatype 'i simplex_state = Simplex_State "'i simplex_state'" datatype 'i simplex_checkpoint = Simplex_Checkpoint "(nat, 'i \ QDelta) mapping \ (nat, 'i \ QDelta) mapping" fun init_simplex where "init_simplex cs = (let tons_cs = to_ns cs in Simplex_State (map snd tons_cs, case preprocess tons_cs of (t, as, trans_v, ui) \ ((create_map as, trans_v, remdups ui), init_state t)))" fun assert_simplex where "assert_simplex i (Simplex_State (cs, (asi, tv, ui), s)) = (if i \ set ui then Inl [i] else case assert_all_s (list_map_to_fun asi i) s of Inl y \ Inl y | Inr s' \ Inr (Simplex_State (cs, (asi, tv, ui), s')))" fun check_simplex where "check_simplex (Simplex_State (cs, asi_tv, s)) = (case check_s s of (s', res) \ (Simplex_State (cs, asi_tv, s'), res))" fun solution_simplex where "solution_simplex (Simplex_State (cs, (asi, tv, ui), s)) = \from_ns (tv (\ s)) cs\" fun checkpoint_simplex where "checkpoint_simplex (Simplex_State (cs, asi_tv, s)) = Simplex_Checkpoint (checkpoint_s s)" fun backtrack_simplex where "backtrack_simplex (Simplex_Checkpoint c) (Simplex_State (cs, asi_tv, s)) = Simplex_State (cs, asi_tv, backtrack_s c s)" subsubsection \Soundness of the incremental simplex implementation\ text \First link the unprimed constants against their primed counterparts.\ lemma init_simplex': "init_simplex cs = Simplex_State (init_simplex' cs)" by (simp add: Let_def Incremental_Simplex.init_cs_def Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.init_nsc_def) lemma assert_simplex': "assert_simplex i (Simplex_State s) = map_sum id Simplex_State (assert_simplex' i s)" by (cases s, cases "fst (snd s)", auto simp add: Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.assert_nsc_def split: sum.splits) lemma check_simplex': "check_simplex (Simplex_State s) = map_prod Simplex_State id (check_simplex' s)" by (cases s, simp add: Incremental_Simplex.check_cs_def Incremental_Atom_Ops_For_NS_Constraint_Ops_Default.check_nsc_def split: prod.splits) lemma solution_simplex': "solution_simplex (Simplex_State s) = solution_simplex' s" by (cases s, auto) lemma checkpoint_simplex': "checkpoint_simplex (Simplex_State s) = Simplex_Checkpoint (checkpoint_simplex' s)" by (cases s, auto split: sum.splits) lemma backtrack_simplex': "backtrack_simplex (Simplex_Checkpoint c) (Simplex_State s) = Simplex_State (backtrack_simplex' c s)" by (cases s, auto split: sum.splits) fun invariant_simplex where "invariant_simplex cs J (Simplex_State s) = invariant_simplex' cs J s" fun weak_invariant_simplex where "weak_invariant_simplex cs J (Simplex_State s) = weak_invariant_simplex' cs J s" fun checked_simplex where "checked_simplex cs J (Simplex_State s) = checked_simplex' cs J s" text \Hide implementation\ declare init_simplex.simps[simp del] declare assert_simplex.simps[simp del] declare check_simplex.simps[simp del] declare solution_simplex.simps[simp del] declare checkpoint_simplex.simps[simp del] declare backtrack_simplex.simps[simp del] text \Soundness lemmas\ lemma init_simplex: "checked_simplex cs {} (init_simplex cs)" using Incremental_Simplex.init_cs by (simp add: init_simplex') lemma assert_simplex_ok: "invariant_simplex cs J s \ assert_simplex j s = Inr s' \ invariant_simplex cs (insert j J) s'" proof (cases s) case s: (Simplex_State ss) show "invariant_simplex cs J s \ assert_simplex j s = Inr s' \ invariant_simplex cs (insert j J) s'" unfolding s invariant_simplex.simps assert_simplex' using Incremental_Simplex.assert_cs_ok[of cs J ss j] by (cases "assert_simplex' j ss", auto) qed lemma assert_simplex_unsat: "invariant_simplex cs J s \ assert_simplex j s = Inl I \ set I \ insert j J \ minimal_unsat_core (set I) cs" proof (cases s) case s: (Simplex_State ss) show "invariant_simplex cs J s \ assert_simplex j s = Inl I \ set I \ insert j J \ minimal_unsat_core (set I) cs" unfolding s invariant_simplex.simps assert_simplex' using Incremental_Simplex.assert_cs_unsat[of cs J ss j] by (cases "assert_simplex' j ss", auto) qed lemma check_simplex_ok: "invariant_simplex cs J s \ check_simplex s = (s',None) \ checked_simplex cs J s'" proof (cases s) case s: (Simplex_State ss) show "invariant_simplex cs J s \ check_simplex s = (s',None) \ checked_simplex cs J s'" unfolding s invariant_simplex.simps check_simplex.simps check_simplex' using Incremental_Simplex.check_cs_ok[of cs J ss] by (cases "check_simplex' ss", auto) qed lemma check_simplex_unsat: "invariant_simplex cs J s \ check_simplex s = (s',Some I) \ weak_invariant_simplex cs J s' \ set I \ J \ minimal_unsat_core (set I) cs" proof (cases s) case s: (Simplex_State ss) show "invariant_simplex cs J s \ check_simplex s = (s',Some I) \ weak_invariant_simplex cs J s' \ set I \ J \ minimal_unsat_core (set I) cs" unfolding s invariant_simplex.simps check_simplex.simps check_simplex' using Incremental_Simplex.check_cs_unsat[of cs J ss _ I] by (cases "check_simplex' ss", auto) qed lemma solution_simplex: "checked_simplex cs J s \ solution_simplex s = v \ (J, v) \\<^sub>i\<^sub>c\<^sub>s set cs" using Incremental_Simplex.solution_cs[of cs J] by (cases s, auto simp: solution_simplex') lemma backtrack_simplex: "checked_simplex cs J s \ checkpoint_simplex s = c \ weak_invariant_simplex cs K s' \ backtrack_simplex c s' = s'' \ J \ K \ invariant_simplex cs J s''" proof - obtain ss where ss: "s = Simplex_State ss" by (cases s, auto) obtain ss' where ss': "s' = Simplex_State ss'" by (cases s', auto) obtain ss'' where ss'': "s'' = Simplex_State ss''" by (cases s'', auto) obtain cc where cc: "c = Simplex_Checkpoint cc" by (cases c, auto) show "checked_simplex cs J s \ checkpoint_simplex s = c \ weak_invariant_simplex cs K s' \ backtrack_simplex c s' = s'' \ J \ K \ invariant_simplex cs J s''" unfolding ss ss' ss'' cc checked_simplex.simps invariant_simplex.simps checkpoint_simplex' backtrack_simplex' using Incremental_Simplex.backtrack_cs[of cs J ss cc K ss' ss''] by simp qed lemma weak_invariant_simplex: "invariant_simplex cs J s \ weak_invariant_simplex cs J s" using Incremental_Simplex.weak_invariant_cs[of cs J] by (cases s, auto) lemma checked_invariant_simplex: "checked_simplex cs J s \ invariant_simplex cs J s" using Incremental_Simplex.checked_invariant_cs[of cs J] by (cases s, auto) declare checked_simplex.simps[simp del] declare invariant_simplex.simps[simp del] declare weak_invariant_simplex.simps[simp del] text \From this point onwards, one should not look into the types @{typ "'i simplex_state"} and @{typ "'i simplex_checkpoint"}.\ text \For convenience: an assert-all function which takes multiple indices.\ fun assert_all_simplex :: "'i list \ 'i simplex_state \ 'i list + 'i simplex_state" where "assert_all_simplex [] s = Inr s" | "assert_all_simplex (j # J) s = (case assert_simplex j s of Unsat I \ Unsat I | Inr s' \ assert_all_simplex J s')" lemma assert_all_simplex_ok: "invariant_simplex cs J s \ assert_all_simplex K s = Inr s' \ invariant_simplex cs (J \ set K) s'" proof (induct K arbitrary: s J) case (Cons k K s J) from Cons(3) obtain s'' where ass: "assert_simplex k s = Inr s''" and rec: "assert_all_simplex K s'' = Inr s'" by (auto split: sum.splits) from Cons(1)[OF assert_simplex_ok[OF Cons(2) ass] rec] show ?case by auto qed auto lemma assert_all_simplex_unsat: "invariant_simplex cs J s \ assert_all_simplex K s = Unsat I \ set I \ set K \ J \ minimal_unsat_core (set I) cs" proof (induct K arbitrary: s J) case (Cons k K s J) show ?case proof (cases "assert_simplex k s") case unsat: (Inl J') with Cons have J': "J' = I" by auto from assert_simplex_unsat[OF Cons(2) unsat] have "set J' \ insert k J" "minimal_unsat_core (set J') cs" by auto then show ?thesis unfolding J' i_satisfies_cs.simps by auto next case (Inr s') from Cons(1)[OF assert_simplex_ok[OF Cons(2) Inr]] Cons(3) Inr show ?thesis by auto qed qed simp text \The collection of soundness lemmas for the incremental simplex algorithm.\ lemmas incremental_simplex = init_simplex assert_simplex_ok assert_simplex_unsat assert_all_simplex_ok assert_all_simplex_unsat check_simplex_ok check_simplex_unsat solution_simplex backtrack_simplex checked_invariant_simplex weak_invariant_simplex subsection \Test Executability and Example for Incremental Interface\ value (code) "let cs = [ (1 :: int, LT (lp_monom 1 1) 4), \ \$x_1 < 4$\ - (2, GTPP (lp_monom 2 1) (lp_monom 1 2)), \ \$2x_1 > x_2$\ - (3, EQPP (lp_monom 1 1) (lp_monom 2 2)), \ \$x_1 = 2x_2$\ + (2, GT (lp_monom 2 1 - lp_monom 1 2) 0), \ \$2x_1 - x_2 > 0$\ + (3, EQ (lp_monom 1 1 - lp_monom 2 2) 0), \ \$x_1 - 2x_2 = 0$\ (4, GT (lp_monom 2 2) 5), \ \$2x_2 > 5$\ (5, GT (lp_monom 3 0) 7), \ \$3x_0 > 7$\ (6, GT (lp_monom 3 3 + lp_monom (1/3) 2) 2)]; \ \$3x_3 + 1/3x_2 > 2$\ s1 = init_simplex cs; \ \initialize\ s2 = (case assert_all_simplex [1,2,3] s1 of Inr s \ s | Unsat _ \ undefined); \ \assert 1,2,3\ s3 = (case check_simplex s2 of (s,None) \ s | _ \ undefined); \ \check that 1,2,3 are sat.\ c123 = checkpoint_simplex s3; \ \after check, store checkpoint for backtracking\ s4 = (case assert_simplex 4 s2 of Inr s \ s | Unsat _ \ undefined); \ \assert 4\ (s5,I) = (case check_simplex s4 of (s,Some I) \ (s,I) | _ \ undefined); \ \checking detects unsat-core 1,3,4\ s6 = backtrack_simplex c123 s5; \ \backtrack to constraints 1,2,3\ s7 = (case assert_all_simplex [5,6] s6 of Inr s \ s | Unsat _ \ undefined); \ \assert 5,6\ s8 = (case check_simplex s7 of (s,None) \ s | _ \ undefined); \ \check that 1,2,3,5,6 are sat.\ sol = solution_simplex s8 \ \solution for 1,2,3,5,6\ in (I, map (\ x. (''x_'', x, ''='', sol x)) [0,1,2,3]) \ \output unsat core and solution\" end \ No newline at end of file