Theory Dense_Linear_Order

Up to index of Isabelle/HOL

theory Dense_Linear_Order
imports Finite_Set
uses Tools/Qelim/qelim.ML Tools/Qelim/langford_data.ML Tools/Qelim/ferrante_rackoff_data.ML (Tools/Qelim/langford.ML) (Tools/Qelim/ferrante_rackoff.ML)
begin

(*
    ID:         $Id: Dense_Linear_Order.thy,v 1.16 2007/10/19 21:21:06 wenzelm Exp $
    Author:     Amine Chaieb, TU Muenchen
*)

header {* Dense linear order without endpoints
  and a quantifier elimination procedure in Ferrante and Rackoff style *}

theory Dense_Linear_Order
imports Finite_Set
uses
  "Tools/Qelim/qelim.ML"
  "Tools/Qelim/langford_data.ML"
  "Tools/Qelim/ferrante_rackoff_data.ML"
  ("Tools/Qelim/langford.ML")
  ("Tools/Qelim/ferrante_rackoff.ML")
begin

setup Langford_Data.setup
setup Ferrante_Rackoff_Data.setup

context linorder
begin

lemma less_not_permute: "¬ (x < y ∧ y < x)" by (simp add: not_less linear)

lemma gather_simps: 
  shows 
  "(∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ U. x < y) ∧ x < u ∧ P x) <-> (∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ (insert u U). x < y) ∧ P x)"
  and "(∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ U. x < y) ∧ l < x ∧ P x) <-> (∃x. (∀y ∈ (insert l L). y < x) ∧ (∀y ∈ U. x < y) ∧ P x)"
  "(∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ U. x < y) ∧ x < u) <-> (∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ (insert u U). x < y))"
  and "(∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ U. x < y) ∧ l < x) <-> (∃x. (∀y ∈ (insert l L). y < x) ∧ (∀y ∈ U. x < y))"  by auto

lemma 
  gather_start: "(∃x. P x) ≡ (∃x. (∀y ∈ {}. y < x) ∧ (∀y∈ {}. x < y) ∧ P x)" 
  by simp

text{* Theorems for @{text "∃z. ∀x. x < z --> (P x <-> P-∞)"}*}
lemma minf_lt:  "∃z . ∀x. x < z --> (x < t <-> True)" by auto
lemma minf_gt: "∃z . ∀x. x < z -->  (t < x <->  False)"
  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)

lemma minf_le: "∃z. ∀x. x < z --> (x ≤ t <-> True)" by (auto simp add: less_le)
lemma minf_ge: "∃z. ∀x. x < z --> (t ≤ x <-> False)"
  by (auto simp add: less_le not_less not_le)
lemma minf_eq: "∃z. ∀x. x < z --> (x = t <-> False)" by auto
lemma minf_neq: "∃z. ∀x. x < z --> (x ≠ t <-> True)" by auto
lemma minf_P: "∃z. ∀x. x < z --> (P <-> P)" by blast

text{* Theorems for @{text "∃z. ∀x. x < z --> (P x <-> P+∞)"}*}
lemma pinf_gt:  "∃z . ∀x. z < x --> (t < x <-> True)" by auto
lemma pinf_lt: "∃z . ∀x. z < x -->  (x < t <->  False)"
  by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le)

lemma pinf_ge: "∃z. ∀x. z < x --> (t ≤ x <-> True)" by (auto simp add: less_le)
lemma pinf_le: "∃z. ∀x. z < x --> (x ≤ t <-> False)"
  by (auto simp add: less_le not_less not_le)
lemma pinf_eq: "∃z. ∀x. z < x --> (x = t <-> False)" by auto
lemma pinf_neq: "∃z. ∀x. z < x --> (x ≠ t <-> True)" by auto
lemma pinf_P: "∃z. ∀x. z < x --> (P <-> P)" by blast

lemma nmi_lt: "t ∈ U ==> ∀x. ¬True ∧ x < t -->  (∃ u∈ U. u ≤ x)" by auto
lemma nmi_gt: "t ∈ U ==> ∀x. ¬False ∧ t < x -->  (∃ u∈ U. u ≤ x)"
  by (auto simp add: le_less)
lemma  nmi_le: "t ∈ U ==> ∀x. ¬True ∧ x≤ t -->  (∃ u∈ U. u ≤ x)" by auto
lemma  nmi_ge: "t ∈ U ==> ∀x. ¬False ∧ t≤ x -->  (∃ u∈ U. u ≤ x)" by auto
lemma  nmi_eq: "t ∈ U ==> ∀x. ¬False ∧  x = t -->  (∃ u∈ U. u ≤ x)" by auto
lemma  nmi_neq: "t ∈ U ==>∀x. ¬True ∧ x ≠ t -->  (∃ u∈ U. u ≤ x)" by auto
lemma  nmi_P: "∀ x. ~P ∧ P -->  (∃ u∈ U. u ≤ x)" by auto
lemma  nmi_conj: "[|∀x. ¬P1' ∧ P1 x -->  (∃ u∈ U. u ≤ x) ;
  ∀x. ¬P2' ∧ P2 x -->  (∃ u∈ U. u ≤ x)|] ==>
  ∀x. ¬(P1' ∧ P2') ∧ (P1 x ∧ P2 x) -->  (∃ u∈ U. u ≤ x)" by auto
lemma  nmi_disj: "[|∀x. ¬P1' ∧ P1 x -->  (∃ u∈ U. u ≤ x) ;
  ∀x. ¬P2' ∧ P2 x -->  (∃ u∈ U. u ≤ x)|] ==>
  ∀x. ¬(P1' ∨ P2') ∧ (P1 x ∨ P2 x) -->  (∃ u∈ U. u ≤ x)" by auto

lemma  npi_lt: "t ∈ U ==> ∀x. ¬False ∧  x < t -->  (∃ u∈ U. x ≤ u)" by (auto simp add: le_less)
lemma  npi_gt: "t ∈ U ==> ∀x. ¬True ∧ t < x -->  (∃ u∈ U. x ≤ u)" by auto
lemma  npi_le: "t ∈ U ==> ∀x. ¬False ∧  x ≤ t -->  (∃ u∈ U. x ≤ u)" by auto
lemma  npi_ge: "t ∈ U ==> ∀x. ¬True ∧ t ≤ x -->  (∃ u∈ U. x ≤ u)" by auto
lemma  npi_eq: "t ∈ U ==> ∀x. ¬False ∧  x = t -->  (∃ u∈ U. x ≤ u)" by auto
lemma  npi_neq: "t ∈ U ==> ∀x. ¬True ∧ x ≠ t -->  (∃ u∈ U. x ≤ u )" by auto
lemma  npi_P: "∀ x. ~P ∧ P -->  (∃ u∈ U. x ≤ u)" by auto
lemma  npi_conj: "[|∀x. ¬P1' ∧ P1 x -->  (∃ u∈ U. x ≤ u) ;  ∀x. ¬P2' ∧ P2 x -->  (∃ u∈ U. x ≤ u)|]
  ==>  ∀x. ¬(P1' ∧ P2') ∧ (P1 x ∧ P2 x) -->  (∃ u∈ U. x ≤ u)" by auto
