src/HOL/Complex/ex/MIR.thy
author haftmann
Mon, 13 Aug 2007 21:22:37 +0200
changeset 24249 1f60b45c5f97
parent 23997 a23d0b4b1c1f
child 24336 fff40259f336
permissions -rw-r--r--
renamed keyword "to" to "module_name"
     1 (*  Title:      Complex/ex/MIR.thy
     2     Author:     Amine Chaieb
     3 *)
     4 
     5 header {* Quatifier elimination for R(0,1,+,floor,<) *}
     6 
     7 theory MIR
     8   imports Real GCD Pretty_Int
     9   uses ("mireif.ML") ("mirtac.ML")
    10   begin
    11 
    12 declare real_of_int_floor_cancel [simp del]
    13 
    14   (* All pairs from two lists *)
    15 
    16 lemma allpairs_set: "set (allpairs Pair xs ys) = {(x,y). x\<in> set xs \<and> y \<in> set ys}"
    17 by (induct xs) auto
    18 
    19 fun alluopairs:: "'a list \<Rightarrow> ('a \<times> 'a) list" where 
    20   "alluopairs [] = []"
    21 | "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)"
    22 
    23 lemma alluopairs_set1: "set (alluopairs xs) \<le> {(x,y). x\<in> set xs \<and> y\<in> set xs}"
    24 by (induct xs, auto)
    25 
    26 lemma alluopairs_set:
    27   "\<lbrakk>x\<in> set xs ; y \<in> set xs\<rbrakk> \<Longrightarrow> (x,y) \<in> set (alluopairs xs) \<or> (y,x) \<in> set (alluopairs xs) "
    28 by (induct xs, auto)
    29 
    30 lemma alluopairs_ex:
    31   assumes Pc: "\<forall> x y. P x y = P y x"
    32   shows "(\<exists> x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists> (x,y) \<in> set (alluopairs xs). P x y)"
    33 proof
    34   assume "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y"
    35   then obtain x y where x: "x \<in> set xs" and y:"y \<in> set xs" and P: "P x y"  by blast
    36   from alluopairs_set[OF x y] P Pc show"\<exists>(x, y)\<in>set (alluopairs xs). P x y" 
    37     by auto
    38 next
    39   assume "\<exists>(x, y)\<in>set (alluopairs xs). P x y"
    40   then obtain "x" and "y"  where xy:"(x,y) \<in> set (alluopairs xs)" and P: "P x y" by blast+
    41   from xy have "x \<in> set xs \<and> y\<in> set xs" using alluopairs_set1 by blast
    42   with P show "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y" by blast
    43 qed
    44 
    45   (* generate a list from i to j*)
    46 consts iupt :: "int \<times> int \<Rightarrow> int list"
    47 recdef iupt "measure (\<lambda> (i,j). nat (j-i +1))" 
    48   "iupt (i,j) = (if j <i then [] else (i# iupt(i+1, j)))"
    49 
    50 lemma iupt_set: "set (iupt(i,j)) = {i .. j}"
    51 proof(induct rule: iupt.induct)
    52   case (1 a b)
    53   show ?case
    54     using prems by (simp add: simp_from_to)
    55 qed
    56 
    57 lemma nth_pos2: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
    58 using Nat.gr0_conv_Suc
    59 by clarsimp
    60 
    61 
    62 lemma myl: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a \<le> b) = (0 \<le> b - a)" 
    63 proof(clarify)
    64   fix x y ::"'a"
    65   have "(x \<le> y) = (x - y \<le> 0)" by (simp only: le_iff_diff_le_0[where a="x" and b="y"])
    66   also have "\<dots> = (- (y - x) \<le> 0)" by simp
    67   also have "\<dots> = (0 \<le> y - x)" by (simp only: neg_le_0_iff_le[where a="y-x"])
    68   finally show "(x \<le> y) = (0 \<le> y - x)" .
    69 qed
    70 
    71 lemma myless: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a < b) = (0 < b - a)" 
    72 proof(clarify)
    73   fix x y ::"'a"
    74   have "(x < y) = (x - y < 0)" by (simp only: less_iff_diff_less_0[where a="x" and b="y"])
    75   also have "\<dots> = (- (y - x) < 0)" by simp
    76   also have "\<dots> = (0 < y - x)" by (simp only: neg_less_0_iff_less[where a="y-x"])
    77   finally show "(x < y) = (0 < y - x)" .
    78 qed
    79 
    80 lemma myeq: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a = b) = (0 = b - a)"
    81   by auto
    82 
    83   (* Maybe should be added to the library \<dots> *)
    84 lemma floor_int_eq: "(real n\<le> x \<and> x < real (n+1)) = (floor x = n)"
    85 proof( auto)
    86   assume lb: "real n \<le> x"
    87     and ub: "x < real n + 1"
    88   have "real (floor x) \<le> x" by simp 
    89   hence "real (floor x) < real (n + 1) " using ub by arith
    90   hence "floor x < n+1" by simp
    91   moreover from lb have "n \<le> floor x" using floor_mono2[where x="real n" and y="x"] 
    92     by simp ultimately show "floor x = n" by simp
    93 qed
    94 
    95 (* Periodicity of dvd *)
    96 lemma dvd_period:
    97   assumes advdd: "(a::int) dvd d"
    98   shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))"
    99   using advdd  
   100 proof-
   101   {fix x k
   102     from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"]  
   103     have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp}
   104   hence "\<forall>x.\<forall>k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))"  by simp
   105   then show ?thesis by simp
   106 qed
   107 
   108   (* The Divisibility relation between reals *)	
   109 definition
   110   rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
   111 where
   112   rdvd_def: "x rdvd y \<longleftrightarrow> (\<exists>k\<Colon>int. y = x * real k)"
   113 
   114 lemma int_rdvd_real: 
   115   shows "real (i::int) rdvd x = (i dvd (floor x) \<and> real (floor x) = x)" (is "?l = ?r")
   116 proof
   117   assume "?l" 
   118   hence th: "\<exists> k. x=real (i*k)" by (simp add: rdvd_def)
   119   hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult)
   120   with th have "\<exists> k. real (floor x) = real (i*k)" by simp
   121   hence "\<exists> k. floor x = i*k" by (simp only: real_of_int_inject)
   122   thus ?r  using th' by (simp add: dvd_def) 
   123 next
   124   assume "?r" hence "(i\<Colon>int) dvd \<lfloor>x\<Colon>real\<rfloor>" ..
   125   hence "\<exists> k. real (floor x) = real (i*k)" 
   126     by (simp only: real_of_int_inject) (simp add: dvd_def)
   127   thus ?l using prems by (simp add: rdvd_def)
   128 qed
   129 
   130 lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)"
   131 by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only :real_of_int_mult[symmetric])
   132 
   133 
   134 lemma rdvd_abs1: 
   135   "(abs (real d) rdvd t) = (real (d ::int) rdvd t)"
   136 proof
   137   assume d: "real d rdvd t"
   138   from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t" by auto
   139 
   140   from iffD2[OF zdvd_abs1] d2 have "(abs d) dvd (floor t)" by blast
   141   with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast 
   142   thus "abs (real d) rdvd t" by simp
   143 next
   144   assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp
   145   with int_rdvd_real[where i="abs d" and x="t"] have d2: "abs d dvd floor t" and ti: "real (floor t) =t" by auto
   146   from iffD1[OF zdvd_abs1] d2 have "d dvd floor t" by blast
   147   with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast
   148 qed
   149 
   150 lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)"
   151   apply (auto simp add: rdvd_def)
   152   apply (rule_tac x="-k" in exI, simp) 
   153   apply (rule_tac x="-k" in exI, simp)
   154 done
   155 
   156 lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)"
   157 by (auto simp add: rdvd_def)
   158 
   159 lemma rdvd_mult: 
   160   assumes knz: "k\<noteq>0"
   161   shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)"
   162 using knz by (simp add:rdvd_def)
   163 
   164 lemma rdvd_trans: assumes mn:"m rdvd n" and  nk:"n rdvd k" 
   165   shows "m rdvd k"
   166 proof-
   167   from rdvd_def mn obtain c where nmc:"n = m * real (c::int)" by auto
   168   from rdvd_def nk obtain c' where nkc:"k = n * real (c'::int)" by auto
   169   hence "k = m * real (c * c')" using nmc by simp
   170   thus ?thesis using rdvd_def by blast
   171 qed
   172 
   173   (*********************************************************************************)
   174   (****                            SHADOW SYNTAX AND SEMANTICS                  ****)
   175   (*********************************************************************************)
   176 
   177 datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num 
   178   | Mul int num | Floor num| CF int num num
   179 
   180   (* A size for num to make inductive proofs simpler*)
   181 fun num_size :: "num \<Rightarrow> nat" where
   182  "num_size (C c) = 1"
   183 | "num_size (Bound n) = 1"
   184 | "num_size (Neg a) = 1 + num_size a"
   185 | "num_size (Add a b) = 1 + num_size a + num_size b"
   186 | "num_size (Sub a b) = 3 + num_size a + num_size b"
   187 | "num_size (CN n c a) = 4 + num_size a "
   188 | "num_size (CF c a b) = 4 + num_size a + num_size b"
   189 | "num_size (Mul c a) = 1 + num_size a"
   190 | "num_size (Floor a) = 1 + num_size a"
   191 
   192   (* Semantics of numeral terms (num) *)
   193 fun Inum :: "real list \<Rightarrow> num \<Rightarrow> real" where
   194   "Inum bs (C c) = (real c)"
   195 | "Inum bs (Bound n) = bs!n"
   196 | "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)"
   197 | "Inum bs (Neg a) = -(Inum bs a)"
   198 | "Inum bs (Add a b) = Inum bs a + Inum bs b"
   199 | "Inum bs (Sub a b) = Inum bs a - Inum bs b"
   200 | "Inum bs (Mul c a) = (real c) * Inum bs a"
   201 | "Inum bs (Floor a) = real (floor (Inum bs a))"
   202 | "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b"
   203 definition "isint t bs \<equiv> real (floor (Inum bs t)) = Inum bs t"
   204 
   205 lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)"
   206 by (simp add: isint_def)
   207 
   208 lemma isint_Floor: "isint (Floor n) bs"
   209   by (simp add: isint_iff)
   210 
   211 lemma isint_Mul: "isint e bs \<Longrightarrow> isint (Mul c e) bs"
   212 proof-
   213   let ?e = "Inum bs e"
   214   let ?fe = "floor ?e"
   215   assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff)
   216   have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp
   217   also have "\<dots> = real (c* ?fe)" by (simp only: floor_real_of_int) 
   218   also have "\<dots> = real c * ?e" using efe by simp
   219   finally show ?thesis using isint_iff by simp
   220 qed
   221 
   222 lemma isint_neg: "isint e bs \<Longrightarrow> isint (Neg e) bs"
   223 proof-
   224   let ?I = "\<lambda> t. Inum bs t"
   225   assume ie: "isint e bs"
   226   hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)  
   227   have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th)
   228   also have "\<dots> = - real (floor (?I e))" by(simp add: floor_minus_real_of_int) 
   229   finally show "isint (Neg e) bs" by (simp add: isint_def th)
   230 qed
   231 
   232 lemma isint_sub: 
   233   assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs"
   234 proof-
   235   let ?I = "\<lambda> t. Inum bs t"
   236   from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)  
   237   have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th)
   238   also have "\<dots> = real (c- floor (?I e))" by(simp add: floor_minus_real_of_int) 
   239   finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th)
   240 qed
   241 
   242 lemma isint_add: assumes
   243   ai:"isint a bs" and bi: "isint b bs" shows "isint (Add a b) bs"
   244 proof-
   245   let ?a = "Inum bs a"
   246   let ?b = "Inum bs b"
   247   from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))" by simp
   248   also have "\<dots> = real (floor ?a) + real (floor ?b)" by simp
   249   also have "\<dots> = ?a + ?b" using ai bi isint_iff by simp
   250   finally show "isint (Add a b) bs" by (simp add: isint_iff)
   251 qed
   252 
   253 lemma isint_c: "isint (C j) bs"
   254   by (simp add: isint_iff)
   255 
   256 
   257     (* FORMULAE *)
   258 datatype fm  = 
   259   T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num|
   260   NOT fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
   261 
   262 
   263   (* A size for fm *)
   264 fun fmsize :: "fm \<Rightarrow> nat" where
   265  "fmsize (NOT p) = 1 + fmsize p"
   266 | "fmsize (And p q) = 1 + fmsize p + fmsize q"
   267 | "fmsize (Or p q) = 1 + fmsize p + fmsize q"
   268 | "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
   269 | "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
   270 | "fmsize (E p) = 1 + fmsize p"
   271 | "fmsize (A p) = 4+ fmsize p"
   272 | "fmsize (Dvd i t) = 2"
   273 | "fmsize (NDvd i t) = 2"
   274 | "fmsize p = 1"
   275   (* several lemmas about fmsize *)
   276 lemma fmsize_pos: "fmsize p > 0"	
   277 by (induct p rule: fmsize.induct) simp_all
   278 
   279   (* Semantics of formulae (fm) *)
   280 fun Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool" where
   281   "Ifm bs T = True"
   282 | "Ifm bs F = False"
   283 | "Ifm bs (Lt a) = (Inum bs a < 0)"
   284 | "Ifm bs (Gt a) = (Inum bs a > 0)"
   285 | "Ifm bs (Le a) = (Inum bs a \<le> 0)"
   286 | "Ifm bs (Ge a) = (Inum bs a \<ge> 0)"
   287 | "Ifm bs (Eq a) = (Inum bs a = 0)"
   288 | "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)"
   289 | "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)"
   290 | "Ifm bs (NDvd i b) = (\<not>(real i rdvd Inum bs b))"
   291 | "Ifm bs (NOT p) = (\<not> (Ifm bs p))"
   292 | "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
   293 | "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)"
   294 | "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))"
   295 | "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)"
   296 | "Ifm bs (E p) = (\<exists> x. Ifm (x#bs) p)"
   297 | "Ifm bs (A p) = (\<forall> x. Ifm (x#bs) p)"
   298 
   299 consts prep :: "fm \<Rightarrow> fm"
   300 recdef prep "measure fmsize"
   301   "prep (E T) = T"
   302   "prep (E F) = F"
   303   "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
   304   "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
   305   "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" 
   306   "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
   307   "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
   308   "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
   309   "prep (E p) = E (prep p)"
   310   "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
   311   "prep (A p) = prep (NOT (E (NOT p)))"
   312   "prep (NOT (NOT p)) = prep p"
   313   "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
   314   "prep (NOT (A p)) = prep (E (NOT p))"
   315   "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
   316   "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
   317   "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
   318   "prep (NOT p) = NOT (prep p)"
   319   "prep (Or p q) = Or (prep p) (prep q)"
   320   "prep (And p q) = And (prep p) (prep q)"
   321   "prep (Imp p q) = prep (Or (NOT p) q)"
   322   "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
   323   "prep p = p"
   324 (hints simp add: fmsize_pos)
   325 lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p"
   326 by (induct p rule: prep.induct, auto)
   327 
   328 
   329   (* Quantifier freeness *)
   330 consts qfree:: "fm \<Rightarrow> bool"
   331 recdef qfree "measure size"
   332   "qfree (E p) = False"
   333   "qfree (A p) = False"
   334   "qfree (NOT p) = qfree p" 
   335   "qfree (And p q) = (qfree p \<and> qfree q)" 
   336   "qfree (Or  p q) = (qfree p \<and> qfree q)" 
   337   "qfree (Imp p q) = (qfree p \<and> qfree q)" 
   338   "qfree (Iff p q) = (qfree p \<and> qfree q)"
   339   "qfree p = True"
   340 
   341   (* Boundedness and substitution *)
   342 consts 
   343   numbound0:: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *)
   344   bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
   345   numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *)
   346   subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *)
   347 primrec
   348   "numbound0 (C c) = True"
   349   "numbound0 (Bound n) = (n>0)"
   350   "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)"
   351   "numbound0 (Neg a) = numbound0 a"
   352   "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
   353   "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)" 
   354   "numbound0 (Mul i a) = numbound0 a"
   355   "numbound0 (Floor a) = numbound0 a"
   356   "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)" 
   357 lemma numbound0_I:
   358   assumes nb: "numbound0 a"
   359   shows "Inum (b#bs) a = Inum (b'#bs) a"
   360 using nb
   361 by (induct a rule: numbound0.induct) (auto simp add: nth_pos2)
   362 
   363 
   364 lemma numbound0_gen: 
   365   assumes nb: "numbound0 t" and ti: "isint t (x#bs)"
   366   shows "\<forall> y. isint t (y#bs)"
   367 using nb ti 
   368 proof(clarify)
   369   fix y
   370   from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def]
   371   show "isint t (y#bs)"
   372     by (simp add: isint_def)
   373 qed
   374 
   375 primrec
   376   "bound0 T = True"
   377   "bound0 F = True"
   378   "bound0 (Lt a) = numbound0 a"
   379   "bound0 (Le a) = numbound0 a"
   380   "bound0 (Gt a) = numbound0 a"
   381   "bound0 (Ge a) = numbound0 a"
   382   "bound0 (Eq a) = numbound0 a"
   383   "bound0 (NEq a) = numbound0 a"
   384   "bound0 (Dvd i a) = numbound0 a"
   385   "bound0 (NDvd i a) = numbound0 a"
   386   "bound0 (NOT p) = bound0 p"
   387   "bound0 (And p q) = (bound0 p \<and> bound0 q)"
   388   "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
   389   "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
   390   "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
   391   "bound0 (E p) = False"
   392   "bound0 (A p) = False"
   393 
   394 lemma bound0_I:
   395   assumes bp: "bound0 p"
   396   shows "Ifm (b#bs) p = Ifm (b'#bs) p"
   397 using bp numbound0_I[where b="b" and bs="bs" and b'="b'"]
   398 by (induct p rule: bound0.induct) (auto simp add: nth_pos2)
   399 
   400 primrec
   401   "numsubst0 t (C c) = (C c)"
   402   "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
   403   "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))"
   404   "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)"
   405   "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
   406   "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
   407   "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" 
   408   "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
   409   "numsubst0 t (Floor a) = Floor (numsubst0 t a)"
   410 
   411 lemma numsubst0_I:
   412   shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t"
   413   by (induct t) (simp_all add: nth_pos2)
   414 
   415 lemma numsubst0_I':
   416   assumes nb: "numbound0 a"
   417   shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t"
   418   by (induct t) (simp_all add: nth_pos2 numbound0_I[OF nb, where b="b" and b'="b'"])
   419 
   420 
   421 primrec
   422   "subst0 t T = T"
   423   "subst0 t F = F"
   424   "subst0 t (Lt a) = Lt (numsubst0 t a)"
   425   "subst0 t (Le a) = Le (numsubst0 t a)"
   426   "subst0 t (Gt a) = Gt (numsubst0 t a)"
   427   "subst0 t (Ge a) = Ge (numsubst0 t a)"
   428   "subst0 t (Eq a) = Eq (numsubst0 t a)"
   429   "subst0 t (NEq a) = NEq (numsubst0 t a)"
   430   "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
   431   "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
   432   "subst0 t (NOT p) = NOT (subst0 t p)"
   433   "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
   434   "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
   435   "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
   436   "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
   437 
   438 lemma subst0_I: assumes qfp: "qfree p"
   439   shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p"
   440   using qfp numsubst0_I[where b="b" and bs="bs" and a="a"]
   441   by (induct p) (simp_all add: nth_pos2 )
   442 
   443 consts 
   444   decrnum:: "num \<Rightarrow> num" 
   445   decr :: "fm \<Rightarrow> fm"
   446 
   447 recdef decrnum "measure size"
   448   "decrnum (Bound n) = Bound (n - 1)"
   449   "decrnum (Neg a) = Neg (decrnum a)"
   450   "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
   451   "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
   452   "decrnum (Mul c a) = Mul c (decrnum a)"
   453   "decrnum (Floor a) = Floor (decrnum a)"
   454   "decrnum (CN n c a) = CN (n - 1) c (decrnum a)"
   455   "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)"
   456   "decrnum a = a"
   457 
   458 recdef decr "measure size"
   459   "decr (Lt a) = Lt (decrnum a)"
   460   "decr (Le a) = Le (decrnum a)"
   461   "decr (Gt a) = Gt (decrnum a)"
   462   "decr (Ge a) = Ge (decrnum a)"
   463   "decr (Eq a) = Eq (decrnum a)"
   464   "decr (NEq a) = NEq (decrnum a)"
   465   "decr (Dvd i a) = Dvd i (decrnum a)"
   466   "decr (NDvd i a) = NDvd i (decrnum a)"
   467   "decr (NOT p) = NOT (decr p)" 
   468   "decr (And p q) = And (decr p) (decr q)"
   469   "decr (Or p q) = Or (decr p) (decr q)"
   470   "decr (Imp p q) = Imp (decr p) (decr q)"
   471   "decr (Iff p q) = Iff (decr p) (decr q)"
   472   "decr p = p"
   473 
   474 lemma decrnum: assumes nb: "numbound0 t"
   475   shows "Inum (x#bs) t = Inum bs (decrnum t)"
   476   using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2)
   477 
   478 lemma decr: assumes nb: "bound0 p"
   479   shows "Ifm (x#bs) p = Ifm bs (decr p)"
   480   using nb 
   481   by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum)
   482 
   483 lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
   484 by (induct p, simp_all)
   485 
   486 consts 
   487   isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
   488 recdef isatom "measure size"
   489   "isatom T = True"
   490   "isatom F = True"
   491   "isatom (Lt a) = True"
   492   "isatom (Le a) = True"
   493   "isatom (Gt a) = True"
   494   "isatom (Ge a) = True"
   495   "isatom (Eq a) = True"
   496   "isatom (NEq a) = True"
   497   "isatom (Dvd i b) = True"
   498   "isatom (NDvd i b) = True"
   499   "isatom p = False"
   500 
   501 lemma numsubst0_numbound0: assumes nb: "numbound0 t"
   502   shows "numbound0 (numsubst0 t a)"
   503 using nb by (induct a rule: numsubst0.induct, auto)
   504 
   505 lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t"
   506   shows "bound0 (subst0 t p)"
   507 using qf numsubst0_numbound0[OF nb] by (induct p  rule: subst0.induct, auto)
   508 
   509 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
   510 by (induct p, simp_all)
   511 
   512 
   513 constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
   514   "djf f p q \<equiv> (if q=T then T else if q=F then f p else 
   515   (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or fp q))"
   516 constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
   517   "evaldjf f ps \<equiv> foldr (djf f) ps F"
   518 
   519 lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)"
   520 by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) 
   521 (cases "f p", simp_all add: Let_def djf_def) 
   522 
   523 lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm bs (f p))"
   524   by(induct ps, simp_all add: evaldjf_def djf_Or)
   525 
   526 lemma evaldjf_bound0: 
   527   assumes nb: "\<forall> x\<in> set xs. bound0 (f x)"
   528   shows "bound0 (evaldjf f xs)"
   529   using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) 
   530 
   531 lemma evaldjf_qf: 
   532   assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
   533   shows "qfree (evaldjf f xs)"
   534   using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) 
   535 
   536 consts 
   537   disjuncts :: "fm \<Rightarrow> fm list" 
   538   conjuncts :: "fm \<Rightarrow> fm list"
   539 recdef disjuncts "measure size"
   540   "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
   541   "disjuncts F = []"
   542   "disjuncts p = [p]"
   543 
   544 recdef conjuncts "measure size"
   545   "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
   546   "conjuncts T = []"
   547   "conjuncts p = [p]"
   548 lemma disjuncts: "(\<exists> q\<in> set (disjuncts p). Ifm bs q) = Ifm bs p"
   549 by(induct p rule: disjuncts.induct, auto)
   550 lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm bs q) = Ifm bs p"
   551 by(induct p rule: conjuncts.induct, auto)
   552 
   553 lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). bound0 q"
   554 proof-
   555   assume nb: "bound0 p"
   556   hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto)
   557   thus ?thesis by (simp only: list_all_iff)
   558 qed
   559 lemma conjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). bound0 q"
   560 proof-
   561   assume nb: "bound0 p"
   562   hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto)
   563   thus ?thesis by (simp only: list_all_iff)
   564 qed
   565 
   566 lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q"
   567 proof-
   568   assume qf: "qfree p"
   569   hence "list_all qfree (disjuncts p)"
   570     by (induct p rule: disjuncts.induct, auto)
   571   thus ?thesis by (simp only: list_all_iff)
   572 qed
   573 lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q"
   574 proof-
   575   assume qf: "qfree p"
   576   hence "list_all qfree (conjuncts p)"
   577     by (induct p rule: conjuncts.induct, auto)
   578   thus ?thesis by (simp only: list_all_iff)
   579 qed
   580 
   581 constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
   582   "DJ f p \<equiv> evaldjf f (disjuncts p)"
   583 
   584 lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)"
   585   and fF: "f F = F"
   586   shows "Ifm bs (DJ f p) = Ifm bs (f p)"
   587 proof-
   588   have "Ifm bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm bs (f q))"
   589     by (simp add: DJ_def evaldjf_ex) 
   590   also have "\<dots> = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto)
   591   finally show ?thesis .
   592 qed
   593 
   594 lemma DJ_qf: assumes 
   595   fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)"
   596   shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "
   597 proof(clarify)
   598   fix  p assume qf: "qfree p"
   599   have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def)
   600   from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" .
   601   with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast
   602   
   603   from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp
   604 qed
   605 
   606 lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
   607   shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm bs ((DJ qe p)) = Ifm bs (E p))"
   608 proof(clarify)
   609   fix p::fm and bs
   610   assume qf: "qfree p"
   611   from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast
   612   from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto
   613   have "Ifm bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (qe q))"
   614     by (simp add: DJ_def evaldjf_ex)
   615   also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto
   616   also have "\<dots> = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto)
   617   finally show "qfree (DJ qe p) \<and> Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast
   618 qed
   619   (* Simplification *)
   620 
   621   (* Algebraic simplifications for nums *)
   622 consts bnds:: "num \<Rightarrow> nat list"
   623   lex_ns:: "nat list \<times> nat list \<Rightarrow> bool"
   624 recdef bnds "measure size"
   625   "bnds (Bound n) = [n]"
   626   "bnds (CN n c a) = n#(bnds a)"
   627   "bnds (Neg a) = bnds a"
   628   "bnds (Add a b) = (bnds a)@(bnds b)"
   629   "bnds (Sub a b) = (bnds a)@(bnds b)"
   630   "bnds (Mul i a) = bnds a"
   631   "bnds (Floor a) = bnds a"
   632   "bnds (CF c a b) = (bnds a)@(bnds b)"
   633   "bnds a = []"
   634 recdef lex_ns "measure (\<lambda> (xs,ys). length xs + length ys)"
   635   "lex_ns ([], ms) = True"
   636   "lex_ns (ns, []) = False"
   637   "lex_ns (n#ns, m#ms) = (n<m \<or> ((n = m) \<and> lex_ns (ns,ms))) "
   638 constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
   639   "lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)"
   640 
   641 consts 
   642   numgcdh:: "num \<Rightarrow> int \<Rightarrow> int"
   643   reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num"
   644   dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
   645 consts maxcoeff:: "num \<Rightarrow> int"
   646 recdef maxcoeff "measure size"
   647   "maxcoeff (C i) = abs i"
   648   "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)"
   649   "maxcoeff (CF c t s) = max (abs c) (maxcoeff s)"
   650   "maxcoeff t = 1"
   651 
   652 lemma maxcoeff_pos: "maxcoeff t \<ge> 0"
   653   apply (induct t rule: maxcoeff.induct, auto) 
   654   done
   655 
   656 recdef numgcdh "measure size"
   657   "numgcdh (C i) = (\<lambda>g. igcd i g)"
   658   "numgcdh (CN n c t) = (\<lambda>g. igcd c (numgcdh t g))"
   659   "numgcdh (CF c s t) = (\<lambda>g. igcd c (numgcdh t g))"
   660   "numgcdh t = (\<lambda>g. 1)"
   661 
   662 definition
   663   numgcd :: "num \<Rightarrow> int"
   664 where
   665   numgcd_def: "numgcd t = numgcdh t (maxcoeff t)"
   666 
   667 recdef reducecoeffh "measure size"
   668   "reducecoeffh (C i) = (\<lambda> g. C (i div g))"
   669   "reducecoeffh (CN n c t) = (\<lambda> g. CN n (c div g) (reducecoeffh t g))"
   670   "reducecoeffh (CF c s t) = (\<lambda> g. CF (c div g)  s (reducecoeffh t g))"
   671   "reducecoeffh t = (\<lambda>g. t)"
   672 
   673 definition
   674   reducecoeff :: "num \<Rightarrow> num"
   675 where
   676   reducecoeff_def: "reducecoeff t =
   677   (let g = numgcd t in 
   678   if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)"
   679 
   680 recdef dvdnumcoeff "measure size"
   681   "dvdnumcoeff (C i) = (\<lambda> g. g dvd i)"
   682   "dvdnumcoeff (CN n c t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
   683   "dvdnumcoeff (CF c s t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
   684   "dvdnumcoeff t = (\<lambda>g. False)"
   685 
   686 lemma dvdnumcoeff_trans: 
   687   assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'"
   688   shows "dvdnumcoeff t g"
   689   using dgt' gdg 
   690   by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg])
   691 
   692 declare zdvd_trans [trans add]
   693 
   694 lemma natabs0: "(nat (abs x) = 0) = (x = 0)"
   695 by arith
   696 
   697 lemma numgcd0:
   698   assumes g0: "numgcd t = 0"
   699   shows "Inum bs t = 0"
   700 proof-
   701   have "\<And>x. numgcdh t x= 0 \<Longrightarrow> Inum bs t = 0"
   702     by (induct t rule: numgcdh.induct, auto simp add: igcd_def gcd_zero natabs0 max_def maxcoeff_pos)
   703   thus ?thesis using g0[simplified numgcd_def] by blast
   704 qed
   705 
   706 lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
   707   using gp
   708   by (induct t rule: numgcdh.induct, auto simp add: igcd_def)
   709 
   710 lemma numgcd_pos: "numgcd t \<ge>0"
   711   by (simp add: numgcd_def numgcdh_pos maxcoeff_pos)
   712 
   713 lemma reducecoeffh:
   714   assumes gt: "dvdnumcoeff t g" and gp: "g > 0" 
   715   shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t"
   716   using gt
   717 proof(induct t rule: reducecoeffh.induct) 
   718   case (1 i) hence gd: "g dvd i" by simp
   719   from gp have gnz: "g \<noteq> 0" by simp
   720   from prems show ?case by (simp add: real_of_int_div[OF gnz gd])
   721 next
   722   case (2 n c t)  hence gd: "g dvd c" by simp
   723   from gp have gnz: "g \<noteq> 0" by simp
   724   from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps)
   725 next
   726   case (3 c s t)  hence gd: "g dvd c" by simp
   727   from gp have gnz: "g \<noteq> 0" by simp
   728   from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) 
   729 qed (auto simp add: numgcd_def gp)
   730 consts ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
   731 recdef ismaxcoeff "measure size"
   732   "ismaxcoeff (C i) = (\<lambda> x. abs i \<le> x)"
   733   "ismaxcoeff (CN n c t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
   734   "ismaxcoeff (CF c s t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
   735   "ismaxcoeff t = (\<lambda>x. True)"
   736 
   737 lemma ismaxcoeff_mono: "ismaxcoeff t c \<Longrightarrow> c \<le> c' \<Longrightarrow> ismaxcoeff t c'"
   738 by (induct t rule: ismaxcoeff.induct, auto)
   739 
   740 lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)"
   741 proof (induct t rule: maxcoeff.induct)
   742   case (2 n c t)
   743   hence H:"ismaxcoeff t (maxcoeff t)" .
   744   have thh: "maxcoeff t \<le> max (abs c) (maxcoeff t)" by (simp add: le_maxI2)
   745   from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1)
   746 next
   747   case (3 c t s) 
   748   hence H1:"ismaxcoeff s (maxcoeff s)" by auto
   749   have thh1: "maxcoeff s \<le> max \<bar>c\<bar> (maxcoeff s)" by (simp add: max_def)
   750   from ismaxcoeff_mono[OF H1 thh1] show ?case by (simp add: le_maxI1)
   751 qed simp_all
   752 
   753 lemma igcd_gt1: "igcd i j > 1 \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))"
   754   apply (unfold igcd_def)
   755   apply (cases "i = 0", simp_all)
   756   apply (cases "j = 0", simp_all)
   757   apply (cases "abs i = 1", simp_all)
   758   apply (cases "abs j = 1", simp_all)
   759   apply auto
   760   done
   761 lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow>  m =0"
   762   by (induct t rule: numgcdh.induct, auto simp add:igcd0)
   763 
   764 lemma dvdnumcoeff_aux:
   765   assumes "ismaxcoeff t m" and mp:"m \<ge> 0" and "numgcdh t m > 1"
   766   shows "dvdnumcoeff t (numgcdh t m)"
   767 using prems
   768 proof(induct t rule: numgcdh.induct)
   769   case (2 n c t) 
   770   let ?g = "numgcdh t m"
   771   from prems have th:"igcd c ?g > 1" by simp
   772   from igcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
   773   have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
   774   moreover {assume "abs c > 1" and gp: "?g > 1" with prems
   775     have th: "dvdnumcoeff t ?g" by simp
   776     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   777     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)}
   778   moreover {assume "abs c = 0 \<and> ?g > 1"
   779     with prems have th: "dvdnumcoeff t ?g" by simp
   780     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   781     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)
   782     hence ?case by simp }
   783   moreover {assume "abs c > 1" and g0:"?g = 0" 
   784     from numgcdh0[OF g0] have "m=0". with prems   have ?case by simp }
   785   ultimately show ?case by blast
   786 next
   787   case (3 c s t) 
   788   let ?g = "numgcdh t m"
   789   from prems have th:"igcd c ?g > 1" by simp
   790   from igcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
   791   have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
   792   moreover {assume "abs c > 1" and gp: "?g > 1" with prems
   793     have th: "dvdnumcoeff t ?g" by simp
   794     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   795     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)}
   796   moreover {assume "abs c = 0 \<and> ?g > 1"
   797     with prems have th: "dvdnumcoeff t ?g" by simp
   798     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   799     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)
   800     hence ?case by simp }
   801   moreover {assume "abs c > 1" and g0:"?g = 0" 
   802     from numgcdh0[OF g0] have "m=0". with prems   have ?case by simp }
   803   ultimately show ?case by blast
   804 qed(auto simp add: igcd_dvd1)
   805 
   806 lemma dvdnumcoeff_aux2:
   807   assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
   808   using prems 
   809 proof (simp add: numgcd_def)
   810   let ?mc = "maxcoeff t"
   811   let ?g = "numgcdh t ?mc"
   812   have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff)
   813   have th2: "?mc \<ge> 0" by (rule maxcoeff_pos)
   814   assume H: "numgcdh t ?mc > 1"
   815   from dvdnumcoeff_aux[OF th1 th2 H]  show "dvdnumcoeff t ?g" .
   816 qed
   817 
   818 lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t"
   819 proof-
   820   let ?g = "numgcd t"
   821   have "?g \<ge> 0"  by (simp add: numgcd_pos)
   822   hence	"?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
   823   moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} 
   824   moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} 
   825   moreover { assume g1:"?g > 1"
   826     from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+
   827     from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis 
   828       by (simp add: reducecoeff_def Let_def)} 
   829   ultimately show ?thesis by blast
   830 qed
   831 
   832 lemma reducecoeffh_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeffh t g)"
   833 by (induct t rule: reducecoeffh.induct, auto)
   834 
   835 lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)"
   836 using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
   837 
   838 consts
   839   simpnum:: "num \<Rightarrow> num"
   840   numadd:: "num \<times> num \<Rightarrow> num"
   841   nummul:: "num \<Rightarrow> int \<Rightarrow> num"
   842 
   843 recdef numadd "measure (\<lambda> (t,s). size t + size s)"
   844   "numadd (CN n1 c1 r1,CN n2 c2 r2) =
   845   (if n1=n2 then 
   846   (let c = c1 + c2
   847   in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2))))
   848   else if n1 \<le> n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2))
   849   else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))"
   850   "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))"  
   851   "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" 
   852   "numadd (CF c1 t1 r1,CF c2 t2 r2) = 
   853    (if t1 = t2 then 
   854     (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s))
   855    else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2))
   856    else CF c2 t2 (numadd(CF c1 t1 r1,r2)))"
   857   "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))"
   858   "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))"
   859   "numadd (C b1, C b2) = C (b1+b2)"
   860   "numadd (a,b) = Add a b"
   861 
   862 lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
   863 apply (induct t s rule: numadd.induct, simp_all add: Let_def)
   864  apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
   865   apply (case_tac "n1 = n2", simp_all add: ring_simps)
   866   apply (simp only: left_distrib[symmetric])
   867  apply simp
   868 apply (case_tac "lex_bnd t1 t2", simp_all)
   869  apply (case_tac "c1+c2 = 0")
   870   by (case_tac "t1 = t2", simp_all add: ring_simps left_distrib[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add left_distrib)
   871 
   872 lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
   873 by (induct t s rule: numadd.induct, auto simp add: Let_def)
   874 
   875 recdef nummul "measure size"
   876   "nummul (C j) = (\<lambda> i. C (i*j))"
   877   "nummul (CN n c t) = (\<lambda> i. CN n (c*i) (nummul t i))"
   878   "nummul (CF c t s) = (\<lambda> i. CF (c*i) t (nummul s i))"
   879   "nummul (Mul c t) = (\<lambda> i. nummul t (i*c))"
   880   "nummul t = (\<lambda> i. Mul i t)"
   881 
   882 lemma nummul[simp]: "\<And> i. Inum bs (nummul t i) = Inum bs (Mul i t)"
   883 by (induct t rule: nummul.induct, auto simp add: ring_simps)
   884 
   885 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
   886 by (induct t rule: nummul.induct, auto)
   887 
   888 constdefs numneg :: "num \<Rightarrow> num"
   889   "numneg t \<equiv> nummul t (- 1)"
   890 
   891 constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
   892   "numsub s t \<equiv> (if s = t then C 0 else numadd (s,numneg t))"
   893 
   894 lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
   895 using numneg_def nummul by simp
   896 
   897 lemma numneg_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numneg t)"
   898 using numneg_def by simp
   899 
   900 lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)"
   901 using numsub_def by simp
   902 
   903 lemma numsub_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numsub t s)"
   904 using numsub_def by simp
   905 
   906 lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs"
   907 proof-
   908   have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor)
   909   
   910   have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def)
   911   also have "\<dots>" by (simp add: isint_add cti si)
   912   finally show ?thesis .
   913 qed
   914 
   915 consts split_int:: "num \<Rightarrow> num\<times>num"
   916 recdef split_int "measure num_size"
   917   "split_int (C c) = (C 0, C c)"
   918   "split_int (CN n c b) = 
   919      (let (bv,bi) = split_int b 
   920        in (CN n c bv, bi))"
   921   "split_int (CF c a b) = 
   922      (let (bv,bi) = split_int b 
   923        in (bv, CF c a bi))"
   924   "split_int a = (a,C 0)"
   925 
   926 lemma split_int:"\<And> tv ti. split_int t = (tv,ti) \<Longrightarrow> (Inum bs (Add tv ti) = Inum bs t) \<and> isint ti bs"
   927 proof (induct t rule: split_int.induct)
   928   case (2 c n b tv ti)
   929   let ?bv = "fst (split_int b)"
   930   let ?bi = "snd (split_int b)"
   931   have "split_int b = (?bv,?bi)" by simp
   932   with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
   933   from prems(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def)
   934   from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def)
   935 next
   936   case (3 c a b tv ti) 
   937   let ?bv = "fst (split_int b)"
   938   let ?bi = "snd (split_int b)"
   939   have "split_int b = (?bv,?bi)" by simp
   940   with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
   941   from prems(2) have tibi: "ti = CF c a ?bi" by (simp add: Let_def split_def)
   942   from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF)
   943 qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def ring_simps)
   944 
   945 lemma split_int_nb: "numbound0 t \<Longrightarrow> numbound0 (fst (split_int t)) \<and> numbound0 (snd (split_int t)) "
   946 by (induct t rule: split_int.induct, auto simp add: Let_def split_def)
   947 
   948 definition
   949   numfloor:: "num \<Rightarrow> num"
   950 where
   951   numfloor_def: "numfloor t = (let (tv,ti) = split_int t in 
   952   (case tv of C i \<Rightarrow> numadd (tv,ti) 
   953   | _ \<Rightarrow> numadd(CF 1 tv (C 0),ti)))"
   954 
   955 lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)")
   956 proof-
   957   let ?tv = "fst (split_int t)"
   958   let ?ti = "snd (split_int t)"
   959   have tvti:"split_int t = (?tv,?ti)" by simp
   960   {assume H: "\<forall> v. ?tv \<noteq> C v"
   961     hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" 
   962       by (cases ?tv, auto simp add: numfloor_def Let_def split_def numadd)
   963     from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
   964     hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp 
   965     also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
   966       by (simp,subst tii[simplified isint_iff, symmetric]) simp
   967     also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
   968     finally have ?thesis using th1 by simp}
   969   moreover {fix v assume H:"?tv = C v" 
   970     from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
   971     hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp 
   972     also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
   973       by (simp,subst tii[simplified isint_iff, symmetric]) simp
   974     also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
   975     finally have ?thesis by (simp add: H numfloor_def Let_def split_def numadd) }
   976   ultimately show ?thesis by auto
   977 qed
   978 
   979 lemma numfloor_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numfloor t)"
   980   using split_int_nb[where t="t"]
   981   by (cases "fst(split_int t)" , auto simp add: numfloor_def Let_def split_def  numadd_nb)
   982 
   983 recdef simpnum "measure num_size"
   984   "simpnum (C j) = C j"
   985   "simpnum (Bound n) = CN n 1 (C 0)"
   986   "simpnum (Neg t) = numneg (simpnum t)"
   987   "simpnum (Add t s) = numadd (simpnum t,simpnum s)"
   988   "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
   989   "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)"
   990   "simpnum (Floor t) = numfloor (simpnum t)"
   991   "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))"
   992   "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)"
   993 
   994 lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
   995 by (induct t rule: simpnum.induct, auto)
   996 
   997 lemma simpnum_numbound0[simp]: 
   998   "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"
   999 by (induct t rule: simpnum.induct, auto)
  1000 
  1001 consts nozerocoeff:: "num \<Rightarrow> bool"
  1002 recdef nozerocoeff "measure size"
  1003   "nozerocoeff (C c) = True"
  1004   "nozerocoeff (CN n c t) = (c\<noteq>0 \<and> nozerocoeff t)"
  1005   "nozerocoeff (CF c s t) = (c \<noteq> 0 \<and> nozerocoeff t)"
  1006   "nozerocoeff (Mul c t) = (c\<noteq>0 \<and> nozerocoeff t)"
  1007   "nozerocoeff t = True"
  1008 
  1009 lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd (a,b))"
  1010 by (induct a b rule: numadd.induct,auto simp add: Let_def)
  1011 
  1012 lemma nummul_nz : "\<And> i. i\<noteq>0 \<Longrightarrow> nozerocoeff a \<Longrightarrow> nozerocoeff (nummul a i)"
  1013   by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz)
  1014 
  1015 lemma numneg_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff (numneg a)"
  1016 by (simp add: numneg_def nummul_nz)
  1017 
  1018 lemma numsub_nz: "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numsub a b)"
  1019 by (simp add: numsub_def numneg_nz numadd_nz)
  1020 
  1021 lemma split_int_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (fst (split_int t)) \<and> nozerocoeff (snd (split_int t))"
  1022 by (induct t rule: split_int.induct,auto simp add: Let_def split_def)
  1023 
  1024 lemma numfloor_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (numfloor t)"
  1025 by (simp add: numfloor_def Let_def split_def)
  1026 (cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz)
  1027 
  1028 lemma simpnum_nz: "nozerocoeff (simpnum t)"
  1029 by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz)
  1030 
  1031 lemma maxcoeff_nz: "nozerocoeff t \<Longrightarrow> maxcoeff t = 0 \<Longrightarrow> t = C 0"
  1032 proof (induct t rule: maxcoeff.induct)
  1033   case (2 n c t)
  1034   hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
  1035   have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1)
  1036   with cnz have "max (abs c) (maxcoeff t) > 0" by arith
  1037   with prems show ?case by simp
  1038 next
  1039   case (3 c s t) 
  1040   hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
  1041   have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1)
  1042   with cnz have "max (abs c) (maxcoeff t) > 0" by arith
  1043   with prems show ?case by simp
  1044 qed auto
  1045 
  1046 lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0"
  1047 proof-
  1048   from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def)
  1049   from numgcdh0[OF th]  have th:"maxcoeff t = 0" .
  1050   from maxcoeff_nz[OF nz th] show ?thesis .
  1051 qed
  1052 
  1053 constdefs simp_num_pair:: "(num \<times> int) \<Rightarrow> num \<times> int"
  1054   "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
  1055    (let t' = simpnum t ; g = numgcd t' in 
  1056       if g > 1 then (let g' = igcd n g in 
  1057         if g' = 1 then (t',n) 
  1058         else (reducecoeffh t' g', n div g')) 
  1059       else (t',n))))"
  1060 
  1061 lemma simp_num_pair_ci:
  1062   shows "((\<lambda> (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\<lambda> (t,n). Inum bs t / real n) (t,n))"
  1063   (is "?lhs = ?rhs")
  1064 proof-
  1065   let ?t' = "simpnum t"
  1066   let ?g = "numgcd ?t'"
  1067   let ?g' = "igcd n ?g"
  1068   {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
  1069   moreover
  1070   { assume nnz: "n \<noteq> 0"
  1071     {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
  1072     moreover
  1073     {assume g1:"?g>1" hence g0: "?g > 0" by simp
  1074       from igcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
  1075       hence g'p: "?g' > 0" using igcd_pos[where i="n" and j="numgcd ?t'"] by arith
  1076       hence "?g'= 1 \<or> ?g' > 1" by arith
  1077       moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
  1078       moreover {assume g'1:"?g'>1"
  1079 	from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
  1080 	let ?tt = "reducecoeffh ?t' ?g'"
  1081 	let ?t = "Inum bs ?tt"
  1082 	have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2)
  1083 	have gpdd: "?g' dvd n" by (simp add: igcd_dvd1) 
  1084 	have gpdgp: "?g' dvd ?g'" by simp
  1085 	from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
  1086 	have th2:"real ?g' * ?t = Inum bs ?t'" by simp
  1087 	from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
  1088 	also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
  1089 	also have "\<dots> = (Inum bs ?t' / real n)"
  1090 	  using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp
  1091 	finally have "?lhs = Inum bs t / real n" by simp
  1092 	then have ?thesis using prems by (simp add: simp_num_pair_def)}
  1093       ultimately have ?thesis by blast}
  1094     ultimately have ?thesis by blast} 
  1095   ultimately show ?thesis by blast
  1096 qed
  1097 
  1098 lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')"
  1099   shows "numbound0 t' \<and> n' >0"
  1100 proof-
  1101     let ?t' = "simpnum t"
  1102   let ?g = "numgcd ?t'"
  1103   let ?g' = "igcd n ?g"
  1104   {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)}
  1105   moreover
  1106   { assume nnz: "n \<noteq> 0"
  1107     {assume "\<not> ?g > 1" hence ?thesis  using prems by (auto simp add: Let_def simp_num_pair_def)}
  1108     moreover
  1109     {assume g1:"?g>1" hence g0: "?g > 0" by simp
  1110       from igcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
  1111       hence g'p: "?g' > 0" using igcd_pos[where i="n" and j="numgcd ?t'"] by arith
  1112       hence "?g'= 1 \<or> ?g' > 1" by arith
  1113       moreover {assume "?g'=1" hence ?thesis using prems 
  1114 	  by (auto simp add: Let_def simp_num_pair_def)}
  1115       moreover {assume g'1:"?g'>1"
  1116 	have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2)
  1117 	have gpdd: "?g' dvd n" by (simp add: igcd_dvd1) 
  1118 	have gpdgp: "?g' dvd ?g'" by simp
  1119 	from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
  1120 	from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
  1121 	have "n div ?g' >0" by simp
  1122 	hence ?thesis using prems 
  1123 	  by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)}
  1124       ultimately have ?thesis by blast}
  1125     ultimately have ?thesis by blast} 
  1126   ultimately show ?thesis by blast
  1127 qed
  1128 
  1129 consts not:: "fm \<Rightarrow> fm"
  1130 recdef not "measure size"
  1131   "not (NOT p) = p"
  1132   "not T = F"
  1133   "not F = T"
  1134   "not (Lt t) = Ge t"
  1135   "not (Le t) = Gt t"
  1136   "not (Gt t) = Le t"
  1137   "not (Ge t) = Lt t"
  1138   "not (Eq t) = NEq t"
  1139   "not (NEq t) = Eq t"
  1140   "not (Dvd i t) = NDvd i t"
  1141   "not (NDvd i t) = Dvd i t"
  1142   "not (And p q) = Or (not p) (not q)"
  1143   "not (Or p q) = And (not p) (not q)"
  1144   "not p = NOT p"
  1145 lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
  1146 by (induct p) auto
  1147 lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
  1148 by (induct p, auto)
  1149 lemma not_nb[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
  1150 by (induct p, auto)
  1151 
  1152 constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1153   "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else 
  1154    if p = q then p else And p q)"
  1155 lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
  1156 by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
  1157 
  1158 lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
  1159 using conj_def by auto 
  1160 lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
  1161 using conj_def by auto 
  1162 
  1163 constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1164   "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
  1165        else if p=q then p else Or p q)"
  1166 
  1167 lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)"
  1168 by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
  1169 lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
  1170 using disj_def by auto 
  1171 lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
  1172 using disj_def by auto 
  1173 
  1174 constdefs   imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1175   "imp p q \<equiv> (if (p = F \<or> q=T \<or> p=q) then T else if p=T then q else if q=F then not p 
  1176     else Imp p q)"
  1177 lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
  1178 by (cases "p=F \<or> q=T",simp_all add: imp_def)
  1179 lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
  1180 using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
  1181 lemma imp_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
  1182 using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def) 
  1183 
  1184 constdefs   iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1185   "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else 
  1186        if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else 
  1187   Iff p q)"
  1188 lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
  1189   by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not) 
  1190 (cases "not p= q", auto simp add:not)
  1191 lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
  1192   by (unfold iff_def,cases "p=q", auto)
  1193 lemma iff_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)"
  1194 using iff_def by (unfold iff_def,cases "p=q", auto)
  1195 
  1196 consts check_int:: "num \<Rightarrow> bool"
  1197 recdef check_int "measure size"
  1198   "check_int (C i) = True"
  1199   "check_int (Floor t) = True"
  1200   "check_int (Mul i t) = check_int t"
  1201   "check_int (Add t s) = (check_int t \<and> check_int s)"
  1202   "check_int (Neg t) = check_int t"
  1203   "check_int (CF c t s) = check_int s"
  1204   "check_int t = False"
  1205 lemma check_int: "check_int t \<Longrightarrow> isint t bs"
  1206 by (induct t, auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF)
  1207 
  1208 lemma rdvd_left1_int: "real \<lfloor>t\<rfloor> = t \<Longrightarrow> 1 rdvd t"
  1209   by (simp add: rdvd_def,rule_tac x="\<lfloor>t\<rfloor>" in exI) simp
  1210 
  1211 lemma rdvd_reduce: 
  1212   assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0"
  1213   shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)"
  1214 proof
  1215   assume d: "real d rdvd real c * t"
  1216   from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto
  1217   from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto
  1218   from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto
  1219   from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp
  1220   hence "real kc * t = real kd * real k" using gp by simp
  1221   hence th:"real kd rdvd real kc * t" using rdvd_def by blast
  1222   from kd_def gp have th':"kd = d div g" by simp
  1223   from kc_def gp have "kc = c div g" by simp
  1224   with th th' show "real (d div g) rdvd real (c div g) * t" by simp
  1225 next
  1226   assume d: "real (d div g) rdvd real (c div g) * t"
  1227   from gp have gnz: "g \<noteq> 0" by simp
  1228   thus "real d rdvd real c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real (c div g) * t"] real_of_int_div[OF gnz gd] real_of_int_div[OF gnz gc] by simp
  1229 qed
  1230 
  1231 constdefs simpdvd:: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)"
  1232   "simpdvd d t \<equiv> 
  1233    (let g = numgcd t in 
  1234       if g > 1 then (let g' = igcd d g in 
  1235         if g' = 1 then (d, t) 
  1236         else (d div g',reducecoeffh t g')) 
  1237       else (d, t))"
  1238 lemma simpdvd: 
  1239   assumes tnz: "nozerocoeff t" and dnz: "d \<noteq> 0"
  1240   shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)"
  1241 proof-
  1242   let ?g = "numgcd t"
  1243   let ?g' = "igcd d ?g"
  1244   {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)}
  1245   moreover
  1246   {assume g1:"?g>1" hence g0: "?g > 0" by simp
  1247     from igcd0 g1 dnz have gp0: "?g' \<noteq> 0" by simp
  1248     hence g'p: "?g' > 0" using igcd_pos[where i="d" and j="numgcd t"] by arith
  1249     hence "?g'= 1 \<or> ?g' > 1" by arith
  1250     moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)}
  1251     moreover {assume g'1:"?g'>1"
  1252       from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" ..
  1253       let ?tt = "reducecoeffh t ?g'"
  1254       let ?t = "Inum bs ?tt"
  1255       have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2)
  1256       have gpdd: "?g' dvd d" by (simp add: igcd_dvd1) 
  1257       have gpdgp: "?g' dvd ?g'" by simp
  1258       from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
  1259       have th2:"real ?g' * ?t = Inum bs t" by simp
  1260       from prems have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)"
  1261 	by (simp add: simpdvd_def Let_def)
  1262       also have "\<dots> = (real d rdvd (Inum bs t))"
  1263 	using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified zdiv_self[OF gp0]] 
  1264 	  th2[symmetric] by simp
  1265       finally have ?thesis by simp  }
  1266     ultimately have ?thesis by blast
  1267   }
  1268   ultimately show ?thesis by blast
  1269 qed
  1270 
  1271 consts simpfm :: "fm \<Rightarrow> fm"
  1272 recdef simpfm "measure fmsize"
  1273   "simpfm (And p q) = conj (simpfm p) (simpfm q)"
  1274   "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
  1275   "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
  1276   "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
  1277   "simpfm (NOT p) = not (simpfm p)"
  1278   "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F 
  1279   | _ \<Rightarrow> Lt (reducecoeff a'))"
  1280   "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<le> 0)  then T else F | _ \<Rightarrow> Le (reducecoeff a'))"
  1281   "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v > 0)  then T else F | _ \<Rightarrow> Gt (reducecoeff a'))"
  1282   "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<ge> 0)  then T else F | _ \<Rightarrow> Ge (reducecoeff a'))"
  1283   "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v = 0)  then T else F | _ \<Rightarrow> Eq (reducecoeff a'))"
  1284   "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<noteq> 0)  then T else F | _ \<Rightarrow> NEq (reducecoeff a'))"
  1285   "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a)
  1286              else if (abs i = 1) \<and> check_int a then T
  1287              else let a' = simpnum a in case a' of C v \<Rightarrow> if (i dvd v)  then T else F | _ \<Rightarrow> (let (d,t) = simpdvd i a' in Dvd d t))"
  1288   "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) 
  1289              else if (abs i = 1) \<and> check_int a then F
  1290              else let a' = simpnum a in case a' of C v \<Rightarrow> if (\<not>(i dvd v)) then T else F | _ \<Rightarrow> (let (d,t) = simpdvd i a' in NDvd d t))"
  1291   "simpfm p = p"
  1292 
  1293 lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p"
  1294 proof(induct p rule: simpfm.induct)
  1295   case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1296   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1297   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1298     let ?g = "numgcd ?sa"
  1299     let ?rsa = "reducecoeff ?sa"
  1300     let ?r = "Inum bs ?rsa"
  1301     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1302     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1303     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1304     hence gp: "real ?g > 0" by simp
  1305     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1306     with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp
  1307     also have "\<dots> = (?r < 0)" using gp
  1308       by (simp only: mult_less_cancel_left) simp
  1309     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1310   ultimately show ?case by blast
  1311 next
  1312   case (7 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1313   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1314   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1315     let ?g = "numgcd ?sa"
  1316     let ?rsa = "reducecoeff ?sa"
  1317     let ?r = "Inum bs ?rsa"
  1318     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1319     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1320     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1321     hence gp: "real ?g > 0" by simp
  1322     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1323     with sa have "Inum bs a \<le> 0 = (real ?g * ?r \<le> real ?g * 0)" by simp
  1324     also have "\<dots> = (?r \<le> 0)" using gp
  1325       by (simp only: mult_le_cancel_left) simp
  1326     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1327   ultimately show ?case by blast
  1328 next
  1329   case (8 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1330   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1331   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1332     let ?g = "numgcd ?sa"
  1333     let ?rsa = "reducecoeff ?sa"
  1334     let ?r = "Inum bs ?rsa"
  1335     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1336     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1337     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1338     hence gp: "real ?g > 0" by simp
  1339     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1340     with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp
  1341     also have "\<dots> = (?r > 0)" using gp
  1342       by (simp only: mult_less_cancel_left) simp
  1343     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1344   ultimately show ?case by blast
  1345 next
  1346   case (9 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1347   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1348   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1349     let ?g = "numgcd ?sa"
  1350     let ?rsa = "reducecoeff ?sa"
  1351     let ?r = "Inum bs ?rsa"
  1352     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1353     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1354     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1355     hence gp: "real ?g > 0" by simp
  1356     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1357     with sa have "Inum bs a \<ge> 0 = (real ?g * ?r \<ge> real ?g * 0)" by simp
  1358     also have "\<dots> = (?r \<ge> 0)" using gp
  1359       by (simp only: mult_le_cancel_left) simp
  1360     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1361   ultimately show ?case by blast
  1362 next
  1363   case (10 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1364   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1365   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1366     let ?g = "numgcd ?sa"
  1367     let ?rsa = "reducecoeff ?sa"
  1368     let ?r = "Inum bs ?rsa"
  1369     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1370     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1371     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1372     hence gp: "real ?g > 0" by simp
  1373     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1374     with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp
  1375     also have "\<dots> = (?r = 0)" using gp
  1376       by (simp add: mult_eq_0_iff)
  1377     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1378   ultimately show ?case by blast
  1379 next
  1380   case (11 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1381   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1382   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1383     let ?g = "numgcd ?sa"
  1384     let ?rsa = "reducecoeff ?sa"
  1385     let ?r = "Inum bs ?rsa"
  1386     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1387     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1388     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1389     hence gp: "real ?g > 0" by simp
  1390     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1391     with sa have "Inum bs a \<noteq> 0 = (real ?g * ?r \<noteq> 0)" by simp
  1392     also have "\<dots> = (?r \<noteq> 0)" using gp
  1393       by (simp add: mult_eq_0_iff)
  1394     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1395   ultimately show ?case by blast
  1396 next
  1397   case (12 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
  1398   have "i=0 \<or> (abs i = 1 \<and> check_int a) \<or> (i\<noteq>0 \<and> ((abs i \<noteq> 1) \<or> (\<not> check_int a)))" by auto
  1399   {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)}
  1400   moreover 
  1401   {assume ai1: "abs i = 1" and ai: "check_int a" 
  1402     hence "i=1 \<or> i= - 1" by arith
  1403     moreover {assume i1: "i = 1" 
  1404       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1405       have ?case using i1 ai by simp }
  1406     moreover {assume i1: "i = - 1" 
  1407       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1408 	rdvd_abs1[where d="- 1" and t="Inum bs a"]
  1409       have ?case using i1 ai by simp }
  1410     ultimately have ?case by blast}
  1411   moreover   
  1412   {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
  1413     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
  1414 	by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
  1415     moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
  1416       hence th: "simpfm (Dvd i a) = Dvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def)
  1417       from simpnum_nz have nz:"nozerocoeff ?sa" by simp
  1418       from simpdvd [OF nz inz] th have ?case using sa by simp}
  1419     ultimately have ?case by blast}
  1420   ultimately show ?case by blast
  1421 next
  1422   case (13 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
  1423   have "i=0 \<or> (abs i = 1 \<and> check_int a) \<or> (i\<noteq>0 \<and> ((abs i \<noteq> 1) \<or> (\<not> check_int a)))" by auto
  1424   {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)}
  1425   moreover 
  1426   {assume ai1: "abs i = 1" and ai: "check_int a" 
  1427     hence "i=1 \<or> i= - 1" by arith
  1428     moreover {assume i1: "i = 1" 
  1429       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1430       have ?case using i1 ai by simp }
  1431     moreover {assume i1: "i = - 1" 
  1432       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1433 	rdvd_abs1[where d="- 1" and t="Inum bs a"]
  1434       have ?case using i1 ai by simp }
  1435     ultimately have ?case by blast}
  1436   moreover   
  1437   {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
  1438     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
  1439 	by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
  1440     moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
  1441       hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond 
  1442 	by (cases ?sa, auto simp add: Let_def split_def)
  1443       from simpnum_nz have nz:"nozerocoeff ?sa" by simp
  1444       from simpdvd [OF nz inz] th have ?case using sa by simp}
  1445     ultimately have ?case by blast}
  1446   ultimately show ?case by blast
  1447 qed (induct p rule: simpfm.induct, simp_all)
  1448 
  1449 lemma simpdvd_numbound0: "numbound0 t \<Longrightarrow> numbound0 (snd (simpdvd d t))"
  1450   by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0)
  1451 
  1452 lemma simpfm_bound0[simp]: "bound0 p \<Longrightarrow> bound0 (simpfm p)"
  1453 proof(induct p rule: simpfm.induct)
  1454   case (6 a) hence nb: "numbound0 a" by simp
  1455   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1456   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1457 next
  1458   case (7 a) hence nb: "numbound0 a" by simp
  1459   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1460   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1461 next
  1462   case (8 a) hence nb: "numbound0 a" by simp
  1463   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1464   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1465 next
  1466   case (9 a) hence nb: "numbound0 a" by simp
  1467   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1468   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1469 next
  1470   case (10 a) hence nb: "numbound0 a" by simp
  1471   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1472   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1473 next
  1474   case (11 a) hence nb: "numbound0 a" by simp
  1475   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1476   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1477 next
  1478   case (12 i a) hence nb: "numbound0 a" by simp
  1479   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1480   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
  1481 next
  1482   case (13 i a) hence nb: "numbound0 a" by simp
  1483   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1484   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
  1485 qed(auto simp add: disj_def imp_def iff_def conj_def)
  1486 
  1487 lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
  1488 by (induct p rule: simpfm.induct, auto simp add: Let_def)
  1489 (case_tac "simpnum a",auto simp add: split_def Let_def)+
  1490 
  1491 
  1492   (* Generic quantifier elimination *)
  1493 
  1494 constdefs list_conj :: "fm list \<Rightarrow> fm"
  1495   "list_conj ps \<equiv> foldr conj ps T"
  1496 lemma list_conj: "Ifm bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm bs p)"
  1497   by (induct ps, auto simp add: list_conj_def)
  1498 lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
  1499   by (induct ps, auto simp add: list_conj_def)
  1500 lemma list_conj_nb: " \<forall>p\<in> set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
  1501   by (induct ps, auto simp add: list_conj_def)
  1502 constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
  1503   "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs
  1504                    in conj (decr (list_conj yes)) (f (list_conj no)))"
  1505 
  1506 lemma CJNB_qe: 
  1507   assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
  1508   shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm bs ((CJNB qe p)) = Ifm bs (E p))"
  1509 proof(clarify)
  1510   fix bs p
  1511   assume qfp: "qfree p"
  1512   let ?cjs = "conjuncts p"
  1513   let ?yes = "fst (partition bound0 ?cjs)"
  1514   let ?no = "snd (partition bound0 ?cjs)"
  1515   let ?cno = "list_conj ?no"
  1516   let ?cyes = "list_conj ?yes"
  1517   have part: "partition bound0 ?cjs = (?yes,?no)" by simp
  1518   from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast 
  1519   hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) 
  1520   hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf)
  1521   from conjuncts_qf[OF qfp] partition_set[OF part] 
  1522   have " \<forall>q\<in> set ?no. qfree q" by auto
  1523   hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
  1524   with qe have cno_qf:"qfree (qe ?cno )" 
  1525     and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+
  1526   from cno_qf yes_qf have qf: "qfree (CJNB qe p)" 
  1527     by (simp add: CJNB_def Let_def conj_qf split_def)
  1528   {fix bs
  1529     from conjuncts have "Ifm bs p = (\<forall>q\<in> set ?cjs. Ifm bs q)" by blast
  1530     also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm bs q) \<and> (\<forall>q\<in> set ?no. Ifm bs q))"
  1531       using partition_set[OF part] by auto
  1532     finally have "Ifm bs p = ((Ifm bs ?cyes) \<and> (Ifm bs ?cno))" using list_conj by simp}
  1533   hence "Ifm bs (E p) = (\<exists>x. (Ifm (x#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))" by simp
  1534   also have "\<dots> = (\<exists>x. (Ifm (y#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))"
  1535     using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
  1536   also have "\<dots> = (Ifm bs (decr ?cyes) \<and> Ifm bs (E ?cno))"
  1537     by (auto simp add: decr[OF yes_nb])
  1538   also have "\<dots> = (Ifm bs (conj (decr ?cyes) (qe ?cno)))"
  1539     using qe[rule_format, OF no_qf] by auto
  1540   finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" 
  1541     by (simp add: Let_def CJNB_def split_def)
  1542   with qf show "qfree (CJNB qe p) \<and> Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast
  1543 qed
  1544 
  1545 consts qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
  1546 recdef qelim "measure fmsize"
  1547   "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
  1548   "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
  1549   "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
  1550   "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))" 
  1551   "qelim (Or  p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))" 
  1552   "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))"
  1553   "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
  1554   "qelim p = (\<lambda> y. simpfm p)"
  1555 
  1556 lemma qelim_ci:
  1557   assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
  1558   shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)"
  1559 using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] 
  1560 by(induct p rule: qelim.induct) 
  1561 (auto simp del: simpfm.simps)
  1562 
  1563 
  1564 text {* The @{text "\<int>"} Part *}
  1565 text{* Linearity for fm where Bound 0 ranges over @{text "\<int>"} *}
  1566 consts
  1567   zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*)
  1568 recdef zsplit0 "measure num_size"
  1569   "zsplit0 (C c) = (0,C c)"
  1570   "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))"
  1571   "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)"
  1572   "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)"
  1573   "zsplit0 (Neg a) = (let (i',a') =  zsplit0 a in (-i', Neg a'))"
  1574   "zsplit0 (Add a b) = (let (ia,a') =  zsplit0 a ; 
  1575                             (ib,b') =  zsplit0 b 
  1576                             in (ia+ib, Add a' b'))"
  1577   "zsplit0 (Sub a b) = (let (ia,a') =  zsplit0 a ; 
  1578                             (ib,b') =  zsplit0 b 
  1579                             in (ia-ib, Sub a' b'))"
  1580   "zsplit0 (Mul i a) = (let (i',a') =  zsplit0 a in (i*i', Mul i a'))"
  1581   "zsplit0 (Floor a) = (let (i',a') =  zsplit0 a in (i',Floor a'))"
  1582 (hints simp add: Let_def)
  1583 
  1584 lemma zsplit0_I:
  1585   shows "\<And> n a. zsplit0 t = (n,a) \<Longrightarrow> (Inum ((real (x::int)) #bs) (CN 0 n a) = Inum (real x #bs) t) \<and> numbound0 a"
  1586   (is "\<And> n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a")
  1587 proof(induct t rule: zsplit0.induct)
  1588   case (1 c n a) thus ?case by auto 
  1589 next
  1590   case (2 m n a) thus ?case by (cases "m=0") auto
  1591 next
  1592   case (3 n i a n a') thus ?case by auto
  1593 next 
  1594   case (4 c a b n a') thus ?case by auto
  1595 next
  1596   case (5 t n a)
  1597   let ?nt = "fst (zsplit0 t)"
  1598   let ?at = "snd (zsplit0 t)"
  1599   have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \<and> n=-?nt" using prems 
  1600     by (simp add: Let_def split_def)
  1601   from abj prems  have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1602   from th2[simplified] th[simplified] show ?case by simp
  1603 next
  1604   case (6 s t n a)
  1605   let ?ns = "fst (zsplit0 s)"
  1606   let ?as = "snd (zsplit0 s)"
  1607   let ?nt = "fst (zsplit0 t)"
  1608   let ?at = "snd (zsplit0 t)"
  1609   have abjs: "zsplit0 s = (?ns,?as)" by simp 
  1610   moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp 
  1611   ultimately have th: "a=Add ?as ?at \<and> n=?ns + ?nt" using prems 
  1612     by (simp add: Let_def split_def)
  1613   from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
  1614   from prems have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \<and> numbound0 xb)" by simp
  1615   with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1616   from abjs prems  have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
  1617   from th3[simplified] th2[simplified] th[simplified] show ?case 
  1618     by (simp add: left_distrib)
  1619 next
  1620   case (7 s t n a)
  1621   let ?ns = "fst (zsplit0 s)"
  1622   let ?as = "snd (zsplit0 s)"
  1623   let ?nt = "fst (zsplit0 t)"
  1624   let ?at = "snd (zsplit0 t)"
  1625   have abjs: "zsplit0 s = (?ns,?as)" by simp 
  1626   moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp 
  1627   ultimately have th: "a=Sub ?as ?at \<and> n=?ns - ?nt" using prems 
  1628     by (simp add: Let_def split_def)
  1629   from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
  1630   from prems have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \<and> numbound0 xb)" by simp
  1631   with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1632   from abjs prems  have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
  1633   from th3[simplified] th2[simplified] th[simplified] show ?case 
  1634     by (simp add: left_diff_distrib)
  1635 next
  1636   case (8 i t n a)
  1637   let ?nt = "fst (zsplit0 t)"
  1638   let ?at = "snd (zsplit0 t)"
  1639   have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \<and> n=i*?nt" using prems 
  1640     by (simp add: Let_def split_def)
  1641   from abj prems  have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1642   hence " ?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp
  1643   also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: right_distrib)
  1644   finally show ?case using th th2 by simp
  1645 next
  1646   case (9 t n a)
  1647   let ?nt = "fst (zsplit0 t)"
  1648   let ?at = "snd (zsplit0 t)"
  1649   have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \<and> n=?nt" using prems 
  1650     by (simp add: Let_def split_def)
  1651   from abj prems  have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1652   hence na: "?N a" using th by simp
  1653   have th': "(real ?nt)*(real x) = real (?nt * x)" by simp
  1654   have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp
  1655   also have "\<dots> = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp
  1656   also have "\<dots> = real (floor (?I x ?at + real (?nt* x)))" by (simp add: add_ac)
  1657   also have "\<dots> = real (floor (?I x ?at) + (?nt* x))" 
  1658     using floor_add[where x="?I x ?at" and a="?nt* x"] by simp 
  1659   also have "\<dots> = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: add_ac)
  1660   finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp
  1661   with na show ?case by simp
  1662 qed
  1663 
  1664 consts
  1665   iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool"   (* Linearity test for fm *)
  1666   zlfm :: "fm \<Rightarrow> fm"       (* Linearity transformation for fm *)
  1667 recdef iszlfm "measure size"
  1668   "iszlfm (And p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" 
  1669   "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" 
  1670   "iszlfm (Eq  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1671   "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1672   "iszlfm (Lt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1673   "iszlfm (Le  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1674   "iszlfm (Gt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1675   "iszlfm (Ge  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1676   "iszlfm (Dvd i (CN 0 c e)) = 
  1677                  (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
  1678   "iszlfm (NDvd i (CN 0 c e))= 
  1679                  (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
  1680   "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 p))"
  1681 
  1682 lemma zlin_qfree: "iszlfm p bs \<Longrightarrow> qfree p"
  1683   by (induct p rule: iszlfm.induct) auto
  1684 
  1685 lemma iszlfm_gen:
  1686   assumes lp: "iszlfm p (x#bs)"
  1687   shows "\<forall> y. iszlfm p (y#bs)"
  1688 proof
  1689   fix y
  1690   show "iszlfm p (y#bs)"
  1691     using lp
  1692   by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"])
  1693 qed
  1694 
  1695 lemma conj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (conj p q) bs"
  1696   using conj_def by (cases p,auto)
  1697 lemma disj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (disj p q) bs"
  1698   using disj_def by (cases p,auto)
  1699 lemma not_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm (not p) bs"
  1700   by (induct p rule:iszlfm.induct ,auto)
  1701 
  1702 recdef zlfm "measure fmsize"
  1703   "zlfm (And p q) = conj (zlfm p) (zlfm q)"
  1704   "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
  1705   "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
  1706   "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
  1707   "zlfm (Lt a) = (let (c,r) = zsplit0 a in 
  1708      if c=0 then Lt r else 
  1709      if c>0 then Or (Lt (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) 
  1710      else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
  1711   "zlfm (Le a) = (let (c,r) = zsplit0 a in 
  1712      if c=0 then Le r else 
  1713      if c>0 then Or (Le (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) 
  1714      else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
  1715   "zlfm (Gt a) = (let (c,r) = zsplit0 a in 
  1716      if c=0 then Gt r else 
  1717      if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) 
  1718      else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
  1719   "zlfm (Ge a) = (let (c,r) = zsplit0 a in 
  1720      if c=0 then Ge r else 
  1721      if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) 
  1722      else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
  1723   "zlfm (Eq a) = (let (c,r) = zsplit0 a in 
  1724               if c=0 then Eq r else 
  1725       if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r)))
  1726       else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))"
  1727   "zlfm (NEq a) = (let (c,r) = zsplit0 a in 
  1728               if c=0 then NEq r else 
  1729       if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r)))
  1730       else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))"
  1731   "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) 
  1732   else (let (c,r) = zsplit0 a in 
  1733               if c=0 then Dvd (abs i) r else 
  1734       if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) 
  1735       else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
  1736   "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) 
  1737   else (let (c,r) = zsplit0 a in 
  1738               if c=0 then NDvd (abs i) r else 
  1739       if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) 
  1740       else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
  1741   "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
  1742   "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
  1743   "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
  1744   "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
  1745   "zlfm (NOT (NOT p)) = zlfm p"
  1746   "zlfm (NOT T) = F"
  1747   "zlfm (NOT F) = T"
  1748   "zlfm (NOT (Lt a)) = zlfm (Ge a)"
  1749   "zlfm (NOT (Le a)) = zlfm (Gt a)"
  1750   "zlfm (NOT (Gt a)) = zlfm (Le a)"
  1751   "zlfm (NOT (Ge a)) = zlfm (Lt a)"
  1752   "zlfm (NOT (Eq a)) = zlfm (NEq a)"
  1753   "zlfm (NOT (NEq a)) = zlfm (Eq a)"
  1754   "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
  1755   "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
  1756   "zlfm p = p" (hints simp add: fmsize_pos)
  1757 
  1758 lemma split_int_less_real: 
  1759   "(real (a::int) < b) = (a < floor b \<or> (a = floor b \<and> real (floor b) < b))"
  1760 proof( auto)
  1761   assume alb: "real a < b" and agb: "\<not> a < floor b"
  1762   from agb have "floor b \<le> a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq)
  1763   from floor_eq[OF alb th] show "a= floor b" by simp 
  1764 next
  1765   assume alb: "a < floor b"
  1766   hence "real a < real (floor b)" by simp
  1767   moreover have "real (floor b) \<le> b" by simp ultimately show  "real a < b" by arith 
  1768 qed
  1769 
  1770 lemma split_int_less_real': 
  1771   "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \<or> (real a - real (floor (-b)) = 0 \<and> real (floor (-b)) + b < 0))"
  1772 proof- 
  1773   have "(real a + b <0) = (real a < -b)" by arith
  1774   with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith  
  1775 qed
  1776 
  1777 lemma split_int_gt_real': 
  1778   "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \<or> (real a + real (floor b) = 0 \<and> real (floor b) - b < 0))"
  1779 proof- 
  1780   have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith
  1781   show ?thesis using myless[rule_format, where b="real (floor b)"] 
  1782     by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) 
  1783     (simp add: ring_simps diff_def[symmetric],arith)
  1784 qed
  1785 
  1786 lemma split_int_le_real: 
  1787   "(real (a::int) \<le> b) = (a \<le> floor b \<or> (a = floor b \<and> real (floor b) < b))"
  1788 proof( auto)
  1789   assume alb: "real a \<le> b" and agb: "\<not> a \<le> floor b"
  1790   from alb have "floor (real a) \<le> floor b " by (simp only: floor_mono2) 
  1791   hence "a \<le> floor b" by simp with agb show "False" by simp
  1792 next
  1793   assume alb: "a \<le> floor b"
  1794   hence "real a \<le> real (floor b)" by (simp only: floor_mono2)
  1795   also have "\<dots>\<le> b" by simp  finally show  "real a \<le> b" . 
  1796 qed
  1797 
  1798 lemma split_int_le_real': 
  1799   "(real (a::int) + b \<le> 0) = (real a - real (floor(-b)) \<le> 0 \<or> (real a - real (floor (-b)) = 0 \<and> real (floor (-b)) + b < 0))"
  1800 proof- 
  1801   have "(real a + b \<le>0) = (real a \<le> -b)" by arith
  1802   with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith  
  1803 qed
  1804 
  1805 lemma split_int_ge_real': 
  1806   "(real (a::int) + b \<ge> 0) = (real a + real (floor b) \<ge> 0 \<or> (real a + real (floor b) = 0 \<and> real (floor b) - b < 0))"
  1807 proof- 
  1808   have th: "(real a + b \<ge>0) = (real (-a) + (-b) \<le> 0)" by arith
  1809   show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"])
  1810     (simp add: ring_simps diff_def[symmetric],arith)
  1811 qed
  1812 
  1813 lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \<and> b = real (floor b))" (is "?l = ?r")
  1814 by auto
  1815 
  1816 lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \<and> real (floor (-b)) + b = 0)" (is "?l = ?r")
  1817 proof-
  1818   have "?l = (real a = -b)" by arith
  1819   with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith
  1820 qed
  1821 
  1822 lemma zlfm_I:
  1823   assumes qfp: "qfree p"
  1824   shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \<and> iszlfm (zlfm p) (real (i::int) #bs)"
  1825   (is "(?I (?l p) = ?I p) \<and> ?L (?l p)")
  1826 using qfp
  1827 proof(induct p rule: zlfm.induct)
  1828   case (5 a) 
  1829   let ?c = "fst (zsplit0 a)"
  1830   let ?r = "snd (zsplit0 a)"
  1831   have spl: "zsplit0 a = (?c,?r)" by simp
  1832   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1833   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1834   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1835   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1836   moreover
  1837   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1838       by (cases "?r", simp_all add: Let_def split_def,case_tac "nat", simp_all)}
  1839   moreover
  1840   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
  1841       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1842     have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
  1843     also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def diff_def)
  1844     finally have ?case using l by simp}
  1845   moreover
  1846   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
  1847       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1848     have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
  1849     also from cn cnz have "\<dots> = (?I (?l (Lt a)))" by (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith)
  1850     finally have ?case using l by simp}
  1851   ultimately show ?case by blast
  1852 next
  1853   case (6 a)
  1854   let ?c = "fst (zsplit0 a)"
  1855   let ?r = "snd (zsplit0 a)"
  1856   have spl: "zsplit0 a = (?c,?r)" by simp
  1857   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1858   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1859   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1860   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1861   moreover
  1862   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1863       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat",simp_all)}
  1864   moreover
  1865   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
  1866       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1867     have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
  1868     also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def)
  1869     finally have ?case using l by simp}
  1870   moreover
  1871   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
  1872       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1873     have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
  1874     also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac ,arith)
  1875     finally have ?case using l by simp}
  1876   ultimately show ?case by blast
  1877 next
  1878   case (7 a) 
  1879   let ?c = "fst (zsplit0 a)"
  1880   let ?r = "snd (zsplit0 a)"
  1881   have spl: "zsplit0 a = (?c,?r)" by simp
  1882   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1883   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1884   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1885   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1886   moreover
  1887   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1888       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1889   moreover
  1890   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
  1891       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1892     have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
  1893     also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def)
  1894     finally have ?case using l by simp}
  1895   moreover
  1896   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
  1897       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1898     have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
  1899     also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith)
  1900     finally have ?case using l by simp}
  1901   ultimately show ?case by blast
  1902 next
  1903   case (8 a)
  1904    let ?c = "fst (zsplit0 a)"
  1905   let ?r = "snd (zsplit0 a)"
  1906   have spl: "zsplit0 a = (?c,?r)" by simp
  1907   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1908   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1909   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1910   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1911   moreover
  1912   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1913       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1914   moreover
  1915   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
  1916       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1917     have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
  1918     also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_def)
  1919     finally have ?case using l by simp}
  1920   moreover
  1921   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
  1922       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1923     have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
  1924     also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_def[symmetric] add_ac, arith)
  1925     finally have ?case using l by simp}
  1926   ultimately show ?case by blast
  1927 next
  1928   case (9 a)
  1929   let ?c = "fst (zsplit0 a)"
  1930   let ?r = "snd (zsplit0 a)"
  1931   have spl: "zsplit0 a = (?c,?r)" by simp
  1932   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1933   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1934   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1935   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1936   moreover
  1937   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1938       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1939   moreover
  1940   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" 
  1941       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1942     have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
  1943     also have "\<dots> = (?I (?l (Eq a)))" using cp cnz  by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult)
  1944     finally have ?case using l by simp}
  1945   moreover
  1946   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" 
  1947       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1948     have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
  1949     also from cn cnz have "\<dots> = (?I (?l (Eq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith)
  1950     finally have ?case using l by simp}
  1951   ultimately show ?case by blast
  1952 next
  1953   case (10 a)
  1954   let ?c = "fst (zsplit0 a)"
  1955   let ?r = "snd (zsplit0 a)"
  1956   have spl: "zsplit0 a = (?c,?r)" by simp
  1957   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1958   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1959   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1960   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1961   moreover
  1962   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1963       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1964   moreover
  1965   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" 
  1966       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1967     have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
  1968     also have "\<dots> = (?I (?l (NEq a)))" using cp cnz  by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult)
  1969     finally have ?case using l by simp}
  1970   moreover
  1971   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" 
  1972       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1973     have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
  1974     also from cn cnz have "\<dots> = (?I (?l (NEq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith)
  1975     finally have ?case using l by simp}
  1976   ultimately show ?case by blast
  1977 next
  1978   case (11 j a)
  1979   let ?c = "fst (zsplit0 a)"
  1980   let ?r = "snd (zsplit0 a)"
  1981   have spl: "zsplit0 a = (?c,?r)" by simp
  1982   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1983   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1984   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1985   have "j=0 \<or> (j\<noteq>0 \<and> ?c = 0) \<or> (j\<noteq>0 \<and> ?c >0 \<and> ?c\<noteq>0) \<or> (j\<noteq> 0 \<and> ?c<0 \<and> ?c\<noteq>0)" by arith
  1986   moreover
  1987   {assume "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) 
  1988     hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)}
  1989   moreover
  1990   {assume "?c=0" and "j\<noteq>0" hence ?case 
  1991       using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
  1992       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1993   moreover
  1994   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
  1995       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1996     have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" 
  1997       using Ia by (simp add: Let_def split_def)
  1998     also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" 
  1999       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  2000     also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  2001        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" 
  2002       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  2003     also have "\<dots> = (?I (?l (Dvd j a)))" using cp cnz jnz  
  2004       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2005 	del: real_of_int_mult) (auto simp add: add_ac)
  2006     finally have ?case using l jnz  by simp }
  2007   moreover
  2008   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
  2009       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  2010     have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" 
  2011       using Ia by (simp add: Let_def split_def)
  2012     also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" 
  2013       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  2014     also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  2015        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" 
  2016       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  2017     also have "\<dots> = (?I (?l (Dvd j a)))" using cn cnz jnz
  2018       using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
  2019       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2020 	del: real_of_int_mult) (auto simp add: add_ac)
  2021     finally have ?case using l jnz by blast }
  2022   ultimately show ?case by blast
  2023 next
  2024   case (12 j a)
  2025   let ?c = "fst (zsplit0 a)"
  2026   let ?r = "snd (zsplit0 a)"
  2027   have spl: "zsplit0 a = (?c,?r)" by simp
  2028   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  2029   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  2030   let ?N = "\<lambda> t. Inum (real i#bs) t"
  2031   have "j=0 \<or> (j\<noteq>0 \<and> ?c = 0) \<or> (j\<noteq>0 \<and> ?c >0 \<and> ?c\<noteq>0) \<or> (j\<noteq> 0 \<and> ?c<0 \<and> ?c\<noteq>0)" by arith
  2032   moreover
  2033   {assume "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) 
  2034     hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)}
  2035   moreover
  2036   {assume "?c=0" and "j\<noteq>0" hence ?case 
  2037       using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
  2038       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  2039   moreover
  2040   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" 
  2041       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  2042     have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" 
  2043       using Ia by (simp add: Let_def split_def)
  2044     also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" 
  2045       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  2046     also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  2047        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" 
  2048       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  2049     also have "\<dots> = (?I (?l (NDvd j a)))" using cp cnz jnz  
  2050       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2051 	del: real_of_int_mult) (auto simp add: add_ac)
  2052     finally have ?case using l jnz  by simp }
  2053   moreover
  2054   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" 
  2055       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  2056     have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" 
  2057       using Ia by (simp add: Let_def split_def)
  2058     also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" 
  2059       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  2060     also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  2061        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" 
  2062       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  2063     also have "\<dots> = (?I (?l (NDvd j a)))" using cn cnz jnz
  2064       using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
  2065       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2066 	del: real_of_int_mult) (auto simp add: add_ac)
  2067     finally have ?case using l jnz by blast }
  2068   ultimately show ?case by blast
  2069 qed auto
  2070 
  2071 text{* plusinf : Virtual substitution of @{text "+\<infinity>"}
  2072        minusinf: Virtual substitution of @{text "-\<infinity>"}
  2073        @{text "\<delta>"} Compute lcm @{text "d| Dvd d  c*x+t \<in> p"}
  2074        @{text "d\<delta>"} checks if a given l divides all the ds above*}
  2075 
  2076 consts 
  2077   plusinf:: "fm \<Rightarrow> fm" 
  2078   minusinf:: "fm \<Rightarrow> fm"
  2079   \<delta> :: "fm \<Rightarrow> int" 
  2080   d\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"
  2081 
  2082 recdef minusinf "measure size"
  2083   "minusinf (And p q) = conj (minusinf p) (minusinf q)" 
  2084   "minusinf (Or p q) = disj (minusinf p) (minusinf q)" 
  2085   "minusinf (Eq  (CN 0 c e)) = F"
  2086   "minusinf (NEq (CN 0 c e)) = T"
  2087   "minusinf (Lt  (CN 0 c e)) = T"
  2088   "minusinf (Le  (CN 0 c e)) = T"
  2089   "minusinf (Gt  (CN 0 c e)) = F"
  2090   "minusinf (Ge  (CN 0 c e)) = F"
  2091   "minusinf p = p"
  2092 
  2093 lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"
  2094   by (induct p rule: minusinf.induct, auto)
  2095 
  2096 recdef plusinf "measure size"
  2097   "plusinf (And p q) = conj (plusinf p) (plusinf q)" 
  2098   "plusinf (Or p q) = disj (plusinf p) (plusinf q)" 
  2099   "plusinf (Eq  (CN 0 c e)) = F"
  2100   "plusinf (NEq (CN 0 c e)) = T"
  2101   "plusinf (Lt  (CN 0 c e)) = F"
  2102   "plusinf (Le  (CN 0 c e)) = F"
  2103   "plusinf (Gt  (CN 0 c e)) = T"
  2104   "plusinf (Ge  (CN 0 c e)) = T"
  2105   "plusinf p = p"
  2106 
  2107 recdef \<delta> "measure size"
  2108   "\<delta> (And p q) = ilcm (\<delta> p) (\<delta> q)" 
  2109   "\<delta> (Or p q) = ilcm (\<delta> p) (\<delta> q)" 
  2110   "\<delta> (Dvd i (CN 0 c e)) = i"
  2111   "\<delta> (NDvd i (CN 0 c e)) = i"
  2112   "\<delta> p = 1"
  2113 
  2114 recdef d\<delta> "measure size"
  2115   "d\<delta> (And p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" 
  2116   "d\<delta> (Or p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" 
  2117   "d\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
  2118   "d\<delta> (NDvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
  2119   "d\<delta> p = (\<lambda> d. True)"
  2120 
  2121 lemma delta_mono: 
  2122   assumes lin: "iszlfm p bs"
  2123   and d: "d dvd d'"
  2124   and ad: "d\<delta> p d"
  2125   shows "d\<delta> p d'"
  2126   using lin ad d
  2127 proof(induct p rule: iszlfm.induct)
  2128   case (9 i c e)  thus ?case using d
  2129     by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"])
  2130 next
  2131   case (10 i c e) thus ?case using d
  2132     by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"])
  2133 qed simp_all
  2134 
  2135 lemma \<delta> : assumes lin:"iszlfm p bs"
  2136   shows "d\<delta> p (\<delta> p) \<and> \<delta> p >0"
  2137 using lin
  2138 proof (induct p rule: iszlfm.induct)
  2139   case (1 p q) 
  2140   let ?d = "\<delta> (And p q)"
  2141   from prems ilcm_pos have dp: "?d >0" by simp
  2142   have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp 
  2143    hence th: "d\<delta> p ?d" 
  2144      using delta_mono prems by (auto simp del: dvd_ilcm_self1)
  2145   have "\<delta> q dvd \<delta> (And p q)" using prems  by simp 
  2146   hence th': "d\<delta> q ?d" using delta_mono prems by (auto simp del: dvd_ilcm_self2)
  2147   from th th' dp show ?case by simp 
  2148 next
  2149   case (2 p q)  
  2150   let ?d = "\<delta> (And p q)"
  2151   from prems ilcm_pos have dp: "?d >0" by simp
  2152   have "\<delta> p dvd \<delta> (And p q)" using prems by simp hence th: "d\<delta> p ?d" using delta_mono prems 
  2153     by (auto simp del: dvd_ilcm_self1)
  2154   have "\<delta> q dvd \<delta> (And p q)" using prems by simp hence th': "d\<delta> q ?d" using delta_mono prems by (auto simp del: dvd_ilcm_self2)
  2155   from th th' dp show ?case by simp 
  2156 qed simp_all
  2157 
  2158 
  2159 lemma minusinf_inf:
  2160   assumes linp: "iszlfm p (a # bs)"
  2161   shows "\<exists> (z::int). \<forall> x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p"
  2162   (is "?P p" is "\<exists> (z::int). \<forall> x < z. ?I x (?M p) = ?I x p")
  2163 using linp
  2164 proof (induct p rule: minusinf.induct)
  2165   case (1 f g)
  2166   from prems have "?P f" by simp
  2167   then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
  2168   from prems have "?P g" by simp
  2169   then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
  2170   let ?z = "min z1 z2"
  2171   from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp
  2172   thus ?case by blast
  2173 next
  2174   case (2 f g)   from prems have "?P f" by simp
  2175   then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
  2176   from prems have "?P g" by simp
  2177   then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
  2178   let ?z = "min z1 z2"
  2179   from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp
  2180   thus ?case by blast
  2181 next
  2182   case (3 c e) 
  2183   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2184   from prems have nbe: "numbound0 e" by simp
  2185   have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))"
  2186   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2187     fix x
  2188     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2189     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2190     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2191       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2192     hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
  2193     thus "real c * real x + Inum (real x # bs) e \<noteq> 0" 
  2194       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"]  by simp
  2195   qed
  2196   thus ?case by blast
  2197 next
  2198   case (4 c e) 
  2199   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2200   from prems have nbe: "numbound0 e" by simp
  2201   have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))"
  2202   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2203     fix x
  2204     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2205     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2206     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2207       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2208     hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
  2209     thus "real c * real x + Inum (real x # bs) e \<noteq> 0" 
  2210       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"]  by simp
  2211   qed
  2212   thus ?case by blast
  2213 next
  2214   case (5 c e) 
  2215   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2216   from prems have nbe: "numbound0 e" by simp
  2217   have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))"
  2218   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2219     fix x
  2220     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2221     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2222     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2223       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2224     thus "real c * real x + Inum (real x # bs) e < 0" 
  2225       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2226   qed
  2227   thus ?case by blast
  2228 next
  2229   case (6 c e) 
  2230   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2231   from prems have nbe: "numbound0 e" by simp
  2232   have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))"
  2233   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2234     fix x
  2235     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2236     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2237     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2238       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2239     thus "real c * real x + Inum (real x # bs) e \<le> 0" 
  2240       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2241   qed
  2242   thus ?case by blast
  2243 next
  2244   case (7 c e) 
  2245   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2246   from prems have nbe: "numbound0 e" by simp
  2247   have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))"
  2248   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2249     fix x
  2250     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2251     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2252     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2253       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2254     thus "\<not> (real c * real x + Inum (real x # bs) e>0)" 
  2255       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2256   qed
  2257   thus ?case by blast
  2258 next
  2259   case (8 c e) 
  2260   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2261   from prems have nbe: "numbound0 e" by simp
  2262   have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))"
  2263   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2264     fix x
  2265     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2266     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2267     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2268       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2269     thus "\<not> real c * real x + Inum (real x # bs) e \<ge> 0" 
  2270       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2271   qed
  2272   thus ?case by blast
  2273 qed simp_all
  2274 
  2275 lemma minusinf_repeats:
  2276   assumes d: "d\<delta> p d" and linp: "iszlfm p (a # bs)"
  2277   shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)"
  2278 using linp d
  2279 proof(induct p rule: iszlfm.induct) 
  2280   case (9 i c e) hence nbe: "numbound0 e"  and id: "i dvd d" by simp+
  2281     hence "\<exists> k. d=i*k" by (simp add: dvd_def)
  2282     then obtain "di" where di_def: "d=i*di" by blast
  2283     show ?case 
  2284     proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI)
  2285       assume 
  2286 	"real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
  2287       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
  2288       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
  2289       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
  2290 	by (simp add: ring_simps di_def)
  2291       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
  2292 	by (simp add: ring_simps)
  2293       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
  2294       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
  2295     next
  2296       assume 
  2297 	"real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
  2298       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
  2299       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
  2300       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
  2301       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps)
  2302       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
  2303 	by blast
  2304       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
  2305     qed
  2306 next
  2307   case (10 i c e) hence nbe: "numbound0 e"  and id: "i dvd d" by simp+
  2308     hence "\<exists> k. d=i*k" by (simp add: dvd_def)
  2309     then obtain "di" where di_def: "d=i*di" by blast
  2310     show ?case 
  2311     proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI)
  2312       assume 
  2313 	"real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
  2314       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
  2315       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
  2316       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
  2317 	by (simp add: ring_simps di_def)
  2318       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
  2319 	by (simp add: ring_simps)
  2320       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
  2321       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
  2322     next
  2323       assume 
  2324 	"real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
  2325       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
  2326       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
  2327       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
  2328       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps)
  2329       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
  2330 	by blast
  2331       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
  2332     qed
  2333 qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="real(x - k*d)" and b'="real x"] simp del: real_of_int_mult real_of_int_diff)
  2334 
  2335 lemma minusinf_ex:
  2336   assumes lin: "iszlfm p (real (a::int) #bs)"
  2337   and exmi: "\<exists> (x::int). Ifm (real x#bs) (minusinf p)" (is "\<exists> x. ?P1 x")
  2338   shows "\<exists> (x::int). Ifm (real x#bs) p" (is "\<exists> x. ?P x")
  2339 proof-
  2340   let ?d = "\<delta> p"
  2341   from \<delta> [OF lin] have dpos: "?d >0" by simp
  2342   from \<delta> [OF lin] have alld: "d\<delta> p ?d" by simp
  2343   from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P1 x = ?P1 (x - (k * ?d))" by simp
  2344   from minusinf_inf[OF lin] have th2:"\<exists> z. \<forall> x. x<z \<longrightarrow> (?P x = ?P1 x)" by blast
  2345   from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast
  2346 qed
  2347 
  2348 lemma minusinf_bex:
  2349   assumes lin: "iszlfm p (real (a::int) #bs)"
  2350   shows "(\<exists> (x::int). Ifm (real x#bs) (minusinf p)) = 
  2351          (\<exists> (x::int)\<in> {1..\<delta> p}. Ifm (real x#bs) (minusinf p))"
  2352   (is "(\<exists> x. ?P x) = _")
  2353 proof-
  2354   let ?d = "\<delta> p"
  2355   from \<delta> [OF lin] have dpos: "?d >0" by simp
  2356   from \<delta> [OF lin] have alld: "d\<delta> p ?d" by simp
  2357   from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P x = ?P (x - (k * ?d))" by simp
  2358   from periodic_finite_ex[OF dpos th1] show ?thesis by blast
  2359 qed
  2360 
  2361 lemma dvd1_eq1: "x >0 \<Longrightarrow> (x::int) dvd 1 = (x = 1)" by auto
  2362 
  2363 consts 
  2364   a\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coeffitients of a formula *)
  2365   d\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*)
  2366   \<zeta>  :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*)
  2367   \<beta> :: "fm \<Rightarrow> num list"
  2368   \<alpha> :: "fm \<Rightarrow> num list"
  2369 
  2370 recdef a\<beta> "measure size"
  2371   "a\<beta> (And p q) = (\<lambda> k. And (a\<beta> p k) (a\<beta> q k))" 
  2372   "a\<beta> (Or p q) = (\<lambda> k. Or (a\<beta> p k) (a\<beta> q k))" 
  2373   "a\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))"
  2374   "a\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))"
  2375   "a\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))"
  2376   "a\<beta> (Le  (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))"
  2377   "a\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))"
  2378   "a\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))"
  2379   "a\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
  2380   "a\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
  2381   "a\<beta> p = (\<lambda> k. p)"
  2382 
  2383 recdef d\<beta> "measure size"
  2384   "d\<beta> (And p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" 
  2385   "d\<beta> (Or p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" 
  2386   "d\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2387   "d\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2388   "d\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2389   "d\<beta> (Le  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2390   "d\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2391   "d\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2392   "d\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)"
  2393   "d\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)"
  2394   "d\<beta> p = (\<lambda> k. True)"
  2395 
  2396 recdef \<zeta> "measure size"
  2397   "\<zeta> (And p q) = ilcm (\<zeta> p) (\<zeta> q)" 
  2398   "\<zeta> (Or p q) = ilcm (\<zeta> p) (\<zeta> q)" 
  2399   "\<zeta> (Eq  (CN 0 c e)) = c"
  2400   "\<zeta> (NEq (CN 0 c e)) = c"
  2401   "\<zeta> (Lt  (CN 0 c e)) = c"
  2402   "\<zeta> (Le  (CN 0 c e)) = c"
  2403   "\<zeta> (Gt  (CN 0 c e)) = c"
  2404   "\<zeta> (Ge  (CN 0 c e)) = c"
  2405   "\<zeta> (Dvd i (CN 0 c e)) = c"
  2406   "\<zeta> (NDvd i (CN 0 c e))= c"
  2407   "\<zeta> p = 1"
  2408 
  2409 recdef \<beta> "measure size"
  2410   "\<beta> (And p q) = (\<beta> p @ \<beta> q)" 
  2411   "\<beta> (Or p q) = (\<beta> p @ \<beta> q)" 
  2412   "\<beta> (Eq  (CN 0 c e)) = [Sub (C -1) e]"
  2413   "\<beta> (NEq (CN 0 c e)) = [Neg e]"
  2414   "\<beta> (Lt  (CN 0 c e)) = []"
  2415   "\<beta> (Le  (CN 0 c e)) = []"
  2416   "\<beta> (Gt  (CN 0 c e)) = [Neg e]"
  2417   "\<beta> (Ge  (CN 0 c e)) = [Sub (C -1) e]"
  2418   "\<beta> p = []"
  2419 
  2420 recdef \<alpha> "measure size"
  2421   "\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)" 
  2422   "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)" 
  2423   "\<alpha> (Eq  (CN 0 c e)) = [Add (C -1) e]"
  2424   "\<alpha> (NEq (CN 0 c e)) = [e]"
  2425   "\<alpha> (Lt  (CN 0 c e)) = [e]"
  2426   "\<alpha> (Le  (CN 0 c e)) = [Add (C -1) e]"
  2427   "\<alpha> (Gt  (CN 0 c e)) = []"
  2428   "\<alpha> (Ge  (CN 0 c e)) = []"
  2429   "\<alpha> p = []"
  2430 consts mirror :: "fm \<Rightarrow> fm"
  2431 recdef mirror "measure size"
  2432   "mirror (And p q) = And (mirror p) (mirror q)" 
  2433   "mirror (Or p q) = Or (mirror p) (mirror q)" 
  2434   "mirror (Eq  (CN 0 c e)) = Eq (CN 0 c (Neg e))"
  2435   "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
  2436   "mirror (Lt  (CN 0 c e)) = Gt (CN 0 c (Neg e))"
  2437   "mirror (Le  (CN 0 c e)) = Ge (CN 0 c (Neg e))"
  2438   "mirror (Gt  (CN 0 c e)) = Lt (CN 0 c (Neg e))"
  2439   "mirror (Ge  (CN 0 c e)) = Le (CN 0 c (Neg e))"
  2440   "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
  2441   "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
  2442   "mirror p = p"
  2443 
  2444 lemma mirror\<alpha>\<beta>:
  2445   assumes lp: "iszlfm p (a#bs)"
  2446   shows "(Inum (real (i::int)#bs)) ` set (\<alpha> p) = (Inum (real i#bs)) ` set (\<beta> (mirror p))"
  2447 using lp
  2448 by (induct p rule: mirror.induct, auto)
  2449 
  2450 lemma mirror: 
  2451   assumes lp: "iszlfm p (a#bs)"
  2452   shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p" 
  2453 using lp
  2454 proof(induct p rule: iszlfm.induct)
  2455   case (9 j c e)
  2456   have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
  2457        (real j rdvd - (real c * real x - Inum (real x # bs) e))"
  2458     by (simp only: rdvd_minus[symmetric])
  2459   from prems show  ?case
  2460     by (simp add: ring_simps th[simplified ring_simps]
  2461       numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
  2462 next
  2463     case (10 j c e)
  2464   have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
  2465        (real j rdvd - (real c * real x - Inum (real x # bs) e))"
  2466     by (simp only: rdvd_minus[symmetric])
  2467   from prems show  ?case
  2468     by (simp add: ring_simps th[simplified ring_simps]
  2469       numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
  2470 qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"] nth_pos2)
  2471 
  2472 lemma mirror_l: "iszlfm p (a#bs) \<Longrightarrow> iszlfm (mirror p) (a#bs)"
  2473 by (induct p rule: mirror.induct, auto simp add: isint_neg)
  2474 
  2475 lemma mirror_d\<beta>: "iszlfm p (a#bs) \<and> d\<beta> p 1 
  2476   \<Longrightarrow> iszlfm (mirror p) (a#bs) \<and> d\<beta> (mirror p) 1"
  2477 by (induct p rule: mirror.induct, auto simp add: isint_neg)
  2478 
  2479 lemma mirror_\<delta>: "iszlfm p (a#bs) \<Longrightarrow> \<delta> (mirror p) = \<delta> p"
  2480 by (induct p rule: mirror.induct,auto)
  2481 
  2482 
  2483 lemma mirror_ex: 
  2484   assumes lp: "iszlfm p (real (i::int)#bs)"
  2485   shows "(\<exists> (x::int). Ifm (real x#bs) (mirror p)) = (\<exists> (x::int). Ifm (real x#bs) p)"
  2486   (is "(\<exists> x. ?I x ?mp) = (\<exists> x. ?I x p)")
  2487 proof(auto)
  2488   fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast
  2489   thus "\<exists> x. ?I x p" by blast
  2490 next
  2491   fix x assume "?I x p" hence "?I (- x) ?mp" 
  2492     using mirror[OF lp, where x="- x", symmetric] by auto
  2493   thus "\<exists> x. ?I x ?mp" by blast
  2494 qed
  2495 
  2496 lemma \<beta>_numbound0: assumes lp: "iszlfm p bs"
  2497   shows "\<forall> b\<in> set (\<beta> p). numbound0 b"
  2498   using lp by (induct p rule: \<beta>.induct,auto)
  2499 
  2500 lemma d\<beta>_mono: 
  2501   assumes linp: "iszlfm p (a #bs)"
  2502   and dr: "d\<beta> p l"
  2503   and d: "l dvd l'"
  2504   shows "d\<beta> p l'"
  2505 using dr linp zdvd_trans[where n="l" and k="l'", simplified d]
  2506 by (induct p rule: iszlfm.induct) simp_all
  2507 
  2508 lemma \<alpha>_l: assumes lp: "iszlfm p (a#bs)"
  2509   shows "\<forall> b\<in> set (\<alpha> p). numbound0 b \<and> isint b (a#bs)"
  2510 using lp
  2511 by(induct p rule: \<alpha>.induct, auto simp add: isint_add isint_c)
  2512 
  2513 lemma \<zeta>: 
  2514   assumes linp: "iszlfm p (a #bs)"
  2515   shows "\<zeta> p > 0 \<and> d\<beta> p (\<zeta> p)"
  2516 using linp
  2517 proof(induct p rule: iszlfm.induct)
  2518   case (1 p q)
  2519   from prems have dl1: "\<zeta> p dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2520   from prems have dl2: "\<zeta> q dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2521   from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2522     d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2523     dl1 dl2 show ?case by (auto simp add: ilcm_pos)
  2524 next
  2525   case (2 p q)
  2526   from prems have dl1: "\<zeta> p dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2527   from prems have dl2: "\<zeta> q dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2528   from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2529     d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2530     dl1 dl2 show ?case by (auto simp add: ilcm_pos)
  2531 qed (auto simp add: ilcm_pos)
  2532 
  2533 lemma a\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d\<beta> p l" and lp: "l > 0"
  2534   shows "iszlfm (a\<beta> p l) (a #bs) \<and> d\<beta> (a\<beta> p l) 1 \<and> (Ifm (real (l * x) #bs) (a\<beta> p l) = Ifm ((real x)#bs) p)"
  2535 using linp d
  2536 proof (induct p rule: iszlfm.induct)
  2537   case (5 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
  2538     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2539     from cp have cnz: "c \<noteq> 0" by simp
  2540     have "c div c\<le> l div c"
  2541       by (simp add: zdiv_mono1[OF clel cp])
  2542     then have ldcp:"0 < l div c" 
  2543       by (simp add: zdiv_self[OF cnz])
  2544     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2545     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2546       by simp
  2547     hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\<Colon>real)) =
  2548           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)"
  2549       by simp
  2550     also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: ring_simps)
  2551     also have "\<dots> = (real c * real x + Inum (real x # bs) e < 0)"
  2552     using mult_less_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
  2553   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be  isint_Mul[OF ei] by simp
  2554 next
  2555   case (6 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
  2556     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2557     from cp have cnz: "c \<noteq> 0" by simp
  2558     have "c div c\<le> l div c"
  2559       by (simp add: zdiv_mono1[OF clel cp])
  2560     then have ldcp:"0 < l div c" 
  2561       by (simp add: zdiv_self[OF cnz])
  2562     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2563     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2564       by simp
  2565     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0\<Colon>real)) =
  2566           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<le> 0)"
  2567       by simp
  2568     also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<le> (real (l div c)) * 0)" by (simp add: ring_simps)
  2569     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<le> 0)"
  2570     using mult_le_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
  2571   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
  2572 next
  2573   case (7 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
  2574     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2575     from cp have cnz: "c \<noteq> 0" by simp
  2576     have "c div c\<le> l div c"
  2577       by (simp add: zdiv_mono1[OF clel cp])
  2578     then have ldcp:"0 < l div c" 
  2579       by (simp add: zdiv_self[OF cnz])
  2580     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2581     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2582       by simp
  2583     hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\<Colon>real)) =
  2584           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)"
  2585       by simp
  2586     also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: ring_simps)
  2587     also have "\<dots> = (real c * real x + Inum (real x # bs) e > 0)"
  2588     using zero_less_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
  2589   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
  2590 next
  2591   case (8 c e) hence cp: "c>0" and be: "numbound0 e"  and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
  2592     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2593     from cp have cnz: "c \<noteq> 0" by simp
  2594     have "c div c\<le> l div c"
  2595       by (simp add: zdiv_mono1[OF clel cp])
  2596     then have ldcp:"0 < l div c" 
  2597       by (simp add: zdiv_self[OF cnz])
  2598     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2599     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2600       by simp
  2601     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0\<Colon>real)) =
  2602           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<ge> 0)"
  2603       by simp
  2604     also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<ge> (real (l div c)) * 0)" by (simp add: ring_simps)
  2605     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<ge> 0)"
  2606     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
  2607   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
  2608 next
  2609   case (3 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
  2610     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2611     from cp have cnz: "c \<noteq> 0" by simp
  2612     have "c div c\<le> l div c"
  2613       by (simp add: zdiv_mono1[OF clel cp])
  2614     then have ldcp:"0 < l div c" 
  2615       by (simp add: zdiv_self[OF cnz])
  2616     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2617     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2618       by simp
  2619     hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\<Colon>real)) =
  2620           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)"
  2621       by simp
  2622     also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: ring_simps)
  2623     also have "\<dots> = (real c * real x + Inum (real x # bs) e = 0)"
  2624     using mult_eq_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
  2625   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
  2626 next
  2627   case (4 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
  2628     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2629     from cp have cnz: "c \<noteq> 0" by simp
  2630     have "c div c\<le> l div c"
  2631       by (simp add: zdiv_mono1[OF clel cp])
  2632     then have ldcp:"0 < l div c" 
  2633       by (simp add: zdiv_self[OF cnz])
  2634     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2635     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2636       by simp
  2637     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0\<Colon>real)) =
  2638           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<noteq> 0)"
  2639       by simp
  2640     also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<noteq> (real (l div c)) * 0)" by (simp add: ring_simps)
  2641     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<noteq> 0)"
  2642     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
  2643   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"]  be  isint_Mul[OF ei] by simp
  2644 next
  2645   case (9 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+
  2646     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2647     from cp have cnz: "c \<noteq> 0" by simp
  2648     have "c div c\<le> l div c"
  2649       by (simp add: zdiv_mono1[OF clel cp])
  2650     then have ldcp:"0 < l div c" 
  2651       by (simp add: zdiv_self[OF cnz])
  2652     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2653     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2654       by simp
  2655     hence "(\<exists> (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\<exists> (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)"  by simp
  2656     also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps)
  2657     also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
  2658     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp
  2659   also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
  2660   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def  be  isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp 
  2661 next
  2662   case (10 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+
  2663     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2664     from cp have cnz: "c \<noteq> 0" by simp
  2665     have "c div c\<le> l div c"
  2666       by (simp add: zdiv_mono1[OF clel cp])
  2667     then have ldcp:"0 < l div c" 
  2668       by (simp add: zdiv_self[OF cnz])
  2669     have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp
  2670     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2671       by simp
  2672     hence "(\<exists> (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\<exists> (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)"  by simp
  2673     also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: ring_simps)
  2674     also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
  2675     using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp
  2676   also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
  2677   finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def  be  isint_Mul[OF ei]  mult_strict_mono[OF ldcp jp ldcp ] by simp
  2678 qed (simp_all add: nth_pos2 numbound0_I[where bs="bs" and b="real (l * x)" and b'="real x"] isint_Mul del: real_of_int_mult)
  2679 
  2680 lemma a\<beta>_ex: assumes linp: "iszlfm p (a#bs)" and d: "d\<beta> p l" and lp: "l>0"
  2681   shows "(\<exists> x. l dvd x \<and> Ifm (real x #bs) (a\<beta> p l)) = (\<exists> (x::int). Ifm (real x#bs) p)"
  2682   (is "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> x. ?P' x)")
  2683 proof-
  2684   have "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> (x::int). ?P (l*x))"
  2685     using unity_coeff_ex[where l="l" and P="?P", simplified] by simp
  2686   also have "\<dots> = (\<exists> (x::int). ?P' x)" using a\<beta>[OF linp d lp] by simp
  2687   finally show ?thesis  . 
  2688 qed
  2689 
  2690 lemma \<beta>:
  2691   assumes lp: "iszlfm p (a#bs)"
  2692   and u: "d\<beta> p 1"
  2693   and d: "d\<delta> p d"
  2694   and dp: "d > 0"
  2695   and nob: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
  2696   and p: "Ifm (real x#bs) p" (is "?P x")
  2697   shows "?P (x - d)"
  2698 using lp u d dp nob p
  2699 proof(induct p rule: iszlfm.induct)
  2700   case (5 c e) hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
  2701     with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems
  2702     show ?case by (simp del: real_of_int_minus)
  2703 next
  2704   case (6 c e)  hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
  2705     with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems
  2706     show ?case by (simp del: real_of_int_minus)
  2707 next
  2708   case (7 c e) hence p: "Ifm (real x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp+
  2709     let ?e = "Inum (real x # bs) e"
  2710     from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"]
  2711       numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]
  2712       by (simp add: isint_iff)
  2713     {assume "real (x-d) +?e > 0" hence ?case using c1 
  2714       numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
  2715 	by (simp del: real_of_int_minus)}
  2716     moreover
  2717     {assume H: "\<not> real (x-d) + ?e > 0" 
  2718       let ?v="Neg e"
  2719       have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp
  2720       from prems(11)[simplified simp_thms Inum.simps \<beta>.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] 
  2721       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e + real j)" by auto 
  2722       from H p have "real x + ?e > 0 \<and> real x + ?e \<le> real d" by (simp add: c1)
  2723       hence "real (x + floor ?e) > real (0::int) \<and> real (x + floor ?e) \<le> real d"
  2724 	using ie by simp
  2725       hence "x + floor ?e \<ge> 1 \<and> x + floor ?e \<le> d"  by simp
  2726       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e" by simp
  2727       hence "\<exists> (j::int) \<in> {1 .. d}. real x = real (- floor ?e + j)" 
  2728 	by (simp only: real_of_int_inject) (simp add: ring_simps)
  2729       hence "\<exists> (j::int) \<in> {1 .. d}. real x = - ?e + real j" 
  2730 	by (simp add: ie[simplified isint_iff])
  2731       with nob have ?case by auto}
  2732     ultimately show ?case by blast
  2733 next
  2734   case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" 
  2735     and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
  2736     let ?e = "Inum (real x # bs) e"
  2737     from ie1 have ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"]
  2738       by (simp add: isint_iff)
  2739     {assume "real (x-d) +?e \<ge> 0" hence ?case using  c1 
  2740       numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
  2741 	by (simp del: real_of_int_minus)}
  2742     moreover
  2743     {assume H: "\<not> real (x-d) + ?e \<ge> 0" 
  2744       let ?v="Sub (C -1) e"
  2745       have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp
  2746       from prems(11)[simplified simp_thms Inum.simps \<beta>.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] 
  2747       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e - 1 + real j)" by auto 
  2748       from H p have "real x + ?e \<ge> 0 \<and> real x + ?e < real d" by (simp add: c1)
  2749       hence "real (x + floor ?e) \<ge> real (0::int) \<and> real (x + floor ?e) < real d"
  2750 	using ie by simp
  2751       hence "x + floor ?e +1 \<ge> 1 \<and> x + floor ?e + 1 \<le> d"  by simp
  2752       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e + 1" by simp
  2753       hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: ring_simps)
  2754       hence "\<exists> (j::int) \<in> {1 .. d}. real x= real (- floor ?e - 1 + j)" 
  2755 	by (simp only: real_of_int_inject)
  2756       hence "\<exists> (j::int) \<in> {1 .. d}. real x= - ?e - 1 + real j" 
  2757 	by (simp add: ie[simplified isint_iff])
  2758       with nob have ?case by simp }
  2759     ultimately show ?case by blast
  2760 next
  2761   case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" 
  2762     and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
  2763     let ?e = "Inum (real x # bs) e"
  2764     let ?v="(Sub (C -1) e)"
  2765     have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
  2766     from p have "real x= - ?e" by (simp add: c1) with prems(11) show ?case using dp
  2767       by simp (erule ballE[where x="1"],
  2768 	simp_all add:ring_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
  2769 next
  2770   case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" 
  2771     and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
  2772     let ?e = "Inum (real x # bs) e"
  2773     let ?v="Neg e"
  2774     have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" by simp
  2775     {assume "real x - real d + Inum ((real (x -d)) # bs) e \<noteq> 0" 
  2776       hence ?case by (simp add: c1)}
  2777     moreover
  2778     {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0"
  2779       hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp
  2780       hence "real x = - Inum (a # bs) e + real d"
  2781 	by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"])
  2782        with prems(11) have ?case using dp by simp}
  2783   ultimately show ?case by blast
  2784 next 
  2785   case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" 
  2786     and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
  2787     let ?e = "Inum (real x # bs) e"
  2788     from prems have "isint e (a #bs)"  by simp 
  2789     hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real x)#bs"] numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"]
  2790       by (simp add: isint_iff)
  2791     from prems have id: "j dvd d" by simp
  2792     from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp
  2793     also have "\<dots> = (j dvd x + floor ?e)" 
  2794       using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
  2795     also have "\<dots> = (j dvd x - d + floor ?e)" 
  2796       using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
  2797     also have "\<dots> = (real j rdvd real (x - d + floor ?e))" 
  2798       using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
  2799       ie by simp
  2800     also have "\<dots> = (real j rdvd real x - real d + ?e)" 
  2801       using ie by simp
  2802     finally show ?case 
  2803       using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
  2804 next
  2805   case (10 j c e) hence p: "Ifm (real x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
  2806     let ?e = "Inum (real x # bs) e"
  2807     from prems have "isint e (a#bs)"  by simp 
  2808     hence ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"]
  2809       by (simp add: isint_iff)
  2810     from prems have id: "j dvd d" by simp
  2811     from c1 ie[symmetric] have "?p x = (\<not> real j rdvd real (x+ floor ?e))" by simp
  2812     also have "\<dots> = (\<not> j dvd x + floor ?e)" 
  2813       using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
  2814     also have "\<dots> = (\<not> j dvd x - d + floor ?e)" 
  2815       using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
  2816     also have "\<dots> = (\<not> real j rdvd real (x - d + floor ?e))" 
  2817       using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
  2818       ie by simp
  2819     also have "\<dots> = (\<not> real j rdvd real x - real d + ?e)" 
  2820       using ie by simp
  2821     finally show ?case using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
  2822 qed (auto simp add: numbound0_I[where bs="bs" and b="real (x - d)" and b'="real x"] nth_pos2 simp del: real_of_int_diff)
  2823 
  2824 lemma \<beta>':   
  2825   assumes lp: "iszlfm p (a #bs)"
  2826   and u: "d\<beta> p 1"
  2827   and d: "d\<delta> p d"
  2828   and dp: "d > 0"
  2829   shows "\<forall> x. \<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> set(\<beta> p). Ifm ((Inum (a#bs) b + real j) #bs) p) \<longrightarrow> Ifm (real x#bs) p \<longrightarrow> Ifm (real (x - d)#bs) p" (is "\<forall> x. ?b \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)")
  2830 proof(clarify)
  2831   fix x 
  2832   assume nb:"?b" and px: "?P x" 
  2833   hence nb2: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
  2834     by auto
  2835   from  \<beta>[OF lp u d dp nb2 px] show "?P (x -d )" .
  2836 qed
  2837 
  2838 lemma \<beta>_int: assumes lp: "iszlfm p bs"
  2839   shows "\<forall> b\<in> set (\<beta> p). isint b bs"
  2840 using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub)
  2841 
  2842 lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x))
  2843 ==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) 
  2844 ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D))))
  2845 ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
  2846 apply(rule iffI)
  2847 prefer 2
  2848 apply(drule minusinfinity)
  2849 apply assumption+
  2850 apply(fastsimp)
  2851 apply clarsimp
  2852 apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)")
  2853 apply(frule_tac x = x and z=z in decr_lemma)
  2854 apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
  2855 prefer 2
  2856 apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
  2857 prefer 2 apply arith
  2858  apply fastsimp
  2859 apply(drule (1)  periodic_finite_ex)
  2860 apply blast
  2861 apply(blast dest:decr_mult_lemma)
  2862 done
  2863 
  2864 
  2865 theorem cp_thm:
  2866   assumes lp: "iszlfm p (a #bs)"
  2867   and u: "d\<beta> p 1"
  2868   and d: "d\<delta> p d"
  2869   and dp: "d > 0"
  2870   shows "(\<exists> (x::int). Ifm (real x #bs) p) = (\<exists> j\<in> {1.. d}. Ifm (real j #bs) (minusinf p) \<or> (\<exists> b \<in> set (\<beta> p). Ifm ((Inum (a#bs) b + real j) #bs) p))"
  2871   (is "(\<exists> (x::int). ?P (real x)) = (\<exists> j\<in> ?D. ?M j \<or> (\<exists> b\<in> ?B. ?P (?I b + real j)))")
  2872 proof-
  2873   from minusinf_inf[OF lp] 
  2874   have th: "\<exists>(z::int). \<forall>x<z. ?P (real x) = ?M x" by blast
  2875   let ?B' = "{floor (?I b) | b. b\<in> ?B}"
  2876   from \<beta>_int[OF lp] isint_iff[where bs="a # bs"] have B: "\<forall> b\<in> ?B. real (floor (?I b)) = ?I b" by simp
  2877   from B[rule_format] 
  2878   have "(\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (?I b + real j)) = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real (floor (?I b)) + real j))" 
  2879     by simp
  2880   also have "\<dots> = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real (floor (?I b) + j)))" by simp
  2881   also have"\<dots> = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j)))"  by blast
  2882   finally have BB': 
  2883     "(\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (?I b + real j)) = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j)))" 
  2884     by blast 
  2885   hence th2: "\<forall> x. \<not> (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j))) \<longrightarrow> ?P (real x) \<longrightarrow> ?P (real (x - d))" using \<beta>'[OF lp u d dp] by blast
  2886   from minusinf_repeats[OF d lp]
  2887   have th3: "\<forall> x k. ?M x = ?M (x-k*d)" by simp
  2888   from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast
  2889 qed
  2890 
  2891     (* Reddy and Loveland *)
  2892 
  2893 
  2894 consts 
  2895   \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*)
  2896   \<sigma>\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*)
  2897   \<alpha>\<rho> :: "fm \<Rightarrow> (num\<times>int) list"
  2898   a\<rho> :: "fm \<Rightarrow> int \<Rightarrow> fm"
  2899 recdef \<rho> "measure size"
  2900   "\<rho> (And p q) = (\<rho> p @ \<rho> q)" 
  2901   "\<rho> (Or p q) = (\<rho> p @ \<rho> q)" 
  2902   "\<rho> (Eq  (CN 0 c e)) = [(Sub (C -1) e,c)]"
  2903   "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]"
  2904   "\<rho> (Lt  (CN 0 c e)) = []"
  2905   "\<rho> (Le  (CN 0 c e)) = []"
  2906   "\<rho> (Gt  (CN 0 c e)) = [(Neg e, c)]"
  2907   "\<rho> (Ge  (CN 0 c e)) = [(Sub (C (-1)) e, c)]"
  2908   "\<rho> p = []"
  2909 
  2910 recdef \<sigma>\<rho> "measure size"
  2911   "\<sigma>\<rho> (And p q) = (\<lambda> (t,k). And (\<sigma>\<rho> p (t,k)) (\<sigma>\<rho> q (t,k)))" 
  2912   "\<sigma>\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>\<rho> p (t,k)) (\<sigma>\<rho> q (t,k)))" 
  2913   "\<sigma>\<rho> (Eq  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) 
  2914                                             else (Eq (Add (Mul c t) (Mul k e))))"
  2915   "\<sigma>\<rho> (NEq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) 
  2916                                             else (NEq (Add (Mul c t) (Mul k e))))"
  2917   "\<sigma>\<rho> (Lt  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) 
  2918                                             else (Lt (Add (Mul c t) (Mul k e))))"
  2919   "\<sigma>\<rho> (Le  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) 
  2920                                             else (Le (Add (Mul c t) (Mul k e))))"
  2921   "\<sigma>\<rho> (Gt  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) 
  2922                                             else (Gt (Add (Mul c t) (Mul k e))))"
  2923   "\<sigma>\<rho> (Ge  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) 
  2924                                             else (Ge (Add (Mul c t) (Mul k e))))"
  2925   "\<sigma>\<rho> (Dvd i (CN 0 c e)) =(\<lambda> (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e)) 
  2926                                             else (Dvd (i*k) (Add (Mul c t) (Mul k e))))"
  2927   "\<sigma>\<rho> (NDvd i (CN 0 c e))=(\<lambda> (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e)) 
  2928                                             else (NDvd (i*k) (Add (Mul c t) (Mul k e))))"
  2929   "\<sigma>\<rho> p = (\<lambda> (t,k). p)"
  2930 
  2931 recdef \<alpha>\<rho> "measure size"
  2932   "\<alpha>\<rho> (And p q) = (\<alpha>\<rho> p @ \<alpha>\<rho> q)" 
  2933   "\<alpha>\<rho> (Or p q) = (\<alpha>\<rho> p @ \<alpha>\<rho> q)" 
  2934   "\<alpha>\<rho> (Eq  (CN 0 c e)) = [(Add (C -1) e,c)]"
  2935   "\<alpha>\<rho> (NEq (CN 0 c e)) = [(e,c)]"
  2936   "\<alpha>\<rho> (Lt  (CN 0 c e)) = [(e,c)]"
  2937   "\<alpha>\<rho> (Le  (CN 0 c e)) = [(Add (C -1) e,c)]"
  2938   "\<alpha>\<rho> p = []"
  2939 
  2940     (* Simulates normal substituion by modifying the formula see correctness theorem *)
  2941 
  2942 recdef a\<rho> "measure size"
  2943   "a\<rho> (And p q) = (\<lambda> k. And (a\<rho> p k) (a\<rho> q k))" 
  2944   "a\<rho> (Or p q) = (\<lambda> k. Or (a\<rho> p k) (a\<rho> q k))" 
  2945   "a\<rho> (Eq (CN 0 c e)) = (\<lambda> k. if k dvd c then (Eq (CN 0 (c div k) e)) 
  2946                                            else (Eq (CN 0 c (Mul k e))))"
  2947   "a\<rho> (NEq (CN 0 c e)) = (\<lambda> k. if k dvd c then (NEq (CN 0 (c div k) e)) 
  2948                                            else (NEq (CN 0 c (Mul k e))))"
  2949   "a\<rho> (Lt (CN 0 c e)) = (\<lambda> k. if k dvd c then (Lt (CN 0 (c div k) e)) 
  2950                                            else (Lt (CN 0 c (Mul k e))))"
  2951   "a\<rho> (Le (CN 0 c e)) = (\<lambda> k. if k dvd c then (Le (CN 0 (c div k) e)) 
  2952                                            else (Le (CN 0 c (Mul k e))))"
  2953   "a\<rho> (Gt (CN 0 c e)) = (\<lambda> k. if k dvd c then (Gt (CN 0 (c div k) e)) 
  2954                                            else (Gt (CN 0 c (Mul k e))))"
  2955   "a\<rho> (Ge (CN 0 c e)) = (\<lambda> k. if k dvd c then (Ge (CN 0 (c div k) e)) 
  2956                                             else (Ge (CN 0 c (Mul k e))))"
  2957   "a\<rho> (Dvd i (CN 0 c e)) = (\<lambda> k. if k dvd c then (Dvd i (CN 0 (c div k) e)) 
  2958                                             else (Dvd (i*k) (CN 0 c (Mul k e))))"
  2959   "a\<rho> (NDvd i (CN 0 c e)) = (\<lambda> k. if k dvd c then (NDvd i (CN 0 (c div k) e)) 
  2960                                             else (NDvd (i*k) (CN 0 c (Mul k e))))"
  2961   "a\<rho> p = (\<lambda> k. p)"
  2962 
  2963 constdefs \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
  2964   "\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>\<rho> p (t,k))"
  2965 
  2966 lemma \<sigma>\<rho>:
  2967   assumes linp: "iszlfm p (real (x::int)#bs)"
  2968   and kpos: "real k > 0"
  2969   and tnb: "numbound0 t"
  2970   and tint: "isint t (real x#bs)"
  2971   and kdt: "k dvd floor (Inum (b'#bs) t)"
  2972   shows "Ifm (real x#bs) (\<sigma>\<rho> p (t,k)) = 
  2973   (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)" 
  2974   (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)")
  2975 using linp kpos tnb
  2976 proof(induct p rule: \<sigma>\<rho>.induct)
  2977   case (3 c e) 
  2978   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  2979     {assume kdc: "k dvd c" 
  2980       from kpos have knz: "k\<noteq>0" by simp
  2981       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  2982       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  2983 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  2984       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  2985     moreover 
  2986     {assume "\<not> k dvd c"
  2987       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  2988       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  2989       from prems have "?I (real x) (?s (Eq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k = 0)"
  2990 	using real_of_int_div[OF knz kdt]
  2991 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  2992 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  2993       also have "\<dots> = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  2994 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  2995 	by (simp add: ti)
  2996       finally have ?case . }
  2997     ultimately show ?case by blast 
  2998 next
  2999   case (4 c e)  
  3000   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3001     {assume kdc: "k dvd c" 
  3002       from kpos have knz: "k\<noteq>0" by simp
  3003       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3004       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3005 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3006       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3007     moreover 
  3008     {assume "\<not> k dvd c"
  3009       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3010       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3011       from prems have "?I (real x) (?s (NEq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<noteq> 0)"
  3012 	using real_of_int_div[OF knz kdt]
  3013 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3014 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3015       also have "\<dots> = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3016 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3017 	by (simp add: ti)
  3018       finally have ?case . }
  3019     ultimately show ?case by blast 
  3020 next
  3021   case (5 c e) 
  3022   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3023     {assume kdc: "k dvd c" 
  3024       from kpos have knz: "k\<noteq>0" by simp
  3025       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3026       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3027 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3028       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3029     moreover 
  3030     {assume "\<not> k dvd c"
  3031       from kpos have knz: "k\<noteq>0" by simp
  3032       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3033       from prems have "?I (real x) (?s (Lt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k < 0)"
  3034 	using real_of_int_div[OF knz kdt]
  3035 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3036 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3037       also have "\<dots> = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3038 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3039 	by (simp add: ti)
  3040       finally have ?case . }
  3041     ultimately show ?case by blast 
  3042 next
  3043   case (6 c e)  
  3044   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3045     {assume kdc: "k dvd c" 
  3046       from kpos have knz: "k\<noteq>0" by simp
  3047       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3048       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3049 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3050       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3051     moreover 
  3052     {assume "\<not> k dvd c"
  3053       from kpos have knz: "k\<noteq>0" by simp
  3054       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3055       from prems have "?I (real x) (?s (Le (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<le> 0)"
  3056 	using real_of_int_div[OF knz kdt]
  3057 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3058 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3059       also have "\<dots> = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3060 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3061 	by (simp add: ti)
  3062       finally have ?case . }
  3063     ultimately show ?case by blast 
  3064 next
  3065   case (7 c e) 
  3066   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3067     {assume kdc: "k dvd c" 
  3068       from kpos have knz: "k\<noteq>0" by simp
  3069       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3070       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3071 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3072       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3073     moreover 
  3074     {assume "\<not> k dvd c"
  3075       from kpos have knz: "k\<noteq>0" by simp
  3076       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3077       from prems have "?I (real x) (?s (Gt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k > 0)"
  3078 	using real_of_int_div[OF knz kdt]
  3079 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3080 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3081       also have "\<dots> = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3082 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3083 	by (simp add: ti)
  3084       finally have ?case . }
  3085     ultimately show ?case by blast 
  3086 next
  3087   case (8 c e)  
  3088   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3089     {assume kdc: "k dvd c" 
  3090       from kpos have knz: "k\<noteq>0" by simp
  3091       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3092       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3093 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3094       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3095     moreover 
  3096     {assume "\<not> k dvd c"
  3097       from kpos have knz: "k\<noteq>0" by simp
  3098       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3099       from prems have "?I (real x) (?s (Ge (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<ge> 0)"
  3100 	using real_of_int_div[OF knz kdt]
  3101 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3102 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3103       also have "\<dots> = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3104 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3105 	by (simp add: ti)
  3106       finally have ?case . }
  3107     ultimately show ?case by blast 
  3108 next
  3109   case (9 i c e)   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3110     {assume kdc: "k dvd c" 
  3111       from kpos have knz: "k\<noteq>0" by simp
  3112       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3113       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3114 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3115       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3116     moreover 
  3117     {assume "\<not> k dvd c"
  3118       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3119       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3120       from prems have "?I (real x) (?s (Dvd i (CN 0 c e))) = (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k)"
  3121 	using real_of_int_div[OF knz kdt]
  3122 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3123 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3124       also have "\<dots> = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3125 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3126 	by (simp add: ti)
  3127       finally have ?case . }
  3128     ultimately show ?case by blast 
  3129 next
  3130   case (10 i c e)    from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3131     {assume kdc: "k dvd c" 
  3132       from kpos have knz: "k\<noteq>0" by simp
  3133       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3134       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3135 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3136       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3137     moreover 
  3138     {assume "\<not> k dvd c"
  3139       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3140       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3141       from prems have "?I (real x) (?s (NDvd i (CN 0 c e))) = (\<not> (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k))"
  3142 	using real_of_int_div[OF knz kdt]
  3143 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3144 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3145       also have "\<dots> = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3146 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3147 	by (simp add: ti)
  3148       finally have ?case . }
  3149     ultimately show ?case by blast 
  3150 qed (simp_all add: nth_pos2 bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"] numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"])
  3151 
  3152 
  3153 lemma a\<rho>: 
  3154   assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "real k > 0" 
  3155   shows "Ifm (real (x*k)#bs) (a\<rho> p k) = Ifm (real x#bs) p" (is "?I (x*k) (?f p k) = ?I x p")
  3156 using lp bound0_I[where bs="bs" and b="real (x*k)" and b'="real x"] numbound0_I[where bs="bs" and b="real (x*k)" and b'="real x"]
  3157 proof(induct p rule: a\<rho>.induct)
  3158   case (3 c e)  
  3159   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3160   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3161     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3162     moreover 
  3163     {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
  3164     ultimately show ?case by blast 
  3165 next
  3166   case (4 c e)   
  3167   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3168   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3169     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3170     moreover 
  3171     {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
  3172     ultimately show ?case by blast 
  3173 next
  3174   case (5 c e)   
  3175   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3176   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3177     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3178     moreover 
  3179     {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_less_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
  3180     ultimately show ?case by blast 
  3181 next
  3182   case (6 c e)    
  3183   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3184   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3185     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3186     moreover 
  3187     {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_le_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
  3188     ultimately show ?case by blast 
  3189 next
  3190   case (7 c e)    
  3191   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3192   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3193     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3194     moreover 
  3195     {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_less_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
  3196     ultimately show ?case by blast 
  3197 next
  3198   case (8 c e)    
  3199   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3200   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3201     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3202     moreover 
  3203     {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_le_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: ring_simps)}
  3204     ultimately show ?case by blast 
  3205 next
  3206   case (9 i c e)
  3207   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3208   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3209   {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3210   moreover 
  3211   {assume "\<not> k dvd c"
  3212     hence "Ifm (real (x*k)#bs) (a\<rho> (Dvd i (CN 0 c e)) k) = 
  3213       (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k)" 
  3214       using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] 
  3215       by (simp add: ring_simps)
  3216     also have "\<dots> = (Ifm (real x#bs) (Dvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"])
  3217     finally have ?case . }
  3218   ultimately show ?case by blast 
  3219 next
  3220   case (10 i c e) 
  3221   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3222   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3223   {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3224   moreover 
  3225   {assume "\<not> k dvd c"
  3226     hence "Ifm (real (x*k)#bs) (a\<rho> (NDvd i (CN 0 c e)) k) = 
  3227       (\<not> (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k))" 
  3228       using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] 
  3229       by (simp add: ring_simps)
  3230     also have "\<dots> = (Ifm (real x#bs) (NDvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"])
  3231     finally have ?case . }
  3232   ultimately show ?case by blast 
  3233 qed (simp_all add: nth_pos2)
  3234 
  3235 lemma a\<rho>_ex: 
  3236   assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0"
  3237   shows "(\<exists> (x::int). real k rdvd real x \<and> Ifm (real x#bs) (a\<rho> p k)) = 
  3238   (\<exists> (x::int). Ifm (real x#bs) p)" (is "(\<exists> x. ?D x \<and> ?P' x) = (\<exists> x. ?P x)")
  3239 proof-
  3240   have "(\<exists> x. ?D x \<and> ?P' x) = (\<exists> x. k dvd x \<and> ?P' x)" using int_rdvd_iff by simp
  3241   also have "\<dots> = (\<exists>x. ?P' (x*k))" using unity_coeff_ex[where P="?P'" and l="k", simplified]
  3242     by (simp add: ring_simps)
  3243   also have "\<dots> = (\<exists> x. ?P x)" using a\<rho> iszlfm_gen[OF lp] kp by auto
  3244   finally show ?thesis .
  3245 qed
  3246 
  3247 lemma \<sigma>\<rho>': assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" and nb: "numbound0 t"
  3248   shows "Ifm (real x#bs) (\<sigma>\<rho> p (t,k)) = Ifm ((Inum (real x#bs) t)#bs) (a\<rho> p k)"
  3249 using lp 
  3250 by(induct p rule: \<sigma>\<rho>.induct, simp_all add: 
  3251   numbound0_I[OF nb, where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] 
  3252   numbound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] 
  3253   bound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] nth_pos2 cong: imp_cong)
  3254 
  3255 lemma \<sigma>\<rho>_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t"
  3256   shows "bound0 (\<sigma>\<rho> p (t,k))"
  3257   using lp
  3258   by (induct p rule: iszlfm.induct, auto simp add: nb)
  3259 
  3260 lemma \<rho>_l:
  3261   assumes lp: "iszlfm p (real (i::int)#bs)"
  3262   shows "\<forall> (b,k) \<in> set (\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
  3263 using lp by (induct p rule: \<rho>.induct, auto simp add: isint_sub isint_neg)
  3264 
  3265 lemma \<alpha>\<rho>_l:
  3266   assumes lp: "iszlfm p (real (i::int)#bs)"
  3267   shows "\<forall> (b,k) \<in> set (\<alpha>\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
  3268 using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"]
  3269  by (induct p rule: \<alpha>\<rho>.induct, auto)
  3270 
  3271 lemma zminusinf_\<rho>:
  3272   assumes lp: "iszlfm p (real (i::int)#bs)"
  3273   and nmi: "\<not> (Ifm (real i#bs) (minusinf p))" (is "\<not> (Ifm (real i#bs) (?M p))")
  3274   and ex: "Ifm (real i#bs) p" (is "?I i p")
  3275   shows "\<exists> (e,c) \<in> set (\<rho> p). real (c*i) > Inum (real i#bs) e" (is "\<exists> (e,c) \<in> ?R p. real (c*i) > ?N i e")
  3276   using lp nmi ex
  3277 by (induct p rule: minusinf.induct, auto)
  3278 
  3279 
  3280 lemma \<sigma>_And: "Ifm bs (\<sigma> (And p q) k t)  = Ifm bs (And (\<sigma> p k t) (\<sigma> q k t))"
  3281 using \<sigma>_def by auto
  3282 lemma \<sigma>_Or: "Ifm bs (\<sigma> (Or p q) k t)  = Ifm bs (Or (\<sigma> p k t) (\<sigma> q k t))"
  3283 using \<sigma>_def by auto
  3284 
  3285 lemma \<rho>: assumes lp: "iszlfm p (real (i::int) #bs)"
  3286   and pi: "Ifm (real i#bs) p"
  3287   and d: "d\<delta> p d"
  3288   and dp: "d > 0"
  3289   and nob: "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> Inum (real i#bs) e + real j"
  3290   (is "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. _ \<noteq> ?N i e + _")
  3291   shows "Ifm (real(i - d)#bs) p"
  3292   using lp pi d nob
  3293 proof(induct p rule: iszlfm.induct)
  3294   case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3295     and pi: "real (c*i) = - 1 -  ?N i e + real (1::int)" and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> -1 - ?N i e + real j"
  3296     by simp+
  3297   from mult_strict_left_mono[OF dp cp]  have one:"1 \<in> {1 .. c*d}" by auto
  3298   from nob[rule_format, where j="1", OF one] pi show ?case by simp
  3299 next
  3300   case (4 c e)  
  3301   hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3302     and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
  3303     by simp+
  3304   {assume "real (c*i) \<noteq> - ?N i e + real (c*d)"
  3305     with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"]
  3306     have ?case by (simp add: ring_simps)}
  3307   moreover
  3308   {assume pi: "real (c*i) = - ?N i e + real (c*d)"
  3309     from mult_strict_left_mono[OF dp cp] have d: "(c*d) \<in> {1 .. c*d}" by simp
  3310     from nob[rule_format, where j="c*d", OF d] pi have ?case by simp }
  3311   ultimately show ?case by blast
  3312 next
  3313   case (5 c e) hence cp: "c > 0" by simp
  3314   from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] 
  3315     real_of_int_mult]
  3316   show ?case using prems dp 
  3317     by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
  3318       ring_simps)
  3319 next
  3320   case (6 c e)  hence cp: "c > 0" by simp
  3321   from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] 
  3322     real_of_int_mult]
  3323   show ?case using prems dp 
  3324     by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
  3325       ring_simps)
  3326 next
  3327   case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3328     and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
  3329     and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0"
  3330     by simp+
  3331   let ?fe = "floor (?N i e)"
  3332   from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: ring_simps)
  3333   from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp
  3334   hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric])
  3335   have "real (c*i) + ?N i e > real (c*d) \<or> real (c*i) + ?N i e \<le> real (c*d)" by auto
  3336   moreover
  3337   {assume "real (c*i) + ?N i e > real (c*d)" hence ?case
  3338       by (simp add: ring_simps 
  3339 	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
  3340   moreover 
  3341   {assume H:"real (c*i) + ?N i e \<le> real (c*d)"
  3342     with ei[simplified isint_iff] have "real (c*i + ?fe) \<le> real (c*d)" by simp
  3343     hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
  3344     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
  3345     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1" 
  3346       by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps)
  3347     with nob  have ?case by blast }
  3348   ultimately show ?case by blast
  3349 next
  3350   case (8 c e)  hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3351     and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - 1 - ?N i e + real j"
  3352     and pi: "real (c*i) + ?N i e \<ge> 0" and cp': "real c >0"
  3353     by simp+
  3354   let ?fe = "floor (?N i e)"
  3355   from pi cp have th:"(real i +?N i e / real c)*real c \<ge> 0" by (simp add: ring_simps)
  3356   from pi ei[simplified isint_iff] have "real (c*i + ?fe) \<ge> real (0::int)" by simp
  3357   hence pi': "c*i + 1 + ?fe \<ge> 1" by (simp only: real_of_int_le_iff[symmetric])
  3358   have "real (c*i) + ?N i e \<ge> real (c*d) \<or> real (c*i) + ?N i e < real (c*d)" by auto
  3359   moreover
  3360   {assume "real (c*i) + ?N i e \<ge> real (c*d)" hence ?case
  3361       by (simp add: ring_simps 
  3362 	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
  3363   moreover 
  3364   {assume H:"real (c*i) + ?N i e < real (c*d)"
  3365     with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp
  3366     hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
  3367     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
  3368     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
  3369       by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] ring_simps real_of_one) 
  3370     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
  3371       by (simp only: ring_simps diff_def[symmetric])
  3372         hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
  3373 	  by (simp only: add_ac diff_def)
  3374     with nob  have ?case by blast }
  3375   ultimately show ?case by blast
  3376 next
  3377   case (9 j c e)  hence p: "real j rdvd real (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e"  by simp+
  3378     let ?e = "Inum (real i # bs) e"
  3379     from prems have "isint e (real i #bs)"  by simp 
  3380     hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"]
  3381       by (simp add: isint_iff)
  3382     from prems have id: "j dvd d" by simp
  3383     from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp
  3384     also have "\<dots> = (j dvd c*i + floor ?e)" 
  3385       using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
  3386     also have "\<dots> = (j dvd c*i - c*d + floor ?e)" 
  3387       using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
  3388     also have "\<dots> = (real j rdvd real (c*i - c*d + floor ?e))" 
  3389       using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
  3390       ie by simp
  3391     also have "\<dots> = (real j rdvd real (c*(i - d)) + ?e)" 
  3392       using ie by (simp add:ring_simps)
  3393     finally show ?case 
  3394       using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
  3395       by (simp add: ring_simps)
  3396 next
  3397   case (10 j c e)   hence p: "\<not> (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e"  by simp+
  3398     let ?e = "Inum (real i # bs) e"
  3399     from prems have "isint e (real i #bs)"  by simp 
  3400     hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"]
  3401       by (simp add: isint_iff)
  3402     from prems have id: "j dvd d" by simp
  3403     from ie[symmetric] have "?p i = (\<not> (real j rdvd real (c*i+ floor ?e)))" by simp
  3404     also have "\<dots> = Not (j dvd c*i + floor ?e)" 
  3405       using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
  3406     also have "\<dots> = Not (j dvd c*i - c*d + floor ?e)" 
  3407       using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
  3408     also have "\<dots> = Not (real j rdvd real (c*i - c*d + floor ?e))" 
  3409       using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
  3410       ie by simp
  3411     also have "\<dots> = Not (real j rdvd real (c*(i - d)) + ?e)" 
  3412       using ie by (simp add:ring_simps)
  3413     finally show ?case 
  3414       using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
  3415       by (simp add: ring_simps)
  3416 qed(auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] nth_pos2)
  3417 
  3418 lemma \<sigma>_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t"
  3419   shows "bound0 (\<sigma> p k t)"
  3420   using \<sigma>\<rho>_nb[OF lp nb] nb by (simp add: \<sigma>_def)
  3421   
  3422 lemma \<rho>':   assumes lp: "iszlfm p (a #bs)"
  3423   and d: "d\<delta> p d"
  3424   and dp: "d > 0"
  3425   shows "\<forall> x. \<not>(\<exists> (e,c) \<in> set(\<rho> p). \<exists>(j::int) \<in> {1 .. c*d}. Ifm (a #bs) (\<sigma> p c (Add e (C j)))) \<longrightarrow> Ifm (real x#bs) p \<longrightarrow> Ifm (real (x - d)#bs) p" (is "\<forall> x. ?b x \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)")
  3426 proof(clarify)
  3427   fix x 
  3428   assume nob1:"?b x" and px: "?P x" 
  3429   from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)".
  3430   have nob: "\<forall>(e, c)\<in>set (\<rho> p). \<forall>j\<in>{1..c * d}. real (c * x) \<noteq> Inum (real x # bs) e + real j" 
  3431   proof(clarify)
  3432     fix e c j assume ecR: "(e,c) \<in> set (\<rho> p)" and jD: "j\<in> {1 .. c*d}"
  3433       and cx: "real (c*x) = Inum (real x#bs) e + real j"
  3434     let ?e = "Inum (real x#bs) e"
  3435     let ?fe = "floor ?e"
  3436     from \<rho>_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e"
  3437       by auto
  3438     from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" .
  3439     from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp
  3440     hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject)
  3441     hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp)
  3442     hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff)
  3443     hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff])
  3444     from cx have "(c*x) div c = (?fe + j) div c" by simp
  3445     with cp have "x = (?fe + j) div c" by simp
  3446     with px have th: "?P ((?fe + j) div c)" by auto
  3447     from cp have cp': "real c > 0" by simp
  3448     from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp
  3449     from nb have nb': "numbound0 (Add e (C j))" by simp
  3450     have ji: "isint (C j) (real x#bs)" by (simp add: isint_def)
  3451     from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" .
  3452     from th \<sigma>\<rho>[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric]
  3453     have "Ifm (real x#bs) (\<sigma>\<rho> p (Add e (C j), c))" by simp
  3454     with rcdej have th: "Ifm (real x#bs) (\<sigma> p c (Add e (C j)))" by (simp add: \<sigma>_def)
  3455     from th bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"]
  3456     have "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" by blast
  3457       with ecR jD nob1    show "False" by blast
  3458   qed
  3459   from \<rho>[OF lp' px d dp nob] show "?P (x -d )" . 
  3460 qed
  3461 
  3462 
  3463 lemma rl_thm: 
  3464   assumes lp: "iszlfm p (real (i::int)#bs)"
  3465   shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> (e,c) \<in> set (\<rho> p). \<exists> j\<in> {1 .. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))))"
  3466   (is "(\<exists>(x::int). ?P x) = ((\<exists> j\<in> {1.. \<delta> p}. ?MP j)\<or>(\<exists> (e,c) \<in> ?R. \<exists> j\<in> _. ?SP c e j))" 
  3467     is "?lhs = (?MD \<or> ?RD)"  is "?lhs = ?rhs")
  3468 proof-
  3469   let ?d= "\<delta> p"
  3470   from \<delta>[OF lp] have d:"d\<delta> p ?d" and dp: "?d > 0" by auto
  3471   { assume H:"?MD" hence th:"\<exists> (x::int). ?MP x" by blast
  3472     from H minusinf_ex[OF lp th] have ?thesis  by blast}
  3473   moreover
  3474   { fix e c j assume exR:"(e,c) \<in> ?R" and jD:"j\<in> {1 .. c*?d}" and spx:"?SP c e j"
  3475     from exR \<rho>_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0"
  3476       by auto
  3477     have "isint (C j) (real i#bs)" by (simp add: isint_iff)
  3478     with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]]
  3479     have eji:"isint (Add e (C j)) (real i#bs)" by simp
  3480     from nb have nb': "numbound0 (Add e (C j))" by simp
  3481     from spx bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"]
  3482     have spx': "Ifm (real i # bs) (\<sigma> p c (Add e (C j)))" by blast
  3483     from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))" 
  3484       and sr:"Ifm (real i#bs) (\<sigma>\<rho> p (Add e (C j),c))" by (simp add: \<sigma>_def)+
  3485     from rcdej eji[simplified isint_iff] 
  3486     have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp
  3487     hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff)
  3488     from cp have cp': "real c > 0" by simp
  3489     from \<sigma>\<rho>[OF lp cp' nb' eji cdej] spx' have "?P (\<lfloor>Inum (real i # bs) (Add e (C j))\<rfloor> div c)"
  3490       by (simp add: \<sigma>_def)
  3491     hence ?lhs by blast
  3492     with exR jD spx have ?thesis by blast}
  3493   moreover
  3494   { fix x assume px: "?P x" and nob: "\<not> ?RD"
  3495     from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" .
  3496     from \<rho>'[OF lp' d dp, rule_format, OF nob] have th:"\<forall> x. ?P x \<longrightarrow> ?P (x - ?d)" by blast
  3497     from minusinf_inf[OF lp] obtain z where z:"\<forall> x<z. ?MP x = ?P x" by blast
  3498     have zp: "abs (x - z) + 1 \<ge> 0" by arith
  3499     from decr_lemma[OF dp,where x="x" and z="z"] 
  3500       decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\<exists> x. ?MP x" by auto
  3501     with minusinf_bex[OF lp] px nob have ?thesis by blast}
  3502   ultimately show ?thesis by blast
  3503 qed
  3504 
  3505 lemma mirror_\<alpha>\<rho>:   assumes lp: "iszlfm p (a#bs)"
  3506   shows "(\<lambda> (t,k). (Inum (a#bs) t, k)) ` set (\<alpha>\<rho> p) = (\<lambda> (t,k). (Inum (a#bs) t,k)) ` set (\<rho> (mirror p))"
  3507 using lp
  3508 by (induct p rule: mirror.induct, simp_all add: split_def image_Un )
  3509   
  3510 text {* The @{text "\<real>"} part*}
  3511 
  3512 text{* Linearity for fm where Bound 0 ranges over @{text "\<real>"}*}
  3513 consts
  3514   isrlfm :: "fm \<Rightarrow> bool"   (* Linearity test for fm *)
  3515 recdef isrlfm "measure size"
  3516   "isrlfm (And p q) = (isrlfm p \<and> isrlfm q)" 
  3517   "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)" 
  3518   "isrlfm (Eq  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3519   "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3520   "isrlfm (Lt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3521   "isrlfm (Le  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3522   "isrlfm (Gt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3523   "isrlfm (Ge  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3524   "isrlfm p = (isatom p \<and> (bound0 p))"
  3525 
  3526 constdefs fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm"
  3527   "fp p n s j \<equiv> (if n > 0 then 
  3528             (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j)))))
  3529                         (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1))))))))
  3530             else 
  3531             (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) 
  3532                         (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))"
  3533 
  3534   (* splits the bounded from the unbounded part*)
  3535 consts rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list" 
  3536 recdef rsplit0 "measure num_size"
  3537   "rsplit0 (Bound 0) = [(T,1,C 0)]"
  3538   "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b 
  3539               in map (\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) (allpairs Pair acs bcs))"
  3540   "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))"
  3541   "rsplit0 (Neg a) = map (\<lambda> (p,n,s). (p,-n,Neg s)) (rsplit0 a)"
  3542   "rsplit0 (Floor a) = foldl (op @) [] (map 
  3543       (\<lambda> (p,n,s). if n=0 then [(p,0,Floor s)]
  3544           else (map (\<lambda> j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then iupt (0,n) else iupt(n,0))))
  3545        (rsplit0 a))"
  3546   "rsplit0 (CN 0 c a) = map (\<lambda> (p,n,s). (p,n+c,s)) (rsplit0 a)"
  3547   "rsplit0 (CN m c a) = map (\<lambda> (p,n,s). (p,n,CN m c s)) (rsplit0 a)"
  3548   "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)"
  3549   "rsplit0 (Mul c a) = map (\<lambda> (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)"
  3550   "rsplit0 t = [(T,0,t)]"
  3551 
  3552 lemma not_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm (not p)"
  3553   by (induct p rule: isrlfm.induct, auto)
  3554 lemma conj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (conj p q)"
  3555   using conj_def by (cases p, auto)
  3556 lemma disj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (disj p q)"
  3557   using disj_def by (cases p, auto)
  3558 
  3559 
  3560 lemma rsplit0_cs:
  3561   shows "\<forall> (p,n,s) \<in> set (rsplit0 t). 
  3562   (Ifm (x#bs) p \<longrightarrow>  (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p" 
  3563   (is "\<forall> (p,n,s) \<in> ?SS t. (?I p \<longrightarrow> ?N t = ?N (CN 0 n s)) \<and> _ \<and> _ ")
  3564 proof(induct t rule: rsplit0.induct)
  3565   case (5 a) 
  3566   let ?p = "\<lambda> (p,n,s) j. fp p n s j"
  3567   let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))"
  3568   let ?J = "\<lambda> n. if n>0 then iupt (0,n) else iupt (n,0)"
  3569   let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
  3570   have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
  3571   have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
  3572     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto
  3573   have U2': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0}. 
  3574     ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))" by auto
  3575   hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
  3576     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). 
  3577     set (map (?f(p,n,s)) (iupt(0,n)))))"
  3578   proof-
  3579     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3580     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3581     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3582       by (auto simp add: split_def)
  3583   qed
  3584   have U3': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(n,0))"
  3585     by auto
  3586   hence U3: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
  3587     (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n<0} (\<lambda>(p,n,s). set (map (?f(p,n,s)) (iupt(n,0)))))"
  3588       proof-
  3589     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3590     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3591     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3592       by (auto simp add: split_def)
  3593   qed
  3594   from foldl_append_map_Nil_set[where xs="rsplit0 a" and f="?ff"]
  3595   have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" by auto
  3596   also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by auto
  3597   also have "\<dots> = 
  3598     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3599     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3600     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" 
  3601     using int_cases[rule_format] by blast
  3602   also have "\<dots> =  
  3603     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un 
  3604    (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n>0} (\<lambda>(p,n,s). set(map(?f(p,n,s)) (iupt(0,n))))) Un 
  3605    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). 
  3606     set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3)
  3607   also have "\<dots> =  
  3608     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3609     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
  3610     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
  3611     by (simp only: set_map iupt_set set.simps)
  3612   also have "\<dots> =   
  3613     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3614     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un 
  3615     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s).  {?f(p,n,s) j| j. j\<in> {n .. 0}})))" by blast
  3616   finally 
  3617   have FS: "?SS (Floor a) =   
  3618     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3619     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un 
  3620     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s).  {?f(p,n,s) j| j. j\<in> {n .. 0}})))"    by blast
  3621   show ?case
  3622     proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -)
  3623       fix p n s
  3624       let ?ths = "(?I p \<longrightarrow> (?N (Floor a) = ?N (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p"
  3625       assume "(\<exists>ba. (p, 0, ba) \<in> set (rsplit0 a) \<and> n = 0 \<and> s = Floor ba) \<or>
  3626        (\<exists>ab ac ba.
  3627            (ab, ac, ba) \<in> set (rsplit0 a) \<and>
  3628            0 < ac \<and>
  3629            (\<exists>j. p = fp ab ac ba j \<and>
  3630                 n = 0 \<and> s = Add (Floor ba) (C j) \<and> 0 \<le> j \<and> j \<le> ac)) \<or>
  3631        (\<exists>ab ac ba.
  3632            (ab, ac, ba) \<in> set (rsplit0 a) \<and>
  3633            ac < 0 \<and>
  3634            (\<exists>j. p = fp ab ac ba j \<and>
  3635                 n = 0 \<and> s = Add (Floor ba) (C j) \<and> ac \<le> j \<and> j \<le> 0))"
  3636       moreover 
  3637       {fix s'
  3638 	assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'"
  3639 	hence ?ths using prems by auto}
  3640       moreover
  3641       {	fix p' n' s' j
  3642 	assume pns: "(p', n', s') \<in> ?SS a" 
  3643 	  and np: "0 < n'" 
  3644 	  and p_def: "p = ?p (p',n',s') j" 
  3645 	  and n0: "n = 0" 
  3646 	  and s_def: "s = (Add (Floor s') (C j))" 
  3647 	  and jp: "0 \<le> j" and jn: "j \<le> n'"
  3648 	from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
  3649           Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
  3650           numbound0 s' \<and> isrlfm p'" by blast
  3651 	hence nb: "numbound0 s'" by simp
  3652 	from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb)
  3653 	let ?nxs = "CN 0 n' s'"
  3654 	let ?l = "floor (?N s') + j"
  3655 	from H 
  3656 	have "?I (?p (p',n',s') j) \<longrightarrow> 
  3657 	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
  3658 	  by (simp add: fp_def np ring_simps numsub numadd numfloor)
  3659 	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
  3660 	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
  3661 	moreover
  3662 	have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
  3663 	ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
  3664 	  by blast
  3665 	with s_def n0 p_def nb nf have ?ths by auto}
  3666       moreover
  3667       {fix p' n' s' j
  3668 	assume pns: "(p', n', s') \<in> ?SS a" 
  3669 	  and np: "n' < 0" 
  3670 	  and p_def: "p = ?p (p',n',s') j" 
  3671 	  and n0: "n = 0" 
  3672 	  and s_def: "s = (Add (Floor s') (C j))" 
  3673 	  and jp: "n' \<le> j" and jn: "j \<le> 0"
  3674 	from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
  3675           Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
  3676           numbound0 s' \<and> isrlfm p'" by blast
  3677 	hence nb: "numbound0 s'" by simp
  3678 	from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb)
  3679 	let ?nxs = "CN 0 n' s'"
  3680 	let ?l = "floor (?N s') + j"
  3681 	from H 
  3682 	have "?I (?p (p',n',s') j) \<longrightarrow> 
  3683 	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
  3684 	  by (simp add: np fp_def ring_simps numneg numfloor numadd numsub)
  3685 	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
  3686 	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
  3687 	moreover
  3688 	have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"  by simp
  3689 	ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
  3690 	  by blast
  3691 	with s_def n0 p_def nb nf have ?ths by auto}
  3692       ultimately show ?ths by auto
  3693     qed
  3694 next
  3695   case (3 a b) thus ?case by auto 
  3696 qed (auto simp add: Let_def allpairs_set split_def ring_simps conj_rl)
  3697 
  3698 lemma real_in_int_intervals: 
  3699   assumes xb: "real m \<le> x \<and> x < real ((n::int) + 1)"
  3700   shows "\<exists> j\<in> {m.. n}. real j \<le> x \<and> x < real (j+1)" (is "\<exists> j\<in> ?N. ?P j")
  3701 by (rule bexI[where P="?P" and x="floor x" and A="?N"]) 
  3702 (auto simp add: floor_less_eq[where x="x" and a="n+1", simplified] xb[simplified] floor_mono2[where x="real m" and y="x", OF conjunct1[OF xb], simplified floor_real_of_int[where n="m"]])
  3703 
  3704 lemma rsplit0_complete:
  3705   assumes xp:"0 \<le> x" and x1:"x < 1"
  3706   shows "\<exists> (p,n,s) \<in> set (rsplit0 t). Ifm (x#bs) p" (is "\<exists> (p,n,s) \<in> ?SS t. ?I p")
  3707 proof(induct t rule: rsplit0.induct)
  3708   case (2 a b) 
  3709   from prems have "\<exists> (pa,na,sa) \<in> ?SS a. ?I pa" by auto
  3710   then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\<in> ?SS a \<and> ?I pa" by blast
  3711   from prems have "\<exists> (pb,nb,sb) \<in> ?SS b. ?I pb" by auto
  3712   then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\<in> ?SS b \<and> ?I pb" by blast
  3713   from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \<in> set (allpairs Pair (rsplit0 a) (rsplit0 b))"
  3714     by (auto simp add: allpairs_set)
  3715   let ?f="(\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))"
  3716   from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \<in> ?SS (Add a b)"
  3717     by (simp add: Let_def)
  3718   hence "(And pa pb, na +nb, Add sa sb) \<in> ?SS (Add a b)" by simp
  3719   moreover from pa pb have "?I (And pa pb)" by simp
  3720   ultimately show ?case by blast
  3721 next
  3722   case (5 a) 
  3723   let ?p = "\<lambda> (p,n,s) j. fp p n s j"
  3724   let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))"
  3725   let ?J = "\<lambda> n. if n>0 then iupt (0,n) else iupt (n,0)"
  3726   let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
  3727   have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
  3728   have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto
  3729   have U2': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))"
  3730     by auto
  3731   hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) (iupt(0,n)))))"
  3732   proof-
  3733     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3734     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3735     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3736       by (auto simp add: split_def)
  3737   qed
  3738   have U3': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0}. ?ff (p,n,s) = map (?f(p,n,s)) (iupt(n,0))"
  3739     by auto
  3740   hence U3: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) (iupt(n,0)))))"
  3741   proof-
  3742     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3743     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3744     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3745       by (auto simp add: split_def)
  3746   qed
  3747   from foldl_append_map_Nil_set[where xs="rsplit0 a" and f="?ff"]
  3748   have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" by auto
  3749   also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by auto
  3750   also have "\<dots> = 
  3751     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3752     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3753     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" 
  3754     using int_cases[rule_format] by blast
  3755   also have "\<dots> =  
  3756     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un 
  3757     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) (iupt(0,n))))) Un 
  3758     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3)
  3759   also have "\<dots> =  
  3760     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3761     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
  3762     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
  3763     by (simp only: set_map iupt_set set.simps)
  3764   also have "\<dots> =   
  3765     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3766     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un 
  3767     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s).  {?f(p,n,s) j| j. j\<in> {n .. 0}})))" by blast
  3768   finally 
  3769   have FS: "?SS (Floor a) =   
  3770     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3771     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un 
  3772     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s).  {?f(p,n,s) j| j. j\<in> {n .. 0}})))"    by blast
  3773   from prems have "\<exists> (p,n,s) \<in> ?SS a. ?I p" by auto
  3774   then obtain "p" "n" "s" where pns: "(p,n,s) \<in> ?SS a \<and> ?I p" by blast
  3775   let ?N = "\<lambda> t. Inum (x#bs) t"
  3776   from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \<and> numbound0 s \<and> isrlfm p"
  3777     by auto
  3778   
  3779   have "n=0 \<or> n >0 \<or> n <0" by arith
  3780   moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto }
  3781   moreover
  3782   {
  3783     assume np: "n > 0"
  3784     from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \<le> ?N s" by simp
  3785     also from mult_left_mono[OF xp] np have "?N s \<le> real n * x + ?N s" by simp
  3786     finally have "?N (Floor s) \<le> real n * x + ?N s" .
  3787     moreover
  3788     {from mult_strict_left_mono[OF x1] np 
  3789       have "real n *x + ?N s < real n + ?N s" by simp
  3790       also from real_of_int_floor_add_one_gt[where r="?N s"] 
  3791       have "\<dots> < real n + ?N (Floor s) + 1" by simp
  3792       finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp}
  3793     ultimately have "?N (Floor s) \<le> real n *x + ?N s\<and> real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp
  3794     hence th: "0 \<le> real n *x + ?N s - ?N (Floor s) \<and> real n *x + ?N s - ?N (Floor s) < real (n+1)" by simp
  3795     from real_in_int_intervals th have  "\<exists> j\<in> {0 .. n}. real j \<le> real n *x + ?N s - ?N (Floor s)\<and> real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp
  3796     
  3797     hence "\<exists> j\<in> {0 .. n}. 0 \<le> real n *x + ?N s - ?N (Floor s) - real j \<and> real n *x + ?N s - ?N (Floor s) - real (j+1) < 0"
  3798       by(simp only: myl[rule_format, where b="real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) 
  3799     hence "\<exists> j\<in> {0.. n}. ?I (?p (p,n,s) j)"
  3800       using pns by (simp add: fp_def np ring_simps numsub numadd)
  3801     then obtain "j" where j_def: "j\<in> {0 .. n} \<and> ?I (?p (p,n,s) j)" by blast
  3802     hence "\<exists>x \<in> {?p (p,n,s) j |j. 0\<le> j \<and> j \<le> n }. ?I x" by auto
  3803     hence ?case using pns 
  3804       by (simp only: FS,simp add: bex_Un) 
  3805     (rule disjI2, rule disjI1,rule exI [where x="p"],
  3806       rule exI [where x="n"],rule exI [where x="s"],simp_all add: np)
  3807   }
  3808   moreover
  3809   { assume nn: "n < 0" hence np: "-n >0" by simp
  3810     from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp
  3811     moreover from mult_left_mono_neg[OF xp] nn have "?N s \<ge> real n * x + ?N s" by simp
  3812     ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith 
  3813     moreover
  3814     {from mult_strict_left_mono_neg[OF x1, where c="real n"] nn
  3815       have "real n *x + ?N s \<ge> real n + ?N s" by simp 
  3816       moreover from real_of_int_floor_le[where r="?N s"]  have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp
  3817       ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n" 
  3818 	by (simp only: ring_simps)}
  3819     ultimately have "?N (Floor s) + real n \<le> real n *x + ?N s\<and> real n *x + ?N s < ?N (Floor s) + real (1::int)" by simp
  3820     hence th: "real n \<le> real n *x + ?N s - ?N (Floor s) \<and> real n *x + ?N s - ?N (Floor s) < real (1::int)" by simp
  3821     have th1: "\<forall> (a::real). (- a > 0) = (a < 0)" by auto
  3822     have th2: "\<forall> (a::real). (0 \<ge> - a) = (a \<ge> 0)" by auto
  3823     from real_in_int_intervals th  have  "\<exists> j\<in> {n .. 0}. real j \<le> real n *x + ?N s - ?N (Floor s)\<and> real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp
  3824     
  3825     hence "\<exists> j\<in> {n .. 0}. 0 \<le> real n *x + ?N s - ?N (Floor s) - real j \<and> real n *x + ?N s - ?N (Floor s) - real (j+1) < 0"
  3826       by(simp only: myl[rule_format, where b="real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) 
  3827     hence "\<exists> j\<in> {n .. 0}. 0 \<ge> - (real n *x + ?N s - ?N (Floor s) - real j) \<and> - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format])
  3828     hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
  3829       using pns by (simp add: fp_def nn diff_def add_ac mult_ac numfloor numadd numneg
  3830 	del: diff_less_0_iff_less diff_le_0_iff_le) 
  3831     then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
  3832     hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
  3833     hence ?case using pns 
  3834       by (simp only: FS,simp add: bex_Un)
  3835     (rule disjI2, rule disjI2,rule exI [where x="p"],
  3836       rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn)
  3837   }
  3838   ultimately show ?case by blast
  3839 qed (auto simp add: Let_def split_def)
  3840 
  3841     (* Linearize a formula where Bound 0 ranges over [0,1) *)
  3842 
  3843 constdefs rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm"
  3844   "rsplit f a \<equiv> foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) (rsplit0 a)) F"
  3845 
  3846 lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\<exists> x \<in> set xs. Ifm bs (f x))"
  3847 by(induct xs, simp_all)
  3848 
  3849 lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\<forall> x \<in> set xs. Ifm bs (f x))"
  3850 by(induct xs, simp_all)
  3851 
  3852 lemma foldr_disj_map_rlfm: 
  3853   assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
  3854   and \<phi>: "\<forall> (\<phi>,n,s) \<in> set xs. numbound0 s \<and> isrlfm \<phi>"
  3855   shows "isrlfm (foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) xs) F)"
  3856 using lf \<phi> by (induct xs, auto)
  3857 
  3858 lemma rsplit_ex: "Ifm bs (rsplit f a) = (\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). Ifm bs (conj \<phi> (f n s)))"
  3859 using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def)
  3860 
  3861 lemma rsplit_l: assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
  3862   shows "isrlfm (rsplit f a)"
  3863 proof-
  3864   from rsplit0_cs[where t="a"] have th: "\<forall> (\<phi>,n,s) \<in> set (rsplit0 a). numbound0 s \<and> isrlfm \<phi>" by blast
  3865   from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp
  3866 qed
  3867 
  3868 lemma rsplit: 
  3869   assumes xp: "x \<ge> 0" and x1: "x < 1"
  3870   and f: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> (Ifm (x#bs) (f n s) = Ifm (x#bs) (g a))"
  3871   shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)"
  3872 proof(auto)
  3873   let ?I = "\<lambda>x p. Ifm (x#bs) p"
  3874   let ?N = "\<lambda> x t. Inum (x#bs) t"
  3875   assume "?I x (rsplit f a)"
  3876   hence "\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). ?I x (And \<phi> (f n s))" using rsplit_ex by simp
  3877   then obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and "?I x (And \<phi> (f n s))" by blast
  3878   hence \<phi>: "?I x \<phi>" and fns: "?I x (f n s)" by auto
  3879   from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \<phi> 
  3880   have th: "(?N x a = ?N x (CN 0 n s)) \<and> numbound0 s" by auto
  3881   from f[rule_format, OF th] fns show "?I x (g a)" by simp
  3882 next
  3883   let ?I = "\<lambda>x p. Ifm (x#bs) p"
  3884   let ?N = "\<lambda> x t. Inum (x#bs) t"
  3885   assume ga: "?I x (g a)"
  3886   from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] 
  3887   obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and fx: "?I x \<phi>" by blast
  3888   from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx
  3889   have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto
  3890   with ga f have "?I x (f n s)" by auto
  3891   with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto
  3892 qed
  3893 
  3894 definition lt :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3895   lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) 
  3896                         else (Gt (CN 0 (-c) (Neg t))))"
  3897 
  3898 definition  le :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3899   le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) 
  3900                         else (Ge (CN 0 (-c) (Neg t))))"
  3901 
  3902 definition  gt :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3903   gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) 
  3904                         else (Lt (CN 0 (-c) (Neg t))))"
  3905 
  3906 definition  ge :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3907   ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) 
  3908                         else (Le (CN 0 (-c) (Neg t))))"
  3909 
  3910 definition  eq :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3911   eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) 
  3912                         else (Eq (CN 0 (-c) (Neg t))))"
  3913 
  3914 definition neq :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3915   neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) 
  3916                         else (NEq (CN 0 (-c) (Neg t))))"
  3917 
  3918 lemma lt_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (lt n s) = Ifm (x#bs) (Lt a)"
  3919   (is "\<forall> a n s . ?N a = ?N (CN 0 n s) \<and> _\<longrightarrow> ?I (lt n s) = ?I (Lt a)")
  3920 proof(clarify)
  3921   fix a n s
  3922   assume H: "?N a = ?N (CN 0 n s)"
  3923   show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def))
  3924   (cases "n > 0", simp_all add: lt_def ring_simps myless[rule_format, where b="0"])
  3925 qed
  3926 
  3927 lemma lt_l: "isrlfm (rsplit lt a)"
  3928   by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def,
  3929     case_tac s, simp_all, case_tac "nat", simp_all)
  3930 
  3931 lemma le_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (le n s) = Ifm (x#bs) (Le a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (le n s) = ?I (Le a)")
  3932 proof(clarify)
  3933   fix a n s
  3934   assume H: "?N a = ?N (CN 0 n s)"
  3935   show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def))
  3936   (cases "n > 0", simp_all add: le_def ring_simps myl[rule_format, where b="0"])
  3937 qed
  3938 
  3939 lemma le_l: "isrlfm (rsplit le a)"
  3940   by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) 
  3941 (case_tac s, simp_all, case_tac "nat",simp_all)
  3942 
  3943 lemma gt_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (gt n s) = Ifm (x#bs) (Gt a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (gt n s) = ?I (Gt a)")
  3944 proof(clarify)
  3945   fix a n s
  3946   assume H: "?N a = ?N (CN 0 n s)"
  3947   show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def))
  3948   (cases "n > 0", simp_all add: gt_def ring_simps myless[rule_format, where b="0"])
  3949 qed
  3950 lemma gt_l: "isrlfm (rsplit gt a)"
  3951   by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) 
  3952 (case_tac s, simp_all, case_tac "nat", simp_all)
  3953 
  3954 lemma ge_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (ge n s) = Ifm (x#bs) (Ge a)" (is "\<forall> a n s . ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (ge n s) = ?I (Ge a)")
  3955 proof(clarify)
  3956   fix a n s 
  3957   assume H: "?N a = ?N (CN 0 n s)"
  3958   show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def))
  3959   (cases "n > 0", simp_all add: ge_def ring_simps myl[rule_format, where b="0"])
  3960 qed
  3961 lemma ge_l: "isrlfm (rsplit ge a)"
  3962   by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) 
  3963 (case_tac s, simp_all, case_tac "nat", simp_all)
  3964 
  3965 lemma eq_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (eq n s) = Ifm (x#bs) (Eq a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (eq n s) = ?I (Eq a)")
  3966 proof(clarify)
  3967   fix a n s 
  3968   assume H: "?N a = ?N (CN 0 n s)"
  3969   show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def ring_simps)
  3970 qed
  3971 lemma eq_l: "isrlfm (rsplit eq a)"
  3972   by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) 
  3973 (case_tac s, simp_all, case_tac"nat", simp_all)
  3974 
  3975 lemma neq_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (neq n s) = Ifm (x#bs) (NEq a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (neq n s) = ?I (NEq a)")
  3976 proof(clarify)
  3977   fix a n s bs
  3978   assume H: "?N a = ?N (CN 0 n s)"
  3979   show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def ring_simps)
  3980 qed
  3981 
  3982 lemma neq_l: "isrlfm (rsplit neq a)"
  3983   by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) 
  3984 (case_tac s, simp_all, case_tac"nat", simp_all)
  3985 
  3986 lemma small_le: 
  3987   assumes u0:"0 \<le> u" and u1: "u < 1"
  3988   shows "(-u \<le> real (n::int)) = (0 \<le> n)"
  3989 using u0 u1  by auto
  3990 
  3991 lemma small_lt: 
  3992   assumes u0:"0 \<le> u" and u1: "u < 1"
  3993   shows "(real (n::int) < real (m::int) - u) = (n < m)"
  3994 using u0 u1  by auto
  3995 
  3996 lemma rdvd01_cs: 
  3997   assumes up: "u \<ge> 0" and u1: "u<1" and np: "real n > 0"
  3998   shows "(real (i::int) rdvd real (n::int) * u - s) = (\<exists> j\<in> {0 .. n - 1}. real n * u = s - real (floor s) + real j \<and> real i rdvd real (j - floor s))" (is "?lhs = ?rhs")
  3999 proof-
  4000   let ?ss = "s - real (floor s)"
  4001   from real_of_int_floor_add_one_gt[where r="s", simplified myless[rule_format,where a="s"]] 
  4002     real_of_int_floor_le[where r="s"]  have ss0:"?ss \<ge> 0" and ss1:"?ss < 1" 
  4003     by (auto simp add: myl[rule_format, where b="s", symmetric] myless[rule_format, where a="?ss"])
  4004   from np have n0: "real n \<ge> 0" by simp
  4005   from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] 
  4006   have nu0:"real n * u - s \<ge> -s" and nun:"real n * u -s < real n - s" by auto  
  4007   from int_rdvd_real[where i="i" and x="real (n::int) * u - s"] 
  4008   have "real i rdvd real n * u - s = 
  4009     (i dvd floor (real n * u -s) \<and> (real (floor (real n * u - s)) = real n * u - s ))" 
  4010     (is "_ = (?DE)" is "_ = (?D \<and> ?E)") by simp
  4011   also have "\<dots> = (?DE \<and> real(floor (real n * u - s) + floor s)\<ge> -?ss 
  4012     \<and> real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \<and>real ?a \<ge> _ \<and> real ?a < _)")
  4013     using nu0 nun  by auto
  4014   also have "\<dots> = (?DE \<and> ?a \<ge> 0 \<and> ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1])
  4015   also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. ?a = j))" by simp
  4016   also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. real (\<lfloor>real n * u - s\<rfloor>) = real j - real \<lfloor>s\<rfloor> ))"
  4017     by (simp only: ring_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff)
  4018   also have "\<dots> = ((\<exists> j\<in> {0 .. (n - 1)}. real n * u - s = real j - real \<lfloor>s\<rfloor> \<and> real i rdvd real n * u - s))" using int_rdvd_iff[where i="i" and t="\<lfloor>real n * u - s\<rfloor>"]
  4019     by (auto cong: conj_cong)
  4020   also have "\<dots> = ?rhs" by(simp cong: conj_cong) (simp add: ring_simps )
  4021   finally show ?thesis .
  4022 qed
  4023 
  4024 definition
  4025   DVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
  4026 where
  4027   DVDJ_def: "DVDJ i n s = (foldr disj (map (\<lambda> j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))) (iupt(0,n - 1))) F)"
  4028 
  4029 definition
  4030   NDVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
  4031 where
  4032   NDVDJ_def: "NDVDJ i n s = (foldr conj (map (\<lambda> j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))) (iupt(0,n - 1))) T)"
  4033 
  4034 lemma DVDJ_DVD: 
  4035   assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
  4036   shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))"
  4037 proof-
  4038   let ?f = "\<lambda> j. conj (Eq(CN 0 n (Add s (Sub(Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))"
  4039   let ?s= "Inum (x#bs) s"
  4040   from foldr_disj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"]
  4041   have "Ifm (x#bs) (DVDJ i n s) = (\<exists> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
  4042     by (simp add: iupt_set np DVDJ_def del: iupt.simps)
  4043   also have "\<dots> = (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s)))" by (simp add: ring_simps diff_def[symmetric])
  4044   also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
  4045   have "\<dots> = (real i rdvd real n * x - (-?s))" by simp
  4046   finally show ?thesis by simp
  4047 qed
  4048 
  4049 lemma NDVDJ_NDVD: 
  4050   assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
  4051   shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))"
  4052 proof-
  4053   let ?f = "\<lambda> j. disj(NEq(CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor(Neg s))))"
  4054   let ?s= "Inum (x#bs) s"
  4055   from foldr_conj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"]
  4056   have "Ifm (x#bs) (NDVDJ i n s) = (\<forall> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
  4057     by (simp add: iupt_set np NDVDJ_def del: iupt.simps)
  4058   also have "\<dots> = (\<not> (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s))))" by (simp add: ring_simps diff_def[symmetric])
  4059   also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
  4060   have "\<dots> = (\<not> (real i rdvd real n * x - (-?s)))" by simp
  4061   finally show ?thesis by simp
  4062 qed  
  4063 
  4064 lemma foldr_disj_map_rlfm2: 
  4065   assumes lf: "\<forall> n . isrlfm (f n)"
  4066   shows "isrlfm (foldr disj (map f xs) F)"
  4067 using lf by (induct xs, auto)
  4068 lemma foldr_And_map_rlfm2: 
  4069   assumes lf: "\<forall> n . isrlfm (f n)"
  4070   shows "isrlfm (foldr conj (map f xs) T)"
  4071 using lf by (induct xs, auto)
  4072 
  4073 lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
  4074   shows "isrlfm (DVDJ i n s)"
  4075 proof-
  4076   let ?f="\<lambda>j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
  4077                          (Dvd i (Sub (C j) (Floor (Neg s))))"
  4078   have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
  4079   from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp 
  4080 qed
  4081 
  4082 lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
  4083   shows "isrlfm (NDVDJ i n s)"
  4084 proof-
  4085   let ?f="\<lambda>j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
  4086                       (NDvd i (Sub (C j) (Floor (Neg s))))"
  4087   have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
  4088   from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto
  4089 qed
  4090 
  4091 definition DVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
  4092   DVD_def: "DVD i c t =
  4093   (if i=0 then eq c t else 
  4094   if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))"
  4095 
  4096 definition  NDVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
  4097   "NDVD i c t =
  4098   (if i=0 then neq c t else 
  4099   if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))"
  4100 
  4101 lemma DVD_mono: 
  4102   assumes xp: "0\<le> x" and x1: "x < 1" 
  4103   shows "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (DVD i n s) = Ifm (x#bs) (Dvd i a)"
  4104   (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (DVD i n s) = ?I (Dvd i a)")
  4105 proof(clarify)
  4106   fix a n s 
  4107   assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
  4108   let ?th = "?I (DVD i n s) = ?I (Dvd i a)"
  4109   have "i=0 \<or> (i\<noteq>0 \<and> n=0) \<or> (i\<noteq>0 \<and> n < 0) \<or> (i\<noteq>0 \<and> n > 0)" by arith
  4110   moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] 
  4111       by (simp add: DVD_def rdvd_left_0_eq)}
  4112   moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H DVD_def) } 
  4113   moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
  4114       by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 
  4115 	rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
  4116   moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)}
  4117   ultimately show ?th by blast
  4118 qed
  4119 
  4120 lemma NDVD_mono:   assumes xp: "0\<le> x" and x1: "x < 1" 
  4121   shows "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (NDVD i n s) = Ifm (x#bs) (NDvd i a)"
  4122   (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (NDVD i n s) = ?I (NDvd i a)")
  4123 proof(clarify)
  4124   fix a n s 
  4125   assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
  4126   let ?th = "?I (NDVD i n s) = ?I (NDvd i a)"
  4127   have "i=0 \<or> (i\<noteq>0 \<and> n=0) \<or> (i\<noteq>0 \<and> n < 0) \<or> (i\<noteq>0 \<and> n > 0)" by arith
  4128   moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] 
  4129       by (simp add: NDVD_def rdvd_left_0_eq)}
  4130   moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H NDVD_def) } 
  4131   moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
  4132       by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 
  4133 	rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
  4134   moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th 
  4135       by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)}
  4136   ultimately show ?th by blast
  4137 qed
  4138 
  4139 lemma DVD_l: "isrlfm (rsplit (DVD i) a)"
  4140   by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) 
  4141 (case_tac s, simp_all, case_tac "nat", simp_all)
  4142 
  4143 lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)"
  4144   by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) 
  4145 (case_tac s, simp_all, case_tac "nat", simp_all)
  4146 
  4147 consts rlfm :: "fm \<Rightarrow> fm"
  4148 recdef rlfm "measure fmsize"
  4149   "rlfm (And p q) = conj (rlfm p) (rlfm q)"
  4150   "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
  4151   "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
  4152   "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
  4153   "rlfm (Lt a) = rsplit lt a"
  4154   "rlfm (Le a) = rsplit le a"
  4155   "rlfm (Gt a) = rsplit gt a"
  4156   "rlfm (Ge a) = rsplit ge a"
  4157   "rlfm (Eq a) = rsplit eq a"
  4158   "rlfm (NEq a) = rsplit neq a"
  4159   "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
  4160   "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
  4161   "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
  4162   "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
  4163   "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
  4164   "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
  4165   "rlfm (NOT (NOT p)) = rlfm p"
  4166   "rlfm (NOT T) = F"
  4167   "rlfm (NOT F) = T"
  4168   "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
  4169   "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
  4170   "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
  4171   "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
  4172   "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
  4173   "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
  4174   "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
  4175   "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
  4176   "rlfm p = p" (hints simp add: fmsize_pos)
  4177 
  4178 lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p"
  4179   by (induct p rule: isrlfm.induct, auto)
  4180 lemma igcd_le1: assumes ip: "0 < i" shows "igcd i j \<le> i"
  4181 proof-
  4182   from igcd_dvd1 have th: "igcd i j dvd i" by blast
  4183   from zdvd_imp_le[OF th ip] show ?thesis .
  4184 qed
  4185 
  4186 
  4187 lemma simpfm_rl: "isrlfm p \<Longrightarrow> isrlfm (simpfm p)"
  4188 proof (induct p)
  4189   case (Lt a) 
  4190   hence "bound0 (Lt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4191     by (cases a,simp_all, case_tac "nat", simp_all)
  4192   moreover
  4193   {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))"  
  4194       using simpfm_bound0 by blast
  4195     have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def)
  4196     with bn bound0at_l have ?case by blast}
  4197   moreover 
  4198   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4199     {
  4200       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4201       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4202       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4203       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4204 	by (simp add: numgcd_def igcd_le1)
  4205       from prems have th': "c\<noteq>0" by auto
  4206       from prems have cp: "c \<ge> 0" by simp
  4207       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4208 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4209     }
  4210     with prems have ?case
  4211       by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
  4212   ultimately show ?case by blast
  4213 next
  4214   case (Le a)   
  4215   hence "bound0 (Le a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4216     by (cases a,simp_all, case_tac "nat", simp_all)
  4217   moreover
  4218   {assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))"  
  4219       using simpfm_bound0 by blast
  4220     have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def)
  4221     with bn bound0at_l have ?case by blast}
  4222   moreover 
  4223   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4224     {
  4225       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4226       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4227       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4228       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4229 	by (simp add: numgcd_def igcd_le1)
  4230       from prems have th': "c\<noteq>0" by auto
  4231       from prems have cp: "c \<ge> 0" by simp
  4232       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4233 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4234     }
  4235     with prems have ?case
  4236       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4237   ultimately show ?case by blast
  4238 next
  4239   case (Gt a)   
  4240   hence "bound0 (Gt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4241     by (cases a,simp_all, case_tac "nat", simp_all)
  4242   moreover
  4243   {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))"  
  4244       using simpfm_bound0 by blast
  4245     have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def)
  4246     with bn bound0at_l have ?case by blast}
  4247   moreover 
  4248   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4249     {
  4250       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4251       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4252       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4253       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4254 	by (simp add: numgcd_def igcd_le1)
  4255       from prems have th': "c\<noteq>0" by auto
  4256       from prems have cp: "c \<ge> 0" by simp
  4257       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4258 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4259     }
  4260     with prems have ?case
  4261       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4262   ultimately show ?case by blast
  4263 next
  4264   case (Ge a)   
  4265   hence "bound0 (Ge a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4266     by (cases a,simp_all, case_tac "nat", simp_all)
  4267   moreover
  4268   {assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))"  
  4269       using simpfm_bound0 by blast
  4270     have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def)
  4271     with bn bound0at_l have ?case by blast}
  4272   moreover 
  4273   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4274     {
  4275       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4276       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4277       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4278       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4279 	by (simp add: numgcd_def igcd_le1)
  4280       from prems have th': "c\<noteq>0" by auto
  4281       from prems have cp: "c \<ge> 0" by simp
  4282       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4283 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4284     }
  4285     with prems have ?case
  4286       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4287   ultimately show ?case by blast
  4288 next
  4289   case (Eq a)   
  4290   hence "bound0 (Eq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4291     by (cases a,simp_all, case_tac "nat", simp_all)
  4292   moreover
  4293   {assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))"  
  4294       using simpfm_bound0 by blast
  4295     have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def)
  4296     with bn bound0at_l have ?case by blast}
  4297   moreover 
  4298   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4299     {
  4300       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4301       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4302       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4303       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4304 	by (simp add: numgcd_def igcd_le1)
  4305       from prems have th': "c\<noteq>0" by auto
  4306       from prems have cp: "c \<ge> 0" by simp
  4307       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4308 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4309     }
  4310     with prems have ?case
  4311       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4312   ultimately show ?case by blast
  4313 next
  4314   case (NEq a)  
  4315   hence "bound0 (NEq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4316     by (cases a,simp_all, case_tac "nat", simp_all)
  4317   moreover
  4318   {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))"  
  4319       using simpfm_bound0 by blast
  4320     have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def)
  4321     with bn bound0at_l have ?case by blast}
  4322   moreover 
  4323   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4324     {
  4325       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4326       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4327       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4328       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4329 	by (simp add: numgcd_def igcd_le1)
  4330       from prems have th': "c\<noteq>0" by auto
  4331       from prems have cp: "c \<ge> 0" by simp
  4332       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4333 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4334     }
  4335     with prems have ?case
  4336       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4337   ultimately show ?case by blast
  4338 next
  4339   case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))"  
  4340     using simpfm_bound0 by blast
  4341   have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
  4342   with bn bound0at_l show ?case by blast
  4343 next
  4344   case (NDvd i a)  hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))"  
  4345     using simpfm_bound0 by blast
  4346   have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
  4347   with bn bound0at_l show ?case by blast
  4348 qed(auto simp add: conj_def imp_def disj_def iff_def Let_def simpfm_bound0 numadd_nb numneg_nb)
  4349 
  4350 lemma rlfm_I:
  4351   assumes qfp: "qfree p"
  4352   and xp: "0 \<le> x" and x1: "x < 1"
  4353   shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \<and> isrlfm (rlfm p)"
  4354   using qfp 
  4355 by (induct p rule: rlfm.induct) 
  4356 (auto simp add: rsplit[OF xp x1 lt_mono] lt_l rsplit[OF xp x1 le_mono] le_l rsplit[OF xp x1 gt_mono] gt_l
  4357                rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l
  4358                rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl)
  4359 lemma rlfm_l:
  4360   assumes qfp: "qfree p"
  4361   shows "isrlfm (rlfm p)"
  4362   using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l 
  4363 by (induct p rule: rlfm.induct,auto simp add: simpfm_rl)
  4364 
  4365     (* Operations needed for Ferrante and Rackoff *)
  4366 lemma rminusinf_inf:
  4367   assumes lp: "isrlfm p"
  4368   shows "\<exists> z. \<forall> x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
  4369 using lp
  4370 proof (induct p rule: minusinf.induct)
  4371   case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
  4372 next
  4373   case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
  4374 next
  4375   case (3 c e) 
  4376   from prems have nb: "numbound0 e" by simp
  4377   from prems have cp: "real c > 0" by simp
  4378   let ?e="Inum (a#bs) e"
  4379   let ?z = "(- ?e) / real c"
  4380   {fix x
  4381     assume xz: "x < ?z"
  4382     hence "(real c * x < - ?e)" 
  4383       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4384     hence "real c * x + ?e < 0" by arith
  4385     hence "real c * x + ?e \<noteq> 0" by simp
  4386     with xz have "?P ?z x (Eq (CN 0 c e))"
  4387       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp  }
  4388   hence "\<forall> x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp
  4389   thus ?case by blast
  4390 next
  4391   case (4 c e)   
  4392   from prems have nb: "numbound0 e" by simp
  4393   from prems have cp: "real c > 0" by simp
  4394   let ?e="Inum (a#bs) e"
  4395   let ?z = "(- ?e) / real c"
  4396   {fix x
  4397     assume xz: "x < ?z"
  4398     hence "(real c * x < - ?e)" 
  4399       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4400     hence "real c * x + ?e < 0" by arith
  4401     hence "real c * x + ?e \<noteq> 0" by simp
  4402     with xz have "?P ?z x (NEq (CN 0 c e))"
  4403       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4404   hence "\<forall> x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp
  4405   thus ?case by blast
  4406 next
  4407   case (5 c e) 
  4408     from prems have nb: "numbound0 e" by simp
  4409   from prems have cp: "real c > 0" by simp
  4410   let ?e="Inum (a#bs) e"
  4411   let ?z = "(- ?e) / real c"
  4412   {fix x
  4413     assume xz: "x < ?z"
  4414     hence "(real c * x < - ?e)" 
  4415       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4416     hence "real c * x + ?e < 0" by arith
  4417     with xz have "?P ?z x (Lt (CN 0 c e))"
  4418       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"]  by simp }
  4419   hence "\<forall> x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp
  4420   thus ?case by blast
  4421 next
  4422   case (6 c e)  
  4423     from prems have nb: "numbound0 e" by simp
  4424   from prems have cp: "real c > 0" by simp
  4425   let ?e="Inum (a#bs) e"
  4426   let ?z = "(- ?e) / real c"
  4427   {fix x
  4428     assume xz: "x < ?z"
  4429     hence "(real c * x < - ?e)" 
  4430       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4431     hence "real c * x + ?e < 0" by arith
  4432     with xz have "?P ?z x (Le (CN 0 c e))"
  4433       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4434   hence "\<forall> x < ?z. ?P ?z x (Le (CN 0 c e))" by simp
  4435   thus ?case by blast
  4436 next
  4437   case (7 c e)  
  4438     from prems have nb: "numbound0 e" by simp
  4439   from prems have cp: "real c > 0" by simp
  4440   let ?e="Inum (a#bs) e"
  4441   let ?z = "(- ?e) / real c"
  4442   {fix x
  4443     assume xz: "x < ?z"
  4444     hence "(real c * x < - ?e)" 
  4445       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4446     hence "real c * x + ?e < 0" by arith
  4447     with xz have "?P ?z x (Gt (CN 0 c e))"
  4448       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4449   hence "\<forall> x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp
  4450   thus ?case by blast
  4451 next
  4452   case (8 c e)  
  4453     from prems have nb: "numbound0 e" by simp
  4454   from prems have cp: "real c > 0" by simp
  4455   let ?e="Inum (a#bs) e"
  4456   let ?z = "(- ?e) / real c"
  4457   {fix x
  4458     assume xz: "x < ?z"
  4459     hence "(real c * x < - ?e)" 
  4460       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4461     hence "real c * x + ?e < 0" by arith
  4462     with xz have "?P ?z x (Ge (CN 0 c e))"
  4463       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4464   hence "\<forall> x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp
  4465   thus ?case by blast
  4466 qed simp_all
  4467 
  4468 lemma rplusinf_inf:
  4469   assumes lp: "isrlfm p"
  4470   shows "\<exists> z. \<forall> x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
  4471 using lp
  4472 proof (induct p rule: isrlfm.induct)
  4473   case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
  4474 next
  4475   case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
  4476 next
  4477   case (3 c e) 
  4478   from prems have nb: "numbound0 e" by simp
  4479   from prems have cp: "real c > 0" by simp
  4480   let ?e="Inum (a#bs) e"
  4481   let ?z = "(- ?e) / real c"
  4482   {fix x
  4483     assume xz: "x > ?z"
  4484     with mult_strict_right_mono [OF xz cp] cp
  4485     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4486     hence "real c * x + ?e > 0" by arith
  4487     hence "real c * x + ?e \<noteq> 0" by simp
  4488     with xz have "?P ?z x (Eq (CN 0 c e))"
  4489       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4490   hence "\<forall> x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp
  4491   thus ?case by blast
  4492 next
  4493   case (4 c e) 
  4494   from prems have nb: "numbound0 e" by simp
  4495   from prems have cp: "real c > 0" by simp
  4496   let ?e="Inum (a#bs) e"
  4497   let ?z = "(- ?e) / real c"
  4498   {fix x
  4499     assume xz: "x > ?z"
  4500     with mult_strict_right_mono [OF xz cp] cp
  4501     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4502     hence "real c * x + ?e > 0" by arith
  4503     hence "real c * x + ?e \<noteq> 0" by simp
  4504     with xz have "?P ?z x (NEq (CN 0 c e))"
  4505       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4506   hence "\<forall> x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp
  4507   thus ?case by blast
  4508 next
  4509   case (5 c e) 
  4510   from prems have nb: "numbound0 e" by simp
  4511   from prems have cp: "real c > 0" by simp
  4512   let ?e="Inum (a#bs) e"
  4513   let ?z = "(- ?e) / real c"
  4514   {fix x
  4515     assume xz: "x > ?z"
  4516     with mult_strict_right_mono [OF xz cp] cp
  4517     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4518     hence "real c * x + ?e > 0" by arith
  4519     with xz have "?P ?z x (Lt (CN 0 c e))"
  4520       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4521   hence "\<forall> x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp
  4522   thus ?case by blast
  4523 next
  4524   case (6 c e) 
  4525   from prems have nb: "numbound0 e" by simp
  4526   from prems have cp: "real c > 0" by simp
  4527   let ?e="Inum (a#bs) e"
  4528   let ?z = "(- ?e) / real c"
  4529   {fix x
  4530     assume xz: "x > ?z"
  4531     with mult_strict_right_mono [OF xz cp] cp
  4532     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4533     hence "real c * x + ?e > 0" by arith
  4534     with xz have "?P ?z x (Le (CN 0 c e))"
  4535       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4536   hence "\<forall> x > ?z. ?P ?z x (Le (CN 0 c e))" by simp
  4537   thus ?case by blast
  4538 next
  4539   case (7 c e) 
  4540   from prems have nb: "numbound0 e" by simp
  4541   from prems have cp: "real c > 0" by simp
  4542   let ?e="Inum (a#bs) e"
  4543   let ?z = "(- ?e) / real c"
  4544   {fix x
  4545     assume xz: "x > ?z"
  4546     with mult_strict_right_mono [OF xz cp] cp
  4547     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4548     hence "real c * x + ?e > 0" by arith
  4549     with xz have "?P ?z x (Gt (CN 0 c e))"
  4550       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4551   hence "\<forall> x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp
  4552   thus ?case by blast
  4553 next
  4554   case (8 c e) 
  4555   from prems have nb: "numbound0 e" by simp
  4556   from prems have cp: "real c > 0" by simp
  4557   let ?e="Inum (a#bs) e"
  4558   let ?z = "(- ?e) / real c"
  4559   {fix x
  4560     assume xz: "x > ?z"
  4561     with mult_strict_right_mono [OF xz cp] cp
  4562     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4563     hence "real c * x + ?e > 0" by arith
  4564     with xz have "?P ?z x (Ge (CN 0 c e))"
  4565       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"]   by simp }
  4566   hence "\<forall> x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp
  4567   thus ?case by blast
  4568 qed simp_all
  4569 
  4570 lemma rminusinf_bound0:
  4571   assumes lp: "isrlfm p"
  4572   shows "bound0 (minusinf p)"
  4573   using lp
  4574   by (induct p rule: minusinf.induct) simp_all
  4575 
  4576 lemma rplusinf_bound0:
  4577   assumes lp: "isrlfm p"
  4578   shows "bound0 (plusinf p)"
  4579   using lp
  4580   by (induct p rule: plusinf.induct) simp_all
  4581 
  4582 lemma rminusinf_ex:
  4583   assumes lp: "isrlfm p"
  4584   and ex: "Ifm (a#bs) (minusinf p)"
  4585   shows "\<exists> x. Ifm (x#bs) p"
  4586 proof-
  4587   from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
  4588   have th: "\<forall> x. Ifm (x#bs) (minusinf p)" by auto
  4589   from rminusinf_inf[OF lp, where bs="bs"] 
  4590   obtain z where z_def: "\<forall>x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p" by blast
  4591   from th have "Ifm ((z - 1)#bs) (minusinf p)" by simp
  4592   moreover have "z - 1 < z" by simp
  4593   ultimately show ?thesis using z_def by auto
  4594 qed
  4595 
  4596 lemma rplusinf_ex:
  4597   assumes lp: "isrlfm p"
  4598   and ex: "Ifm (a#bs) (plusinf p)"
  4599   shows "\<exists> x. Ifm (x#bs) p"
  4600 proof-
  4601   from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
  4602   have th: "\<forall> x. Ifm (x#bs) (plusinf p)" by auto
  4603   from rplusinf_inf[OF lp, where bs="bs"] 
  4604   obtain z where z_def: "\<forall>x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast
  4605   from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp
  4606   moreover have "z + 1 > z" by simp
  4607   ultimately show ?thesis using z_def by auto
  4608 qed
  4609 
  4610 consts 
  4611   \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list"
  4612   \<upsilon> :: "fm \<Rightarrow> (num \<times> int) \<Rightarrow> fm "
  4613 recdef \<Upsilon> "measure size"
  4614   "\<Upsilon> (And p q) = (\<Upsilon> p @ \<Upsilon> q)" 
  4615   "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)" 
  4616   "\<Upsilon> (Eq  (CN 0 c e)) = [(Neg e,c)]"
  4617   "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]"
  4618   "\<Upsilon> (Lt  (CN 0 c e)) = [(Neg e,c)]"
  4619   "\<Upsilon> (Le  (CN 0 c e)) = [(Neg e,c)]"
  4620   "\<Upsilon> (Gt  (CN 0 c e)) = [(Neg e,c)]"
  4621   "\<Upsilon> (Ge  (CN 0 c e)) = [(Neg e,c)]"
  4622   "\<Upsilon> p = []"
  4623 
  4624 recdef \<upsilon> "measure size"
  4625   "\<upsilon> (And p q) = (\<lambda> (t,n). And (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
  4626   "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
  4627   "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))"
  4628   "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))"
  4629   "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))"
  4630   "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))"
  4631   "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))"
  4632   "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))"
  4633   "\<upsilon> p = (\<lambda> (t,n). p)"
  4634 
  4635 lemma \<upsilon>_I: assumes lp: "isrlfm p"
  4636   and np: "real n > 0" and nbt: "numbound0 t"
  4637   shows "(Ifm (x#bs) (\<upsilon> p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) \<and> bound0 (\<upsilon> p (t,n))" (is "(?I x (\<upsilon> p (t,n)) = ?I ?u p) \<and> ?B p" is "(_ = ?I (?t/?n) p) \<and> _" is "(_ = ?I (?N x t /_) p) \<and> _")
  4638   using lp
  4639 proof(induct p rule: \<upsilon>.induct)
  4640   case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4641   have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)"
  4642     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4643   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)"
  4644     by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4645       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4646   also have "\<dots> = (real c *?t + ?n* (?N x e) < 0)"
  4647     using np by simp 
  4648   finally show ?case using nbt nb by (simp add: ring_simps)
  4649 next
  4650   case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4651   have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<le> 0)"
  4652     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4653   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<le> 0)"
  4654     by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4655       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4656   also have "\<dots> = (real c *?t + ?n* (?N x e) \<le> 0)"
  4657     using np by simp 
  4658   finally show ?case using nbt nb by (simp add: ring_simps)
  4659 next
  4660   case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4661   have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)"
  4662     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4663   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)"
  4664     by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4665       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4666   also have "\<dots> = (real c *?t + ?n* (?N x e) > 0)"
  4667     using np by simp 
  4668   finally show ?case using nbt nb by (simp add: ring_simps)
  4669 next
  4670   case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4671   have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<ge> 0)"
  4672     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4673   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<ge> 0)"
  4674     by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4675       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4676   also have "\<dots> = (real c *?t + ?n* (?N x e) \<ge> 0)"
  4677     using np by simp 
  4678   finally show ?case using nbt nb by (simp add: ring_simps)
  4679 next
  4680   case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4681   from np have np: "real n \<noteq> 0" by simp
  4682   have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)"
  4683     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4684   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)"
  4685     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4686       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4687   also have "\<dots> = (real c *?t + ?n* (?N x e) = 0)"
  4688     using np by simp 
  4689   finally show ?case using nbt nb by (simp add: ring_simps)
  4690 next
  4691   case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4692   from np have np: "real n \<noteq> 0" by simp
  4693   have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<noteq> 0)"
  4694     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4695   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)"
  4696     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4697       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4698   also have "\<dots> = (real c *?t + ?n* (?N x e) \<noteq> 0)"
  4699     using np by simp 
  4700   finally show ?case using nbt nb by (simp add: ring_simps)
  4701 qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2)
  4702 
  4703 lemma \<Upsilon>_l:
  4704   assumes lp: "isrlfm p"
  4705   shows "\<forall> (t,k) \<in> set (\<Upsilon> p). numbound0 t \<and> k >0"
  4706 using lp
  4707 by(induct p rule: \<Upsilon>.induct)  auto
  4708 
  4709 lemma rminusinf_\<Upsilon>:
  4710   assumes lp: "isrlfm p"
  4711   and nmi: "\<not> (Ifm (a#bs) (minusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
  4712   and ex: "Ifm (x#bs) p" (is "?I x p")
  4713   shows "\<exists> (s,m) \<in> set (\<Upsilon> p). x \<ge> Inum (a#bs) s / real m" (is "\<exists> (s,m) \<in> ?U p. x \<ge> ?N a s / real m")
  4714 proof-
  4715   have "\<exists> (s,m) \<in> set (\<Upsilon> p). real m * x \<ge> Inum (a#bs) s " (is "\<exists> (s,m) \<in> ?U p. real m *x \<ge> ?N a s")
  4716     using lp nmi ex
  4717     by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2)
  4718   then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<ge> ?N a s" by blast
  4719   from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
  4720   from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \<ge> ?N a s / real m" 
  4721     by (auto simp add: mult_commute)
  4722   thus ?thesis using smU by auto
  4723 qed
  4724 
  4725 lemma rplusinf_\<Upsilon>:
  4726   assumes lp: "isrlfm p"
  4727   and nmi: "\<not> (Ifm (a#bs) (plusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
  4728   and ex: "Ifm (x#bs) p" (is "?I x p")
  4729   shows "\<exists> (s,m) \<in> set (\<Upsilon> p). x \<le> Inum (a#bs) s / real m" (is "\<exists> (s,m) \<in> ?U p. x \<le> ?N a s / real m")
  4730 proof-
  4731   have "\<exists> (s,m) \<in> set (\<Upsilon> p). real m * x \<le> Inum (a#bs) s " (is "\<exists> (s,m) \<in> ?U p. real m *x \<le> ?N a s")
  4732     using lp nmi ex
  4733     by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2)
  4734   then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<le> ?N a s" by blast
  4735   from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
  4736   from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \<le> ?N a s / real m" 
  4737     by (auto simp add: mult_commute)
  4738   thus ?thesis using smU by auto
  4739 qed
  4740 
  4741 lemma lin_dense: 
  4742   assumes lp: "isrlfm p"
  4743   and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (t,n). Inum (x#bs) t / real n) ` set (\<Upsilon> p)" 
  4744   (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (t,n). ?N x t / real n ) ` (?U p)")
  4745   and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p"
  4746   and ly: "l < y" and yu: "y < u"
  4747   shows "Ifm (y#bs) p"
  4748 using lp px noS
  4749 proof (induct p rule: isrlfm.induct)
  4750   case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4751     from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps)
  4752     hence pxc: "x < (- ?N x e) / real c" 
  4753       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"])
  4754     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4755     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4756     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4757     moreover {assume y: "y < (-?N x e)/ real c"
  4758       hence "y * real c < - ?N x e"
  4759 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4760       hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
  4761       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4762     moreover {assume y: "y > (- ?N x e) / real c" 
  4763       with yu have eu: "u > (- ?N x e) / real c" by auto
  4764       with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
  4765       with lx pxc have "False" by auto
  4766       hence ?case by simp }
  4767     ultimately show ?case by blast
  4768 next
  4769   case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp +
  4770     from prems have "x * real c + ?N x e \<le> 0" by (simp add: ring_simps)
  4771     hence pxc: "x \<le> (- ?N x e) / real c" 
  4772       by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"])
  4773     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4774     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4775     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4776     moreover {assume y: "y < (-?N x e)/ real c"
  4777       hence "y * real c < - ?N x e"
  4778 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4779       hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
  4780       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4781     moreover {assume y: "y > (- ?N x e) / real c" 
  4782       with yu have eu: "u > (- ?N x e) / real c" by auto
  4783       with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
  4784       with lx pxc have "False" by auto
  4785       hence ?case by simp }
  4786     ultimately show ?case by blast
  4787 next
  4788   case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4789     from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps)
  4790     hence pxc: "x > (- ?N x e) / real c" 
  4791       by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"])
  4792     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4793     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4794     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4795     moreover {assume y: "y > (-?N x e)/ real c"
  4796       hence "y * real c > - ?N x e"
  4797 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4798       hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
  4799       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4800     moreover {assume y: "y < (- ?N x e) / real c" 
  4801       with ly have eu: "l < (- ?N x e) / real c" by auto
  4802       with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
  4803       with xu pxc have "False" by auto
  4804       hence ?case by simp }
  4805     ultimately show ?case by blast
  4806 next
  4807   case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4808     from prems have "x * real c + ?N x e \<ge> 0" by (simp add: ring_simps)
  4809     hence pxc: "x \<ge> (- ?N x e) / real c" 
  4810       by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"])
  4811     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4812     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4813     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4814     moreover {assume y: "y > (-?N x e)/ real c"
  4815       hence "y * real c > - ?N x e"
  4816 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4817       hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
  4818       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4819     moreover {assume y: "y < (- ?N x e) / real c" 
  4820       with ly have eu: "l < (- ?N x e) / real c" by auto
  4821       with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
  4822       with xu pxc have "False" by auto
  4823       hence ?case by simp }
  4824     ultimately show ?case by blast
  4825 next
  4826   case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4827     from cp have cnz: "real c \<noteq> 0" by simp
  4828     from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps)
  4829     hence pxc: "x = (- ?N x e) / real c" 
  4830       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"])
  4831     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4832     with lx xu have yne: "x \<noteq> - ?N x e / real c" by auto
  4833     with pxc show ?case by simp
  4834 next
  4835   case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4836     from cp have cnz: "real c \<noteq> 0" by simp
  4837     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4838     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4839     hence "y* real c \<noteq> -?N x e"      
  4840       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp
  4841     hence "y* real c + ?N x e \<noteq> 0" by (simp add: ring_simps)
  4842     thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] 
  4843       by (simp add: ring_simps)
  4844 qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"])
  4845 
  4846 lemma finite_set_intervals:
  4847   assumes px: "P (x::real)" 
  4848   and lx: "l \<le> x" and xu: "x \<le> u"
  4849   and linS: "l\<in> S" and uinS: "u \<in> S"
  4850   and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
  4851   shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x"
  4852 proof-
  4853   let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
  4854   let ?xM = "{y. y\<in> S \<and> x \<le> y}"
  4855   let ?a = "Max ?Mx"
  4856   let ?b = "Min ?xM"
  4857   have MxS: "?Mx \<subseteq> S" by blast
  4858   hence fMx: "finite ?Mx" using fS finite_subset by auto
  4859   from lx linS have linMx: "l \<in> ?Mx" by blast
  4860   hence Mxne: "?Mx \<noteq> {}" by blast
  4861   have xMS: "?xM \<subseteq> S" by blast
  4862   hence fxM: "finite ?xM" using fS finite_subset by auto
  4863   from xu uinS have linxM: "u \<in> ?xM" by blast
  4864   hence xMne: "?xM \<noteq> {}" by blast
  4865   have ax:"?a \<le> x" using Mxne fMx by auto
  4866   have xb:"x \<le> ?b" using xMne fxM by auto
  4867   have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
  4868   have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
  4869   have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
  4870   proof(clarsimp)
  4871     fix y
  4872     assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
  4873     from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by auto
  4874     moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by simp}
  4875     moreover {assume "y \<in> ?xM" hence "y \<ge> ?b" using xMne fxM by auto with yb have "False" by simp}
  4876     ultimately show "False" by blast
  4877   qed
  4878   from ainS binS noy ax xb px show ?thesis by blast
  4879 qed
  4880 
  4881 lemma finite_set_intervals2:
  4882   assumes px: "P (x::real)" 
  4883   and lx: "l \<le> x" and xu: "x \<le> u"
  4884   and linS: "l\<in> S" and uinS: "u \<in> S"
  4885   and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
  4886   shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)"
  4887 proof-
  4888   from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
  4889   obtain a and b where 
  4890     as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S" and axb: "a \<le> x \<and> x \<le> b \<and> P x"  by auto
  4891   from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by auto
  4892   thus ?thesis using px as bs noS by blast 
  4893 qed
  4894 
  4895 lemma rinf_\<Upsilon>:
  4896   assumes lp: "isrlfm p"
  4897   and nmi: "\<not> (Ifm (x#bs) (minusinf p))" (is "\<not> (Ifm (x#bs) (?M p))")
  4898   and npi: "\<not> (Ifm (x#bs) (plusinf p))" (is "\<not> (Ifm (x#bs) (?P p))")
  4899   and ex: "\<exists> x.  Ifm (x#bs) p" (is "\<exists> x. ?I x p")
  4900   shows "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" 
  4901 proof-
  4902   let ?N = "\<lambda> x t. Inum (x#bs) t"
  4903   let ?U = "set (\<Upsilon> p)"
  4904   from ex obtain a where pa: "?I a p" by blast
  4905   from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi
  4906   have nmi': "\<not> (?I a (?M p))" by simp
  4907   from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi
  4908   have npi': "\<not> (?I a (?P p))" by simp
  4909   have "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). ?I ((?N a l/real n + ?N a s /real m) / 2) p"
  4910   proof-
  4911     let ?M = "(\<lambda> (t,c). ?N a t / real c) ` ?U"
  4912     have fM: "finite ?M" by auto
  4913     from rminusinf_\<Upsilon>[OF lp nmi pa] rplusinf_\<Upsilon>[OF lp npi pa] 
  4914     have "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). a \<le> ?N x l / real n \<and> a \<ge> ?N x s / real m" by blast
  4915     then obtain "t" "n" "s" "m" where 
  4916       tnU: "(t,n) \<in> ?U" and smU: "(s,m) \<in> ?U" 
  4917       and xs1: "a \<le> ?N x s / real m" and tx1: "a \<ge> ?N x t / real n" by blast
  4918     from \<Upsilon>_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \<le> ?N a s / real m" and tx: "a \<ge> ?N a t / real n" by auto
  4919     from tnU have Mne: "?M \<noteq> {}" by auto
  4920     hence Une: "?U \<noteq> {}" by simp
  4921     let ?l = "Min ?M"
  4922     let ?u = "Max ?M"
  4923     have linM: "?l \<in> ?M" using fM Mne by simp
  4924     have uinM: "?u \<in> ?M" using fM Mne by simp
  4925     have tnM: "?N a t / real n \<in> ?M" using tnU by auto
  4926     have smM: "?N a s / real m \<in> ?M" using smU by auto 
  4927     have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto
  4928     have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
  4929     have "?l \<le> ?N a t / real n" using tnM Mne by simp hence lx: "?l \<le> a" using tx by simp
  4930     have "?N a s / real m \<le> ?u" using smM Mne by simp hence xu: "a \<le> ?u" using xs by simp
  4931     from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
  4932     have "(\<exists> s\<in> ?M. ?I s p) \<or> 
  4933       (\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p)" .
  4934     moreover { fix u assume um: "u\<in> ?M" and pu: "?I u p"
  4935       hence "\<exists> (tu,nu) \<in> ?U. u = ?N a tu / real nu" by auto
  4936       then obtain "tu" "nu" where tuU: "(tu,nu) \<in> ?U" and tuu:"u= ?N a tu / real nu" by blast
  4937       have "(u + u) / 2 = u" by auto with pu tuu 
  4938       have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp
  4939       with tuU have ?thesis by blast}
  4940     moreover{
  4941       assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
  4942       then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" 
  4943 	and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
  4944 	by blast
  4945       from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real t1n" by auto
  4946       then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast
  4947       from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real t2n" by auto
  4948       then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \<in> ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast
  4949       from t1x xt2 have t1t2: "t1 < t2" by simp
  4950       let ?u = "(t1 + t2) / 2"
  4951       from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto
  4952       from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" .
  4953       with t1uU t2uU t1u t2u have ?thesis by blast}
  4954     ultimately show ?thesis by blast
  4955   qed
  4956   then obtain "l" "n" "s"  "m" where lnU: "(l,n) \<in> ?U" and smU:"(s,m) \<in> ?U" 
  4957     and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast
  4958   from lnU smU \<Upsilon>_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto
  4959   from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] 
  4960     numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
  4961   have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp
  4962   with lnU smU
  4963   show ?thesis by auto
  4964 qed
  4965     (* The Ferrante - Rackoff Theorem *)
  4966 
  4967 theorem fr_eq: 
  4968   assumes lp: "isrlfm p"
  4969   shows "(\<exists> x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \<or> (Ifm (x#bs) (plusinf p)) \<or> (\<exists> (t,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). Ifm ((((Inum (x#bs) t)/  real n + (Inum (x#bs) s) / real m) /2)#bs) p))"
  4970   (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
  4971 proof
  4972   assume px: "\<exists> x. ?I x p"
  4973   have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
  4974   moreover {assume "?M \<or> ?P" hence "?D" by blast}
  4975   moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
  4976     from rinf_\<Upsilon>[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast}
  4977   ultimately show "?D" by blast
  4978 next
  4979   assume "?D" 
  4980   moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
  4981   moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
  4982   moreover {assume f:"?F" hence "?E" by blast}
  4983   ultimately show "?E" by blast
  4984 qed
  4985 
  4986 
  4987 lemma fr_eq\<upsilon>: 
  4988   assumes lp: "isrlfm p"
  4989   shows "(\<exists> x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \<or> (Ifm (x#bs) (plusinf p)) \<or> (\<exists> (t,k) \<in> set (\<Upsilon> p). \<exists> (s,l) \<in> set (\<Upsilon> p). Ifm (x#bs) (\<upsilon> p (Add(Mul l t) (Mul k s) , 2*k*l))))"
  4990   (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
  4991 proof
  4992   assume px: "\<exists> x. ?I x p"
  4993   have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
  4994   moreover {assume "?M \<or> ?P" hence "?D" by blast}
  4995   moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
  4996     let ?f ="\<lambda> (t,n). Inum (x#bs) t / real n"
  4997     let ?N = "\<lambda> t. Inum (x#bs) t"
  4998     {fix t n s m assume "(t,n)\<in> set (\<Upsilon> p)" and "(s,m) \<in> set (\<Upsilon> p)"
  4999       with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0"
  5000 	by auto
  5001       let ?st = "Add (Mul m t) (Mul n s)"
  5002       from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  5003 	by (simp add: mult_commute)
  5004       from tnb snb have st_nb: "numbound0 ?st" by simp
  5005       have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5006 	using mnp mp np by (simp add: ring_simps add_divide_distrib)
  5007       from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"] 
  5008       have "?I x (\<upsilon> p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])}
  5009     with rinf_\<Upsilon>[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
  5010   ultimately show "?D" by blast
  5011 next
  5012   assume "?D" 
  5013   moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
  5014   moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
  5015   moreover {fix t k s l assume "(t,k) \<in> set (\<Upsilon> p)" and "(s,l) \<in> set (\<Upsilon> p)" 
  5016     and px:"?I x (\<upsilon> p (Add (Mul l t) (Mul k s), 2*k*l))"
  5017     with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto
  5018     let ?st = "Add (Mul l t) (Mul k s)"
  5019     from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" 
  5020       by (simp add: mult_commute)
  5021     from tnb snb have st_nb: "numbound0 ?st" by simp
  5022     from \<upsilon>_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto}
  5023   ultimately show "?E" by blast
  5024 qed
  5025 
  5026 text{* The overall Part *}
  5027 
  5028 lemma real_ex_int_real01:
  5029   shows "(\<exists> (x::real). P x) = (\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))"
  5030 proof(auto)
  5031   fix x
  5032   assume Px: "P x"
  5033   let ?i = "floor x"
  5034   let ?u = "x - real ?i"
  5035   have "x = real ?i + ?u" by simp
  5036   hence "P (real ?i + ?u)" using Px by simp
  5037   moreover have "real ?i \<le> x" using real_of_int_floor_le by simp hence "0 \<le> ?u" by arith
  5038   moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith 
  5039   ultimately show "(\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))" by blast
  5040 qed
  5041 
  5042 consts exsplitnum :: "num \<Rightarrow> num"
  5043   exsplit :: "fm \<Rightarrow> fm"
  5044 recdef exsplitnum "measure size"
  5045   "exsplitnum (C c) = (C c)"
  5046   "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)"
  5047   "exsplitnum (Bound n) = Bound (n+1)"
  5048   "exsplitnum (Neg a) = Neg (exsplitnum a)"
  5049   "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) "
  5050   "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) "
  5051   "exsplitnum (Mul c a) = Mul c (exsplitnum a)"
  5052   "exsplitnum (Floor a) = Floor (exsplitnum a)"
  5053   "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))"
  5054   "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)"
  5055   "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)"
  5056 
  5057 recdef exsplit "measure size"
  5058   "exsplit (Lt a) = Lt (exsplitnum a)"
  5059   "exsplit (Le a) = Le (exsplitnum a)"
  5060   "exsplit (Gt a) = Gt (exsplitnum a)"
  5061   "exsplit (Ge a) = Ge (exsplitnum a)"
  5062   "exsplit (Eq a) = Eq (exsplitnum a)"
  5063   "exsplit (NEq a) = NEq (exsplitnum a)"
  5064   "exsplit (Dvd i a) = Dvd i (exsplitnum a)"
  5065   "exsplit (NDvd i a) = NDvd i (exsplitnum a)"
  5066   "exsplit (And p q) = And (exsplit p) (exsplit q)"
  5067   "exsplit (Or p q) = Or (exsplit p) (exsplit q)"
  5068   "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)"
  5069   "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)"
  5070   "exsplit (NOT p) = NOT (exsplit p)"
  5071   "exsplit p = p"
  5072 
  5073 lemma exsplitnum: 
  5074   "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t"
  5075   by(induct t rule: exsplitnum.induct) (simp_all add: ring_simps)
  5076 
  5077 lemma exsplit: 
  5078   assumes qfp: "qfree p"
  5079   shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p"
  5080 using qfp exsplitnum[where x="x" and y="y" and bs="bs"]
  5081 by(induct p rule: exsplit.induct) simp_all
  5082 
  5083 lemma splitex:
  5084   assumes qf: "qfree p"
  5085   shows "(Ifm bs (E p)) = (\<exists> (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs")
  5086 proof-
  5087   have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))"
  5088     by (simp add: myless[rule_format, where b="1"] myless[rule_format, where b="0"] add_ac diff_def)
  5089   also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)"
  5090     by (simp only: exsplit[OF qf] add_ac)
  5091   also have "\<dots> = (\<exists> x. Ifm (x#bs) p)" 
  5092     by (simp only: real_ex_int_real01[where P="\<lambda> x. Ifm (x#bs) p"])
  5093   finally show ?thesis by simp
  5094 qed
  5095 
  5096     (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *)
  5097 
  5098 constdefs ferrack01:: "fm \<Rightarrow> fm"
  5099   "ferrack01 p \<equiv> (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p);
  5100                     U = remdups(map simp_num_pair 
  5101                      (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m))
  5102                            (alluopairs (\<Upsilon> p')))) 
  5103   in decr (evaldjf (\<upsilon> p') U ))"
  5104 
  5105 lemma fr_eq_01: 
  5106   assumes qf: "qfree p"
  5107   shows "(\<exists> x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (\<exists> (t,n) \<in> set (\<Upsilon> (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). \<exists> (s,m) \<in> set (\<Upsilon> (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). Ifm (x#bs) (\<upsilon> (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) (Add (Mul m t) (Mul n s), 2*n*m)))"
  5108   (is "(\<exists> x. ?I x ?q) = ?F")
  5109 proof-
  5110   let ?rq = "rlfm ?q"
  5111   let ?M = "?I x (minusinf ?rq)"
  5112   let ?P = "?I x (plusinf ?rq)"
  5113   have MF: "?M = False"
  5114     apply (simp add: Let_def reducecoeff_def numgcd_def igcd_def rsplit_def ge_def lt_def conj_def disj_def)
  5115     by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
  5116   have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def igcd_def rsplit_def ge_def lt_def conj_def disj_def)
  5117     by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
  5118   have "(\<exists> x. ?I x ?q ) = 
  5119     ((?I x (minusinf ?rq)) \<or> (?I x (plusinf ?rq )) \<or> (\<exists> (t,n) \<in> set (\<Upsilon> ?rq). \<exists> (s,m) \<in> set (\<Upsilon> ?rq ). ?I x (\<upsilon> ?rq (Add (Mul m t) (Mul n s), 2*n*m))))"
  5120     (is "(\<exists> x. ?I x ?q) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
  5121   proof
  5122     assume "\<exists> x. ?I x ?q"  
  5123     then obtain x where qx: "?I x ?q" by blast
  5124     hence xp: "0\<le> x" and x1: "x< 1" and px: "?I x p" 
  5125       by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf])
  5126     from qx have "?I x ?rq " 
  5127       by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
  5128     hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto
  5129     from qf have qfq:"isrlfm ?rq"  
  5130       by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
  5131     with lqx fr_eq\<upsilon>[OF qfq] show "?M \<or> ?P \<or> ?F" by blast
  5132   next
  5133     assume D: "?D"
  5134     let ?U = "set (\<Upsilon> ?rq )"
  5135     from MF PF D have "?F" by auto
  5136     then obtain t n s m where aU:"(t,n) \<in> ?U" and bU:"(s,m)\<in> ?U" and rqx: "?I x (\<upsilon> ?rq (Add (Mul m t) (Mul n s), 2*n*m))" by blast
  5137     from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] 
  5138       by (auto simp add: rsplit_def lt_def ge_def)
  5139     from aU bU \<Upsilon>_l[OF lrq] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by (auto simp add: split_def)
  5140     let ?st = "Add (Mul m t) (Mul n s)"
  5141     from tnb snb have stnb: "numbound0 ?st" by simp
  5142     from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  5143       by (simp add: mult_commute)
  5144     from conjunct1[OF \<upsilon>_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx
  5145     have "\<exists> x. ?I x ?rq" by auto
  5146     thus "?E" 
  5147       using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def)
  5148   qed
  5149   with MF PF show ?thesis by blast
  5150 qed
  5151 
  5152 lemma \<Upsilon>_cong_aux:
  5153   assumes Ul: "\<forall> (t,n) \<in> set U. numbound0 t \<and> n >0"
  5154   shows "((\<lambda> (t,n). Inum (x#bs) t /real n) ` (set (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\<lambda> ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U \<times> set U))"
  5155   (is "?lhs = ?rhs")
  5156 proof(auto)
  5157   fix t n s m
  5158   assume "((t,n),(s,m)) \<in> set (alluopairs U)"
  5159   hence th: "((t,n),(s,m)) \<in> (set U \<times> set U)"
  5160     using alluopairs_set1[where xs="U"] by blast
  5161   let ?N = "\<lambda> t. Inum (x#bs) t"
  5162   let ?st= "Add (Mul m t) (Mul n s)"
  5163   from Ul th have mnz: "m \<noteq> 0" by auto
  5164   from Ul th have  nnz: "n \<noteq> 0" by auto  
  5165   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5166    using mnz nnz by (simp add: ring_simps add_divide_distrib)
  5167  
  5168   thus "(real m *  Inum (x # bs) t + real n * Inum (x # bs) s) /
  5169        (2 * real n * real m)
  5170        \<in> (\<lambda>((t, n), s, m).
  5171              (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
  5172          (set U \<times> set U)"using mnz nnz th  
  5173     apply (auto simp add: th add_divide_distrib ring_simps split_def image_def)
  5174     by (rule_tac x="(s,m)" in bexI,simp_all) 
  5175   (rule_tac x="(t,n)" in bexI,simp_all)
  5176 next
  5177   fix t n s m
  5178   assume tnU: "(t,n) \<in> set U" and smU:"(s,m) \<in> set U" 
  5179   let ?N = "\<lambda> t. Inum (x#bs) t"
  5180   let ?st= "Add (Mul m t) (Mul n s)"
  5181   from Ul smU have mnz: "m \<noteq> 0" by auto
  5182   from Ul tnU have  nnz: "n \<noteq> 0" by auto  
  5183   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5184    using mnz nnz by (simp add: ring_simps add_divide_distrib)
  5185  let ?P = "\<lambda> (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2"
  5186  have Pc:"\<forall> a b. ?P a b = ?P b a"
  5187    by auto
  5188  from Ul alluopairs_set1 have Up:"\<forall> ((t,n),(s,m)) \<in> set (alluopairs U). n \<noteq> 0 \<and> m \<noteq> 0" by blast
  5189  from alluopairs_ex[OF Pc, where xs="U"] tnU smU
  5190  have th':"\<exists> ((t',n'),(s',m')) \<in> set (alluopairs U). ?P (t',n') (s',m')"
  5191    by blast
  5192  then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \<in> set (alluopairs U)" 
  5193    and Pts': "?P (t',n') (s',m')" by blast
  5194  from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0" by auto
  5195  let ?st' = "Add (Mul m' t') (Mul n' s')"
  5196    have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')"
  5197    using mnz' nnz' by (simp add: ring_simps add_divide_distrib)
  5198  from Pts' have 
  5199    "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp
  5200  also have "\<dots> = ((\<lambda>(t, n). Inum (x # bs) t / real n) ((\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st')
  5201  finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2
  5202           \<in> (\<lambda>(t, n). Inum (x # bs) t / real n) `
  5203             (\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) `
  5204             set (alluopairs U)"
  5205    using ts'_U by blast
  5206 qed
  5207 
  5208 lemma \<Upsilon>_cong:
  5209   assumes lp: "isrlfm p"
  5210   and UU': "((\<lambda> (t,n). Inum (x#bs) t /real n) ` U') = ((\<lambda> ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U \<times> U))" (is "?f ` U' = ?g ` (U\<times>U)")
  5211   and U: "\<forall> (t,n) \<in> U. numbound0 t \<and> n > 0"
  5212   and U': "\<forall> (t,n) \<in> U'. numbound0 t \<and> n > 0"
  5213   shows "(\<exists> (t,n) \<in> U. \<exists> (s,m) \<in> U. Ifm (x#bs) (\<upsilon> p (Add (Mul m t) (Mul n s),2*n*m))) = (\<exists> (t,n) \<in> U'. Ifm (x#bs) (\<upsilon> p (t,n)))"
  5214   (is "?lhs = ?rhs")
  5215 proof
  5216   assume ?lhs
  5217   then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and 
  5218     Pst: "Ifm (x#bs) (\<upsilon> p (Add (Mul m t) (Mul n s),2*n*m))" by blast
  5219   let ?N = "\<lambda> t. Inum (x#bs) t"
  5220   from tnU smU U have tnb: "numbound0 t" and np: "n > 0" 
  5221     and snb: "numbound0 s" and mp:"m > 0"  by auto
  5222   let ?st= "Add (Mul m t) (Mul n s)"
  5223   from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  5224       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
  5225     from tnb snb have stnb: "numbound0 ?st" by simp
  5226   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5227    using mp np by (simp add: ring_simps add_divide_distrib)
  5228   from tnU smU UU' have "?g ((t,n),(s,m)) \<in> ?f ` U'" by blast
  5229   hence "\<exists> (t',n') \<in> U'. ?g ((t,n),(s,m)) = ?f (t',n')"
  5230     by auto (rule_tac x="(a,b)" in bexI, auto)
  5231   then obtain t' n' where tnU': "(t',n') \<in> U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast
  5232   from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
  5233   from \<upsilon>_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst 
  5234   have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
  5235   from conjunct1[OF \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]]
  5236   have "Ifm (x # bs) (\<upsilon> p (t', n')) " by (simp only: st) 
  5237   then show ?rhs using tnU' by auto 
  5238 next
  5239   assume ?rhs
  5240   then obtain t' n' where tnU': "(t',n') \<in> U'" and Pt': "Ifm (x # bs) (\<upsilon> p (t', n'))" 
  5241     by blast
  5242   from tnU' UU' have "?f (t',n') \<in> ?g ` (U\<times>U)" by blast
  5243   hence "\<exists> ((t,n),(s,m)) \<in> (U\<times>U). ?f (t',n') = ?g ((t,n),(s,m))" 
  5244     by auto (rule_tac x="(a,b)" in bexI, auto)
  5245   then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and 
  5246     th: "?f (t',n') = ?g((t,n),(s,m)) "by blast
  5247     let ?N = "\<lambda> t. Inum (x#bs) t"
  5248   from tnU smU U have tnb: "numbound0 t" and np: "n > 0" 
  5249     and snb: "numbound0 s" and mp:"m > 0"  by auto
  5250   let ?st= "Add (Mul m t) (Mul n s)"
  5251   from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  5252       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
  5253     from tnb snb have stnb: "numbound0 ?st" by simp
  5254   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5255    using mp np by (simp add: ring_simps add_divide_distrib)
  5256   from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
  5257   from \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt'
  5258   have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
  5259   with \<upsilon>_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast
  5260 qed
  5261   
  5262 lemma ferrack01: 
  5263   assumes qf: "qfree p"
  5264   shows "((\<exists> x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (Ifm bs (ferrack01 p))) \<and> qfree (ferrack01 p)" (is "(?lhs = ?rhs) \<and> _")
  5265 proof-
  5266   let ?I = "\<lambda> x p. Ifm (x#bs) p"
  5267   let ?N = "\<lambda> t. Inum (x#bs) t"
  5268   let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)"
  5269   let ?U = "\<Upsilon> ?q"
  5270   let ?Up = "alluopairs ?U"
  5271   let ?g = "\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)"
  5272   let ?S = "map ?g ?Up"
  5273   let ?SS = "map simp_num_pair ?S"
  5274   let ?Y = "remdups ?SS"
  5275   let ?f= "(\<lambda> (t,n). ?N t / real n)"
  5276   let ?h = "\<lambda> ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2"
  5277   let ?F = "\<lambda> p. \<exists> a \<in> set (\<Upsilon> p). \<exists> b \<in> set (\<Upsilon> p). ?I x (\<upsilon> p (?g(a,b)))"
  5278   let ?ep = "evaldjf (\<upsilon> ?q) ?Y"
  5279   from rlfm_l[OF qf] have lq: "isrlfm ?q" 
  5280     by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def igcd_def)
  5281   from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \<le> (set ?U \<times> set ?U)" by simp
  5282   from \<Upsilon>_l[OF lq] have U_l: "\<forall> (t,n) \<in> set ?U. numbound0 t \<and> n > 0" .
  5283   from U_l UpU 
  5284   have Up_: "\<forall> ((t,n),(s,m)) \<in> set ?Up. numbound0 t \<and> n> 0 \<and> numbound0 s \<and> m > 0" by auto
  5285   hence Snb: "\<forall> (t,n) \<in> set ?S. numbound0 t \<and> n > 0 "
  5286     by (auto simp add: mult_pos_pos)
  5287   have Y_l: "\<forall> (t,n) \<in> set ?Y. numbound0 t \<and> n > 0" 
  5288   proof-
  5289     { fix t n assume tnY: "(t,n) \<in> set ?Y" 
  5290       hence "(t,n) \<in> set ?SS" by simp
  5291       hence "\<exists> (t',n') \<in> set ?S. simp_num_pair (t',n') = (t,n)"
  5292 	by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
  5293       then obtain t' n' where tn'S: "(t',n') \<in> set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast
  5294       from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto
  5295       from simp_num_pair_l[OF tnb np tns]
  5296       have "numbound0 t \<and> n > 0" . }
  5297     thus ?thesis by blast
  5298   qed
  5299 
  5300   have YU: "(?f ` set ?Y) = (?h ` (set ?U \<times> set ?U))"
  5301   proof-
  5302      from simp_num_pair_ci[where bs="x#bs"] have 
  5303     "\<forall>x. (?f o simp_num_pair) x = ?f x" by auto
  5304      hence th: "?f o simp_num_pair = ?f" using ext by blast
  5305     have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose)
  5306     also have "\<dots> = (?f ` set ?S)" by (simp add: th)
  5307     also have "\<dots> = ((?f o ?g) ` set ?Up)" 
  5308       by (simp only: set_map o_def image_compose[symmetric])
  5309     also have "\<dots> = (?h ` (set ?U \<times> set ?U))"
  5310       using \<Upsilon>_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast
  5311     finally show ?thesis .
  5312   qed
  5313   have "\<forall> (t,n) \<in> set ?Y. bound0 (\<upsilon> ?q (t,n))"
  5314   proof-
  5315     { fix t n assume tnY: "(t,n) \<in> set ?Y"
  5316       with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto
  5317       from \<upsilon>_I[OF lq np tnb]
  5318     have "bound0 (\<upsilon> ?q (t,n))"  by simp}
  5319     thus ?thesis by blast
  5320   qed
  5321   hence ep_nb: "bound0 ?ep"  using evaldjf_bound0[where xs="?Y" and f="\<upsilon> ?q"]
  5322     by auto
  5323 
  5324   from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q"
  5325     by (simp only: split_def fst_conv snd_conv)
  5326   also have "\<dots> = (\<exists> (t,n) \<in> set ?Y. ?I x (\<upsilon> ?q (t,n)))" using \<Upsilon>_cong[OF lq YU U_l Y_l]
  5327     by (simp only: split_def fst_conv snd_conv) 
  5328   also have "\<dots> = (Ifm (x#bs) ?ep)" 
  5329     using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\<upsilon> ?q",symmetric]
  5330     by (simp only: split_def pair_collapse)
  5331   also have "\<dots> = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast
  5332   finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def)
  5333   from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def)
  5334   with lr show ?thesis by blast
  5335 qed
  5336 
  5337 lemma cp_thm': 
  5338   assumes lp: "iszlfm p (real (i::int)#bs)"
  5339   and up: "d\<beta> p 1" and dd: "d\<delta> p d" and dp: "d > 0"
  5340   shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (real i#bs)) ` set (\<beta> p). Ifm ((b+real j)#bs) p))"
  5341   using cp_thm[OF lp up dd dp] by auto
  5342 
  5343 constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int"
  5344   "unit p \<equiv> (let p' = zlfm p ; l = \<zeta> p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\<beta> p' l); d = \<delta> q;
  5345              B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
  5346              in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
  5347 
  5348 lemma unit: assumes qf: "qfree p"
  5349   shows "\<And> q B d. unit p = (q,B,d) \<Longrightarrow> ((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and> (Inum (real i#bs)) ` set B = (Inum (real i#bs)) ` set (\<beta> q) \<and> d\<beta> q 1 \<and> d\<delta> q d \<and> d >0 \<and> iszlfm q (real (i::int)#bs) \<and> (\<forall> b\<in> set B. numbound0 b)"
  5350 proof-
  5351   fix q B d 
  5352   assume qBd: "unit p = (q,B,d)"
  5353   let ?thes = "((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and>
  5354     Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\<beta> q) \<and>
  5355     d\<beta> q 1 \<and> d\<delta> q d \<and> 0 < d \<and> iszlfm q (real i # bs) \<and> (\<forall> b\<in> set B. numbound0 b)"
  5356   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5357   let ?p' = "zlfm p"
  5358   let ?l = "\<zeta> ?p'"
  5359   let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a\<beta> ?p' ?l)"
  5360   let ?d = "\<delta> ?q"
  5361   let ?B = "set (\<beta> ?q)"
  5362   let ?B'= "remdups (map simpnum (\<beta> ?q))"
  5363   let ?A = "set (\<alpha> ?q)"
  5364   let ?A'= "remdups (map simpnum (\<alpha> ?q))"
  5365   from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] 
  5366   have pp': "\<forall> i. ?I i ?p' = ?I i p" by auto
  5367   from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]]
  5368   have lp': "\<forall> (i::int). iszlfm ?p' (real i#bs)" by simp 
  5369   hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp
  5370   from lp' \<zeta>[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d\<beta> ?p' ?l" by auto
  5371   from a\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp'
  5372   have pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by (simp add: int_rdvd_iff) 
  5373   from lp'' lp a\<beta>[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d\<beta> ?q 1" 
  5374     by (auto simp add: isint_def)
  5375   from \<delta>[OF lq] have dp:"?d >0" and dd: "d\<delta> ?q ?d" by blast+
  5376   let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
  5377   have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_compose) 
  5378   also have "\<dots> = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto
  5379   finally have BB': "?N ` set ?B' = ?N ` ?B" .
  5380   have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_compose) 
  5381   also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto
  5382   finally have AA': "?N ` set ?A' = ?N ` ?A" .
  5383   from \<beta>_numbound0[OF lq] have B_nb:"\<forall> b\<in> set ?B'. numbound0 b"
  5384     by (simp add: simpnum_numbound0)
  5385   from \<alpha>_l[OF lq] have A_nb: "\<forall> b\<in> set ?A'. numbound0 b"
  5386     by (simp add: simpnum_numbound0)
  5387     {assume "length ?B' \<le> length ?A'"
  5388     hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
  5389       using qBd by (auto simp add: Let_def unit_def)
  5390     with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<beta> q)" 
  5391       and bn: "\<forall>b\<in> set B. numbound0 b" by simp+ 
  5392   with pq_ex dp uq dd lq q d have ?thes by simp}
  5393   moreover 
  5394   {assume "\<not> (length ?B' \<le> length ?A')"
  5395     hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
  5396       using qBd by (auto simp add: Let_def unit_def)
  5397     with AA' mirror\<alpha>\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)" 
  5398       and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
  5399     from mirror_ex[OF lq] pq_ex q 
  5400     have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
  5401     from lq uq q mirror_d\<beta> [where p="?q" and bs="bs" and a="real i"]
  5402     have lq': "iszlfm q (real i#bs)" and uq: "d\<beta> q 1" by auto
  5403     from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq:"d\<delta> q d " by auto
  5404     from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp
  5405   }
  5406   ultimately show ?thes by blast
  5407 qed
  5408     (* Cooper's Algorithm *)
  5409 
  5410 constdefs cooper :: "fm \<Rightarrow> fm"
  5411   "cooper p \<equiv> 
  5412   (let (q,B,d) = unit p; js = iupt (1,d);
  5413        mq = simpfm (minusinf q);
  5414        md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) js
  5415    in if md = T then T else
  5416     (let qd = evaldjf (\<lambda> t. simpfm (subst0 t q)) 
  5417                                (remdups (map (\<lambda> (b,j). simpnum (Add b (C j))) 
  5418                                             (allpairs Pair B js)))
  5419      in decr (disj md qd)))"
  5420 lemma cooper: assumes qf: "qfree p"
  5421   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \<and> qfree (cooper p)" 
  5422   (is "(?lhs = ?rhs) \<and> _")
  5423 proof-
  5424 
  5425   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5426   let ?q = "fst (unit p)"
  5427   let ?B = "fst (snd(unit p))"
  5428   let ?d = "snd (snd (unit p))"
  5429   let ?js = "iupt (1,?d)"
  5430   let ?mq = "minusinf ?q"
  5431   let ?smq = "simpfm ?mq"
  5432   let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
  5433   let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
  5434   let ?bjs = "allpairs Pair ?B ?js"
  5435   let ?sbjs = "map (\<lambda> (b,j). simpnum (Add b (C j))) (allpairs Pair ?B ?js)"
  5436   let ?qd = "evaldjf (\<lambda> t. simpfm (subst0 t ?q)) (remdups ?sbjs)"
  5437   have qbf:"unit p = (?q,?B,?d)" by simp
  5438   from unit[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and 
  5439     B:"?N ` set ?B = ?N ` set (\<beta> ?q)" and 
  5440     uq:"d\<beta> ?q 1" and dd: "d\<delta> ?q ?d" and dp: "?d > 0" and 
  5441     lq: "iszlfm ?q (real i#bs)" and 
  5442     Bn: "\<forall> b\<in> set ?B. numbound0 b" by auto
  5443   from zlin_qfree[OF lq] have qfq: "qfree ?q" .
  5444   from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
  5445   have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
  5446   hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
  5447     by (auto simp only: subst0_bound0[OF qfmq])
  5448   hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
  5449     by (auto simp add: simpfm_bound0)
  5450   from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
  5451   from Bn jsnb have "\<forall> (b,j) \<in> set (allpairs Pair ?B ?js). numbound0 (Add b (C j))"
  5452     by (simp add: allpairs_set)
  5453   hence "\<forall> (b,j) \<in> set (allpairs Pair ?B ?js). numbound0 (simpnum (Add b (C j)))"
  5454     using simpnum_numbound0 by blast 
  5455   hence "\<forall> t \<in> set ?sbjs. numbound0 t" by simp
  5456   hence "\<forall> t \<in> set (remdups ?sbjs). bound0 (subst0 t ?q)"
  5457     using subst0_bound0[OF qfq] by auto 
  5458   hence th': "\<forall> t \<in> set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))"
  5459     using simpfm_bound0 by blast
  5460   from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp
  5461   from mdb qdb 
  5462   have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
  5463   from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B
  5464   have "?lhs = (\<exists> j\<in> {1.. ?d}. ?I j ?mq \<or> (\<exists> b\<in> ?N ` set ?B. Ifm ((b+ real j)#bs) ?q))" by auto
  5465   also have "\<dots> = ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> (b,j) \<in> (?N ` set ?B \<times> set ?js). Ifm ((b+ real j)#bs) ?q))" apply (simp only: iupt_set simpfm) by auto
  5466   also have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> (\<lambda> (b,j). ?N (Add b (C j))) ` set (allpairs Pair ?B ?js). Ifm (t #bs) ?q))" by (simp only: allpairs_set) simp
  5467   also have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> (\<lambda> (b,j). ?N (simpnum (Add b (C j)))) ` set (allpairs Pair ?B ?js). Ifm (t #bs) ?q))" by (simp only: simpnum_ci)
  5468   also  have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> set ?sbjs. Ifm (?N t #bs) ?q))" 
  5469     by (auto simp add: split_def) 
  5470   also have "\<dots> = ((\<exists> j\<in> set ?js. (\<lambda> j. ?I i (simpfm (subst0 (C j) ?smq))) j) \<or> (\<exists> t \<in> set (remdups ?sbjs). (\<lambda> t. ?I i (simpfm (subst0 t ?q))) t))" by (simp only: simpfm subst0_I[OF qfq] simpfm Inum.simps subst0_I[OF qfmq] set_remdups)
  5471   also have "\<dots> = ((?I i (evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js)) \<or> (?I i (evaldjf (\<lambda> t. simpfm (subst0 t ?q)) (remdups ?sbjs))))" by (simp only: evaldjf_ex)
  5472   finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by (simp add: disj) 
  5473   hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp
  5474   {assume mdT: "?md = T"
  5475     hence cT:"cooper p = T" 
  5476       by (simp only: cooper_def unit_def split_def Let_def if_True) simp
  5477     from mdT mdqd have lhs:"?lhs" by (auto simp add: disj)
  5478     from mdT have "?rhs" by (simp add: cooper_def unit_def split_def)
  5479     with lhs cT have ?thesis by simp }
  5480   moreover
  5481   {assume mdT: "?md \<noteq> T" hence "cooper p = decr (disj ?md ?qd)" 
  5482       by (simp only: cooper_def unit_def split_def Let_def if_False) 
  5483     with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
  5484   ultimately show ?thesis by blast
  5485 qed
  5486 
  5487 lemma DJcooper: 
  5488   assumes qf: "qfree p"
  5489   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \<and> qfree (DJ cooper p)"
  5490 proof-
  5491   from cooper have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (cooper p)" by  blast
  5492   from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast
  5493   have "Ifm bs (DJ cooper p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (cooper q))" 
  5494      by (simp add: DJ_def evaldjf_ex)
  5495   also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs)  q)" 
  5496     using cooper disjuncts_qf[OF qf] by blast
  5497   also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
  5498   finally show ?thesis using thqf by blast
  5499 qed
  5500 
  5501     (* Redy and Loveland *)
  5502 
  5503 lemma \<sigma>\<rho>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
  5504   shows "Ifm (a#bs) (\<sigma>\<rho> p (t,c)) = Ifm (a#bs) (\<sigma>\<rho> p (t',c))"
  5505   using lp 
  5506   by (induct p rule: iszlfm.induct, auto simp add: tt')
  5507 
  5508 lemma \<sigma>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
  5509   shows "Ifm (a#bs) (\<sigma> p c t) = Ifm (a#bs) (\<sigma> p c t')"
  5510   by (simp add: \<sigma>_def tt' \<sigma>\<rho>_cong[OF lp tt'])
  5511 
  5512 lemma \<rho>_cong: assumes lp: "iszlfm p (a#bs)" 
  5513   and RR: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R =  (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
  5514   shows "(\<exists> (e,c) \<in> R. \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))) = (\<exists> (e,c) \<in> set (\<rho> p). \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j))))"
  5515   (is "?lhs = ?rhs")
  5516 proof
  5517   let ?d = "\<delta> p"
  5518   assume ?lhs then obtain e c j where ecR: "(e,c) \<in> R" and jD:"j \<in> {1 .. c*?d}" 
  5519     and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
  5520   from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" by auto
  5521   hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" using RR by simp
  5522   hence "\<exists> (e',c') \<in> set (\<rho> p). Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
  5523   then obtain e' c' where ecRo:"(e',c') \<in> set (\<rho> p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
  5524     and cc':"c = c'" by blast
  5525   from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
  5526   
  5527   from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
  5528   from ecRo jD px' cc'  show ?rhs apply auto 
  5529     by (rule_tac x="(e', c')" in bexI,simp_all)
  5530   (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
  5531 next
  5532   let ?d = "\<delta> p"
  5533   assume ?rhs then obtain e c j where ecR: "(e,c) \<in> set (\<rho> p)" and jD:"j \<in> {1 .. c*?d}" 
  5534     and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
  5535   from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" by auto
  5536   hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp
  5537   hence "\<exists> (e',c') \<in> R. Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
  5538   then obtain e' c' where ecRo:"(e',c') \<in> R" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
  5539     and cc':"c = c'" by blast
  5540   from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
  5541   from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
  5542   from ecRo jD px' cc'  show ?lhs apply auto 
  5543     by (rule_tac x="(e', c')" in bexI,simp_all)
  5544   (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
  5545 qed
  5546 
  5547 lemma rl_thm': 
  5548   assumes lp: "iszlfm p (real (i::int)#bs)" 
  5549   and R: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R =  (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
  5550   shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> (e,c) \<in> R. \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))))"
  5551   using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp 
  5552 
  5553 constdefs chooset:: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int"
  5554   "chooset p \<equiv> (let q = zlfm p ; d = \<delta> q;
  5555              B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ; 
  5556              a = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<alpha>\<rho> q))
  5557              in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
  5558 
  5559 lemma chooset: assumes qf: "qfree p"
  5560   shows "\<And> q B d. chooset p = (q,B,d) \<Longrightarrow> ((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and> ((\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set B = (\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set (\<rho> q)) \<and> (\<delta> q = d) \<and> d >0 \<and> iszlfm q (real (i::int)#bs) \<and> (\<forall> (e,c)\<in> set B. numbound0 e \<and> c>0)"
  5561 proof-
  5562   fix q B d 
  5563   assume qBd: "chooset p = (q,B,d)"
  5564   let ?thes = "((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and> ((\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set B = (\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set (\<rho> q)) \<and> (\<delta> q = d) \<and> d >0 \<and> iszlfm q (real (i::int)#bs) \<and> (\<forall> (e,c)\<in> set B. numbound0 e \<and> c>0)" 
  5565   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5566   let ?q = "zlfm p"
  5567   let ?d = "\<delta> ?q"
  5568   let ?B = "set (\<rho> ?q)"
  5569   let ?f = "\<lambda> (t,k). (simpnum t,k)"
  5570   let ?B'= "remdups (map ?f (\<rho> ?q))"
  5571   let ?A = "set (\<alpha>\<rho> ?q)"
  5572   let ?A'= "remdups (map ?f (\<alpha>\<rho> ?q))"
  5573   from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] 
  5574   have pp': "\<forall> i. ?I i ?q = ?I i p" by auto
  5575   hence pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by simp 
  5576   from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"]
  5577   have lq: "iszlfm ?q (real (i::int)#bs)" . 
  5578   from \<delta>[OF lq] have dp:"?d >0" by blast
  5579   let ?N = "\<lambda> (t,c). (Inum (real (i::int)#bs) t,c)"
  5580   have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_compose)
  5581   also have "\<dots> = ?N ` ?B"
  5582     by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def)
  5583   finally have BB': "?N ` set ?B' = ?N ` ?B" .
  5584   have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_compose) 
  5585   also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"]
  5586     by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) 
  5587   finally have AA': "?N ` set ?A' = ?N ` ?A" .
  5588   from \<rho>_l[OF lq] have B_nb:"\<forall> (e,c)\<in> set ?B'. numbound0 e \<and> c > 0"
  5589     by (simp add: simpnum_numbound0 split_def)
  5590   from \<alpha>\<rho>_l[OF lq] have A_nb: "\<forall> (e,c)\<in> set ?A'. numbound0 e \<and> c > 0"
  5591     by (simp add: simpnum_numbound0 split_def)
  5592     {assume "length ?B' \<le> length ?A'"
  5593     hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
  5594       using qBd by (auto simp add: Let_def chooset_def)
  5595     with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<rho> q)" 
  5596       and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto
  5597   with pq_ex dp lq q d have ?thes by simp}
  5598   moreover 
  5599   {assume "\<not> (length ?B' \<le> length ?A')"
  5600     hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
  5601       using qBd by (auto simp add: Let_def chooset_def)
  5602     with AA' mirror_\<alpha>\<rho>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<rho> q)" 
  5603       and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto 
  5604     from mirror_ex[OF lq] pq_ex q 
  5605     have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
  5606     from lq q mirror_l [where p="?q" and bs="bs" and a="real i"]
  5607     have lq': "iszlfm q (real i#bs)" by auto
  5608     from mirror_\<delta>[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp
  5609   }
  5610   ultimately show ?thes by blast
  5611 qed
  5612 
  5613 constdefs stage:: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm"
  5614   "stage p d \<equiv> (\<lambda> (e,c). evaldjf (\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))) (iupt (1,c*d)))"
  5615 lemma stage:
  5616   shows "Ifm bs (stage p d (e,c)) = (\<exists> j\<in>{1 .. c*d}. Ifm bs (\<sigma> p c (Add e (C j))))"
  5617   by (unfold stage_def split_def ,simp only: evaldjf_ex iupt_set simpfm) simp
  5618 
  5619 lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e"
  5620   shows "bound0 (stage p d (e,c))"
  5621 proof-
  5622   let ?f = "\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))"
  5623   have th: "\<forall> j\<in> set (iupt(1,c*d)). bound0 (?f j)"
  5624   proof
  5625     fix j
  5626     from nb have nb':"numbound0 (Add e (C j))" by simp
  5627     from simpfm_bound0[OF \<sigma>_nb[OF lp nb', where k="c"]]
  5628     show "bound0 (simpfm (\<sigma> p c (Add e (C j))))" .
  5629   qed
  5630   from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp
  5631 qed
  5632 
  5633 constdefs redlove:: "fm \<Rightarrow> fm"
  5634   "redlove p \<equiv> 
  5635   (let (q,B,d) = chooset p;
  5636        mq = simpfm (minusinf q);
  5637        md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) (iupt (1,d))
  5638    in if md = T then T else
  5639     (let qd = evaldjf (stage q d) B
  5640      in decr (disj md qd)))"
  5641 
  5642 lemma redlove: assumes qf: "qfree p"
  5643   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \<and> qfree (redlove p)" 
  5644   (is "(?lhs = ?rhs) \<and> _")
  5645 proof-
  5646 
  5647   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5648   let ?q = "fst (chooset p)"
  5649   let ?B = "fst (snd(chooset p))"
  5650   let ?d = "snd (snd (chooset p))"
  5651   let ?js = "iupt (1,?d)"
  5652   let ?mq = "minusinf ?q"
  5653   let ?smq = "simpfm ?mq"
  5654   let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
  5655   let ?N = "\<lambda> (t,k). (Inum (real (i::int)#bs) t,k)"
  5656   let ?qd = "evaldjf (stage ?q ?d) ?B"
  5657   have qbf:"chooset p = (?q,?B,?d)" by simp
  5658   from chooset[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and 
  5659     B:"?N ` set ?B = ?N ` set (\<rho> ?q)" and dd: "\<delta> ?q = ?d" and dp: "?d > 0" and 
  5660     lq: "iszlfm ?q (real i#bs)" and 
  5661     Bn: "\<forall> (e,c)\<in> set ?B. numbound0 e \<and> c > 0" by auto
  5662   from zlin_qfree[OF lq] have qfq: "qfree ?q" .
  5663   from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
  5664   have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
  5665   hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
  5666     by (auto simp only: subst0_bound0[OF qfmq])
  5667   hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
  5668     by (auto simp add: simpfm_bound0)
  5669   from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
  5670   from Bn stage_nb[OF lq] have th:"\<forall> x \<in> set ?B. bound0 (stage ?q ?d x)" by auto
  5671   from evaldjf_bound0[OF th]  have qdb: "bound0 ?qd" .
  5672   from mdb qdb 
  5673   have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
  5674   from trans [OF pq_ex rl_thm'[OF lq B]] dd
  5675   have "?lhs = ((\<exists> j\<in> {1.. ?d}. ?I j ?mq) \<or> (\<exists> (e,c)\<in> set ?B. \<exists> j\<in> {1 .. c*?d}. Ifm (real i#bs) (\<sigma> ?q c (Add e (C j)))))" by auto
  5676   also have "\<dots> = ((\<exists> j\<in> {1.. ?d}. ?I j ?smq) \<or> (\<exists> (e,c)\<in> set ?B. ?I i (stage ?q ?d (e,c) )))" 
  5677     by (simp add: simpfm stage split_def)
  5678   also have "\<dots> = ((\<exists> j\<in> {1 .. ?d}. ?I i (subst0 (C j) ?smq))  \<or> ?I i ?qd)"
  5679     by (simp add: evaldjf_ex subst0_I[OF qfmq])
  5680   finally have mdqd:"?lhs = (?I i ?md \<or> ?I i ?qd)" by (simp only: evaldjf_ex iupt_set simpfm) 
  5681   also have "\<dots> = (?I i (disj ?md ?qd))" by (simp add: disj)
  5682   also have "\<dots> = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) 
  5683   finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . 
  5684   {assume mdT: "?md = T"
  5685     hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def)
  5686     from mdT have lhs:"?lhs" using mdqd by simp 
  5687     from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def)
  5688     with lhs cT have ?thesis by simp }
  5689   moreover
  5690   {assume mdT: "?md \<noteq> T" hence "redlove p = decr (disj ?md ?qd)" 
  5691       by (simp add: redlove_def chooset_def split_def Let_def)
  5692     with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
  5693   ultimately show ?thesis by blast
  5694 qed
  5695 
  5696 lemma DJredlove: 
  5697   assumes qf: "qfree p"
  5698   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \<and> qfree (DJ redlove p)"
  5699 proof-
  5700   from redlove have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (redlove p)" by  blast
  5701   from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast
  5702   have "Ifm bs (DJ redlove p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (redlove q))" 
  5703      by (simp add: DJ_def evaldjf_ex)
  5704   also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs)  q)" 
  5705     using redlove disjuncts_qf[OF qf] by blast
  5706   also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
  5707   finally show ?thesis using thqf by blast
  5708 qed
  5709 
  5710 
  5711 lemma exsplit_qf: assumes qf: "qfree p"
  5712   shows "qfree (exsplit p)"
  5713 using qf by (induct p rule: exsplit.induct, auto)
  5714 
  5715 constdefs mircfr :: "fm \<Rightarrow> fm"
  5716 "mircfr \<equiv> (DJ cooper) o ferrack01 o simpfm o exsplit"
  5717 
  5718 constdefs mirlfr :: "fm \<Rightarrow> fm"
  5719 "mirlfr \<equiv> (DJ redlove) o ferrack01 o simpfm o exsplit"
  5720 
  5721 
  5722 lemma mircfr: "\<forall> bs p. qfree p \<longrightarrow> qfree (mircfr p) \<and> Ifm bs (mircfr p) = Ifm bs (E p)"
  5723 proof(clarsimp simp del: Ifm.simps)
  5724   fix bs p
  5725   assume qf: "qfree p"
  5726   show "qfree (mircfr p)\<and>(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
  5727   proof-
  5728     let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
  5729     have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" 
  5730       using splitex[OF qf] by simp
  5731     with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\<exists> (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+
  5732     with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def)
  5733   qed
  5734 qed
  5735   
  5736 lemma mirlfr: "\<forall> bs p. qfree p \<longrightarrow> qfree(mirlfr p) \<and> Ifm bs (mirlfr p) = Ifm bs (E p)"
  5737 proof(clarsimp simp del: Ifm.simps)
  5738   fix bs p
  5739   assume qf: "qfree p"
  5740   show "qfree (mirlfr p)\<and>(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
  5741   proof-
  5742     let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
  5743     have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" 
  5744       using splitex[OF qf] by simp
  5745     with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\<exists> (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+
  5746     with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def)
  5747   qed
  5748 qed
  5749   
  5750 constdefs mircfrqe:: "fm \<Rightarrow> fm"
  5751   "mircfrqe \<equiv> (\<lambda> p. qelim (prep p) mircfr)"
  5752 
  5753 constdefs mirlfrqe:: "fm \<Rightarrow> fm"
  5754   "mirlfrqe \<equiv> (\<lambda> p. qelim (prep p) mirlfr)"
  5755 
  5756 theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \<and> qfree (mircfrqe p)"
  5757   using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def)
  5758 
  5759 theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \<and> qfree (mirlfrqe p)"
  5760   using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def)
  5761 
  5762 declare zdvd_iff_zmod_eq_0 [code]
  5763 declare max_def [code unfold]
  5764 
  5765 definition
  5766   "test1 (u\<Colon>unit) = mircfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))"
  5767 
  5768 definition
  5769   "test2 (u\<Colon>unit) = mircfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))"
  5770 
  5771 definition
  5772   "test3 (u\<Colon>unit) = mirlfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))"
  5773 
  5774 definition
  5775   "test4 (u\<Colon>unit) = mirlfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))"
  5776 
  5777 definition
  5778   "test5 (u\<Colon>unit) = mircfrqe (A(E(And (Ge(Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg(Bound 0))))))))"
  5779 
  5780 code_gen mircfrqe mirlfrqe test1 test2 test3 test4 test5
  5781   in SML module_name Mir
  5782 
  5783 (*code_gen mircfrqe mirlfrqe
  5784   in SML module_name Mir file "raw_mir.ML"*)
  5785 
  5786 ML "set Toplevel.timing"
  5787 ML "Mir.test1 ()"
  5788 ML "Mir.test2 ()"
  5789 ML "Mir.test3 ()"
  5790 ML "Mir.test4 ()"
  5791 ML "Mir.test5 ()"
  5792 ML "reset Toplevel.timing"
  5793 
  5794 use "mireif.ML"
  5795 oracle mircfr_oracle ("term") = ReflectedMir.mircfr_oracle
  5796 oracle mirlfr_oracle ("term") = ReflectedMir.mirlfr_oracle
  5797 use "mirtac.ML"
  5798 setup "MirTac.setup"
  5799 
  5800 ML "set Toplevel.timing"
  5801 
  5802 lemma "ALL (x::real). (\<lfloor>x\<rfloor> = \<lceil>x\<rceil> = (x = real \<lfloor>x\<rfloor>))"
  5803 apply mir
  5804 done
  5805 
  5806 lemma "ALL (x::real). real (2::int)*x - (real (1::int)) < real \<lfloor>x\<rfloor> + real \<lceil>x\<rceil> \<and> real \<lfloor>x\<rfloor> + real \<lceil>x\<rceil>  \<le> real (2::int)*x + (real (1::int))"
  5807 apply mir
  5808 done
  5809 
  5810 lemma "ALL (x::real). 2*\<lfloor>x\<rfloor> \<le> \<lfloor>2*x\<rfloor> \<and> \<lfloor>2*x\<rfloor> \<le> 2*\<lfloor>x+1\<rfloor>"
  5811 apply mir 
  5812 done
  5813 
  5814 lemma "ALL (x::real). \<exists>y \<le> x. (\<lfloor>x\<rfloor> = \<lceil>y\<rceil>)"
  5815 apply mir
  5816 done
  5817 
  5818 ML "reset Toplevel.timing"
  5819 
  5820 end