lemma  npi_disj: "[|∀x. ¬P1' ∧ P1 x -->  (∃ u∈ U. x ≤ u) ; ∀x. ¬P2' ∧ P2 x -->  (∃ u∈ U. x ≤ u)|]
  ==> ∀x. ¬(P1' ∨ P2') ∧ (P1 x ∨ P2 x) -->  (∃ u∈ U. x ≤ u)" by auto

lemma lin_dense_lt: "t ∈ U ==> ∀x l u. (∀ t. l < t ∧ t < u --> t ∉ U) ∧ l< x ∧ x < u ∧ x < t --> (∀ y. l < y ∧ y < u --> y < t)"
proof(clarsimp)
  fix x l u y  assume tU: "t ∈ U" and noU: "∀t. l < t ∧ t < u --> t ∉ U" and lx: "l < x"
    and xu: "x<u"  and px: "x < t" and ly: "l<y" and yu:"y < u"
  from tU noU ly yu have tny: "t≠y" by auto
  {assume H: "t < y"
    from less_trans[OF lx px] less_trans[OF H yu]
    have "l < t ∧ t < u"  by simp
    with tU noU have "False" by auto}
  hence "¬ t < y"  by auto hence "y ≤ t" by (simp add: not_less)
  thus "y < t" using tny by (simp add: less_le)
qed

lemma lin_dense_gt: "t ∈ U ==> ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l < x ∧ x < u ∧ t < x --> (∀ y. l < y ∧ y < u --> t < y)"
proof(clarsimp)
  fix x l u y
  assume tU: "t ∈ U" and noU: "∀t. l < t ∧ t < u --> t ∉ U" and lx: "l < x" and xu: "x<u"
  and px: "t < x" and ly: "l<y" and yu:"y < u"
  from tU noU ly yu have tny: "t≠y" by auto
  {assume H: "y< t"
    from less_trans[OF ly H] less_trans[OF px xu] have "l < t ∧ t < u" by simp
    with tU noU have "False" by auto}
  hence "¬ y<t"  by auto hence "t ≤ y" by (auto simp add: not_less)
  thus "t < y" using tny by (simp add:less_le)
qed

lemma lin_dense_le: "t ∈ U ==> ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ x ≤ t --> (∀ y. l < y ∧ y < u --> y≤ t)"
proof(clarsimp)
  fix x l u y
  assume tU: "t ∈ U" and noU: "∀t. l < t ∧ t < u --> t ∉ U" and lx: "l < x" and xu: "x<u"
  and px: "x ≤ t" and ly: "l<y" and yu:"y < u"
  from tU noU ly yu have tny: "t≠y" by auto
  {assume H: "t < y"
    from less_le_trans[OF lx px] less_trans[OF H yu]
    have "l < t ∧ t < u" by simp
    with tU noU have "False" by auto}
  hence "¬ t < y"  by auto thus "y ≤ t" by (simp add: not_less)
qed

lemma lin_dense_ge: "t ∈ U ==> ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ t ≤ x --> (∀ y. l < y ∧ y < u --> t ≤ y)"
proof(clarsimp)
  fix x l u y
  assume tU: "t ∈ U" and noU: "∀t. l < t ∧ t < u --> t ∉ U" and lx: "l < x" and xu: "x<u"
  and px: "t ≤ x" and ly: "l<y" and yu:"y < u"
  from tU noU ly yu have tny: "t≠y" by auto
  {assume H: "y< t"
    from less_trans[OF ly H] le_less_trans[OF px xu]
    have "l < t ∧ t < u" by simp
    with tU noU have "False" by auto}
  hence "¬ y<t"  by auto thus "t ≤ y" by (simp add: not_less)
qed
lemma lin_dense_eq: "t ∈ U ==> ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ x = t   --> (∀ y. l < y ∧ y < u --> y= t)"  by auto
lemma lin_dense_neq: "t ∈ U ==> ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ x ≠ t   --> (∀ y. l < y ∧ y < u --> y≠ t)"  by auto
lemma lin_dense_P: "∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ P   --> (∀ y. l < y ∧ y < u --> P)"  by auto

lemma lin_dense_conj:
  "[|∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ P1 x
  --> (∀ y. l < y ∧ y < u --> P1 y) ;
  ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ P2 x
  --> (∀ y. l < y ∧ y < u --> P2 y)|] ==>
  ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ (P1 x ∧ P2 x)
  --> (∀ y. l < y ∧ y < u --> (P1 y ∧ P2 y))"
  by blast
lemma lin_dense_disj:
  "[|∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ P1 x
  --> (∀ y. l < y ∧ y < u --> P1 y) ;
  ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ P2 x
  --> (∀ y. l < y ∧ y < u --> P2 y)|] ==>
  ∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ (P1 x ∨ P2 x)
  --> (∀ y. l < y ∧ y < u --> (P1 y ∨ P2 y))"
  by blast

lemma npmibnd: "[|∀x. ¬ MP ∧ P x --> (∃ u∈ U. u ≤ x); ∀x. ¬PP ∧ P x --> (∃ u∈ U. x ≤ u)|]
  ==> ∀x. ¬ MP ∧ ¬PP ∧ P x --> (∃ u∈ U. ∃ u' ∈ U. u ≤ x ∧ x ≤ u')"
by auto

lemma finite_set_intervals:
  assumes px: "P x" and lx: "l ≤ x" and xu: "x ≤ u" and linS: "l∈ S"
  and uinS: "u ∈ S" and fS:"finite S" and lS: "∀ x∈ S. l ≤ x" and Su: "∀ x∈ S. x ≤ u"
  shows "∃ a ∈ S. ∃ b ∈ S. (∀ y. a < y ∧ y < b --> y ∉ S) ∧ a ≤ x ∧ x ≤ b ∧ P x"
proof-
  let ?Mx = "{y. y∈ S ∧ y ≤ x}"
  let ?xM = "{y. y∈ S ∧ x ≤ y}"
  let ?a = "Max ?Mx"
  let ?b = "Min ?xM"
  have MxS: "?Mx ⊆ S" by blast
  hence fMx: "finite ?Mx" using fS finite_subset by auto
  from lx linS have linMx: "l ∈ ?Mx" by blast
  hence Mxne: "?Mx ≠ {}" by blast
  have xMS: "?xM ⊆ S" by blast
  hence fxM: "finite ?xM" using fS finite_subset by auto
  from xu uinS have linxM: "u ∈ ?xM" by blast
  hence xMne: "?xM ≠ {}" by blast
  have ax:"?a ≤ x" using Mxne fMx by auto
  have xb:"x ≤ ?b" using xMne fxM by auto
  have "?a ∈ ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a ∈ S" using MxS by blast
  have "?b ∈ ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b ∈ S" using xMS by blast
  have noy:"∀ y. ?a < y ∧ y < ?b --> y ∉ S"
  proof(clarsimp)
    fix y   assume ay: "?a < y" and yb: "y < ?b" and yS: "y ∈ S"
    from yS have "y∈ ?Mx ∨ y∈ ?xM" by (auto simp add: linear)
    moreover {assume "y ∈ ?Mx" hence "y ≤ ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])}
    moreover {assume "y ∈ ?xM" hence "?b ≤ y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])}
    ultimately show "False" by blast
  qed
  from ainS binS noy ax xb px show ?thesis by blast
qed

lemma finite_set_intervals2:
  assumes px: "P x" and lx: "l ≤ x" and xu: "x ≤ u" and linS: "l∈ S"
  and uinS: "u ∈ S" and fS:"finite S" and lS: "∀ x∈ S. l ≤ x" and Su: "∀ x∈ S. x ≤ u"
  shows "(∃ s∈ S. P s) ∨ (∃ a ∈ S. ∃ b ∈ S. (∀ y. a < y ∧ y < b --> y ∉ S) ∧ a < x ∧ x < b ∧ P x)"
proof-
  from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
  obtain a and b where
    as: "a∈ S" and bs: "b∈ S" and noS:"∀y. a < y ∧ y < b --> y ∉ S"
    and axb: "a ≤ x ∧ x ≤ b ∧ P x"  by auto
  from axb have "x= a ∨ x= b ∨ (a < x ∧ x < b)" by (auto simp add: le_less)
  thus ?thesis using px as bs noS by blast
qed

end

section {* The classical QE after Langford for dense linear orders *}

context dense_linear_order
begin

lemma dlo_qe_bnds: 
  assumes ne: "L ≠ {}" and neU: "U ≠ {}" and fL: "finite L" and fU: "finite U"
  shows "(∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ U. x < y)) ≡ (∀ l ∈ L. ∀u ∈ U. l < u)"
proof (simp only: atomize_eq, rule iffI)
  assume H: "∃x. (∀y∈L. y < x) ∧ (∀y∈U. x < y)"
  then obtain x where xL: "∀y∈L. y < x" and xU: "∀y∈U. x < y" by blast
  {fix l u assume l: "l ∈ L" and u: "u ∈ U"
    have "l < x" using xL l by blast
    also have "x < u" using xU u by blast
    finally (less_trans) have "l < u" .}
  thus "∀l∈L. ∀u∈U. l < u" by blast
next
  assume H: "∀l∈L. ∀u∈U. l < u"
  let ?ML = "Max L"
  let ?MU = "Min U"  
  from fL ne have th1: "?ML ∈ L" and th1': "∀l∈L. l ≤ ?ML" by auto
  from fU neU have th2: "?MU ∈ U" and th2': "∀u∈U. ?MU ≤ u" by auto
  from th1 th2 H have "?ML < ?MU" by auto
  with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast
  from th3 th1' have "∀l ∈ L. l < w" by auto
  moreover from th4 th2' have "∀u ∈ U. w < u" by auto
  ultimately show "∃x. (∀y∈L. y < x) ∧ (∀y∈U. x < y)" by auto
qed

lemma dlo_qe_noub: 
  assumes ne: "L ≠ {}" and fL: "finite L"
  shows "(∃x. (∀y ∈ L. y < x) ∧ (∀y ∈ {}. x < y)) ≡ True"
proof(simp add: atomize_eq)
  from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast
  from ne fL have "∀x ∈ L. x ≤ Max L" by simp
  with M have "∀x∈L. x < M" by (auto intro: le_less_trans)
  thus "∃x. ∀y∈L. y < x" by blast
qed

lemma dlo_qe_nolb: 
  assumes ne: "U ≠ {}" and fU: "finite U"
  shows "(∃x. (∀y ∈ {}. y < x) ∧ (∀y ∈ U. x < y)) ≡ True"
proof(simp add: atomize_eq)
  from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast
  from ne fU have "∀x ∈ U. Min U ≤ x" by simp
  with M have "∀x∈U. M < x" by (auto intro: less_le_trans)
  thus "∃x. ∀y∈U. x < y" by blast
qed

lemma exists_neq: "∃(x::'a). x ≠ t" "∃(x::'a). t ≠ x" 
  using gt_ex[of t] by auto

lemmas dlo_simps = order_refl less_irrefl not_less not_le exists_neq 
  le_less neq_iff linear less_not_permute

lemma axiom: "dense_linear_order (op ≤) (op <)" by fact
lemma atoms:
  includes meta_term_syntax
  shows "TERM (less :: 'a => _)"
    and "TERM (less_eq :: 'a => _)"
    and "TERM (op = :: 'a => _)" .

declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms]
declare dlo_simps[langfordsimp]

end

(* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *)
lemma dnf:
  "(P & (Q | R)) = ((P&Q) | (P&R))" 
  "((Q | R) & P) = ((Q&P) | (R&P))"
  by blast+

lemmas weak_dnf_simps = simp_thms dnf

lemma nnf_simps:
    "(¬(P ∧ Q)) = (¬P ∨ ¬Q)" "(¬(P ∨ Q)) = (¬P ∧ ¬Q)" "(P --> Q) = (¬P ∨ Q)"
    "(P = Q) = ((P ∧ Q) ∨ (¬P ∧ ¬ Q))" "(¬ ¬(P)) = P"
  by blast+

lemma ex_distrib: "(∃x. P x ∨ Q x) <-> ((∃x. P x) ∨ (∃x. Q x))" by blast

lemmas dnf_simps = weak_dnf_simps nnf_simps ex_distrib

use "Tools/Qelim/langford.ML"
method_setup dlo = {*
  Method.ctxt_args (Method.SIMPLE_METHOD' o LangfordQE.dlo_tac)
*} "Langford's algorithm for quantifier elimination in dense linear orders"


section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"} *}

text {* Linear order without upper bounds *}

class linorder_no_ub = linorder +
  assumes gt_ex: "∃y. x < y"
begin

lemma ge_ex: "∃y. x ≤ y" using gt_ex by auto

text {* Theorems for @{text "∃z. ∀x. z < x --> (P x <-> P+∞)"} *}
lemma pinf_conj:
  assumes ex1: "∃z1. ∀x. z1 < x --> (P1 x <-> P1')"
  and ex2: "∃z2. ∀x. z2 < x --> (P2 x <-> P2')"
  shows "∃z. ∀x. z <  x --> ((P1 x ∧ P2 x) <-> (P1' ∧ P2'))"
proof-
  from ex1 ex2 obtain z1 and z2 where z1: "∀x. z1 < x --> (P1 x <-> P1')"
     and z2: "∀x. z2 < x --> (P2 x <-> P2')" by blast
  from gt_ex obtain z where z:"max z1 z2 < z" by blast
  from z have zz1: "z1 < z" and zz2: "z2 < z" by simp_all
  {fix x assume H: "z < x"
    from less_trans[OF zz1 H] less_trans[OF zz2 H]
    have "(P1 x ∧ P2 x) <-> (P1' ∧ P2')"  using z1 zz1 z2 zz2 by auto
  }
  thus ?thesis by blast
qed

lemma pinf_disj:
  assumes ex1: "∃z1. ∀x. z1 < x --> (P1 x <-> P1')"
  and ex2: "∃z2. ∀x. z2 < x --> (P2 x <-> P2')"
  shows "∃z. ∀x. z <  x --> ((P1 x ∨ P2 x) <-> (P1' ∨ P2'))"
proof-
  from ex1 ex2 obtain z1 and z2 where z1: "∀x. z1 < x --> (P1 x <-> P1')"
     and z2: "∀x. z2 < x --> (P2 x <-> P2')" by blast
  from gt_ex obtain z where z:"max z1 z2 < z" by blast
  from z have zz1: "z1 < z" and zz2: "z2 < z" by simp_all
  {fix x assume H: "z < x"
    from less_trans[OF zz1 H] less_trans[OF zz2 H]
    have "(P1 x ∨ P2 x) <-> (P1' ∨ P2')"  using z1 zz1 z2 zz2 by auto
  }
  thus ?thesis by blast
qed

lemma pinf_ex: assumes ex:"∃z. ∀x. z < x --> (P x <-> P1)" and p1: P1 shows "∃ x. P x"
proof-
  from ex obtain z where z: "∀x. z < x --> (P x <-> P1)" by blast
  from gt_ex obtain x where x: "z < x" by blast
  from z x p1 show ?thesis by blast
qed

end

text {* Linear order without upper bounds *}

class linorder_no_lb = linorder +
  assumes lt_ex: "∃y. y < x"
begin

lemma le_ex: "∃y. y ≤ x" using lt_ex by auto


text {* Theorems for @{text "∃z. ∀x. x < z --> (P x <-> P-∞)"} *}
lemma minf_conj:
  assumes ex1: "∃z1. ∀x. x < z1 --> (P1 x <-> P1')"
  and ex2: "∃z2. ∀x. x < z2 --> (P2 x <-> P2')"
  shows "∃z. ∀x. x <  z --> ((P1 x ∧ P2 x) <-> (P1' ∧ P2'))"
proof-
  from ex1 ex2 obtain z1 and z2 where z1: "∀x. x < z1 --> (P1 x <-> P1')"and z2: "∀x. x < z2 --> (P2 x <-> P2')" by blast
  from lt_ex obtain z where z:"z < min z1 z2" by blast
  from z have zz1: "z < z1" and zz2: "z < z2" by simp_all
  {fix x assume H: "x < z"
    from less_trans[OF H zz1] less_trans[OF H zz2]
    have "(P1 x ∧ P2 x) <-> (P1' ∧ P2')"  using z1 zz1 z2 zz2 by auto
  }
  thus ?thesis by blast
qed

lemma minf_disj:
  assumes ex1: "∃z1. ∀x. x < z1 --> (P1 x <-> P1')"
  and ex2: "∃z2. ∀x. x < z2 --> (P2 x <-> P2')"
  shows "∃z. ∀x. x <  z --> ((P1 x ∨ P2 x) <-> (P1' ∨ P2'))"
proof-
  from ex1 ex2 obtain z1 and z2 where z1: "∀x. x < z1 --> (P1 x <-> P1')"and z2: "∀x. x < z2 --> (P2 x <-> P2')" by blast
  from lt_ex obtain z where z:"z < min z1 z2" by blast
  from z have zz1: "z < z1" and zz2: "z < z2" by simp_all
  {fix x assume H: "x < z"
    from less_trans[OF H zz1] less_trans[OF H zz2]
    have "(P1 x ∨ P2 x) <-> (P1' ∨ P2')"  using z1 zz1 z2 zz2 by auto
  }
  thus ?thesis by blast
qed

lemma minf_ex: assumes ex:"∃z. ∀x. x < z --> (P x <-> P1)" and p1: P1 shows "∃ x. P x"
proof-
  from ex obtain z where z: "∀x. x < z --> (P x <-> P1)" by blast
  from lt_ex obtain x where x: "x < z" by blast
  from z x p1 show ?thesis by blast
qed

end


class constr_dense_linear_order = linorder_no_lb + linorder_no_ub +
  fixes between
  assumes between_less: "x < y ==> x < between x y ∧ between x y < y"
     and  between_same: "between x x = x"
begin

subclass dense_linear_order
  apply unfold_locales
  using gt_ex lt_ex between_less
    by (auto, rule_tac x="between x y" in exI, simp)

lemma rinf_U:
  assumes fU: "finite U"
  and lin_dense: "∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ P x
  --> (∀ y. l < y ∧ y < u --> P y )"
  and nmpiU: "∀x. ¬ MP ∧ ¬PP ∧ P x --> (∃ u∈ U. ∃ u' ∈ U. u ≤ x ∧ x ≤ u')"
  and nmi: "¬ MP"  and npi: "¬ PP"  and ex: "∃ x.  P x"
  shows "∃ u∈ U. ∃ u' ∈ U. P (between u u')"
proof-
  from ex obtain x where px: "P x" by blast
  from px nmi npi nmpiU have "∃ u∈ U. ∃ u' ∈ U. u ≤ x ∧ x ≤ u'" by auto
  then obtain u and u' where uU:"u∈ U" and uU': "u' ∈ U" and ux:"u ≤ x" and xu':"x ≤ u'" by auto
  from uU have Une: "U ≠ {}" by auto
  let ?l = "Min U"
  let ?u = "Max U"
  have linM: "?l ∈ U" using fU Une by simp
  have uinM: "?u ∈ U" using fU Une by simp
  have lM: "∀ t∈ U. ?l ≤ t" using Une fU by auto
  have Mu: "∀ t∈ U. t ≤ ?u" using Une fU by auto
  have th:"?l ≤ u" using uU Une lM by auto
  from order_trans[OF th ux] have lx: "?l ≤ x" .
  have th: "u' ≤ ?u" using uU' Une Mu by simp
  from order_trans[OF xu' th] have xu: "x ≤ ?u" .
  from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu]
  have "(∃ s∈ U. P s) ∨
      (∃ t1∈ U. ∃ t2 ∈ U. (∀ y. t1 < y ∧ y < t2 --> y ∉ U) ∧ t1 < x ∧ x < t2 ∧ P x)" .
  moreover { fix u assume um: "u∈U" and pu: "P u"
    have "between u u = u" by (simp add: between_same)
    with um pu have "P (between u u)" by simp
    with um have ?thesis by blast}
  moreover{
    assume "∃ t1∈ U. ∃ t2 ∈ U. (∀ y. t1 < y ∧ y < t2 --> y ∉ U) ∧ t1 < x ∧ x < t2 ∧ P x"
      then obtain t1 and t2 where t1M: "t1 ∈ U" and t2M: "t2∈ U"
        and noM: "∀ y. t1 < y ∧ y < t2 --> y ∉ U" and t1x: "t1 < x" and xt2: "x < t2" and px: "P x"
        by blast
      from less_trans[OF t1x xt2] have t1t2: "t1 < t2" .
      let ?u = "between t1 t2"
      from between_less t1t2 have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto
      from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast
      with t1M t2M have ?thesis by blast}
    ultimately show ?thesis by blast
  qed

theorem fr_eq:
  assumes fU: "finite U"
  and lin_dense: "∀x l u. (∀ t. l < t ∧ t< u --> t ∉ U) ∧ l< x ∧ x < u ∧ P x
   --> (∀ y. l < y ∧ y < u --> P y )"
  and nmibnd: "∀x. ¬ MP ∧ P x --> (∃ u∈ U. u ≤ x)"
  and npibnd: "∀x. ¬PP ∧ P x --> (∃ u∈ U. x ≤ u)"
  and mi: "∃z. ∀x. x < z --> (P x = MP)"  and pi: "∃z. ∀x. z < x --> (P x = PP)"
  shows "(∃ x. P x) ≡ (MP ∨ PP ∨ (∃ u ∈ U. ∃ u'∈ U. P (between u u')))"
  (is "_ ≡ (_ ∨ _ ∨ ?F)" is "?E ≡ ?D")
proof-
 {
   assume px: "∃ x. P x"
   have "MP ∨ PP ∨ (¬ MP ∧ ¬ PP)" by blast
   moreover {assume "MP ∨ PP" hence "?D" by blast}
   moreover {assume nmi: "¬ MP" and npi: "¬ PP"
     from npmibnd[OF nmibnd npibnd]
     have nmpiU: "∀x. ¬ MP ∧ ¬PP ∧ P x --> (∃ u∈ U. ∃ u' ∈ U. u ≤ x ∧ x ≤ u')" .
     from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast}
   ultimately have "?D" by blast}
 moreover
 { assume "?D"
   moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .}
   moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . }
   moreover {assume f:"?F" hence "?E" by blast}
   ultimately have "?E" by blast}
 ultimately have "?E = ?D" by blast thus "?E ≡ ?D" by simp
qed

lemmas minf_thms = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P
lemmas pinf_thms = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P

lemmas nmi_thms = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P
lemmas npi_thms = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P
lemmas lin_dense_thms = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P

lemma ferrack_axiom: "constr_dense_linear_order less_eq less between" by fact
lemma atoms:
  includes meta_term_syntax
  shows "TERM (less :: 'a => _)"
    and "TERM (less_eq :: 'a => _)"
    and "TERM (op = :: 'a => _)" .

declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms
    nmi: nmi_thms npi: npi_thms lindense:
    lin_dense_thms qe: fr_eq atoms: atoms]

declaration {*
let
fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}]
fun generic_whatis phi =
 let
  val [lt, le] = map (Morphism.term phi) [@{term "op <"}, @{term "op ≤"}]
  fun h x t =
   case term_of t of
     Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
                            else Ferrante_Rackoff_Data.Nox
   | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
                            else Ferrante_Rackoff_Data.Nox
   | b$y$z => if Term.could_unify (b, lt) then
                 if term_of x aconv y then Ferrante_Rackoff_Data.Lt
                 else if term_of x aconv z then Ferrante_Rackoff_Data.Gt
                 else Ferrante_Rackoff_Data.Nox
             else if Term.could_unify (b, le) then
                 if term_of x aconv y then Ferrante_Rackoff_Data.Le
                 else if term_of x aconv z then Ferrante_Rackoff_Data.Ge
                 else Ferrante_Rackoff_Data.Nox
             else Ferrante_Rackoff_Data.Nox
   | _ => Ferrante_Rackoff_Data.Nox
 in h end
 fun ss phi = HOL_ss addsimps (simps phi)
in
 Ferrante_Rackoff_Data.funs  @{thm "ferrack_axiom"}
  {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss}
end
*}

end

use "Tools/Qelim/ferrante_rackoff.ML"

method_setup ferrack = {*
  Method.ctxt_args (Method.SIMPLE_METHOD' o FerranteRackoff.dlo_tac)
*} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders"

end 

lemma less_not_permute:

  ¬ (x < yy < x)

lemma gather_simps(1):

  (∃x. (∀yL. y < x) ∧ (∀yU. x < y) ∧ x < uP x) =
  (∃x. (∀yL. y < x) ∧ (∀y∈insert u U. x < y) ∧ P x)

and gather_simps(2-3):

  (∃x. (∀yL. y < x) ∧ (∀yU. x < y) ∧ l < xP x) =
  (∃x. (∀y∈insert l L. y < x) ∧ (∀yU. x < y) ∧ P x)
  (∃x. (∀yL. y < x) ∧ (∀yU. x < y) ∧ x < u) =
  (∃x. (∀yL. y < x) ∧ (∀y∈insert u U. x < y))

and gather_simps(4):

  (∃x. (∀yL. y < x) ∧ (∀yU. x < y) ∧ l < x) =
  (∃x. (∀y∈insert l L. y < x) ∧ (∀yU. x < y))

lemma gather_start:

  x. P x == ∃x. (∀y∈{}. y < x) ∧ (∀y∈{}. x < y) ∧ P x

lemma minf_lt:

  z. ∀x<z. (x < t) = True

lemma minf_gt:

  z. ∀x<z. (t < x) = False

lemma minf_le:

  z. ∀x<z. (x  t) = True

lemma minf_ge:

  z. ∀x<z. (t  x) = False

lemma minf_eq:

  z. ∀x<z. (x = t) = False

lemma minf_neq:

  z. ∀x<z. (x  t) = True

lemma minf_P:

  z. ∀x<z. P = P

lemma pinf_gt:

  z. ∀x>z. (t < x) = True

lemma pinf_lt:

  z. ∀x>z. (x < t) = False

lemma pinf_ge:

  z. ∀x>z. (t  x) = True

lemma pinf_le:

  z. ∀x>z. (x  t) = False

lemma pinf_eq:

  z. ∀x>z. (x = t) = False

lemma pinf_neq:

  z. ∀x>z. (x  t) = True

lemma pinf_P:

  z. ∀x>z. P = P

lemma nmi_lt:

  tU ==> ∀x. ¬ True ∧ x < t --> (∃uU. u  x)

lemma nmi_gt:

  tU ==> ∀x. ¬ False ∧ t < x --> (∃uU. u  x)

lemma nmi_le:

  tU ==> ∀x. ¬ True ∧ x  t --> (∃uU. u  x)

lemma nmi_ge:

  tU ==> ∀x. ¬ False ∧ t  x --> (∃uU. u  x)

lemma nmi_eq:

  tU ==> ∀x. ¬ False ∧ x = t --> (∃uU. u  x)

lemma nmi_neq:

  tU ==> ∀x. ¬ True ∧ x  t --> (∃uU. u  x)

lemma nmi_P:

  x. ¬ PP --> (∃uU. u  x)

lemma nmi_conj:

  [| ∀x. ¬ P1'P1.0 x --> (∃uU. u  x); ∀x. ¬ P2'P2.0 x --> (∃uU. u  x) |]
  ==> ∀x. ¬ (P1'P2') ∧ P1.0 xP2.0 x --> (∃uU. u  x)

lemma nmi_disj:

  [| ∀x. ¬ P1'P1.0 x --> (∃uU. u  x); ∀x. ¬ P2'P2.0 x --> (∃uU. u  x) |]
  ==> ∀x. ¬ (P1'P2') ∧ (P1.0 xP2.0 x) --> (∃uU. u  x)

lemma npi_lt:

  tU ==> ∀x. ¬ False ∧ x < t --> (∃uU. x  u)

lemma npi_gt:

  tU ==> ∀x. ¬ True ∧ t < x --> (∃uU. x  u)

lemma npi_le:

  tU ==> ∀x. ¬ False ∧ x  t --> (∃uU. x  u)

lemma npi_ge:

  tU ==> ∀x. ¬ True ∧ t  x --> (∃uU. x  u)

lemma npi_eq:

  tU ==> ∀x. ¬ False ∧ x = t --> (∃uU. x  u)

lemma npi_neq:

  tU ==> ∀x. ¬ True ∧ x  t --> (∃uU. x  u)

lemma npi_P:

  x. ¬ PP --> (∃uU. x  u)

lemma npi_conj:

  [| ∀x. ¬ P1'P1.0 x --> (∃uU. x  u); ∀x. ¬ P2'P2.0 x --> (∃uU. x  u) |]
  ==> ∀x. ¬ (P1'P2') ∧ P1.0 xP2.0 x --> (∃uU. x  u)

lemma npi_disj:

  [| ∀x. ¬ P1'P1.0 x --> (∃uU. x  u); ∀x. ¬ P2'P2.0 x --> (∃uU. x  u) |]
  ==> ∀x. ¬ (P1'P2') ∧ (P1.0 xP2.0 x) --> (∃uU. x  u)

lemma lin_dense_lt:

  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux < t -->
         (∀y. l < yy < u --> y < t)

lemma lin_dense_gt:

  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ut < x -->
         (∀y. l < yy < u --> t < y)

lemma lin_dense_le:

  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux  t -->
         (∀y. l < yy < u --> y  t)

lemma lin_dense_ge:

  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ut  x -->
         (∀y. l < yy < u --> t  y)

lemma lin_dense_eq:

  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux = t -->
         (∀y. l < yy < u --> y = t)

lemma lin_dense_neq:

  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux  t -->
         (∀y. l < yy < u --> y  t)

lemma lin_dense_P:

  x l u.
     (∀t. l < tt < u --> t  U) ∧ l < xx < uP -->
     (∀y. l < yy < u --> P)

lemma lin_dense_conj:

  [| ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP1.0 x -->
        (∀y. l < yy < u --> P1.0 y);
     ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP2.0 x -->
        (∀y. l < yy < u --> P2.0 y) |]
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < uP1.0 xP2.0 x -->
         (∀y. l < yy < u --> P1.0 yP2.0 y)

lemma lin_dense_disj:

  [| ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP1.0 x -->
        (∀y. l < yy < u --> P1.0 y);
     ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP2.0 x -->
        (∀y. l < yy < u --> P2.0 y) |]
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < u ∧ (P1.0 xP2.0 x) -->
         (∀y. l < yy < u --> P1.0 yP2.0 y)

lemma npmibnd:

  [| ∀x. ¬ MPP x --> (∃uU. u  x); ∀x. ¬ PPP x --> (∃uU. x  u) |]
  ==> ∀x. ¬ MP ∧ ¬ PPP x --> (∃uU. ∃u'U. u  xx  u')

lemma finite_set_intervals:

  [| P x; l  x; x  u; lS; uS; finite S; ∀xS. l  x; ∀xS. x  u |]
  ==> ∃aS. ∃bS. (∀y. a < yy < b --> y  S) ∧ a  xx  bP x

lemma finite_set_intervals2:

  [| P x; l  x; x  u; lS; uS; finite S; ∀xS. l  x; ∀xS. x  u |]
  ==> (∃sS. P s) ∨
      (∃aS. ∃bS. (∀y. a < yy < b --> y  S) ∧ a < xx < bP x)

The classical QE after Langford for dense linear orders

lemma dlo_qe_bnds:

  [| L  {}; U  {}; finite L; finite U |]
  ==> ∃x. (∀yL. y < x) ∧ (∀yU. x < y) == ∀lL. ∀uU. l < u

lemma dlo_qe_noub:

  [| L  {}; finite L |] ==> ∃x. (∀yL. y < x) ∧ (∀y∈{}. x < y) == True

lemma dlo_qe_nolb:

  [| U  {}; finite U |] ==> ∃x. (∀y∈{}. y < x) ∧ (∀yU. x < y) == True

lemma exists_neq:

  x. x  t
  x. t  x

lemma dlo_simps:

  x  x
  ¬ x < x
  x < y) = (y  x)
  x  y) = (y < x)
  x. x  t
  x. t  x
  (x  y) = (x < yx = y)
  (x  y) = (x < yy < x)
  x  yy  x
  ¬ (x < yy < x)

lemma axiom:

  dense_linear_order op  op <

lemma atoms(1):

  TERM op <

and atoms(2):

  TERM op 

and atoms(3):

  TERM op =

lemma dnf:

  (P ∧ (QR)) = (PQPR)
  ((QR) ∧ P) = (QPRP)

lemma weak_dnf_simps:

  (¬ ¬ P) = P
  ((¬ P) = (¬ Q)) = (P = Q)
  (P  Q) = (P = (¬ Q))
  (P ∨ ¬ P) = True
  PP) = True
  (x = x) = True
  (¬ True) = False
  (¬ False) = True
  P)  P
  P P)
  (True = P) = P
  (P = True) = P
  (False = P) = (¬ P)
  (P = False) = (¬ P)
  (True --> P) = P
  (False --> P) = True
  (P --> True) = True
  (P --> P) = True
  (P --> False) = (¬ P)
  (P --> ¬ P) = (¬ P)
  (P ∧ True) = P
  (True ∧ P) = P
  (P ∧ False) = False
  (False ∧ P) = False
  (PP) = P
  (PPQ) = (PQ)
  (P ∧ ¬ P) = False
  PP) = False
  (P ∨ True) = True
  (True ∨ P) = True
  (P ∨ False) = P
  (False ∨ P) = P
  (PP) = P
  (PPQ) = (PQ)
  (∀x. P) = P
  (∃x. P) = P
  x. x = t
  x. t = x
  (∃x. x = tP x) = P t
  (∃x. t = xP x) = P t
  (∀x. x = t --> P x) = P t
  (∀x. t = x --> P x) = P t
  (P ∧ (QR)) = (PQPR)
  ((QR) ∧ P) = (QPRP)

lemma nnf_simps:

  (¬ (PQ)) = (¬ P ∨ ¬ Q)
  (¬ (PQ)) = (¬ P ∧ ¬ Q)
  (P --> Q) = (¬ PQ)
  (P = Q) = (PQ ∨ ¬ P ∧ ¬ Q)
  (¬ ¬ P) = P

lemma ex_distrib:

  (∃x. P xQ x) = ((∃x. P x) ∨ (∃x. Q x))

lemma dnf_simps:

  (¬ ¬ P) = P
  ((¬ P) = (¬ Q)) = (P = Q)
  (P  Q) = (P = (¬ Q))
  (P ∨ ¬ P) = True
  PP) = True
  (x = x) = True
  (¬ True) = False
  (¬ False) = True
  P)  P
  P P)
  (True = P) = P
  (P = True) = P
  (False = P) = (¬ P)
  (P = False) = (¬ P)
  (True --> P) = P
  (False --> P) = True
  (P --> True) = True
  (P --> P) = True
  (P --> False) = (¬ P)
  (P --> ¬ P) = (¬ P)
  (P ∧ True) = P
  (True ∧ P) = P
  (P ∧ False) = False
  (False ∧ P) = False
  (PP) = P
  (PPQ) = (PQ)
  (P ∧ ¬ P) = False
  PP) = False
  (P ∨ True) = True
  (True ∨ P) = True
  (P ∨ False) = P
  (False ∨ P) = P
  (PP) = P
  (PPQ) = (PQ)
  (∀x. P) = P
  (∃x. P) = P
  x. x = t
  x. t = x
  (∃x. x = tP x) = P t
  (∃x. t = xP x) = P t
  (∀x. x = t --> P x) = P t
  (∀x. t = x --> P x) = P t
  (P ∧ (QR)) = (PQPR)
  ((QR) ∧ P) = (QPRP)
  (¬ (PQ)) = (¬ P ∨ ¬ Q)
  (¬ (PQ)) = (¬ P ∧ ¬ Q)
  (P --> Q) = (¬ PQ)
  (P = Q) = (PQ ∨ ¬ P ∧ ¬ Q)
  (¬ ¬ P) = P
  (∃x. P xQ x) = ((∃x. P x) ∨ (∃x. Q x))

Contructive dense linear orders yield QE for linear arithmetic over ordered Fields -- see @{text "Arith_Tools.thy"}

lemma ge_ex:

  y. x  y

lemma pinf_conj:

  [| ∃z1. ∀x>z1. P1.0 x = P1'; ∃z2. ∀x>z2. P2.0 x = P2' |]
  ==> ∃z. ∀x>z. (P1.0 xP2.0 x) = (P1'P2')

lemma pinf_disj:

  [| ∃z1. ∀x>z1. P1.0 x = P1'; ∃z2. ∀x>z2. P2.0 x = P2' |]
  ==> ∃z. ∀x>z. (P1.0 xP2.0 x) = (P1'P2')

lemma pinf_ex:

  [| ∃z. ∀x>z. P x = P1.0; P1.0 |] ==> ∃x. P x

lemma le_ex:

  y. y  x

lemma minf_conj:

  [| ∃z1. ∀x<z1. P1.0 x = P1'; ∃z2. ∀x<z2. P2.0 x = P2' |]
  ==> ∃z. ∀x<z. (P1.0 xP2.0 x) = (P1'P2')

lemma minf_disj:

  [| ∃z1. ∀x<z1. P1.0 x = P1'; ∃z2. ∀x<z2. P2.0 x = P2' |]
  ==> ∃z. ∀x<z. (P1.0 xP2.0 x) = (P1'P2')

lemma minf_ex:

  [| ∃z. ∀x<z. P x = P1.0; P1.0 |] ==> ∃x. P x

lemma rinf_U:

  [| finite U;
     ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP x -->
        (∀y. l < yy < u --> P y);
     ∀x. ¬ MP ∧ ¬ PPP x --> (∃uU. ∃u'U. u  xx  u'); ¬ MP; ¬ PP;
     ∃x. P x |]
  ==> ∃uU. ∃u'U. P (between u u')

theorem fr_eq:

  [| finite U;
     ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP x -->
        (∀y. l < yy < u --> P y);
     ∀x. ¬ MPP x --> (∃uU. u  x); ∀x. ¬ PPP x --> (∃uU. x  u);
     ∃z. ∀x<z. P x = MP; ∃z. ∀x>z. P x = PP |]
  ==> ∃x. P x == MPPP ∨ (∃uU. ∃u'U. P (between u u'))

lemma minf_thms:

  [| ∃z1. ∀x<z1. P1.0 x = P1'; ∃z2. ∀x<z2. P2.0 x = P2' |]
  ==> ∃z. ∀x<z. (P1.0 xP2.0 x) = (P1'P2')
  [| ∃z1. ∀x<z1. P1.0 x = P1'; ∃z2. ∀x<z2. P2.0 x = P2' |]
  ==> ∃z. ∀x<z. (P1.0 xP2.0 x) = (P1'P2')
  z. ∀x<z. (x = t) = False
  z. ∀x<z. (x  t) = True
  z. ∀x<z. (x < t) = True
  z. ∀x<z. (x  t) = True
  z. ∀x<z. (t < x) = False
  z. ∀x<z. (t  x) = False
  z. ∀x<z. P = P

lemma pinf_thms:

  [| ∃z1. ∀x>z1. P1.0 x = P1'; ∃z2. ∀x>z2. P2.0 x = P2' |]
  ==> ∃z. ∀x>z. (P1.0 xP2.0 x) = (P1'P2')
  [| ∃z1. ∀x>z1. P1.0 x = P1'; ∃z2. ∀x>z2. P2.0 x = P2' |]
  ==> ∃z. ∀x>z. (P1.0 xP2.0 x) = (P1'P2')
  z. ∀x>z. (x = t) = False
  z. ∀x>z. (x  t) = True
  z. ∀x>z. (x < t) = False
  z. ∀x>z. (x  t) = False
  z. ∀x>z. (t < x) = True
  z. ∀x>z. (t  x) = True
  z. ∀x>z. P = P

lemma nmi_thms:

  [| ∀x. ¬ P1'P1.0 x --> (∃uU. u  x); ∀x. ¬ P2'P2.0 x --> (∃uU. u  x) |]
  ==> ∀x. ¬ (P1'P2') ∧ P1.0 xP2.0 x --> (∃uU. u  x)
  [| ∀x. ¬ P1'P1.0 x --> (∃uU. u  x); ∀x. ¬ P2'P2.0 x --> (∃uU. u  x) |]
  ==> ∀x. ¬ (P1'P2') ∧ (P1.0 xP2.0 x) --> (∃uU. u  x)
  tU ==> ∀x. ¬ False ∧ x = t --> (∃uU. u  x)
  tU ==> ∀x. ¬ True ∧ x  t --> (∃uU. u  x)
  tU ==> ∀x. ¬ True ∧ x < t --> (∃uU. u  x)
  tU ==> ∀x. ¬ True ∧ x  t --> (∃uU. u  x)
  tU ==> ∀x. ¬ False ∧ t < x --> (∃uU. u  x)
  tU ==> ∀x. ¬ False ∧ t  x --> (∃uU. u  x)
  x. ¬ PP --> (∃uU. u  x)

lemma npi_thms:

  [| ∀x. ¬ P1'P1.0 x --> (∃uU. x  u); ∀x. ¬ P2'P2.0 x --> (∃uU. x  u) |]
  ==> ∀x. ¬ (P1'P2') ∧ P1.0 xP2.0 x --> (∃uU. x  u)
  [| ∀x. ¬ P1'P1.0 x --> (∃uU. x  u); ∀x. ¬ P2'P2.0 x --> (∃uU. x  u) |]
  ==> ∀x. ¬ (P1'P2') ∧ (P1.0 xP2.0 x) --> (∃uU. x  u)
  tU ==> ∀x. ¬ False ∧ x = t --> (∃uU. x  u)
  tU ==> ∀x. ¬ True ∧ x  t --> (∃uU. x  u)
  tU ==> ∀x. ¬ False ∧ x < t --> (∃uU. x  u)
  tU ==> ∀x. ¬ False ∧ x  t --> (∃uU. x  u)
  tU ==> ∀x. ¬ True ∧ t < x --> (∃uU. x  u)
  tU ==> ∀x. ¬ True ∧ t  x --> (∃uU. x  u)
  x. ¬ PP --> (∃uU. x  u)

lemma lin_dense_thms:

  [| ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP1.0 x -->
        (∀y. l < yy < u --> P1.0 y);
     ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP2.0 x -->
        (∀y. l < yy < u --> P2.0 y) |]
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < uP1.0 xP2.0 x -->
         (∀y. l < yy < u --> P1.0 yP2.0 y)
  [| ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP1.0 x -->
        (∀y. l < yy < u --> P1.0 y);
     ∀x l u.
        (∀t. l < tt < u --> t  U) ∧ l < xx < uP2.0 x -->
        (∀y. l < yy < u --> P2.0 y) |]
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < u ∧ (P1.0 xP2.0 x) -->
         (∀y. l < yy < u --> P1.0 yP2.0 y)
  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux = t -->
         (∀y. l < yy < u --> y = t)
  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux  t -->
         (∀y. l < yy < u --> y  t)
  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux < t -->
         (∀y. l < yy < u --> y < t)
  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ux  t -->
         (∀y. l < yy < u --> y  t)
  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ut < x -->
         (∀y. l < yy < u --> t < y)
  tU
  ==> ∀x l u.
         (∀t. l < tt < u --> t  U) ∧ l < xx < ut  x -->
         (∀y. l < yy < u --> t  y)
  x l u.
     (∀t. l < tt < u --> t  U) ∧ l < xx < uP -->
     (∀y. l < yy < u --> P)

lemma ferrack_axiom:

  constr_dense_linear_order op  op < between

lemma atoms(1):

  TERM op <

and atoms(2):

  TERM op 

and atoms(3):

  TERM op =