src/HOL/Complex/ex/MIR.thy
author nipkow
Mon, 20 Aug 2007 11:18:18 +0200
changeset 24336 fff40259f336
parent 24249 1f60b45c5f97
child 24348 c708ea5b109a
permissions -rw-r--r--
removed allpairs
     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 fun alluopairs:: "'a list \<Rightarrow> ('a \<times> 'a) list" where 
    15   "alluopairs [] = []"
    16 | "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)"
    17 
    18 lemma alluopairs_set1: "set (alluopairs xs) \<le> {(x,y). x\<in> set xs \<and> y\<in> set xs}"
    19 by (induct xs, auto)
    20 
    21 lemma alluopairs_set:
    22   "\<lbrakk>x\<in> set xs ; y \<in> set xs\<rbrakk> \<Longrightarrow> (x,y) \<in> set (alluopairs xs) \<or> (y,x) \<in> set (alluopairs xs) "
    23 by (induct xs, auto)
    24 
    25 lemma alluopairs_ex:
    26   assumes Pc: "\<forall> x y. P x y = P y x"
    27   shows "(\<exists> x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists> (x,y) \<in> set (alluopairs xs). P x y)"
    28 proof
    29   assume "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y"
    30   then obtain x y where x: "x \<in> set xs" and y:"y \<in> set xs" and P: "P x y"  by blast
    31   from alluopairs_set[OF x y] P Pc show"\<exists>(x, y)\<in>set (alluopairs xs). P x y" 
    32     by auto
    33 next
    34   assume "\<exists>(x, y)\<in>set (alluopairs xs). P x y"
    35   then obtain "x" and "y"  where xy:"(x,y) \<in> set (alluopairs xs)" and P: "P x y" by blast+
    36   from xy have "x \<in> set xs \<and> y\<in> set xs" using alluopairs_set1 by blast
    37   with P show "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y" by blast
    38 qed
    39 
    40   (* generate a list from i to j*)
    41 consts iupt :: "int \<times> int \<Rightarrow> int list"
    42 recdef iupt "measure (\<lambda> (i,j). nat (j-i +1))" 
    43   "iupt (i,j) = (if j <i then [] else (i# iupt(i+1, j)))"
    44 
    45 lemma iupt_set: "set (iupt(i,j)) = {i .. j}"
    46 proof(induct rule: iupt.induct)
    47   case (1 a b)
    48   show ?case
    49     using prems by (simp add: simp_from_to)
    50 qed
    51 
    52 lemma nth_pos2: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
    53 using Nat.gr0_conv_Suc
    54 by clarsimp
    55 
    56 
    57 lemma myl: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a \<le> b) = (0 \<le> b - a)" 
    58 proof(clarify)
    59   fix x y ::"'a"
    60   have "(x \<le> y) = (x - y \<le> 0)" by (simp only: le_iff_diff_le_0[where a="x" and b="y"])
    61   also have "\<dots> = (- (y - x) \<le> 0)" by simp
    62   also have "\<dots> = (0 \<le> y - x)" by (simp only: neg_le_0_iff_le[where a="y-x"])
    63   finally show "(x \<le> y) = (0 \<le> y - x)" .
    64 qed
    65 
    66 lemma myless: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a < b) = (0 < b - a)" 
    67 proof(clarify)
    68   fix x y ::"'a"
    69   have "(x < y) = (x - y < 0)" by (simp only: less_iff_diff_less_0[where a="x" and b="y"])
    70   also have "\<dots> = (- (y - x) < 0)" by simp
    71   also have "\<dots> = (0 < y - x)" by (simp only: neg_less_0_iff_less[where a="y-x"])
    72   finally show "(x < y) = (0 < y - x)" .
    73 qed
    74 
    75 lemma myeq: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a = b) = (0 = b - a)"
    76   by auto
    77 
    78   (* Maybe should be added to the library \<dots> *)
    79 lemma floor_int_eq: "(real n\<le> x \<and> x < real (n+1)) = (floor x = n)"
    80 proof( auto)
    81   assume lb: "real n \<le> x"
    82     and ub: "x < real n + 1"
    83   have "real (floor x) \<le> x" by simp 
    84   hence "real (floor x) < real (n + 1) " using ub by arith
    85   hence "floor x < n+1" by simp
    86   moreover from lb have "n \<le> floor x" using floor_mono2[where x="real n" and y="x"] 
    87     by simp ultimately show "floor x = n" by simp
    88 qed
    89 
    90 (* Periodicity of dvd *)
    91 lemma dvd_period:
    92   assumes advdd: "(a::int) dvd d"
    93   shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))"
    94   using advdd  
    95 proof-
    96   {fix x k
    97     from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"]  
    98     have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp}
    99   hence "\<forall>x.\<forall>k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))"  by simp
   100   then show ?thesis by simp
   101 qed
   102 
   103   (* The Divisibility relation between reals *)	
   104 definition
   105   rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
   106 where
   107   rdvd_def: "x rdvd y \<longleftrightarrow> (\<exists>k\<Colon>int. y = x * real k)"
   108 
   109 lemma int_rdvd_real: 
   110   shows "real (i::int) rdvd x = (i dvd (floor x) \<and> real (floor x) = x)" (is "?l = ?r")
   111 proof
   112   assume "?l" 
   113   hence th: "\<exists> k. x=real (i*k)" by (simp add: rdvd_def)
   114   hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult)
   115   with th have "\<exists> k. real (floor x) = real (i*k)" by simp
   116   hence "\<exists> k. floor x = i*k" by (simp only: real_of_int_inject)
   117   thus ?r  using th' by (simp add: dvd_def) 
   118 next
   119   assume "?r" hence "(i\<Colon>int) dvd \<lfloor>x\<Colon>real\<rfloor>" ..
   120   hence "\<exists> k. real (floor x) = real (i*k)" 
   121     by (simp only: real_of_int_inject) (simp add: dvd_def)
   122   thus ?l using prems by (simp add: rdvd_def)
   123 qed
   124 
   125 lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)"
   126 by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only :real_of_int_mult[symmetric])
   127 
   128 
   129 lemma rdvd_abs1: 
   130   "(abs (real d) rdvd t) = (real (d ::int) rdvd t)"
   131 proof
   132   assume d: "real d rdvd t"
   133   from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t" by auto
   134 
   135   from iffD2[OF zdvd_abs1] d2 have "(abs d) dvd (floor t)" by blast
   136   with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast 
   137   thus "abs (real d) rdvd t" by simp
   138 next
   139   assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp
   140   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
   141   from iffD1[OF zdvd_abs1] d2 have "d dvd floor t" by blast
   142   with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast
   143 qed
   144 
   145 lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)"
   146   apply (auto simp add: rdvd_def)
   147   apply (rule_tac x="-k" in exI, simp) 
   148   apply (rule_tac x="-k" in exI, simp)
   149 done
   150 
   151 lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)"
   152 by (auto simp add: rdvd_def)
   153 
   154 lemma rdvd_mult: 
   155   assumes knz: "k\<noteq>0"
   156   shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)"
   157 using knz by (simp add:rdvd_def)
   158 
   159 lemma rdvd_trans: assumes mn:"m rdvd n" and  nk:"n rdvd k" 
   160   shows "m rdvd k"
   161 proof-
   162   from rdvd_def mn obtain c where nmc:"n = m * real (c::int)" by auto
   163   from rdvd_def nk obtain c' where nkc:"k = n * real (c'::int)" by auto
   164   hence "k = m * real (c * c')" using nmc by simp
   165   thus ?thesis using rdvd_def by blast
   166 qed
   167 
   168   (*********************************************************************************)
   169   (****                            SHADOW SYNTAX AND SEMANTICS                  ****)
   170   (*********************************************************************************)
   171 
   172 datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num 
   173   | Mul int num | Floor num| CF int num num
   174 
   175   (* A size for num to make inductive proofs simpler*)
   176 fun num_size :: "num \<Rightarrow> nat" where
   177  "num_size (C c) = 1"
   178 | "num_size (Bound n) = 1"
   179 | "num_size (Neg a) = 1 + num_size a"
   180 | "num_size (Add a b) = 1 + num_size a + num_size b"
   181 | "num_size (Sub a b) = 3 + num_size a + num_size b"
   182 | "num_size (CN n c a) = 4 + num_size a "
   183 | "num_size (CF c a b) = 4 + num_size a + num_size b"
   184 | "num_size (Mul c a) = 1 + num_size a"
   185 | "num_size (Floor a) = 1 + num_size a"
   186 
   187   (* Semantics of numeral terms (num) *)
   188 fun Inum :: "real list \<Rightarrow> num \<Rightarrow> real" where
   189   "Inum bs (C c) = (real c)"
   190 | "Inum bs (Bound n) = bs!n"
   191 | "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)"
   192 | "Inum bs (Neg a) = -(Inum bs a)"
   193 | "Inum bs (Add a b) = Inum bs a + Inum bs b"
   194 | "Inum bs (Sub a b) = Inum bs a - Inum bs b"
   195 | "Inum bs (Mul c a) = (real c) * Inum bs a"
   196 | "Inum bs (Floor a) = real (floor (Inum bs a))"
   197 | "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b"
   198 definition "isint t bs \<equiv> real (floor (Inum bs t)) = Inum bs t"
   199 
   200 lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)"
   201 by (simp add: isint_def)
   202 
   203 lemma isint_Floor: "isint (Floor n) bs"
   204   by (simp add: isint_iff)
   205 
   206 lemma isint_Mul: "isint e bs \<Longrightarrow> isint (Mul c e) bs"
   207 proof-
   208   let ?e = "Inum bs e"
   209   let ?fe = "floor ?e"
   210   assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff)
   211   have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp
   212   also have "\<dots> = real (c* ?fe)" by (simp only: floor_real_of_int) 
   213   also have "\<dots> = real c * ?e" using efe by simp
   214   finally show ?thesis using isint_iff by simp
   215 qed
   216 
   217 lemma isint_neg: "isint e bs \<Longrightarrow> isint (Neg e) bs"
   218 proof-
   219   let ?I = "\<lambda> t. Inum bs t"
   220   assume ie: "isint e bs"
   221   hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)  
   222   have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th)
   223   also have "\<dots> = - real (floor (?I e))" by(simp add: floor_minus_real_of_int) 
   224   finally show "isint (Neg e) bs" by (simp add: isint_def th)
   225 qed
   226 
   227 lemma isint_sub: 
   228   assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs"
   229 proof-
   230   let ?I = "\<lambda> t. Inum bs t"
   231   from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)  
   232   have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th)
   233   also have "\<dots> = real (c- floor (?I e))" by(simp add: floor_minus_real_of_int) 
   234   finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th)
   235 qed
   236 
   237 lemma isint_add: assumes
   238   ai:"isint a bs" and bi: "isint b bs" shows "isint (Add a b) bs"
   239 proof-
   240   let ?a = "Inum bs a"
   241   let ?b = "Inum bs b"
   242   from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))" by simp
   243   also have "\<dots> = real (floor ?a) + real (floor ?b)" by simp
   244   also have "\<dots> = ?a + ?b" using ai bi isint_iff by simp
   245   finally show "isint (Add a b) bs" by (simp add: isint_iff)
   246 qed
   247 
   248 lemma isint_c: "isint (C j) bs"
   249   by (simp add: isint_iff)
   250 
   251 
   252     (* FORMULAE *)
   253 datatype fm  = 
   254   T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num|
   255   NOT fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
   256 
   257 
   258   (* A size for fm *)
   259 fun fmsize :: "fm \<Rightarrow> nat" where
   260  "fmsize (NOT p) = 1 + fmsize p"
   261 | "fmsize (And p q) = 1 + fmsize p + fmsize q"
   262 | "fmsize (Or p q) = 1 + fmsize p + fmsize q"
   263 | "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
   264 | "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
   265 | "fmsize (E p) = 1 + fmsize p"
   266 | "fmsize (A p) = 4+ fmsize p"
   267 | "fmsize (Dvd i t) = 2"
   268 | "fmsize (NDvd i t) = 2"
   269 | "fmsize p = 1"
   270   (* several lemmas about fmsize *)
   271 lemma fmsize_pos: "fmsize p > 0"	
   272 by (induct p rule: fmsize.induct) simp_all
   273 
   274   (* Semantics of formulae (fm) *)
   275 fun Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool" where
   276   "Ifm bs T = True"
   277 | "Ifm bs F = False"
   278 | "Ifm bs (Lt a) = (Inum bs a < 0)"
   279 | "Ifm bs (Gt a) = (Inum bs a > 0)"
   280 | "Ifm bs (Le a) = (Inum bs a \<le> 0)"
   281 | "Ifm bs (Ge a) = (Inum bs a \<ge> 0)"
   282 | "Ifm bs (Eq a) = (Inum bs a = 0)"
   283 | "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)"
   284 | "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)"
   285 | "Ifm bs (NDvd i b) = (\<not>(real i rdvd Inum bs b))"
   286 | "Ifm bs (NOT p) = (\<not> (Ifm bs p))"
   287 | "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
   288 | "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)"
   289 | "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))"
   290 | "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)"
   291 | "Ifm bs (E p) = (\<exists> x. Ifm (x#bs) p)"
   292 | "Ifm bs (A p) = (\<forall> x. Ifm (x#bs) p)"
   293 
   294 consts prep :: "fm \<Rightarrow> fm"
   295 recdef prep "measure fmsize"
   296   "prep (E T) = T"
   297   "prep (E F) = F"
   298   "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
   299   "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
   300   "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" 
   301   "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
   302   "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
   303   "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
   304   "prep (E p) = E (prep p)"
   305   "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
   306   "prep (A p) = prep (NOT (E (NOT p)))"
   307   "prep (NOT (NOT p)) = prep p"
   308   "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
   309   "prep (NOT (A p)) = prep (E (NOT p))"
   310   "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
   311   "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
   312   "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
   313   "prep (NOT p) = NOT (prep p)"
   314   "prep (Or p q) = Or (prep p) (prep q)"
   315   "prep (And p q) = And (prep p) (prep q)"
   316   "prep (Imp p q) = prep (Or (NOT p) q)"
   317   "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
   318   "prep p = p"
   319 (hints simp add: fmsize_pos)
   320 lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p"
   321 by (induct p rule: prep.induct, auto)
   322 
   323 
   324   (* Quantifier freeness *)
   325 consts qfree:: "fm \<Rightarrow> bool"
   326 recdef qfree "measure size"
   327   "qfree (E p) = False"
   328   "qfree (A p) = False"
   329   "qfree (NOT p) = qfree p" 
   330   "qfree (And p q) = (qfree p \<and> qfree q)" 
   331   "qfree (Or  p q) = (qfree p \<and> qfree q)" 
   332   "qfree (Imp p q) = (qfree p \<and> qfree q)" 
   333   "qfree (Iff p q) = (qfree p \<and> qfree q)"
   334   "qfree p = True"
   335 
   336   (* Boundedness and substitution *)
   337 consts 
   338   numbound0:: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *)
   339   bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
   340   numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *)
   341   subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *)
   342 primrec
   343   "numbound0 (C c) = True"
   344   "numbound0 (Bound n) = (n>0)"
   345   "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)"
   346   "numbound0 (Neg a) = numbound0 a"
   347   "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
   348   "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)" 
   349   "numbound0 (Mul i a) = numbound0 a"
   350   "numbound0 (Floor a) = numbound0 a"
   351   "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)" 
   352 lemma numbound0_I:
   353   assumes nb: "numbound0 a"
   354   shows "Inum (b#bs) a = Inum (b'#bs) a"
   355 using nb
   356 by (induct a rule: numbound0.induct) (auto simp add: nth_pos2)
   357 
   358 
   359 lemma numbound0_gen: 
   360   assumes nb: "numbound0 t" and ti: "isint t (x#bs)"
   361   shows "\<forall> y. isint t (y#bs)"
   362 using nb ti 
   363 proof(clarify)
   364   fix y
   365   from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def]
   366   show "isint t (y#bs)"
   367     by (simp add: isint_def)
   368 qed
   369 
   370 primrec
   371   "bound0 T = True"
   372   "bound0 F = True"
   373   "bound0 (Lt a) = numbound0 a"
   374   "bound0 (Le a) = numbound0 a"
   375   "bound0 (Gt a) = numbound0 a"
   376   "bound0 (Ge a) = numbound0 a"
   377   "bound0 (Eq a) = numbound0 a"
   378   "bound0 (NEq a) = numbound0 a"
   379   "bound0 (Dvd i a) = numbound0 a"
   380   "bound0 (NDvd i a) = numbound0 a"
   381   "bound0 (NOT p) = bound0 p"
   382   "bound0 (And p q) = (bound0 p \<and> bound0 q)"
   383   "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
   384   "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
   385   "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
   386   "bound0 (E p) = False"
   387   "bound0 (A p) = False"
   388 
   389 lemma bound0_I:
   390   assumes bp: "bound0 p"
   391   shows "Ifm (b#bs) p = Ifm (b'#bs) p"
   392 using bp numbound0_I[where b="b" and bs="bs" and b'="b'"]
   393 by (induct p rule: bound0.induct) (auto simp add: nth_pos2)
   394 
   395 primrec
   396   "numsubst0 t (C c) = (C c)"
   397   "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
   398   "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))"
   399   "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)"
   400   "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
   401   "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
   402   "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" 
   403   "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
   404   "numsubst0 t (Floor a) = Floor (numsubst0 t a)"
   405 
   406 lemma numsubst0_I:
   407   shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t"
   408   by (induct t) (simp_all add: nth_pos2)
   409 
   410 lemma numsubst0_I':
   411   assumes nb: "numbound0 a"
   412   shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t"
   413   by (induct t) (simp_all add: nth_pos2 numbound0_I[OF nb, where b="b" and b'="b'"])
   414 
   415 
   416 primrec
   417   "subst0 t T = T"
   418   "subst0 t F = F"
   419   "subst0 t (Lt a) = Lt (numsubst0 t a)"
   420   "subst0 t (Le a) = Le (numsubst0 t a)"
   421   "subst0 t (Gt a) = Gt (numsubst0 t a)"
   422   "subst0 t (Ge a) = Ge (numsubst0 t a)"
   423   "subst0 t (Eq a) = Eq (numsubst0 t a)"
   424   "subst0 t (NEq a) = NEq (numsubst0 t a)"
   425   "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
   426   "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
   427   "subst0 t (NOT p) = NOT (subst0 t p)"
   428   "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
   429   "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
   430   "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
   431   "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
   432 
   433 lemma subst0_I: assumes qfp: "qfree p"
   434   shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p"
   435   using qfp numsubst0_I[where b="b" and bs="bs" and a="a"]
   436   by (induct p) (simp_all add: nth_pos2 )
   437 
   438 consts 
   439   decrnum:: "num \<Rightarrow> num" 
   440   decr :: "fm \<Rightarrow> fm"
   441 
   442 recdef decrnum "measure size"
   443   "decrnum (Bound n) = Bound (n - 1)"
   444   "decrnum (Neg a) = Neg (decrnum a)"
   445   "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
   446   "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
   447   "decrnum (Mul c a) = Mul c (decrnum a)"
   448   "decrnum (Floor a) = Floor (decrnum a)"
   449   "decrnum (CN n c a) = CN (n - 1) c (decrnum a)"
   450   "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)"
   451   "decrnum a = a"
   452 
   453 recdef decr "measure size"
   454   "decr (Lt a) = Lt (decrnum a)"
   455   "decr (Le a) = Le (decrnum a)"
   456   "decr (Gt a) = Gt (decrnum a)"
   457   "decr (Ge a) = Ge (decrnum a)"
   458   "decr (Eq a) = Eq (decrnum a)"
   459   "decr (NEq a) = NEq (decrnum a)"
   460   "decr (Dvd i a) = Dvd i (decrnum a)"
   461   "decr (NDvd i a) = NDvd i (decrnum a)"
   462   "decr (NOT p) = NOT (decr p)" 
   463   "decr (And p q) = And (decr p) (decr q)"
   464   "decr (Or p q) = Or (decr p) (decr q)"
   465   "decr (Imp p q) = Imp (decr p) (decr q)"
   466   "decr (Iff p q) = Iff (decr p) (decr q)"
   467   "decr p = p"
   468 
   469 lemma decrnum: assumes nb: "numbound0 t"
   470   shows "Inum (x#bs) t = Inum bs (decrnum t)"
   471   using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2)
   472 
   473 lemma decr: assumes nb: "bound0 p"
   474   shows "Ifm (x#bs) p = Ifm bs (decr p)"
   475   using nb 
   476   by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum)
   477 
   478 lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
   479 by (induct p, simp_all)
   480 
   481 consts 
   482   isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
   483 recdef isatom "measure size"
   484   "isatom T = True"
   485   "isatom F = True"
   486   "isatom (Lt a) = True"
   487   "isatom (Le a) = True"
   488   "isatom (Gt a) = True"
   489   "isatom (Ge a) = True"
   490   "isatom (Eq a) = True"
   491   "isatom (NEq a) = True"
   492   "isatom (Dvd i b) = True"
   493   "isatom (NDvd i b) = True"
   494   "isatom p = False"
   495 
   496 lemma numsubst0_numbound0: assumes nb: "numbound0 t"
   497   shows "numbound0 (numsubst0 t a)"
   498 using nb by (induct a rule: numsubst0.induct, auto)
   499 
   500 lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t"
   501   shows "bound0 (subst0 t p)"
   502 using qf numsubst0_numbound0[OF nb] by (induct p  rule: subst0.induct, auto)
   503 
   504 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
   505 by (induct p, simp_all)
   506 
   507 
   508 constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
   509   "djf f p q \<equiv> (if q=T then T else if q=F then f p else 
   510   (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or fp q))"
   511 constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
   512   "evaldjf f ps \<equiv> foldr (djf f) ps F"
   513 
   514 lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)"
   515 by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) 
   516 (cases "f p", simp_all add: Let_def djf_def) 
   517 
   518 lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm bs (f p))"
   519   by(induct ps, simp_all add: evaldjf_def djf_Or)
   520 
   521 lemma evaldjf_bound0: 
   522   assumes nb: "\<forall> x\<in> set xs. bound0 (f x)"
   523   shows "bound0 (evaldjf f xs)"
   524   using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) 
   525 
   526 lemma evaldjf_qf: 
   527   assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
   528   shows "qfree (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 consts 
   532   disjuncts :: "fm \<Rightarrow> fm list" 
   533   conjuncts :: "fm \<Rightarrow> fm list"
   534 recdef disjuncts "measure size"
   535   "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
   536   "disjuncts F = []"
   537   "disjuncts p = [p]"
   538 
   539 recdef conjuncts "measure size"
   540   "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
   541   "conjuncts T = []"
   542   "conjuncts p = [p]"
   543 lemma disjuncts: "(\<exists> q\<in> set (disjuncts p). Ifm bs q) = Ifm bs p"
   544 by(induct p rule: disjuncts.induct, auto)
   545 lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm bs q) = Ifm bs p"
   546 by(induct p rule: conjuncts.induct, auto)
   547 
   548 lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). bound0 q"
   549 proof-
   550   assume nb: "bound0 p"
   551   hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto)
   552   thus ?thesis by (simp only: list_all_iff)
   553 qed
   554 lemma conjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). bound0 q"
   555 proof-
   556   assume nb: "bound0 p"
   557   hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto)
   558   thus ?thesis by (simp only: list_all_iff)
   559 qed
   560 
   561 lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q"
   562 proof-
   563   assume qf: "qfree p"
   564   hence "list_all qfree (disjuncts p)"
   565     by (induct p rule: disjuncts.induct, auto)
   566   thus ?thesis by (simp only: list_all_iff)
   567 qed
   568 lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q"
   569 proof-
   570   assume qf: "qfree p"
   571   hence "list_all qfree (conjuncts p)"
   572     by (induct p rule: conjuncts.induct, auto)
   573   thus ?thesis by (simp only: list_all_iff)
   574 qed
   575 
   576 constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
   577   "DJ f p \<equiv> evaldjf f (disjuncts p)"
   578 
   579 lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)"
   580   and fF: "f F = F"
   581   shows "Ifm bs (DJ f p) = Ifm bs (f p)"
   582 proof-
   583   have "Ifm bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm bs (f q))"
   584     by (simp add: DJ_def evaldjf_ex) 
   585   also have "\<dots> = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto)
   586   finally show ?thesis .
   587 qed
   588 
   589 lemma DJ_qf: assumes 
   590   fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)"
   591   shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "
   592 proof(clarify)
   593   fix  p assume qf: "qfree p"
   594   have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def)
   595   from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" .
   596   with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast
   597   
   598   from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp
   599 qed
   600 
   601 lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
   602   shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm bs ((DJ qe p)) = Ifm bs (E p))"
   603 proof(clarify)
   604   fix p::fm and bs
   605   assume qf: "qfree p"
   606   from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast
   607   from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto
   608   have "Ifm bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (qe q))"
   609     by (simp add: DJ_def evaldjf_ex)
   610   also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto
   611   also have "\<dots> = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto)
   612   finally show "qfree (DJ qe p) \<and> Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast
   613 qed
   614   (* Simplification *)
   615 
   616   (* Algebraic simplifications for nums *)
   617 consts bnds:: "num \<Rightarrow> nat list"
   618   lex_ns:: "nat list \<times> nat list \<Rightarrow> bool"
   619 recdef bnds "measure size"
   620   "bnds (Bound n) = [n]"
   621   "bnds (CN n c a) = n#(bnds a)"
   622   "bnds (Neg a) = bnds a"
   623   "bnds (Add a b) = (bnds a)@(bnds b)"
   624   "bnds (Sub a b) = (bnds a)@(bnds b)"
   625   "bnds (Mul i a) = bnds a"
   626   "bnds (Floor a) = bnds a"
   627   "bnds (CF c a b) = (bnds a)@(bnds b)"
   628   "bnds a = []"
   629 recdef lex_ns "measure (\<lambda> (xs,ys). length xs + length ys)"
   630   "lex_ns ([], ms) = True"
   631   "lex_ns (ns, []) = False"
   632   "lex_ns (n#ns, m#ms) = (n<m \<or> ((n = m) \<and> lex_ns (ns,ms))) "
   633 constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
   634   "lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)"
   635 
   636 consts 
   637   numgcdh:: "num \<Rightarrow> int \<Rightarrow> int"
   638   reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num"
   639   dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
   640 consts maxcoeff:: "num \<Rightarrow> int"
   641 recdef maxcoeff "measure size"
   642   "maxcoeff (C i) = abs i"
   643   "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)"
   644   "maxcoeff (CF c t s) = max (abs c) (maxcoeff s)"
   645   "maxcoeff t = 1"
   646 
   647 lemma maxcoeff_pos: "maxcoeff t \<ge> 0"
   648   apply (induct t rule: maxcoeff.induct, auto) 
   649   done
   650 
   651 recdef numgcdh "measure size"
   652   "numgcdh (C i) = (\<lambda>g. igcd i g)"
   653   "numgcdh (CN n c t) = (\<lambda>g. igcd c (numgcdh t g))"
   654   "numgcdh (CF c s t) = (\<lambda>g. igcd c (numgcdh t g))"
   655   "numgcdh t = (\<lambda>g. 1)"
   656 
   657 definition
   658   numgcd :: "num \<Rightarrow> int"
   659 where
   660   numgcd_def: "numgcd t = numgcdh t (maxcoeff t)"
   661 
   662 recdef reducecoeffh "measure size"
   663   "reducecoeffh (C i) = (\<lambda> g. C (i div g))"
   664   "reducecoeffh (CN n c t) = (\<lambda> g. CN n (c div g) (reducecoeffh t g))"
   665   "reducecoeffh (CF c s t) = (\<lambda> g. CF (c div g)  s (reducecoeffh t g))"
   666   "reducecoeffh t = (\<lambda>g. t)"
   667 
   668 definition
   669   reducecoeff :: "num \<Rightarrow> num"
   670 where
   671   reducecoeff_def: "reducecoeff t =
   672   (let g = numgcd t in 
   673   if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)"
   674 
   675 recdef dvdnumcoeff "measure size"
   676   "dvdnumcoeff (C i) = (\<lambda> g. g dvd i)"
   677   "dvdnumcoeff (CN n c t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
   678   "dvdnumcoeff (CF c s t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
   679   "dvdnumcoeff t = (\<lambda>g. False)"
   680 
   681 lemma dvdnumcoeff_trans: 
   682   assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'"
   683   shows "dvdnumcoeff t g"
   684   using dgt' gdg 
   685   by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg])
   686 
   687 declare zdvd_trans [trans add]
   688 
   689 lemma natabs0: "(nat (abs x) = 0) = (x = 0)"
   690 by arith
   691 
   692 lemma numgcd0:
   693   assumes g0: "numgcd t = 0"
   694   shows "Inum bs t = 0"
   695 proof-
   696   have "\<And>x. numgcdh t x= 0 \<Longrightarrow> Inum bs t = 0"
   697     by (induct t rule: numgcdh.induct, auto simp add: igcd_def gcd_zero natabs0 max_def maxcoeff_pos)
   698   thus ?thesis using g0[simplified numgcd_def] by blast
   699 qed
   700 
   701 lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
   702   using gp
   703   by (induct t rule: numgcdh.induct, auto simp add: igcd_def)
   704 
   705 lemma numgcd_pos: "numgcd t \<ge>0"
   706   by (simp add: numgcd_def numgcdh_pos maxcoeff_pos)
   707 
   708 lemma reducecoeffh:
   709   assumes gt: "dvdnumcoeff t g" and gp: "g > 0" 
   710   shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t"
   711   using gt
   712 proof(induct t rule: reducecoeffh.induct) 
   713   case (1 i) hence gd: "g dvd i" by simp
   714   from gp have gnz: "g \<noteq> 0" by simp
   715   from prems show ?case by (simp add: real_of_int_div[OF gnz gd])
   716 next
   717   case (2 n c t)  hence gd: "g dvd c" by simp
   718   from gp have gnz: "g \<noteq> 0" by simp
   719   from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps)
   720 next
   721   case (3 c s t)  hence gd: "g dvd c" by simp
   722   from gp have gnz: "g \<noteq> 0" by simp
   723   from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) 
   724 qed (auto simp add: numgcd_def gp)
   725 consts ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
   726 recdef ismaxcoeff "measure size"
   727   "ismaxcoeff (C i) = (\<lambda> x. abs i \<le> x)"
   728   "ismaxcoeff (CN n c t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
   729   "ismaxcoeff (CF c s t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
   730   "ismaxcoeff t = (\<lambda>x. True)"
   731 
   732 lemma ismaxcoeff_mono: "ismaxcoeff t c \<Longrightarrow> c \<le> c' \<Longrightarrow> ismaxcoeff t c'"
   733 by (induct t rule: ismaxcoeff.induct, auto)
   734 
   735 lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)"
   736 proof (induct t rule: maxcoeff.induct)
   737   case (2 n c t)
   738   hence H:"ismaxcoeff t (maxcoeff t)" .
   739   have thh: "maxcoeff t \<le> max (abs c) (maxcoeff t)" by (simp add: le_maxI2)
   740   from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1)
   741 next
   742   case (3 c t s) 
   743   hence H1:"ismaxcoeff s (maxcoeff s)" by auto
   744   have thh1: "maxcoeff s \<le> max \<bar>c\<bar> (maxcoeff s)" by (simp add: max_def)
   745   from ismaxcoeff_mono[OF H1 thh1] show ?case by (simp add: le_maxI1)
   746 qed simp_all
   747 
   748 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))"
   749   apply (unfold igcd_def)
   750   apply (cases "i = 0", simp_all)
   751   apply (cases "j = 0", simp_all)
   752   apply (cases "abs i = 1", simp_all)
   753   apply (cases "abs j = 1", simp_all)
   754   apply auto
   755   done
   756 lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow>  m =0"
   757   by (induct t rule: numgcdh.induct, auto simp add:igcd0)
   758 
   759 lemma dvdnumcoeff_aux:
   760   assumes "ismaxcoeff t m" and mp:"m \<ge> 0" and "numgcdh t m > 1"
   761   shows "dvdnumcoeff t (numgcdh t m)"
   762 using prems
   763 proof(induct t rule: numgcdh.induct)
   764   case (2 n c t) 
   765   let ?g = "numgcdh t m"
   766   from prems have th:"igcd c ?g > 1" by simp
   767   from igcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
   768   have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
   769   moreover {assume "abs c > 1" and gp: "?g > 1" with prems
   770     have th: "dvdnumcoeff t ?g" by simp
   771     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   772     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)}
   773   moreover {assume "abs c = 0 \<and> ?g > 1"
   774     with prems have th: "dvdnumcoeff t ?g" by simp
   775     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   776     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)
   777     hence ?case by simp }
   778   moreover {assume "abs c > 1" and g0:"?g = 0" 
   779     from numgcdh0[OF g0] have "m=0". with prems   have ?case by simp }
   780   ultimately show ?case by blast
   781 next
   782   case (3 c s t) 
   783   let ?g = "numgcdh t m"
   784   from prems have th:"igcd c ?g > 1" by simp
   785   from igcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
   786   have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
   787   moreover {assume "abs c > 1" and gp: "?g > 1" with prems
   788     have th: "dvdnumcoeff t ?g" by simp
   789     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   790     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)}
   791   moreover {assume "abs c = 0 \<and> ?g > 1"
   792     with prems have th: "dvdnumcoeff t ?g" by simp
   793     have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2)
   794     from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)
   795     hence ?case by simp }
   796   moreover {assume "abs c > 1" and g0:"?g = 0" 
   797     from numgcdh0[OF g0] have "m=0". with prems   have ?case by simp }
   798   ultimately show ?case by blast
   799 qed(auto simp add: igcd_dvd1)
   800 
   801 lemma dvdnumcoeff_aux2:
   802   assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
   803   using prems 
   804 proof (simp add: numgcd_def)
   805   let ?mc = "maxcoeff t"
   806   let ?g = "numgcdh t ?mc"
   807   have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff)
   808   have th2: "?mc \<ge> 0" by (rule maxcoeff_pos)
   809   assume H: "numgcdh t ?mc > 1"
   810   from dvdnumcoeff_aux[OF th1 th2 H]  show "dvdnumcoeff t ?g" .
   811 qed
   812 
   813 lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t"
   814 proof-
   815   let ?g = "numgcd t"
   816   have "?g \<ge> 0"  by (simp add: numgcd_pos)
   817   hence	"?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
   818   moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} 
   819   moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} 
   820   moreover { assume g1:"?g > 1"
   821     from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+
   822     from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis 
   823       by (simp add: reducecoeff_def Let_def)} 
   824   ultimately show ?thesis by blast
   825 qed
   826 
   827 lemma reducecoeffh_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeffh t g)"
   828 by (induct t rule: reducecoeffh.induct, auto)
   829 
   830 lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)"
   831 using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
   832 
   833 consts
   834   simpnum:: "num \<Rightarrow> num"
   835   numadd:: "num \<times> num \<Rightarrow> num"
   836   nummul:: "num \<Rightarrow> int \<Rightarrow> num"
   837 
   838 recdef numadd "measure (\<lambda> (t,s). size t + size s)"
   839   "numadd (CN n1 c1 r1,CN n2 c2 r2) =
   840   (if n1=n2 then 
   841   (let c = c1 + c2
   842   in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2))))
   843   else if n1 \<le> n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2))
   844   else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))"
   845   "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))"  
   846   "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" 
   847   "numadd (CF c1 t1 r1,CF c2 t2 r2) = 
   848    (if t1 = t2 then 
   849     (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s))
   850    else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2))
   851    else CF c2 t2 (numadd(CF c1 t1 r1,r2)))"
   852   "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))"
   853   "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))"
   854   "numadd (C b1, C b2) = C (b1+b2)"
   855   "numadd (a,b) = Add a b"
   856 
   857 lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
   858 apply (induct t s rule: numadd.induct, simp_all add: Let_def)
   859  apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
   860   apply (case_tac "n1 = n2", simp_all add: ring_simps)
   861   apply (simp only: left_distrib[symmetric])
   862  apply simp
   863 apply (case_tac "lex_bnd t1 t2", simp_all)
   864  apply (case_tac "c1+c2 = 0")
   865   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)
   866 
   867 lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
   868 by (induct t s rule: numadd.induct, auto simp add: Let_def)
   869 
   870 recdef nummul "measure size"
   871   "nummul (C j) = (\<lambda> i. C (i*j))"
   872   "nummul (CN n c t) = (\<lambda> i. CN n (c*i) (nummul t i))"
   873   "nummul (CF c t s) = (\<lambda> i. CF (c*i) t (nummul s i))"
   874   "nummul (Mul c t) = (\<lambda> i. nummul t (i*c))"
   875   "nummul t = (\<lambda> i. Mul i t)"
   876 
   877 lemma nummul[simp]: "\<And> i. Inum bs (nummul t i) = Inum bs (Mul i t)"
   878 by (induct t rule: nummul.induct, auto simp add: ring_simps)
   879 
   880 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
   881 by (induct t rule: nummul.induct, auto)
   882 
   883 constdefs numneg :: "num \<Rightarrow> num"
   884   "numneg t \<equiv> nummul t (- 1)"
   885 
   886 constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
   887   "numsub s t \<equiv> (if s = t then C 0 else numadd (s,numneg t))"
   888 
   889 lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
   890 using numneg_def nummul by simp
   891 
   892 lemma numneg_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numneg t)"
   893 using numneg_def by simp
   894 
   895 lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)"
   896 using numsub_def by simp
   897 
   898 lemma numsub_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numsub t s)"
   899 using numsub_def by simp
   900 
   901 lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs"
   902 proof-
   903   have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor)
   904   
   905   have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def)
   906   also have "\<dots>" by (simp add: isint_add cti si)
   907   finally show ?thesis .
   908 qed
   909 
   910 consts split_int:: "num \<Rightarrow> num\<times>num"
   911 recdef split_int "measure num_size"
   912   "split_int (C c) = (C 0, C c)"
   913   "split_int (CN n c b) = 
   914      (let (bv,bi) = split_int b 
   915        in (CN n c bv, bi))"
   916   "split_int (CF c a b) = 
   917      (let (bv,bi) = split_int b 
   918        in (bv, CF c a bi))"
   919   "split_int a = (a,C 0)"
   920 
   921 lemma split_int:"\<And> tv ti. split_int t = (tv,ti) \<Longrightarrow> (Inum bs (Add tv ti) = Inum bs t) \<and> isint ti bs"
   922 proof (induct t rule: split_int.induct)
   923   case (2 c n b tv ti)
   924   let ?bv = "fst (split_int b)"
   925   let ?bi = "snd (split_int b)"
   926   have "split_int b = (?bv,?bi)" by simp
   927   with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
   928   from prems(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def)
   929   from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def)
   930 next
   931   case (3 c a b tv ti) 
   932   let ?bv = "fst (split_int b)"
   933   let ?bi = "snd (split_int b)"
   934   have "split_int b = (?bv,?bi)" by simp
   935   with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
   936   from prems(2) have tibi: "ti = CF c a ?bi" by (simp add: Let_def split_def)
   937   from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF)
   938 qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def ring_simps)
   939 
   940 lemma split_int_nb: "numbound0 t \<Longrightarrow> numbound0 (fst (split_int t)) \<and> numbound0 (snd (split_int t)) "
   941 by (induct t rule: split_int.induct, auto simp add: Let_def split_def)
   942 
   943 definition
   944   numfloor:: "num \<Rightarrow> num"
   945 where
   946   numfloor_def: "numfloor t = (let (tv,ti) = split_int t in 
   947   (case tv of C i \<Rightarrow> numadd (tv,ti) 
   948   | _ \<Rightarrow> numadd(CF 1 tv (C 0),ti)))"
   949 
   950 lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)")
   951 proof-
   952   let ?tv = "fst (split_int t)"
   953   let ?ti = "snd (split_int t)"
   954   have tvti:"split_int t = (?tv,?ti)" by simp
   955   {assume H: "\<forall> v. ?tv \<noteq> C v"
   956     hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" 
   957       by (cases ?tv, auto simp add: numfloor_def Let_def split_def numadd)
   958     from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
   959     hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp 
   960     also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
   961       by (simp,subst tii[simplified isint_iff, symmetric]) simp
   962     also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
   963     finally have ?thesis using th1 by simp}
   964   moreover {fix v assume H:"?tv = C v" 
   965     from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
   966     hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp 
   967     also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
   968       by (simp,subst tii[simplified isint_iff, symmetric]) simp
   969     also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
   970     finally have ?thesis by (simp add: H numfloor_def Let_def split_def numadd) }
   971   ultimately show ?thesis by auto
   972 qed
   973 
   974 lemma numfloor_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numfloor t)"
   975   using split_int_nb[where t="t"]
   976   by (cases "fst(split_int t)" , auto simp add: numfloor_def Let_def split_def  numadd_nb)
   977 
   978 recdef simpnum "measure num_size"
   979   "simpnum (C j) = C j"
   980   "simpnum (Bound n) = CN n 1 (C 0)"
   981   "simpnum (Neg t) = numneg (simpnum t)"
   982   "simpnum (Add t s) = numadd (simpnum t,simpnum s)"
   983   "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
   984   "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)"
   985   "simpnum (Floor t) = numfloor (simpnum t)"
   986   "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))"
   987   "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)"
   988 
   989 lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
   990 by (induct t rule: simpnum.induct, auto)
   991 
   992 lemma simpnum_numbound0[simp]: 
   993   "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"
   994 by (induct t rule: simpnum.induct, auto)
   995 
   996 consts nozerocoeff:: "num \<Rightarrow> bool"
   997 recdef nozerocoeff "measure size"
   998   "nozerocoeff (C c) = True"
   999   "nozerocoeff (CN n c t) = (c\<noteq>0 \<and> nozerocoeff t)"
  1000   "nozerocoeff (CF c s t) = (c \<noteq> 0 \<and> nozerocoeff t)"
  1001   "nozerocoeff (Mul c t) = (c\<noteq>0 \<and> nozerocoeff t)"
  1002   "nozerocoeff t = True"
  1003 
  1004 lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd (a,b))"
  1005 by (induct a b rule: numadd.induct,auto simp add: Let_def)
  1006 
  1007 lemma nummul_nz : "\<And> i. i\<noteq>0 \<Longrightarrow> nozerocoeff a \<Longrightarrow> nozerocoeff (nummul a i)"
  1008   by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz)
  1009 
  1010 lemma numneg_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff (numneg a)"
  1011 by (simp add: numneg_def nummul_nz)
  1012 
  1013 lemma numsub_nz: "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numsub a b)"
  1014 by (simp add: numsub_def numneg_nz numadd_nz)
  1015 
  1016 lemma split_int_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (fst (split_int t)) \<and> nozerocoeff (snd (split_int t))"
  1017 by (induct t rule: split_int.induct,auto simp add: Let_def split_def)
  1018 
  1019 lemma numfloor_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (numfloor t)"
  1020 by (simp add: numfloor_def Let_def split_def)
  1021 (cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz)
  1022 
  1023 lemma simpnum_nz: "nozerocoeff (simpnum t)"
  1024 by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz)
  1025 
  1026 lemma maxcoeff_nz: "nozerocoeff t \<Longrightarrow> maxcoeff t = 0 \<Longrightarrow> t = C 0"
  1027 proof (induct t rule: maxcoeff.induct)
  1028   case (2 n c t)
  1029   hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
  1030   have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1)
  1031   with cnz have "max (abs c) (maxcoeff t) > 0" by arith
  1032   with prems show ?case by simp
  1033 next
  1034   case (3 c s t) 
  1035   hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
  1036   have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1)
  1037   with cnz have "max (abs c) (maxcoeff t) > 0" by arith
  1038   with prems show ?case by simp
  1039 qed auto
  1040 
  1041 lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0"
  1042 proof-
  1043   from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def)
  1044   from numgcdh0[OF th]  have th:"maxcoeff t = 0" .
  1045   from maxcoeff_nz[OF nz th] show ?thesis .
  1046 qed
  1047 
  1048 constdefs simp_num_pair:: "(num \<times> int) \<Rightarrow> num \<times> int"
  1049   "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
  1050    (let t' = simpnum t ; g = numgcd t' in 
  1051       if g > 1 then (let g' = igcd n g in 
  1052         if g' = 1 then (t',n) 
  1053         else (reducecoeffh t' g', n div g')) 
  1054       else (t',n))))"
  1055 
  1056 lemma simp_num_pair_ci:
  1057   shows "((\<lambda> (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\<lambda> (t,n). Inum bs t / real n) (t,n))"
  1058   (is "?lhs = ?rhs")
  1059 proof-
  1060   let ?t' = "simpnum t"
  1061   let ?g = "numgcd ?t'"
  1062   let ?g' = "igcd n ?g"
  1063   {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
  1064   moreover
  1065   { assume nnz: "n \<noteq> 0"
  1066     {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
  1067     moreover
  1068     {assume g1:"?g>1" hence g0: "?g > 0" by simp
  1069       from igcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
  1070       hence g'p: "?g' > 0" using igcd_pos[where i="n" and j="numgcd ?t'"] by arith
  1071       hence "?g'= 1 \<or> ?g' > 1" by arith
  1072       moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
  1073       moreover {assume g'1:"?g'>1"
  1074 	from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
  1075 	let ?tt = "reducecoeffh ?t' ?g'"
  1076 	let ?t = "Inum bs ?tt"
  1077 	have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2)
  1078 	have gpdd: "?g' dvd n" by (simp add: igcd_dvd1) 
  1079 	have gpdgp: "?g' dvd ?g'" by simp
  1080 	from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
  1081 	have th2:"real ?g' * ?t = Inum bs ?t'" by simp
  1082 	from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
  1083 	also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
  1084 	also have "\<dots> = (Inum bs ?t' / real n)"
  1085 	  using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp
  1086 	finally have "?lhs = Inum bs t / real n" by simp
  1087 	then have ?thesis using prems by (simp add: simp_num_pair_def)}
  1088       ultimately have ?thesis by blast}
  1089     ultimately have ?thesis by blast} 
  1090   ultimately show ?thesis by blast
  1091 qed
  1092 
  1093 lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')"
  1094   shows "numbound0 t' \<and> n' >0"
  1095 proof-
  1096     let ?t' = "simpnum t"
  1097   let ?g = "numgcd ?t'"
  1098   let ?g' = "igcd n ?g"
  1099   {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)}
  1100   moreover
  1101   { assume nnz: "n \<noteq> 0"
  1102     {assume "\<not> ?g > 1" hence ?thesis  using prems by (auto simp add: Let_def simp_num_pair_def)}
  1103     moreover
  1104     {assume g1:"?g>1" hence g0: "?g > 0" by simp
  1105       from igcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp
  1106       hence g'p: "?g' > 0" using igcd_pos[where i="n" and j="numgcd ?t'"] by arith
  1107       hence "?g'= 1 \<or> ?g' > 1" by arith
  1108       moreover {assume "?g'=1" hence ?thesis using prems 
  1109 	  by (auto simp add: Let_def simp_num_pair_def)}
  1110       moreover {assume g'1:"?g'>1"
  1111 	have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2)
  1112 	have gpdd: "?g' dvd n" by (simp add: igcd_dvd1) 
  1113 	have gpdgp: "?g' dvd ?g'" by simp
  1114 	from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
  1115 	from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]]
  1116 	have "n div ?g' >0" by simp
  1117 	hence ?thesis using prems 
  1118 	  by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)}
  1119       ultimately have ?thesis by blast}
  1120     ultimately have ?thesis by blast} 
  1121   ultimately show ?thesis by blast
  1122 qed
  1123 
  1124 consts not:: "fm \<Rightarrow> fm"
  1125 recdef not "measure size"
  1126   "not (NOT p) = p"
  1127   "not T = F"
  1128   "not F = T"
  1129   "not (Lt t) = Ge t"
  1130   "not (Le t) = Gt t"
  1131   "not (Gt t) = Le t"
  1132   "not (Ge t) = Lt t"
  1133   "not (Eq t) = NEq t"
  1134   "not (NEq t) = Eq t"
  1135   "not (Dvd i t) = NDvd i t"
  1136   "not (NDvd i t) = Dvd i t"
  1137   "not (And p q) = Or (not p) (not q)"
  1138   "not (Or p q) = And (not p) (not q)"
  1139   "not p = NOT p"
  1140 lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
  1141 by (induct p) auto
  1142 lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
  1143 by (induct p, auto)
  1144 lemma not_nb[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
  1145 by (induct p, auto)
  1146 
  1147 constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1148   "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 
  1149    if p = q then p else And p q)"
  1150 lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
  1151 by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
  1152 
  1153 lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
  1154 using conj_def by auto 
  1155 lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
  1156 using conj_def by auto 
  1157 
  1158 constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1159   "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
  1160        else if p=q then p else Or p q)"
  1161 
  1162 lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)"
  1163 by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
  1164 lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
  1165 using disj_def by auto 
  1166 lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
  1167 using disj_def by auto 
  1168 
  1169 constdefs   imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1170   "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 
  1171     else Imp p q)"
  1172 lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
  1173 by (cases "p=F \<or> q=T",simp_all add: imp_def)
  1174 lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
  1175 using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
  1176 lemma imp_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
  1177 using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def) 
  1178 
  1179 constdefs   iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
  1180   "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else 
  1181        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 
  1182   Iff p q)"
  1183 lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
  1184   by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not) 
  1185 (cases "not p= q", auto simp add:not)
  1186 lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
  1187   by (unfold iff_def,cases "p=q", auto)
  1188 lemma iff_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)"
  1189 using iff_def by (unfold iff_def,cases "p=q", auto)
  1190 
  1191 consts check_int:: "num \<Rightarrow> bool"
  1192 recdef check_int "measure size"
  1193   "check_int (C i) = True"
  1194   "check_int (Floor t) = True"
  1195   "check_int (Mul i t) = check_int t"
  1196   "check_int (Add t s) = (check_int t \<and> check_int s)"
  1197   "check_int (Neg t) = check_int t"
  1198   "check_int (CF c t s) = check_int s"
  1199   "check_int t = False"
  1200 lemma check_int: "check_int t \<Longrightarrow> isint t bs"
  1201 by (induct t, auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF)
  1202 
  1203 lemma rdvd_left1_int: "real \<lfloor>t\<rfloor> = t \<Longrightarrow> 1 rdvd t"
  1204   by (simp add: rdvd_def,rule_tac x="\<lfloor>t\<rfloor>" in exI) simp
  1205 
  1206 lemma rdvd_reduce: 
  1207   assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0"
  1208   shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)"
  1209 proof
  1210   assume d: "real d rdvd real c * t"
  1211   from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto
  1212   from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto
  1213   from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto
  1214   from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp
  1215   hence "real kc * t = real kd * real k" using gp by simp
  1216   hence th:"real kd rdvd real kc * t" using rdvd_def by blast
  1217   from kd_def gp have th':"kd = d div g" by simp
  1218   from kc_def gp have "kc = c div g" by simp
  1219   with th th' show "real (d div g) rdvd real (c div g) * t" by simp
  1220 next
  1221   assume d: "real (d div g) rdvd real (c div g) * t"
  1222   from gp have gnz: "g \<noteq> 0" by simp
  1223   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
  1224 qed
  1225 
  1226 constdefs simpdvd:: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)"
  1227   "simpdvd d t \<equiv> 
  1228    (let g = numgcd t in 
  1229       if g > 1 then (let g' = igcd d g in 
  1230         if g' = 1 then (d, t) 
  1231         else (d div g',reducecoeffh t g')) 
  1232       else (d, t))"
  1233 lemma simpdvd: 
  1234   assumes tnz: "nozerocoeff t" and dnz: "d \<noteq> 0"
  1235   shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)"
  1236 proof-
  1237   let ?g = "numgcd t"
  1238   let ?g' = "igcd d ?g"
  1239   {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)}
  1240   moreover
  1241   {assume g1:"?g>1" hence g0: "?g > 0" by simp
  1242     from igcd0 g1 dnz have gp0: "?g' \<noteq> 0" by simp
  1243     hence g'p: "?g' > 0" using igcd_pos[where i="d" and j="numgcd t"] by arith
  1244     hence "?g'= 1 \<or> ?g' > 1" by arith
  1245     moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)}
  1246     moreover {assume g'1:"?g'>1"
  1247       from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" ..
  1248       let ?tt = "reducecoeffh t ?g'"
  1249       let ?t = "Inum bs ?tt"
  1250       have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2)
  1251       have gpdd: "?g' dvd d" by (simp add: igcd_dvd1) 
  1252       have gpdgp: "?g' dvd ?g'" by simp
  1253       from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
  1254       have th2:"real ?g' * ?t = Inum bs t" by simp
  1255       from prems have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)"
  1256 	by (simp add: simpdvd_def Let_def)
  1257       also have "\<dots> = (real d rdvd (Inum bs t))"
  1258 	using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified zdiv_self[OF gp0]] 
  1259 	  th2[symmetric] by simp
  1260       finally have ?thesis by simp  }
  1261     ultimately have ?thesis by blast
  1262   }
  1263   ultimately show ?thesis by blast
  1264 qed
  1265 
  1266 consts simpfm :: "fm \<Rightarrow> fm"
  1267 recdef simpfm "measure fmsize"
  1268   "simpfm (And p q) = conj (simpfm p) (simpfm q)"
  1269   "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
  1270   "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
  1271   "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
  1272   "simpfm (NOT p) = not (simpfm p)"
  1273   "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F 
  1274   | _ \<Rightarrow> Lt (reducecoeff a'))"
  1275   "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'))"
  1276   "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'))"
  1277   "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'))"
  1278   "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'))"
  1279   "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'))"
  1280   "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a)
  1281              else if (abs i = 1) \<and> check_int a then T
  1282              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))"
  1283   "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) 
  1284              else if (abs i = 1) \<and> check_int a then F
  1285              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))"
  1286   "simpfm p = p"
  1287 
  1288 lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p"
  1289 proof(induct p rule: simpfm.induct)
  1290   case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1291   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1292   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1293     let ?g = "numgcd ?sa"
  1294     let ?rsa = "reducecoeff ?sa"
  1295     let ?r = "Inum bs ?rsa"
  1296     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1297     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1298     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1299     hence gp: "real ?g > 0" by simp
  1300     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1301     with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp
  1302     also have "\<dots> = (?r < 0)" using gp
  1303       by (simp only: mult_less_cancel_left) simp
  1304     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1305   ultimately show ?case by blast
  1306 next
  1307   case (7 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1308   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1309   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1310     let ?g = "numgcd ?sa"
  1311     let ?rsa = "reducecoeff ?sa"
  1312     let ?r = "Inum bs ?rsa"
  1313     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1314     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1315     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1316     hence gp: "real ?g > 0" by simp
  1317     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1318     with sa have "Inum bs a \<le> 0 = (real ?g * ?r \<le> real ?g * 0)" by simp
  1319     also have "\<dots> = (?r \<le> 0)" using gp
  1320       by (simp only: mult_le_cancel_left) simp
  1321     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1322   ultimately show ?case by blast
  1323 next
  1324   case (8 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1325   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1326   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1327     let ?g = "numgcd ?sa"
  1328     let ?rsa = "reducecoeff ?sa"
  1329     let ?r = "Inum bs ?rsa"
  1330     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1331     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1332     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1333     hence gp: "real ?g > 0" by simp
  1334     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1335     with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp
  1336     also have "\<dots> = (?r > 0)" using gp
  1337       by (simp only: mult_less_cancel_left) simp
  1338     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1339   ultimately show ?case by blast
  1340 next
  1341   case (9 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1342   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1343   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1344     let ?g = "numgcd ?sa"
  1345     let ?rsa = "reducecoeff ?sa"
  1346     let ?r = "Inum bs ?rsa"
  1347     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1348     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1349     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1350     hence gp: "real ?g > 0" by simp
  1351     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1352     with sa have "Inum bs a \<ge> 0 = (real ?g * ?r \<ge> real ?g * 0)" by simp
  1353     also have "\<dots> = (?r \<ge> 0)" using gp
  1354       by (simp only: mult_le_cancel_left) simp
  1355     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1356   ultimately show ?case by blast
  1357 next
  1358   case (10 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1359   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1360   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1361     let ?g = "numgcd ?sa"
  1362     let ?rsa = "reducecoeff ?sa"
  1363     let ?r = "Inum bs ?rsa"
  1364     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1365     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1366     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1367     hence gp: "real ?g > 0" by simp
  1368     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1369     with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp
  1370     also have "\<dots> = (?r = 0)" using gp
  1371       by (simp add: mult_eq_0_iff)
  1372     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1373   ultimately show ?case by blast
  1374 next
  1375   case (11 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
  1376   {fix v assume "?sa = C v" hence ?case using sa by simp }
  1377   moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
  1378     let ?g = "numgcd ?sa"
  1379     let ?rsa = "reducecoeff ?sa"
  1380     let ?r = "Inum bs ?rsa"
  1381     have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
  1382     {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
  1383     with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
  1384     hence gp: "real ?g > 0" by simp
  1385     have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
  1386     with sa have "Inum bs a \<noteq> 0 = (real ?g * ?r \<noteq> 0)" by simp
  1387     also have "\<dots> = (?r \<noteq> 0)" using gp
  1388       by (simp add: mult_eq_0_iff)
  1389     finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
  1390   ultimately show ?case by blast
  1391 next
  1392   case (12 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
  1393   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
  1394   {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)}
  1395   moreover 
  1396   {assume ai1: "abs i = 1" and ai: "check_int a" 
  1397     hence "i=1 \<or> i= - 1" by arith
  1398     moreover {assume i1: "i = 1" 
  1399       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1400       have ?case using i1 ai by simp }
  1401     moreover {assume i1: "i = - 1" 
  1402       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1403 	rdvd_abs1[where d="- 1" and t="Inum bs a"]
  1404       have ?case using i1 ai by simp }
  1405     ultimately have ?case by blast}
  1406   moreover   
  1407   {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
  1408     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
  1409 	by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
  1410     moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
  1411       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)
  1412       from simpnum_nz have nz:"nozerocoeff ?sa" by simp
  1413       from simpdvd [OF nz inz] th have ?case using sa by simp}
  1414     ultimately have ?case by blast}
  1415   ultimately show ?case by blast
  1416 next
  1417   case (13 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
  1418   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
  1419   {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)}
  1420   moreover 
  1421   {assume ai1: "abs i = 1" and ai: "check_int a" 
  1422     hence "i=1 \<or> i= - 1" by arith
  1423     moreover {assume i1: "i = 1" 
  1424       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1425       have ?case using i1 ai by simp }
  1426     moreover {assume i1: "i = - 1" 
  1427       from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
  1428 	rdvd_abs1[where d="- 1" and t="Inum bs a"]
  1429       have ?case using i1 ai by simp }
  1430     ultimately have ?case by blast}
  1431   moreover   
  1432   {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
  1433     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
  1434 	by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
  1435     moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
  1436       hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond 
  1437 	by (cases ?sa, auto simp add: Let_def split_def)
  1438       from simpnum_nz have nz:"nozerocoeff ?sa" by simp
  1439       from simpdvd [OF nz inz] th have ?case using sa by simp}
  1440     ultimately have ?case by blast}
  1441   ultimately show ?case by blast
  1442 qed (induct p rule: simpfm.induct, simp_all)
  1443 
  1444 lemma simpdvd_numbound0: "numbound0 t \<Longrightarrow> numbound0 (snd (simpdvd d t))"
  1445   by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0)
  1446 
  1447 lemma simpfm_bound0[simp]: "bound0 p \<Longrightarrow> bound0 (simpfm p)"
  1448 proof(induct p rule: simpfm.induct)
  1449   case (6 a) hence nb: "numbound0 a" by simp
  1450   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1451   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1452 next
  1453   case (7 a) hence nb: "numbound0 a" by simp
  1454   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1455   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1456 next
  1457   case (8 a) hence nb: "numbound0 a" by simp
  1458   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1459   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1460 next
  1461   case (9 a) hence nb: "numbound0 a" by simp
  1462   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1463   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1464 next
  1465   case (10 a) hence nb: "numbound0 a" by simp
  1466   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1467   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1468 next
  1469   case (11 a) hence nb: "numbound0 a" by simp
  1470   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1471   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
  1472 next
  1473   case (12 i a) hence nb: "numbound0 a" by simp
  1474   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1475   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
  1476 next
  1477   case (13 i a) hence nb: "numbound0 a" by simp
  1478   hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
  1479   thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
  1480 qed(auto simp add: disj_def imp_def iff_def conj_def)
  1481 
  1482 lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
  1483 by (induct p rule: simpfm.induct, auto simp add: Let_def)
  1484 (case_tac "simpnum a",auto simp add: split_def Let_def)+
  1485 
  1486 
  1487   (* Generic quantifier elimination *)
  1488 
  1489 constdefs list_conj :: "fm list \<Rightarrow> fm"
  1490   "list_conj ps \<equiv> foldr conj ps T"
  1491 lemma list_conj: "Ifm bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm bs p)"
  1492   by (induct ps, auto simp add: list_conj_def)
  1493 lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
  1494   by (induct ps, auto simp add: list_conj_def)
  1495 lemma list_conj_nb: " \<forall>p\<in> set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
  1496   by (induct ps, auto simp add: list_conj_def)
  1497 constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
  1498   "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs
  1499                    in conj (decr (list_conj yes)) (f (list_conj no)))"
  1500 
  1501 lemma CJNB_qe: 
  1502   assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
  1503   shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm bs ((CJNB qe p)) = Ifm bs (E p))"
  1504 proof(clarify)
  1505   fix bs p
  1506   assume qfp: "qfree p"
  1507   let ?cjs = "conjuncts p"
  1508   let ?yes = "fst (partition bound0 ?cjs)"
  1509   let ?no = "snd (partition bound0 ?cjs)"
  1510   let ?cno = "list_conj ?no"
  1511   let ?cyes = "list_conj ?yes"
  1512   have part: "partition bound0 ?cjs = (?yes,?no)" by simp
  1513   from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast 
  1514   hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) 
  1515   hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf)
  1516   from conjuncts_qf[OF qfp] partition_set[OF part] 
  1517   have " \<forall>q\<in> set ?no. qfree q" by auto
  1518   hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
  1519   with qe have cno_qf:"qfree (qe ?cno )" 
  1520     and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+
  1521   from cno_qf yes_qf have qf: "qfree (CJNB qe p)" 
  1522     by (simp add: CJNB_def Let_def conj_qf split_def)
  1523   {fix bs
  1524     from conjuncts have "Ifm bs p = (\<forall>q\<in> set ?cjs. Ifm bs q)" by blast
  1525     also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm bs q) \<and> (\<forall>q\<in> set ?no. Ifm bs q))"
  1526       using partition_set[OF part] by auto
  1527     finally have "Ifm bs p = ((Ifm bs ?cyes) \<and> (Ifm bs ?cno))" using list_conj by simp}
  1528   hence "Ifm bs (E p) = (\<exists>x. (Ifm (x#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))" by simp
  1529   also have "\<dots> = (\<exists>x. (Ifm (y#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))"
  1530     using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
  1531   also have "\<dots> = (Ifm bs (decr ?cyes) \<and> Ifm bs (E ?cno))"
  1532     by (auto simp add: decr[OF yes_nb])
  1533   also have "\<dots> = (Ifm bs (conj (decr ?cyes) (qe ?cno)))"
  1534     using qe[rule_format, OF no_qf] by auto
  1535   finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" 
  1536     by (simp add: Let_def CJNB_def split_def)
  1537   with qf show "qfree (CJNB qe p) \<and> Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast
  1538 qed
  1539 
  1540 consts qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
  1541 recdef qelim "measure fmsize"
  1542   "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
  1543   "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
  1544   "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
  1545   "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))" 
  1546   "qelim (Or  p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))" 
  1547   "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))"
  1548   "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
  1549   "qelim p = (\<lambda> y. simpfm p)"
  1550 
  1551 lemma qelim_ci:
  1552   assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
  1553   shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)"
  1554 using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] 
  1555 by(induct p rule: qelim.induct) 
  1556 (auto simp del: simpfm.simps)
  1557 
  1558 
  1559 text {* The @{text "\<int>"} Part *}
  1560 text{* Linearity for fm where Bound 0 ranges over @{text "\<int>"} *}
  1561 consts
  1562   zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*)
  1563 recdef zsplit0 "measure num_size"
  1564   "zsplit0 (C c) = (0,C c)"
  1565   "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))"
  1566   "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)"
  1567   "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)"
  1568   "zsplit0 (Neg a) = (let (i',a') =  zsplit0 a in (-i', Neg a'))"
  1569   "zsplit0 (Add a b) = (let (ia,a') =  zsplit0 a ; 
  1570                             (ib,b') =  zsplit0 b 
  1571                             in (ia+ib, Add a' b'))"
  1572   "zsplit0 (Sub a b) = (let (ia,a') =  zsplit0 a ; 
  1573                             (ib,b') =  zsplit0 b 
  1574                             in (ia-ib, Sub a' b'))"
  1575   "zsplit0 (Mul i a) = (let (i',a') =  zsplit0 a in (i*i', Mul i a'))"
  1576   "zsplit0 (Floor a) = (let (i',a') =  zsplit0 a in (i',Floor a'))"
  1577 (hints simp add: Let_def)
  1578 
  1579 lemma zsplit0_I:
  1580   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"
  1581   (is "\<And> n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a")
  1582 proof(induct t rule: zsplit0.induct)
  1583   case (1 c n a) thus ?case by auto 
  1584 next
  1585   case (2 m n a) thus ?case by (cases "m=0") auto
  1586 next
  1587   case (3 n i a n a') thus ?case by auto
  1588 next 
  1589   case (4 c a b n a') thus ?case by auto
  1590 next
  1591   case (5 t n a)
  1592   let ?nt = "fst (zsplit0 t)"
  1593   let ?at = "snd (zsplit0 t)"
  1594   have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \<and> n=-?nt" using prems 
  1595     by (simp add: Let_def split_def)
  1596   from abj prems  have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1597   from th2[simplified] th[simplified] show ?case by simp
  1598 next
  1599   case (6 s t n a)
  1600   let ?ns = "fst (zsplit0 s)"
  1601   let ?as = "snd (zsplit0 s)"
  1602   let ?nt = "fst (zsplit0 t)"
  1603   let ?at = "snd (zsplit0 t)"
  1604   have abjs: "zsplit0 s = (?ns,?as)" by simp 
  1605   moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp 
  1606   ultimately have th: "a=Add ?as ?at \<and> n=?ns + ?nt" using prems 
  1607     by (simp add: Let_def split_def)
  1608   from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
  1609   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
  1610   with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1611   from abjs prems  have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
  1612   from th3[simplified] th2[simplified] th[simplified] show ?case 
  1613     by (simp add: left_distrib)
  1614 next
  1615   case (7 s t n a)
  1616   let ?ns = "fst (zsplit0 s)"
  1617   let ?as = "snd (zsplit0 s)"
  1618   let ?nt = "fst (zsplit0 t)"
  1619   let ?at = "snd (zsplit0 t)"
  1620   have abjs: "zsplit0 s = (?ns,?as)" by simp 
  1621   moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp 
  1622   ultimately have th: "a=Sub ?as ?at \<and> n=?ns - ?nt" using prems 
  1623     by (simp add: Let_def split_def)
  1624   from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
  1625   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
  1626   with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1627   from abjs prems  have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
  1628   from th3[simplified] th2[simplified] th[simplified] show ?case 
  1629     by (simp add: left_diff_distrib)
  1630 next
  1631   case (8 i t n a)
  1632   let ?nt = "fst (zsplit0 t)"
  1633   let ?at = "snd (zsplit0 t)"
  1634   have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \<and> n=i*?nt" using prems 
  1635     by (simp add: Let_def split_def)
  1636   from abj prems  have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1637   hence " ?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp
  1638   also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: right_distrib)
  1639   finally show ?case using th th2 by simp
  1640 next
  1641   case (9 t n a)
  1642   let ?nt = "fst (zsplit0 t)"
  1643   let ?at = "snd (zsplit0 t)"
  1644   have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \<and> n=?nt" using prems 
  1645     by (simp add: Let_def split_def)
  1646   from abj prems  have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
  1647   hence na: "?N a" using th by simp
  1648   have th': "(real ?nt)*(real x) = real (?nt * x)" by simp
  1649   have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp
  1650   also have "\<dots> = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp
  1651   also have "\<dots> = real (floor (?I x ?at + real (?nt* x)))" by (simp add: add_ac)
  1652   also have "\<dots> = real (floor (?I x ?at) + (?nt* x))" 
  1653     using floor_add[where x="?I x ?at" and a="?nt* x"] by simp 
  1654   also have "\<dots> = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: add_ac)
  1655   finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp
  1656   with na show ?case by simp
  1657 qed
  1658 
  1659 consts
  1660   iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool"   (* Linearity test for fm *)
  1661   zlfm :: "fm \<Rightarrow> fm"       (* Linearity transformation for fm *)
  1662 recdef iszlfm "measure size"
  1663   "iszlfm (And p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" 
  1664   "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" 
  1665   "iszlfm (Eq  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1666   "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1667   "iszlfm (Lt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1668   "iszlfm (Le  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1669   "iszlfm (Gt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1670   "iszlfm (Ge  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
  1671   "iszlfm (Dvd i (CN 0 c e)) = 
  1672                  (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
  1673   "iszlfm (NDvd i (CN 0 c e))= 
  1674                  (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
  1675   "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 p))"
  1676 
  1677 lemma zlin_qfree: "iszlfm p bs \<Longrightarrow> qfree p"
  1678   by (induct p rule: iszlfm.induct) auto
  1679 
  1680 lemma iszlfm_gen:
  1681   assumes lp: "iszlfm p (x#bs)"
  1682   shows "\<forall> y. iszlfm p (y#bs)"
  1683 proof
  1684   fix y
  1685   show "iszlfm p (y#bs)"
  1686     using lp
  1687   by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"])
  1688 qed
  1689 
  1690 lemma conj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (conj p q) bs"
  1691   using conj_def by (cases p,auto)
  1692 lemma disj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (disj p q) bs"
  1693   using disj_def by (cases p,auto)
  1694 lemma not_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm (not p) bs"
  1695   by (induct p rule:iszlfm.induct ,auto)
  1696 
  1697 recdef zlfm "measure fmsize"
  1698   "zlfm (And p q) = conj (zlfm p) (zlfm q)"
  1699   "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
  1700   "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
  1701   "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
  1702   "zlfm (Lt a) = (let (c,r) = zsplit0 a in 
  1703      if c=0 then Lt r else 
  1704      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))) 
  1705      else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
  1706   "zlfm (Le a) = (let (c,r) = zsplit0 a in 
  1707      if c=0 then Le r else 
  1708      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))) 
  1709      else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
  1710   "zlfm (Gt a) = (let (c,r) = zsplit0 a in 
  1711      if c=0 then Gt r else 
  1712      if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) 
  1713      else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
  1714   "zlfm (Ge a) = (let (c,r) = zsplit0 a in 
  1715      if c=0 then Ge r else 
  1716      if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) 
  1717      else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
  1718   "zlfm (Eq a) = (let (c,r) = zsplit0 a in 
  1719               if c=0 then Eq r else 
  1720       if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r)))
  1721       else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))"
  1722   "zlfm (NEq a) = (let (c,r) = zsplit0 a in 
  1723               if c=0 then NEq r else 
  1724       if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r)))
  1725       else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))"
  1726   "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) 
  1727   else (let (c,r) = zsplit0 a in 
  1728               if c=0 then Dvd (abs i) r else 
  1729       if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) 
  1730       else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
  1731   "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) 
  1732   else (let (c,r) = zsplit0 a in 
  1733               if c=0 then NDvd (abs i) r else 
  1734       if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) 
  1735       else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
  1736   "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
  1737   "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
  1738   "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
  1739   "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
  1740   "zlfm (NOT (NOT p)) = zlfm p"
  1741   "zlfm (NOT T) = F"
  1742   "zlfm (NOT F) = T"
  1743   "zlfm (NOT (Lt a)) = zlfm (Ge a)"
  1744   "zlfm (NOT (Le a)) = zlfm (Gt a)"
  1745   "zlfm (NOT (Gt a)) = zlfm (Le a)"
  1746   "zlfm (NOT (Ge a)) = zlfm (Lt a)"
  1747   "zlfm (NOT (Eq a)) = zlfm (NEq a)"
  1748   "zlfm (NOT (NEq a)) = zlfm (Eq a)"
  1749   "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
  1750   "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
  1751   "zlfm p = p" (hints simp add: fmsize_pos)
  1752 
  1753 lemma split_int_less_real: 
  1754   "(real (a::int) < b) = (a < floor b \<or> (a = floor b \<and> real (floor b) < b))"
  1755 proof( auto)
  1756   assume alb: "real a < b" and agb: "\<not> a < floor b"
  1757   from agb have "floor b \<le> a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq)
  1758   from floor_eq[OF alb th] show "a= floor b" by simp 
  1759 next
  1760   assume alb: "a < floor b"
  1761   hence "real a < real (floor b)" by simp
  1762   moreover have "real (floor b) \<le> b" by simp ultimately show  "real a < b" by arith 
  1763 qed
  1764 
  1765 lemma split_int_less_real': 
  1766   "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \<or> (real a - real (floor (-b)) = 0 \<and> real (floor (-b)) + b < 0))"
  1767 proof- 
  1768   have "(real a + b <0) = (real a < -b)" by arith
  1769   with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith  
  1770 qed
  1771 
  1772 lemma split_int_gt_real': 
  1773   "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \<or> (real a + real (floor b) = 0 \<and> real (floor b) - b < 0))"
  1774 proof- 
  1775   have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith
  1776   show ?thesis using myless[rule_format, where b="real (floor b)"] 
  1777     by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) 
  1778     (simp add: ring_simps diff_def[symmetric],arith)
  1779 qed
  1780 
  1781 lemma split_int_le_real: 
  1782   "(real (a::int) \<le> b) = (a \<le> floor b \<or> (a = floor b \<and> real (floor b) < b))"
  1783 proof( auto)
  1784   assume alb: "real a \<le> b" and agb: "\<not> a \<le> floor b"
  1785   from alb have "floor (real a) \<le> floor b " by (simp only: floor_mono2) 
  1786   hence "a \<le> floor b" by simp with agb show "False" by simp
  1787 next
  1788   assume alb: "a \<le> floor b"
  1789   hence "real a \<le> real (floor b)" by (simp only: floor_mono2)
  1790   also have "\<dots>\<le> b" by simp  finally show  "real a \<le> b" . 
  1791 qed
  1792 
  1793 lemma split_int_le_real': 
  1794   "(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))"
  1795 proof- 
  1796   have "(real a + b \<le>0) = (real a \<le> -b)" by arith
  1797   with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith  
  1798 qed
  1799 
  1800 lemma split_int_ge_real': 
  1801   "(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))"
  1802 proof- 
  1803   have th: "(real a + b \<ge>0) = (real (-a) + (-b) \<le> 0)" by arith
  1804   show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"])
  1805     (simp add: ring_simps diff_def[symmetric],arith)
  1806 qed
  1807 
  1808 lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \<and> b = real (floor b))" (is "?l = ?r")
  1809 by auto
  1810 
  1811 lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \<and> real (floor (-b)) + b = 0)" (is "?l = ?r")
  1812 proof-
  1813   have "?l = (real a = -b)" by arith
  1814   with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith
  1815 qed
  1816 
  1817 lemma zlfm_I:
  1818   assumes qfp: "qfree p"
  1819   shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \<and> iszlfm (zlfm p) (real (i::int) #bs)"
  1820   (is "(?I (?l p) = ?I p) \<and> ?L (?l p)")
  1821 using qfp
  1822 proof(induct p rule: zlfm.induct)
  1823   case (5 a) 
  1824   let ?c = "fst (zsplit0 a)"
  1825   let ?r = "snd (zsplit0 a)"
  1826   have spl: "zsplit0 a = (?c,?r)" by simp
  1827   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1828   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1829   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1830   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1831   moreover
  1832   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1833       by (cases "?r", simp_all add: Let_def split_def,case_tac "nat", simp_all)}
  1834   moreover
  1835   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
  1836       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1837     have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
  1838     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)
  1839     finally have ?case using l by simp}
  1840   moreover
  1841   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
  1842       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1843     have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
  1844     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)
  1845     finally have ?case using l by simp}
  1846   ultimately show ?case by blast
  1847 next
  1848   case (6 a)
  1849   let ?c = "fst (zsplit0 a)"
  1850   let ?r = "snd (zsplit0 a)"
  1851   have spl: "zsplit0 a = (?c,?r)" by simp
  1852   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1853   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1854   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1855   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1856   moreover
  1857   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1858       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat",simp_all)}
  1859   moreover
  1860   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
  1861       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1862     have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
  1863     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)
  1864     finally have ?case using l by simp}
  1865   moreover
  1866   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
  1867       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1868     have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
  1869     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)
  1870     finally have ?case using l by simp}
  1871   ultimately show ?case by blast
  1872 next
  1873   case (7 a) 
  1874   let ?c = "fst (zsplit0 a)"
  1875   let ?r = "snd (zsplit0 a)"
  1876   have spl: "zsplit0 a = (?c,?r)" by simp
  1877   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1878   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1879   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1880   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1881   moreover
  1882   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1883       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1884   moreover
  1885   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
  1886       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1887     have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
  1888     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)
  1889     finally have ?case using l by simp}
  1890   moreover
  1891   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
  1892       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1893     have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
  1894     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)
  1895     finally have ?case using l by simp}
  1896   ultimately show ?case by blast
  1897 next
  1898   case (8 a)
  1899    let ?c = "fst (zsplit0 a)"
  1900   let ?r = "snd (zsplit0 a)"
  1901   have spl: "zsplit0 a = (?c,?r)" by simp
  1902   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1903   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1904   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1905   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1906   moreover
  1907   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1908       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1909   moreover
  1910   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
  1911       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1912     have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
  1913     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)
  1914     finally have ?case using l by simp}
  1915   moreover
  1916   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
  1917       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1918     have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
  1919     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)
  1920     finally have ?case using l by simp}
  1921   ultimately show ?case by blast
  1922 next
  1923   case (9 a)
  1924   let ?c = "fst (zsplit0 a)"
  1925   let ?r = "snd (zsplit0 a)"
  1926   have spl: "zsplit0 a = (?c,?r)" by simp
  1927   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1928   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1929   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1930   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1931   moreover
  1932   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1933       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1934   moreover
  1935   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" 
  1936       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1937     have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
  1938     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)
  1939     finally have ?case using l by simp}
  1940   moreover
  1941   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" 
  1942       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1943     have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
  1944     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)
  1945     finally have ?case using l by simp}
  1946   ultimately show ?case by blast
  1947 next
  1948   case (10 a)
  1949   let ?c = "fst (zsplit0 a)"
  1950   let ?r = "snd (zsplit0 a)"
  1951   have spl: "zsplit0 a = (?c,?r)" by simp
  1952   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1953   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1954   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1955   have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
  1956   moreover
  1957   {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1958       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1959   moreover
  1960   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" 
  1961       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1962     have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
  1963     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)
  1964     finally have ?case using l by simp}
  1965   moreover
  1966   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" 
  1967       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1968     have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
  1969     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)
  1970     finally have ?case using l by simp}
  1971   ultimately show ?case by blast
  1972 next
  1973   case (11 j a)
  1974   let ?c = "fst (zsplit0 a)"
  1975   let ?r = "snd (zsplit0 a)"
  1976   have spl: "zsplit0 a = (?c,?r)" by simp
  1977   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  1978   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  1979   let ?N = "\<lambda> t. Inum (real i#bs) t"
  1980   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
  1981   moreover
  1982   {assume "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) 
  1983     hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)}
  1984   moreover
  1985   {assume "?c=0" and "j\<noteq>0" hence ?case 
  1986       using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
  1987       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  1988   moreover
  1989   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
  1990       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  1991     have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" 
  1992       using Ia by (simp add: Let_def split_def)
  1993     also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" 
  1994       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  1995     also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  1996        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" 
  1997       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  1998     also have "\<dots> = (?I (?l (Dvd j a)))" using cp cnz jnz  
  1999       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2000 	del: real_of_int_mult) (auto simp add: add_ac)
  2001     finally have ?case using l jnz  by simp }
  2002   moreover
  2003   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
  2004       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  2005     have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" 
  2006       using Ia by (simp add: Let_def split_def)
  2007     also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" 
  2008       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  2009     also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  2010        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" 
  2011       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  2012     also have "\<dots> = (?I (?l (Dvd j a)))" using cn cnz jnz
  2013       using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
  2014       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2015 	del: real_of_int_mult) (auto simp add: add_ac)
  2016     finally have ?case using l jnz by blast }
  2017   ultimately show ?case by blast
  2018 next
  2019   case (12 j a)
  2020   let ?c = "fst (zsplit0 a)"
  2021   let ?r = "snd (zsplit0 a)"
  2022   have spl: "zsplit0 a = (?c,?r)" by simp
  2023   from zsplit0_I[OF spl, where x="i" and bs="bs"] 
  2024   have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
  2025   let ?N = "\<lambda> t. Inum (real i#bs) t"
  2026   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
  2027   moreover
  2028   {assume "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) 
  2029     hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)}
  2030   moreover
  2031   {assume "?c=0" and "j\<noteq>0" hence ?case 
  2032       using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
  2033       by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
  2034   moreover
  2035   {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" 
  2036       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  2037     have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" 
  2038       using Ia by (simp add: Let_def split_def)
  2039     also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" 
  2040       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  2041     also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  2042        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" 
  2043       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  2044     also have "\<dots> = (?I (?l (NDvd j a)))" using cp cnz jnz  
  2045       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2046 	del: real_of_int_mult) (auto simp add: add_ac)
  2047     finally have ?case using l jnz  by simp }
  2048   moreover
  2049   {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" 
  2050       by (simp add: nb Let_def split_def isint_Floor isint_neg)
  2051     have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" 
  2052       using Ia by (simp add: Let_def split_def)
  2053     also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" 
  2054       by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
  2055     also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
  2056        (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" 
  2057       by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
  2058     also have "\<dots> = (?I (?l (NDvd j a)))" using cn cnz jnz
  2059       using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
  2060       by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
  2061 	del: real_of_int_mult) (auto simp add: add_ac)
  2062     finally have ?case using l jnz by blast }
  2063   ultimately show ?case by blast
  2064 qed auto
  2065 
  2066 text{* plusinf : Virtual substitution of @{text "+\<infinity>"}
  2067        minusinf: Virtual substitution of @{text "-\<infinity>"}
  2068        @{text "\<delta>"} Compute lcm @{text "d| Dvd d  c*x+t \<in> p"}
  2069        @{text "d\<delta>"} checks if a given l divides all the ds above*}
  2070 
  2071 consts 
  2072   plusinf:: "fm \<Rightarrow> fm" 
  2073   minusinf:: "fm \<Rightarrow> fm"
  2074   \<delta> :: "fm \<Rightarrow> int" 
  2075   d\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"
  2076 
  2077 recdef minusinf "measure size"
  2078   "minusinf (And p q) = conj (minusinf p) (minusinf q)" 
  2079   "minusinf (Or p q) = disj (minusinf p) (minusinf q)" 
  2080   "minusinf (Eq  (CN 0 c e)) = F"
  2081   "minusinf (NEq (CN 0 c e)) = T"
  2082   "minusinf (Lt  (CN 0 c e)) = T"
  2083   "minusinf (Le  (CN 0 c e)) = T"
  2084   "minusinf (Gt  (CN 0 c e)) = F"
  2085   "minusinf (Ge  (CN 0 c e)) = F"
  2086   "minusinf p = p"
  2087 
  2088 lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"
  2089   by (induct p rule: minusinf.induct, auto)
  2090 
  2091 recdef plusinf "measure size"
  2092   "plusinf (And p q) = conj (plusinf p) (plusinf q)" 
  2093   "plusinf (Or p q) = disj (plusinf p) (plusinf q)" 
  2094   "plusinf (Eq  (CN 0 c e)) = F"
  2095   "plusinf (NEq (CN 0 c e)) = T"
  2096   "plusinf (Lt  (CN 0 c e)) = F"
  2097   "plusinf (Le  (CN 0 c e)) = F"
  2098   "plusinf (Gt  (CN 0 c e)) = T"
  2099   "plusinf (Ge  (CN 0 c e)) = T"
  2100   "plusinf p = p"
  2101 
  2102 recdef \<delta> "measure size"
  2103   "\<delta> (And p q) = ilcm (\<delta> p) (\<delta> q)" 
  2104   "\<delta> (Or p q) = ilcm (\<delta> p) (\<delta> q)" 
  2105   "\<delta> (Dvd i (CN 0 c e)) = i"
  2106   "\<delta> (NDvd i (CN 0 c e)) = i"
  2107   "\<delta> p = 1"
  2108 
  2109 recdef d\<delta> "measure size"
  2110   "d\<delta> (And p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" 
  2111   "d\<delta> (Or p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" 
  2112   "d\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
  2113   "d\<delta> (NDvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
  2114   "d\<delta> p = (\<lambda> d. True)"
  2115 
  2116 lemma delta_mono: 
  2117   assumes lin: "iszlfm p bs"
  2118   and d: "d dvd d'"
  2119   and ad: "d\<delta> p d"
  2120   shows "d\<delta> p d'"
  2121   using lin ad d
  2122 proof(induct p rule: iszlfm.induct)
  2123   case (9 i c e)  thus ?case using d
  2124     by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"])
  2125 next
  2126   case (10 i c e) thus ?case using d
  2127     by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"])
  2128 qed simp_all
  2129 
  2130 lemma \<delta> : assumes lin:"iszlfm p bs"
  2131   shows "d\<delta> p (\<delta> p) \<and> \<delta> p >0"
  2132 using lin
  2133 proof (induct p rule: iszlfm.induct)
  2134   case (1 p q) 
  2135   let ?d = "\<delta> (And p q)"
  2136   from prems ilcm_pos have dp: "?d >0" by simp
  2137   have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp 
  2138    hence th: "d\<delta> p ?d" 
  2139      using delta_mono prems by (auto simp del: dvd_ilcm_self1)
  2140   have "\<delta> q dvd \<delta> (And p q)" using prems  by simp 
  2141   hence th': "d\<delta> q ?d" using delta_mono prems by (auto simp del: dvd_ilcm_self2)
  2142   from th th' dp show ?case by simp 
  2143 next
  2144   case (2 p q)  
  2145   let ?d = "\<delta> (And p q)"
  2146   from prems ilcm_pos have dp: "?d >0" by simp
  2147   have "\<delta> p dvd \<delta> (And p q)" using prems by simp hence th: "d\<delta> p ?d" using delta_mono prems 
  2148     by (auto simp del: dvd_ilcm_self1)
  2149   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)
  2150   from th th' dp show ?case by simp 
  2151 qed simp_all
  2152 
  2153 
  2154 lemma minusinf_inf:
  2155   assumes linp: "iszlfm p (a # bs)"
  2156   shows "\<exists> (z::int). \<forall> x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p"
  2157   (is "?P p" is "\<exists> (z::int). \<forall> x < z. ?I x (?M p) = ?I x p")
  2158 using linp
  2159 proof (induct p rule: minusinf.induct)
  2160   case (1 f g)
  2161   from prems have "?P f" by simp
  2162   then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
  2163   from prems have "?P g" by simp
  2164   then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
  2165   let ?z = "min z1 z2"
  2166   from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp
  2167   thus ?case by blast
  2168 next
  2169   case (2 f g)   from prems have "?P f" by simp
  2170   then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
  2171   from prems have "?P g" by simp
  2172   then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
  2173   let ?z = "min z1 z2"
  2174   from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp
  2175   thus ?case by blast
  2176 next
  2177   case (3 c e) 
  2178   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2179   from prems have nbe: "numbound0 e" by simp
  2180   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))"
  2181   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2182     fix x
  2183     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2184     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2185     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2186       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2187     hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
  2188     thus "real c * real x + Inum (real x # bs) e \<noteq> 0" 
  2189       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"]  by simp
  2190   qed
  2191   thus ?case by blast
  2192 next
  2193   case (4 c e) 
  2194   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2195   from prems have nbe: "numbound0 e" by simp
  2196   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))"
  2197   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2198     fix x
  2199     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2200     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2201     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2202       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2203     hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
  2204     thus "real c * real x + Inum (real x # bs) e \<noteq> 0" 
  2205       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"]  by simp
  2206   qed
  2207   thus ?case by blast
  2208 next
  2209   case (5 c e) 
  2210   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2211   from prems have nbe: "numbound0 e" by simp
  2212   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))"
  2213   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2214     fix x
  2215     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2216     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2217     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2218       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2219     thus "real c * real x + Inum (real x # bs) e < 0" 
  2220       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2221   qed
  2222   thus ?case by blast
  2223 next
  2224   case (6 c e) 
  2225   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2226   from prems have nbe: "numbound0 e" by simp
  2227   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))"
  2228   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2229     fix x
  2230     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2231     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2232     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2233       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2234     thus "real c * real x + Inum (real x # bs) e \<le> 0" 
  2235       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2236   qed
  2237   thus ?case by blast
  2238 next
  2239   case (7 c e) 
  2240   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2241   from prems have nbe: "numbound0 e" by simp
  2242   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))"
  2243   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2244     fix x
  2245     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2246     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2247     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2248       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2249     thus "\<not> (real c * real x + Inum (real x # bs) e>0)" 
  2250       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2251   qed
  2252   thus ?case by blast
  2253 next
  2254   case (8 c e) 
  2255   from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp
  2256   from prems have nbe: "numbound0 e" by simp
  2257   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))"
  2258   proof (simp add: less_floor_eq , rule allI, rule impI) 
  2259     fix x
  2260     assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
  2261     hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
  2262     with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
  2263       by (simp only:  real_mult_less_mono2[OF rcpos th1])
  2264     thus "\<not> real c * real x + Inum (real x # bs) e \<ge> 0" 
  2265       using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
  2266   qed
  2267   thus ?case by blast
  2268 qed simp_all
  2269 
  2270 lemma minusinf_repeats:
  2271   assumes d: "d\<delta> p d" and linp: "iszlfm p (a # bs)"
  2272   shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)"
  2273 using linp d
  2274 proof(induct p rule: iszlfm.induct) 
  2275   case (9 i c e) hence nbe: "numbound0 e"  and id: "i dvd d" by simp+
  2276     hence "\<exists> k. d=i*k" by (simp add: dvd_def)
  2277     then obtain "di" where di_def: "d=i*di" by blast
  2278     show ?case 
  2279     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)
  2280       assume 
  2281 	"real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
  2282       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
  2283       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
  2284       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
  2285 	by (simp add: ring_simps di_def)
  2286       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
  2287 	by (simp add: ring_simps)
  2288       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
  2289       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
  2290     next
  2291       assume 
  2292 	"real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
  2293       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
  2294       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
  2295       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)
  2296       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps)
  2297       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
  2298 	by blast
  2299       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
  2300     qed
  2301 next
  2302   case (10 i c e) hence nbe: "numbound0 e"  and id: "i dvd d" by simp+
  2303     hence "\<exists> k. d=i*k" by (simp add: dvd_def)
  2304     then obtain "di" where di_def: "d=i*di" by blast
  2305     show ?case 
  2306     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)
  2307       assume 
  2308 	"real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
  2309       (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
  2310       hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
  2311       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
  2312 	by (simp add: ring_simps di_def)
  2313       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
  2314 	by (simp add: ring_simps)
  2315       hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
  2316       thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
  2317     next
  2318       assume 
  2319 	"real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
  2320       hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
  2321       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
  2322       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)
  2323       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: ring_simps)
  2324       hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
  2325 	by blast
  2326       thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
  2327     qed
  2328 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)
  2329 
  2330 lemma minusinf_ex:
  2331   assumes lin: "iszlfm p (real (a::int) #bs)"
  2332   and exmi: "\<exists> (x::int). Ifm (real x#bs) (minusinf p)" (is "\<exists> x. ?P1 x")
  2333   shows "\<exists> (x::int). Ifm (real x#bs) p" (is "\<exists> x. ?P x")
  2334 proof-
  2335   let ?d = "\<delta> p"
  2336   from \<delta> [OF lin] have dpos: "?d >0" by simp
  2337   from \<delta> [OF lin] have alld: "d\<delta> p ?d" by simp
  2338   from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P1 x = ?P1 (x - (k * ?d))" by simp
  2339   from minusinf_inf[OF lin] have th2:"\<exists> z. \<forall> x. x<z \<longrightarrow> (?P x = ?P1 x)" by blast
  2340   from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast
  2341 qed
  2342 
  2343 lemma minusinf_bex:
  2344   assumes lin: "iszlfm p (real (a::int) #bs)"
  2345   shows "(\<exists> (x::int). Ifm (real x#bs) (minusinf p)) = 
  2346          (\<exists> (x::int)\<in> {1..\<delta> p}. Ifm (real x#bs) (minusinf p))"
  2347   (is "(\<exists> x. ?P x) = _")
  2348 proof-
  2349   let ?d = "\<delta> p"
  2350   from \<delta> [OF lin] have dpos: "?d >0" by simp
  2351   from \<delta> [OF lin] have alld: "d\<delta> p ?d" by simp
  2352   from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P x = ?P (x - (k * ?d))" by simp
  2353   from periodic_finite_ex[OF dpos th1] show ?thesis by blast
  2354 qed
  2355 
  2356 lemma dvd1_eq1: "x >0 \<Longrightarrow> (x::int) dvd 1 = (x = 1)" by auto
  2357 
  2358 consts 
  2359   a\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coeffitients of a formula *)
  2360   d\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*)
  2361   \<zeta>  :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*)
  2362   \<beta> :: "fm \<Rightarrow> num list"
  2363   \<alpha> :: "fm \<Rightarrow> num list"
  2364 
  2365 recdef a\<beta> "measure size"
  2366   "a\<beta> (And p q) = (\<lambda> k. And (a\<beta> p k) (a\<beta> q k))" 
  2367   "a\<beta> (Or p q) = (\<lambda> k. Or (a\<beta> p k) (a\<beta> q k))" 
  2368   "a\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))"
  2369   "a\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))"
  2370   "a\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))"
  2371   "a\<beta> (Le  (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))"
  2372   "a\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))"
  2373   "a\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))"
  2374   "a\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
  2375   "a\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
  2376   "a\<beta> p = (\<lambda> k. p)"
  2377 
  2378 recdef d\<beta> "measure size"
  2379   "d\<beta> (And p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" 
  2380   "d\<beta> (Or p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" 
  2381   "d\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2382   "d\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2383   "d\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2384   "d\<beta> (Le  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2385   "d\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2386   "d\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. c dvd k)"
  2387   "d\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)"
  2388   "d\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)"
  2389   "d\<beta> p = (\<lambda> k. True)"
  2390 
  2391 recdef \<zeta> "measure size"
  2392   "\<zeta> (And p q) = ilcm (\<zeta> p) (\<zeta> q)" 
  2393   "\<zeta> (Or p q) = ilcm (\<zeta> p) (\<zeta> q)" 
  2394   "\<zeta> (Eq  (CN 0 c e)) = c"
  2395   "\<zeta> (NEq (CN 0 c e)) = c"
  2396   "\<zeta> (Lt  (CN 0 c e)) = c"
  2397   "\<zeta> (Le  (CN 0 c e)) = c"
  2398   "\<zeta> (Gt  (CN 0 c e)) = c"
  2399   "\<zeta> (Ge  (CN 0 c e)) = c"
  2400   "\<zeta> (Dvd i (CN 0 c e)) = c"
  2401   "\<zeta> (NDvd i (CN 0 c e))= c"
  2402   "\<zeta> p = 1"
  2403 
  2404 recdef \<beta> "measure size"
  2405   "\<beta> (And p q) = (\<beta> p @ \<beta> q)" 
  2406   "\<beta> (Or p q) = (\<beta> p @ \<beta> q)" 
  2407   "\<beta> (Eq  (CN 0 c e)) = [Sub (C -1) e]"
  2408   "\<beta> (NEq (CN 0 c e)) = [Neg e]"
  2409   "\<beta> (Lt  (CN 0 c e)) = []"
  2410   "\<beta> (Le  (CN 0 c e)) = []"
  2411   "\<beta> (Gt  (CN 0 c e)) = [Neg e]"
  2412   "\<beta> (Ge  (CN 0 c e)) = [Sub (C -1) e]"
  2413   "\<beta> p = []"
  2414 
  2415 recdef \<alpha> "measure size"
  2416   "\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)" 
  2417   "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)" 
  2418   "\<alpha> (Eq  (CN 0 c e)) = [Add (C -1) e]"
  2419   "\<alpha> (NEq (CN 0 c e)) = [e]"
  2420   "\<alpha> (Lt  (CN 0 c e)) = [e]"
  2421   "\<alpha> (Le  (CN 0 c e)) = [Add (C -1) e]"
  2422   "\<alpha> (Gt  (CN 0 c e)) = []"
  2423   "\<alpha> (Ge  (CN 0 c e)) = []"
  2424   "\<alpha> p = []"
  2425 consts mirror :: "fm \<Rightarrow> fm"
  2426 recdef mirror "measure size"
  2427   "mirror (And p q) = And (mirror p) (mirror q)" 
  2428   "mirror (Or p q) = Or (mirror p) (mirror q)" 
  2429   "mirror (Eq  (CN 0 c e)) = Eq (CN 0 c (Neg e))"
  2430   "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
  2431   "mirror (Lt  (CN 0 c e)) = Gt (CN 0 c (Neg e))"
  2432   "mirror (Le  (CN 0 c e)) = Ge (CN 0 c (Neg e))"
  2433   "mirror (Gt  (CN 0 c e)) = Lt (CN 0 c (Neg e))"
  2434   "mirror (Ge  (CN 0 c e)) = Le (CN 0 c (Neg e))"
  2435   "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
  2436   "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
  2437   "mirror p = p"
  2438 
  2439 lemma mirror\<alpha>\<beta>:
  2440   assumes lp: "iszlfm p (a#bs)"
  2441   shows "(Inum (real (i::int)#bs)) ` set (\<alpha> p) = (Inum (real i#bs)) ` set (\<beta> (mirror p))"
  2442 using lp
  2443 by (induct p rule: mirror.induct, auto)
  2444 
  2445 lemma mirror: 
  2446   assumes lp: "iszlfm p (a#bs)"
  2447   shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p" 
  2448 using lp
  2449 proof(induct p rule: iszlfm.induct)
  2450   case (9 j c e)
  2451   have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
  2452        (real j rdvd - (real c * real x - Inum (real x # bs) e))"
  2453     by (simp only: rdvd_minus[symmetric])
  2454   from prems show  ?case
  2455     by (simp add: ring_simps th[simplified ring_simps]
  2456       numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
  2457 next
  2458     case (10 j c e)
  2459   have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
  2460        (real j rdvd - (real c * real x - Inum (real x # bs) e))"
  2461     by (simp only: rdvd_minus[symmetric])
  2462   from prems show  ?case
  2463     by (simp add: ring_simps th[simplified ring_simps]
  2464       numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
  2465 qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"] nth_pos2)
  2466 
  2467 lemma mirror_l: "iszlfm p (a#bs) \<Longrightarrow> iszlfm (mirror p) (a#bs)"
  2468 by (induct p rule: mirror.induct, auto simp add: isint_neg)
  2469 
  2470 lemma mirror_d\<beta>: "iszlfm p (a#bs) \<and> d\<beta> p 1 
  2471   \<Longrightarrow> iszlfm (mirror p) (a#bs) \<and> d\<beta> (mirror p) 1"
  2472 by (induct p rule: mirror.induct, auto simp add: isint_neg)
  2473 
  2474 lemma mirror_\<delta>: "iszlfm p (a#bs) \<Longrightarrow> \<delta> (mirror p) = \<delta> p"
  2475 by (induct p rule: mirror.induct,auto)
  2476 
  2477 
  2478 lemma mirror_ex: 
  2479   assumes lp: "iszlfm p (real (i::int)#bs)"
  2480   shows "(\<exists> (x::int). Ifm (real x#bs) (mirror p)) = (\<exists> (x::int). Ifm (real x#bs) p)"
  2481   (is "(\<exists> x. ?I x ?mp) = (\<exists> x. ?I x p)")
  2482 proof(auto)
  2483   fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast
  2484   thus "\<exists> x. ?I x p" by blast
  2485 next
  2486   fix x assume "?I x p" hence "?I (- x) ?mp" 
  2487     using mirror[OF lp, where x="- x", symmetric] by auto
  2488   thus "\<exists> x. ?I x ?mp" by blast
  2489 qed
  2490 
  2491 lemma \<beta>_numbound0: assumes lp: "iszlfm p bs"
  2492   shows "\<forall> b\<in> set (\<beta> p). numbound0 b"
  2493   using lp by (induct p rule: \<beta>.induct,auto)
  2494 
  2495 lemma d\<beta>_mono: 
  2496   assumes linp: "iszlfm p (a #bs)"
  2497   and dr: "d\<beta> p l"
  2498   and d: "l dvd l'"
  2499   shows "d\<beta> p l'"
  2500 using dr linp zdvd_trans[where n="l" and k="l'", simplified d]
  2501 by (induct p rule: iszlfm.induct) simp_all
  2502 
  2503 lemma \<alpha>_l: assumes lp: "iszlfm p (a#bs)"
  2504   shows "\<forall> b\<in> set (\<alpha> p). numbound0 b \<and> isint b (a#bs)"
  2505 using lp
  2506 by(induct p rule: \<alpha>.induct, auto simp add: isint_add isint_c)
  2507 
  2508 lemma \<zeta>: 
  2509   assumes linp: "iszlfm p (a #bs)"
  2510   shows "\<zeta> p > 0 \<and> d\<beta> p (\<zeta> p)"
  2511 using linp
  2512 proof(induct p rule: iszlfm.induct)
  2513   case (1 p q)
  2514   from prems have dl1: "\<zeta> p dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2515   from prems have dl2: "\<zeta> q dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2516   from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2517     d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2518     dl1 dl2 show ?case by (auto simp add: ilcm_pos)
  2519 next
  2520   case (2 p q)
  2521   from prems have dl1: "\<zeta> p dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2522   from prems have dl2: "\<zeta> q dvd ilcm (\<zeta> p) (\<zeta> q)" by simp
  2523   from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2524     d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="ilcm (\<zeta> p) (\<zeta> q)"] 
  2525     dl1 dl2 show ?case by (auto simp add: ilcm_pos)
  2526 qed (auto simp add: ilcm_pos)
  2527 
  2528 lemma a\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d\<beta> p l" and lp: "l > 0"
  2529   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)"
  2530 using linp d
  2531 proof (induct p rule: iszlfm.induct)
  2532   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+
  2533     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2534     from cp have cnz: "c \<noteq> 0" by simp
  2535     have "c div c\<le> l div c"
  2536       by (simp add: zdiv_mono1[OF clel cp])
  2537     then have ldcp:"0 < l div c" 
  2538       by (simp add: zdiv_self[OF cnz])
  2539     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
  2540     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2541       by simp
  2542     hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\<Colon>real)) =
  2543           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)"
  2544       by simp
  2545     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)
  2546     also have "\<dots> = (real c * real x + Inum (real x # bs) e < 0)"
  2547     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
  2548   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
  2549 next
  2550   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+
  2551     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2552     from cp have cnz: "c \<noteq> 0" by simp
  2553     have "c div c\<le> l div c"
  2554       by (simp add: zdiv_mono1[OF clel cp])
  2555     then have ldcp:"0 < l div c" 
  2556       by (simp add: zdiv_self[OF cnz])
  2557     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
  2558     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2559       by simp
  2560     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0\<Colon>real)) =
  2561           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<le> 0)"
  2562       by simp
  2563     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)
  2564     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<le> 0)"
  2565     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
  2566   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
  2567 next
  2568   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+
  2569     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2570     from cp have cnz: "c \<noteq> 0" by simp
  2571     have "c div c\<le> l div c"
  2572       by (simp add: zdiv_mono1[OF clel cp])
  2573     then have ldcp:"0 < l div c" 
  2574       by (simp add: zdiv_self[OF cnz])
  2575     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
  2576     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2577       by simp
  2578     hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\<Colon>real)) =
  2579           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)"
  2580       by simp
  2581     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)
  2582     also have "\<dots> = (real c * real x + Inum (real x # bs) e > 0)"
  2583     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
  2584   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
  2585 next
  2586   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+
  2587     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2588     from cp have cnz: "c \<noteq> 0" by simp
  2589     have "c div c\<le> l div c"
  2590       by (simp add: zdiv_mono1[OF clel cp])
  2591     then have ldcp:"0 < l div c" 
  2592       by (simp add: zdiv_self[OF cnz])
  2593     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
  2594     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2595       by simp
  2596     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0\<Colon>real)) =
  2597           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<ge> 0)"
  2598       by simp
  2599     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)
  2600     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<ge> 0)"
  2601     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
  2602   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
  2603 next
  2604   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+
  2605     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2606     from cp have cnz: "c \<noteq> 0" by simp
  2607     have "c div c\<le> l div c"
  2608       by (simp add: zdiv_mono1[OF clel cp])
  2609     then have ldcp:"0 < l div c" 
  2610       by (simp add: zdiv_self[OF cnz])
  2611     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
  2612     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2613       by simp
  2614     hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\<Colon>real)) =
  2615           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)"
  2616       by simp
  2617     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)
  2618     also have "\<dots> = (real c * real x + Inum (real x # bs) e = 0)"
  2619     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
  2620   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
  2621 next
  2622   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+
  2623     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2624     from cp have cnz: "c \<noteq> 0" by simp
  2625     have "c div c\<le> l div c"
  2626       by (simp add: zdiv_mono1[OF clel cp])
  2627     then have ldcp:"0 < l div c" 
  2628       by (simp add: zdiv_self[OF cnz])
  2629     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
  2630     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2631       by simp
  2632     hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0\<Colon>real)) =
  2633           (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<noteq> 0)"
  2634       by simp
  2635     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)
  2636     also have "\<dots> = (real c * real x + Inum (real x # bs) e \<noteq> 0)"
  2637     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
  2638   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
  2639 next
  2640   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+
  2641     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2642     from cp have cnz: "c \<noteq> 0" by simp
  2643     have "c div c\<le> l div c"
  2644       by (simp add: zdiv_mono1[OF clel cp])
  2645     then have ldcp:"0 < l div c" 
  2646       by (simp add: zdiv_self[OF cnz])
  2647     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
  2648     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2649       by simp
  2650     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
  2651     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)
  2652     also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
  2653     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
  2654   also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
  2655   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 
  2656 next
  2657   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+
  2658     from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
  2659     from cp have cnz: "c \<noteq> 0" by simp
  2660     have "c div c\<le> l div c"
  2661       by (simp add: zdiv_mono1[OF clel cp])
  2662     then have ldcp:"0 < l div c" 
  2663       by (simp add: zdiv_self[OF cnz])
  2664     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
  2665     hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
  2666       by simp
  2667     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
  2668     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)
  2669     also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
  2670     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
  2671   also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
  2672   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
  2673 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)
  2674 
  2675 lemma a\<beta>_ex: assumes linp: "iszlfm p (a#bs)" and d: "d\<beta> p l" and lp: "l>0"
  2676   shows "(\<exists> x. l dvd x \<and> Ifm (real x #bs) (a\<beta> p l)) = (\<exists> (x::int). Ifm (real x#bs) p)"
  2677   (is "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> x. ?P' x)")
  2678 proof-
  2679   have "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> (x::int). ?P (l*x))"
  2680     using unity_coeff_ex[where l="l" and P="?P", simplified] by simp
  2681   also have "\<dots> = (\<exists> (x::int). ?P' x)" using a\<beta>[OF linp d lp] by simp
  2682   finally show ?thesis  . 
  2683 qed
  2684 
  2685 lemma \<beta>:
  2686   assumes lp: "iszlfm p (a#bs)"
  2687   and u: "d\<beta> p 1"
  2688   and d: "d\<delta> p d"
  2689   and dp: "d > 0"
  2690   and nob: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
  2691   and p: "Ifm (real x#bs) p" (is "?P x")
  2692   shows "?P (x - d)"
  2693 using lp u d dp nob p
  2694 proof(induct p rule: iszlfm.induct)
  2695   case (5 c e) hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
  2696     with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems
  2697     show ?case by (simp del: real_of_int_minus)
  2698 next
  2699   case (6 c e)  hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
  2700     with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems
  2701     show ?case by (simp del: real_of_int_minus)
  2702 next
  2703   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+
  2704     let ?e = "Inum (real x # bs) e"
  2705     from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"]
  2706       numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]
  2707       by (simp add: isint_iff)
  2708     {assume "real (x-d) +?e > 0" hence ?case using c1 
  2709       numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
  2710 	by (simp del: real_of_int_minus)}
  2711     moreover
  2712     {assume H: "\<not> real (x-d) + ?e > 0" 
  2713       let ?v="Neg e"
  2714       have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp
  2715       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"]] 
  2716       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e + real j)" by auto 
  2717       from H p have "real x + ?e > 0 \<and> real x + ?e \<le> real d" by (simp add: c1)
  2718       hence "real (x + floor ?e) > real (0::int) \<and> real (x + floor ?e) \<le> real d"
  2719 	using ie by simp
  2720       hence "x + floor ?e \<ge> 1 \<and> x + floor ?e \<le> d"  by simp
  2721       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e" by simp
  2722       hence "\<exists> (j::int) \<in> {1 .. d}. real x = real (- floor ?e + j)" 
  2723 	by (simp only: real_of_int_inject) (simp add: ring_simps)
  2724       hence "\<exists> (j::int) \<in> {1 .. d}. real x = - ?e + real j" 
  2725 	by (simp add: ie[simplified isint_iff])
  2726       with nob have ?case by auto}
  2727     ultimately show ?case by blast
  2728 next
  2729   case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" 
  2730     and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
  2731     let ?e = "Inum (real x # bs) e"
  2732     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"]
  2733       by (simp add: isint_iff)
  2734     {assume "real (x-d) +?e \<ge> 0" hence ?case using  c1 
  2735       numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
  2736 	by (simp del: real_of_int_minus)}
  2737     moreover
  2738     {assume H: "\<not> real (x-d) + ?e \<ge> 0" 
  2739       let ?v="Sub (C -1) e"
  2740       have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp
  2741       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"]] 
  2742       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e - 1 + real j)" by auto 
  2743       from H p have "real x + ?e \<ge> 0 \<and> real x + ?e < real d" by (simp add: c1)
  2744       hence "real (x + floor ?e) \<ge> real (0::int) \<and> real (x + floor ?e) < real d"
  2745 	using ie by simp
  2746       hence "x + floor ?e +1 \<ge> 1 \<and> x + floor ?e + 1 \<le> d"  by simp
  2747       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e + 1" by simp
  2748       hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: ring_simps)
  2749       hence "\<exists> (j::int) \<in> {1 .. d}. real x= real (- floor ?e - 1 + j)" 
  2750 	by (simp only: real_of_int_inject)
  2751       hence "\<exists> (j::int) \<in> {1 .. d}. real x= - ?e - 1 + real j" 
  2752 	by (simp add: ie[simplified isint_iff])
  2753       with nob have ?case by simp }
  2754     ultimately show ?case by blast
  2755 next
  2756   case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" 
  2757     and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
  2758     let ?e = "Inum (real x # bs) e"
  2759     let ?v="(Sub (C -1) e)"
  2760     have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
  2761     from p have "real x= - ?e" by (simp add: c1) with prems(11) show ?case using dp
  2762       by simp (erule ballE[where x="1"],
  2763 	simp_all add:ring_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
  2764 next
  2765   case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" 
  2766     and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
  2767     let ?e = "Inum (real x # bs) e"
  2768     let ?v="Neg e"
  2769     have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" by simp
  2770     {assume "real x - real d + Inum ((real (x -d)) # bs) e \<noteq> 0" 
  2771       hence ?case by (simp add: c1)}
  2772     moreover
  2773     {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0"
  2774       hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp
  2775       hence "real x = - Inum (a # bs) e + real d"
  2776 	by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"])
  2777        with prems(11) have ?case using dp by simp}
  2778   ultimately show ?case by blast
  2779 next 
  2780   case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" 
  2781     and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
  2782     let ?e = "Inum (real x # bs) e"
  2783     from prems have "isint e (a #bs)"  by simp 
  2784     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"]
  2785       by (simp add: isint_iff)
  2786     from prems have id: "j dvd d" by simp
  2787     from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp
  2788     also have "\<dots> = (j dvd x + floor ?e)" 
  2789       using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
  2790     also have "\<dots> = (j dvd x - d + floor ?e)" 
  2791       using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
  2792     also have "\<dots> = (real j rdvd real (x - d + floor ?e))" 
  2793       using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
  2794       ie by simp
  2795     also have "\<dots> = (real j rdvd real x - real d + ?e)" 
  2796       using ie by simp
  2797     finally show ?case 
  2798       using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
  2799 next
  2800   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+
  2801     let ?e = "Inum (real x # bs) e"
  2802     from prems have "isint e (a#bs)"  by simp 
  2803     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"]
  2804       by (simp add: isint_iff)
  2805     from prems have id: "j dvd d" by simp
  2806     from c1 ie[symmetric] have "?p x = (\<not> real j rdvd real (x+ floor ?e))" by simp
  2807     also have "\<dots> = (\<not> j dvd x + floor ?e)" 
  2808       using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
  2809     also have "\<dots> = (\<not> j dvd x - d + floor ?e)" 
  2810       using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
  2811     also have "\<dots> = (\<not> real j rdvd real (x - d + floor ?e))" 
  2812       using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
  2813       ie by simp
  2814     also have "\<dots> = (\<not> real j rdvd real x - real d + ?e)" 
  2815       using ie by simp
  2816     finally show ?case using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
  2817 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)
  2818 
  2819 lemma \<beta>':   
  2820   assumes lp: "iszlfm p (a #bs)"
  2821   and u: "d\<beta> p 1"
  2822   and d: "d\<delta> p d"
  2823   and dp: "d > 0"
  2824   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)")
  2825 proof(clarify)
  2826   fix x 
  2827   assume nb:"?b" and px: "?P x" 
  2828   hence nb2: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
  2829     by auto
  2830   from  \<beta>[OF lp u d dp nb2 px] show "?P (x -d )" .
  2831 qed
  2832 
  2833 lemma \<beta>_int: assumes lp: "iszlfm p bs"
  2834   shows "\<forall> b\<in> set (\<beta> p). isint b bs"
  2835 using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub)
  2836 
  2837 lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x))
  2838 ==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) 
  2839 ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D))))
  2840 ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
  2841 apply(rule iffI)
  2842 prefer 2
  2843 apply(drule minusinfinity)
  2844 apply assumption+
  2845 apply(fastsimp)
  2846 apply clarsimp
  2847 apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)")
  2848 apply(frule_tac x = x and z=z in decr_lemma)
  2849 apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
  2850 prefer 2
  2851 apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
  2852 prefer 2 apply arith
  2853  apply fastsimp
  2854 apply(drule (1)  periodic_finite_ex)
  2855 apply blast
  2856 apply(blast dest:decr_mult_lemma)
  2857 done
  2858 
  2859 
  2860 theorem cp_thm:
  2861   assumes lp: "iszlfm p (a #bs)"
  2862   and u: "d\<beta> p 1"
  2863   and d: "d\<delta> p d"
  2864   and dp: "d > 0"
  2865   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))"
  2866   (is "(\<exists> (x::int). ?P (real x)) = (\<exists> j\<in> ?D. ?M j \<or> (\<exists> b\<in> ?B. ?P (?I b + real j)))")
  2867 proof-
  2868   from minusinf_inf[OF lp] 
  2869   have th: "\<exists>(z::int). \<forall>x<z. ?P (real x) = ?M x" by blast
  2870   let ?B' = "{floor (?I b) | b. b\<in> ?B}"
  2871   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
  2872   from B[rule_format] 
  2873   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))" 
  2874     by simp
  2875   also have "\<dots> = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real (floor (?I b) + j)))" by simp
  2876   also have"\<dots> = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j)))"  by blast
  2877   finally have BB': 
  2878     "(\<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)))" 
  2879     by blast 
  2880   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
  2881   from minusinf_repeats[OF d lp]
  2882   have th3: "\<forall> x k. ?M x = ?M (x-k*d)" by simp
  2883   from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast
  2884 qed
  2885 
  2886     (* Reddy and Loveland *)
  2887 
  2888 
  2889 consts 
  2890   \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*)
  2891   \<sigma>\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*)
  2892   \<alpha>\<rho> :: "fm \<Rightarrow> (num\<times>int) list"
  2893   a\<rho> :: "fm \<Rightarrow> int \<Rightarrow> fm"
  2894 recdef \<rho> "measure size"
  2895   "\<rho> (And p q) = (\<rho> p @ \<rho> q)" 
  2896   "\<rho> (Or p q) = (\<rho> p @ \<rho> q)" 
  2897   "\<rho> (Eq  (CN 0 c e)) = [(Sub (C -1) e,c)]"
  2898   "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]"
  2899   "\<rho> (Lt  (CN 0 c e)) = []"
  2900   "\<rho> (Le  (CN 0 c e)) = []"
  2901   "\<rho> (Gt  (CN 0 c e)) = [(Neg e, c)]"
  2902   "\<rho> (Ge  (CN 0 c e)) = [(Sub (C (-1)) e, c)]"
  2903   "\<rho> p = []"
  2904 
  2905 recdef \<sigma>\<rho> "measure size"
  2906   "\<sigma>\<rho> (And p q) = (\<lambda> (t,k). And (\<sigma>\<rho> p (t,k)) (\<sigma>\<rho> q (t,k)))" 
  2907   "\<sigma>\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>\<rho> p (t,k)) (\<sigma>\<rho> q (t,k)))" 
  2908   "\<sigma>\<rho> (Eq  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) 
  2909                                             else (Eq (Add (Mul c t) (Mul k e))))"
  2910   "\<sigma>\<rho> (NEq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) 
  2911                                             else (NEq (Add (Mul c t) (Mul k e))))"
  2912   "\<sigma>\<rho> (Lt  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) 
  2913                                             else (Lt (Add (Mul c t) (Mul k e))))"
  2914   "\<sigma>\<rho> (Le  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) 
  2915                                             else (Le (Add (Mul c t) (Mul k e))))"
  2916   "\<sigma>\<rho> (Gt  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) 
  2917                                             else (Gt (Add (Mul c t) (Mul k e))))"
  2918   "\<sigma>\<rho> (Ge  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) 
  2919                                             else (Ge (Add (Mul c t) (Mul k e))))"
  2920   "\<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)) 
  2921                                             else (Dvd (i*k) (Add (Mul c t) (Mul k e))))"
  2922   "\<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)) 
  2923                                             else (NDvd (i*k) (Add (Mul c t) (Mul k e))))"
  2924   "\<sigma>\<rho> p = (\<lambda> (t,k). p)"
  2925 
  2926 recdef \<alpha>\<rho> "measure size"
  2927   "\<alpha>\<rho> (And p q) = (\<alpha>\<rho> p @ \<alpha>\<rho> q)" 
  2928   "\<alpha>\<rho> (Or p q) = (\<alpha>\<rho> p @ \<alpha>\<rho> q)" 
  2929   "\<alpha>\<rho> (Eq  (CN 0 c e)) = [(Add (C -1) e,c)]"
  2930   "\<alpha>\<rho> (NEq (CN 0 c e)) = [(e,c)]"
  2931   "\<alpha>\<rho> (Lt  (CN 0 c e)) = [(e,c)]"
  2932   "\<alpha>\<rho> (Le  (CN 0 c e)) = [(Add (C -1) e,c)]"
  2933   "\<alpha>\<rho> p = []"
  2934 
  2935     (* Simulates normal substituion by modifying the formula see correctness theorem *)
  2936 
  2937 recdef a\<rho> "measure size"
  2938   "a\<rho> (And p q) = (\<lambda> k. And (a\<rho> p k) (a\<rho> q k))" 
  2939   "a\<rho> (Or p q) = (\<lambda> k. Or (a\<rho> p k) (a\<rho> q k))" 
  2940   "a\<rho> (Eq (CN 0 c e)) = (\<lambda> k. if k dvd c then (Eq (CN 0 (c div k) e)) 
  2941                                            else (Eq (CN 0 c (Mul k e))))"
  2942   "a\<rho> (NEq (CN 0 c e)) = (\<lambda> k. if k dvd c then (NEq (CN 0 (c div k) e)) 
  2943                                            else (NEq (CN 0 c (Mul k e))))"
  2944   "a\<rho> (Lt (CN 0 c e)) = (\<lambda> k. if k dvd c then (Lt (CN 0 (c div k) e)) 
  2945                                            else (Lt (CN 0 c (Mul k e))))"
  2946   "a\<rho> (Le (CN 0 c e)) = (\<lambda> k. if k dvd c then (Le (CN 0 (c div k) e)) 
  2947                                            else (Le (CN 0 c (Mul k e))))"
  2948   "a\<rho> (Gt (CN 0 c e)) = (\<lambda> k. if k dvd c then (Gt (CN 0 (c div k) e)) 
  2949                                            else (Gt (CN 0 c (Mul k e))))"
  2950   "a\<rho> (Ge (CN 0 c e)) = (\<lambda> k. if k dvd c then (Ge (CN 0 (c div k) e)) 
  2951                                             else (Ge (CN 0 c (Mul k e))))"
  2952   "a\<rho> (Dvd i (CN 0 c e)) = (\<lambda> k. if k dvd c then (Dvd i (CN 0 (c div k) e)) 
  2953                                             else (Dvd (i*k) (CN 0 c (Mul k e))))"
  2954   "a\<rho> (NDvd i (CN 0 c e)) = (\<lambda> k. if k dvd c then (NDvd i (CN 0 (c div k) e)) 
  2955                                             else (NDvd (i*k) (CN 0 c (Mul k e))))"
  2956   "a\<rho> p = (\<lambda> k. p)"
  2957 
  2958 constdefs \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
  2959   "\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>\<rho> p (t,k))"
  2960 
  2961 lemma \<sigma>\<rho>:
  2962   assumes linp: "iszlfm p (real (x::int)#bs)"
  2963   and kpos: "real k > 0"
  2964   and tnb: "numbound0 t"
  2965   and tint: "isint t (real x#bs)"
  2966   and kdt: "k dvd floor (Inum (b'#bs) t)"
  2967   shows "Ifm (real x#bs) (\<sigma>\<rho> p (t,k)) = 
  2968   (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)" 
  2969   (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)")
  2970 using linp kpos tnb
  2971 proof(induct p rule: \<sigma>\<rho>.induct)
  2972   case (3 c e) 
  2973   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  2974     {assume kdc: "k dvd c" 
  2975       from kpos have knz: "k\<noteq>0" by simp
  2976       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  2977       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  2978 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  2979       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  2980     moreover 
  2981     {assume "\<not> k dvd c"
  2982       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  2983       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  2984       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)"
  2985 	using real_of_int_div[OF knz kdt]
  2986 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  2987 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  2988       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"]
  2989 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  2990 	by (simp add: ti)
  2991       finally have ?case . }
  2992     ultimately show ?case by blast 
  2993 next
  2994   case (4 c e)  
  2995   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  2996     {assume kdc: "k dvd c" 
  2997       from kpos have knz: "k\<noteq>0" by simp
  2998       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  2999       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3000 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3001       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3002     moreover 
  3003     {assume "\<not> k dvd c"
  3004       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3005       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3006       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)"
  3007 	using real_of_int_div[OF knz kdt]
  3008 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3009 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3010       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"]
  3011 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3012 	by (simp add: ti)
  3013       finally have ?case . }
  3014     ultimately show ?case by blast 
  3015 next
  3016   case (5 c e) 
  3017   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3018     {assume kdc: "k dvd c" 
  3019       from kpos have knz: "k\<noteq>0" by simp
  3020       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3021       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3022 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3023       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3024     moreover 
  3025     {assume "\<not> k dvd c"
  3026       from kpos have knz: "k\<noteq>0" by simp
  3027       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3028       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)"
  3029 	using real_of_int_div[OF knz kdt]
  3030 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3031 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3032       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"]
  3033 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3034 	by (simp add: ti)
  3035       finally have ?case . }
  3036     ultimately show ?case by blast 
  3037 next
  3038   case (6 c e)  
  3039   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3040     {assume kdc: "k dvd c" 
  3041       from kpos have knz: "k\<noteq>0" by simp
  3042       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3043       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3044 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3045       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3046     moreover 
  3047     {assume "\<not> k dvd c"
  3048       from kpos have knz: "k\<noteq>0" by simp
  3049       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3050       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)"
  3051 	using real_of_int_div[OF knz kdt]
  3052 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3053 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3054       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"]
  3055 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3056 	by (simp add: ti)
  3057       finally have ?case . }
  3058     ultimately show ?case by blast 
  3059 next
  3060   case (7 c e) 
  3061   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3062     {assume kdc: "k dvd c" 
  3063       from kpos have knz: "k\<noteq>0" by simp
  3064       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3065       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3066 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3067       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3068     moreover 
  3069     {assume "\<not> k dvd c"
  3070       from kpos have knz: "k\<noteq>0" by simp
  3071       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3072       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)"
  3073 	using real_of_int_div[OF knz kdt]
  3074 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3075 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3076       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"]
  3077 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3078 	by (simp add: ti)
  3079       finally have ?case . }
  3080     ultimately show ?case by blast 
  3081 next
  3082   case (8 c e)  
  3083   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3084     {assume kdc: "k dvd c" 
  3085       from kpos have knz: "k\<noteq>0" by simp
  3086       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3087       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3088 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3089       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3090     moreover 
  3091     {assume "\<not> k dvd c"
  3092       from kpos have knz: "k\<noteq>0" by simp
  3093       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3094       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)"
  3095 	using real_of_int_div[OF knz kdt]
  3096 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3097 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3098       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"]
  3099 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3100 	by (simp add: ti)
  3101       finally have ?case . }
  3102     ultimately show ?case by blast 
  3103 next
  3104   case (9 i c e)   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3105     {assume kdc: "k dvd c" 
  3106       from kpos have knz: "k\<noteq>0" by simp
  3107       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3108       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3109 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3110       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3111     moreover 
  3112     {assume "\<not> k dvd c"
  3113       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3114       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3115       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)"
  3116 	using real_of_int_div[OF knz kdt]
  3117 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3118 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3119       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"]
  3120 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3121 	by (simp add: ti)
  3122       finally have ?case . }
  3123     ultimately show ?case by blast 
  3124 next
  3125   case (10 i c e)    from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3126     {assume kdc: "k dvd c" 
  3127       from kpos have knz: "k\<noteq>0" by simp
  3128       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3129       from prems have  ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt]
  3130 	numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3131       numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
  3132     moreover 
  3133     {assume "\<not> k dvd c"
  3134       from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3135       from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
  3136       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))"
  3137 	using real_of_int_div[OF knz kdt]
  3138 	  numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
  3139 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti ring_simps)
  3140       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"]
  3141 	  numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
  3142 	by (simp add: ti)
  3143       finally have ?case . }
  3144     ultimately show ?case by blast 
  3145 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"])
  3146 
  3147 
  3148 lemma a\<rho>: 
  3149   assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "real k > 0" 
  3150   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")
  3151 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"]
  3152 proof(induct p rule: a\<rho>.induct)
  3153   case (3 c e)  
  3154   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3155   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3156     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3157     moreover 
  3158     {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)}
  3159     ultimately show ?case by blast 
  3160 next
  3161   case (4 c e)   
  3162   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3163   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3164     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3165     moreover 
  3166     {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)}
  3167     ultimately show ?case by blast 
  3168 next
  3169   case (5 c e)   
  3170   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3171   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3172     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3173     moreover 
  3174     {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)}
  3175     ultimately show ?case by blast 
  3176 next
  3177   case (6 c e)    
  3178   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3179   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3180     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3181     moreover 
  3182     {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)}
  3183     ultimately show ?case by blast 
  3184 next
  3185   case (7 c e)    
  3186   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3187   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3188     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3189     moreover 
  3190     {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)}
  3191     ultimately show ?case by blast 
  3192 next
  3193   case (8 c e)    
  3194   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3195   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3196     {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3197     moreover 
  3198     {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)}
  3199     ultimately show ?case by blast 
  3200 next
  3201   case (9 i c e)
  3202   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3203   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3204   {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3205   moreover 
  3206   {assume "\<not> k dvd c"
  3207     hence "Ifm (real (x*k)#bs) (a\<rho> (Dvd i (CN 0 c e)) k) = 
  3208       (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k)" 
  3209       using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] 
  3210       by (simp add: ring_simps)
  3211     also have "\<dots> = (Ifm (real x#bs) (Dvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"])
  3212     finally have ?case . }
  3213   ultimately show ?case by blast 
  3214 next
  3215   case (10 i c e) 
  3216   from prems have cp: "c > 0" and nb: "numbound0 e" by auto
  3217   from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
  3218   {assume kdc: "k dvd c" from prems have  ?case using real_of_int_div[OF knz kdc] by simp } 
  3219   moreover 
  3220   {assume "\<not> k dvd c"
  3221     hence "Ifm (real (x*k)#bs) (a\<rho> (NDvd i (CN 0 c e)) k) = 
  3222       (\<not> (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k))" 
  3223       using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] 
  3224       by (simp add: ring_simps)
  3225     also have "\<dots> = (Ifm (real x#bs) (NDvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"])
  3226     finally have ?case . }
  3227   ultimately show ?case by blast 
  3228 qed (simp_all add: nth_pos2)
  3229 
  3230 lemma a\<rho>_ex: 
  3231   assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0"
  3232   shows "(\<exists> (x::int). real k rdvd real x \<and> Ifm (real x#bs) (a\<rho> p k)) = 
  3233   (\<exists> (x::int). Ifm (real x#bs) p)" (is "(\<exists> x. ?D x \<and> ?P' x) = (\<exists> x. ?P x)")
  3234 proof-
  3235   have "(\<exists> x. ?D x \<and> ?P' x) = (\<exists> x. k dvd x \<and> ?P' x)" using int_rdvd_iff by simp
  3236   also have "\<dots> = (\<exists>x. ?P' (x*k))" using unity_coeff_ex[where P="?P'" and l="k", simplified]
  3237     by (simp add: ring_simps)
  3238   also have "\<dots> = (\<exists> x. ?P x)" using a\<rho> iszlfm_gen[OF lp] kp by auto
  3239   finally show ?thesis .
  3240 qed
  3241 
  3242 lemma \<sigma>\<rho>': assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" and nb: "numbound0 t"
  3243   shows "Ifm (real x#bs) (\<sigma>\<rho> p (t,k)) = Ifm ((Inum (real x#bs) t)#bs) (a\<rho> p k)"
  3244 using lp 
  3245 by(induct p rule: \<sigma>\<rho>.induct, simp_all add: 
  3246   numbound0_I[OF nb, where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] 
  3247   numbound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] 
  3248   bound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] nth_pos2 cong: imp_cong)
  3249 
  3250 lemma \<sigma>\<rho>_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t"
  3251   shows "bound0 (\<sigma>\<rho> p (t,k))"
  3252   using lp
  3253   by (induct p rule: iszlfm.induct, auto simp add: nb)
  3254 
  3255 lemma \<rho>_l:
  3256   assumes lp: "iszlfm p (real (i::int)#bs)"
  3257   shows "\<forall> (b,k) \<in> set (\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
  3258 using lp by (induct p rule: \<rho>.induct, auto simp add: isint_sub isint_neg)
  3259 
  3260 lemma \<alpha>\<rho>_l:
  3261   assumes lp: "iszlfm p (real (i::int)#bs)"
  3262   shows "\<forall> (b,k) \<in> set (\<alpha>\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
  3263 using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"]
  3264  by (induct p rule: \<alpha>\<rho>.induct, auto)
  3265 
  3266 lemma zminusinf_\<rho>:
  3267   assumes lp: "iszlfm p (real (i::int)#bs)"
  3268   and nmi: "\<not> (Ifm (real i#bs) (minusinf p))" (is "\<not> (Ifm (real i#bs) (?M p))")
  3269   and ex: "Ifm (real i#bs) p" (is "?I i p")
  3270   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")
  3271   using lp nmi ex
  3272 by (induct p rule: minusinf.induct, auto)
  3273 
  3274 
  3275 lemma \<sigma>_And: "Ifm bs (\<sigma> (And p q) k t)  = Ifm bs (And (\<sigma> p k t) (\<sigma> q k t))"
  3276 using \<sigma>_def by auto
  3277 lemma \<sigma>_Or: "Ifm bs (\<sigma> (Or p q) k t)  = Ifm bs (Or (\<sigma> p k t) (\<sigma> q k t))"
  3278 using \<sigma>_def by auto
  3279 
  3280 lemma \<rho>: assumes lp: "iszlfm p (real (i::int) #bs)"
  3281   and pi: "Ifm (real i#bs) p"
  3282   and d: "d\<delta> p d"
  3283   and dp: "d > 0"
  3284   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"
  3285   (is "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. _ \<noteq> ?N i e + _")
  3286   shows "Ifm (real(i - d)#bs) p"
  3287   using lp pi d nob
  3288 proof(induct p rule: iszlfm.induct)
  3289   case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3290     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"
  3291     by simp+
  3292   from mult_strict_left_mono[OF dp cp]  have one:"1 \<in> {1 .. c*d}" by auto
  3293   from nob[rule_format, where j="1", OF one] pi show ?case by simp
  3294 next
  3295   case (4 c e)  
  3296   hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3297     and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
  3298     by simp+
  3299   {assume "real (c*i) \<noteq> - ?N i e + real (c*d)"
  3300     with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"]
  3301     have ?case by (simp add: ring_simps)}
  3302   moreover
  3303   {assume pi: "real (c*i) = - ?N i e + real (c*d)"
  3304     from mult_strict_left_mono[OF dp cp] have d: "(c*d) \<in> {1 .. c*d}" by simp
  3305     from nob[rule_format, where j="c*d", OF d] pi have ?case by simp }
  3306   ultimately show ?case by blast
  3307 next
  3308   case (5 c e) hence cp: "c > 0" by simp
  3309   from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] 
  3310     real_of_int_mult]
  3311   show ?case using prems dp 
  3312     by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
  3313       ring_simps)
  3314 next
  3315   case (6 c e)  hence cp: "c > 0" by simp
  3316   from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] 
  3317     real_of_int_mult]
  3318   show ?case using prems dp 
  3319     by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
  3320       ring_simps)
  3321 next
  3322   case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3323     and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
  3324     and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0"
  3325     by simp+
  3326   let ?fe = "floor (?N i e)"
  3327   from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: ring_simps)
  3328   from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp
  3329   hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric])
  3330   have "real (c*i) + ?N i e > real (c*d) \<or> real (c*i) + ?N i e \<le> real (c*d)" by auto
  3331   moreover
  3332   {assume "real (c*i) + ?N i e > real (c*d)" hence ?case
  3333       by (simp add: ring_simps 
  3334 	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
  3335   moreover 
  3336   {assume H:"real (c*i) + ?N i e \<le> real (c*d)"
  3337     with ei[simplified isint_iff] have "real (c*i + ?fe) \<le> real (c*d)" by simp
  3338     hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
  3339     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
  3340     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1" 
  3341       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)
  3342     with nob  have ?case by blast }
  3343   ultimately show ?case by blast
  3344 next
  3345   case (8 c e)  hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
  3346     and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - 1 - ?N i e + real j"
  3347     and pi: "real (c*i) + ?N i e \<ge> 0" and cp': "real c >0"
  3348     by simp+
  3349   let ?fe = "floor (?N i e)"
  3350   from pi cp have th:"(real i +?N i e / real c)*real c \<ge> 0" by (simp add: ring_simps)
  3351   from pi ei[simplified isint_iff] have "real (c*i + ?fe) \<ge> real (0::int)" by simp
  3352   hence pi': "c*i + 1 + ?fe \<ge> 1" by (simp only: real_of_int_le_iff[symmetric])
  3353   have "real (c*i) + ?N i e \<ge> real (c*d) \<or> real (c*i) + ?N i e < real (c*d)" by auto
  3354   moreover
  3355   {assume "real (c*i) + ?N i e \<ge> real (c*d)" hence ?case
  3356       by (simp add: ring_simps 
  3357 	numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
  3358   moreover 
  3359   {assume H:"real (c*i) + ?N i e < real (c*d)"
  3360     with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp
  3361     hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
  3362     with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
  3363     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
  3364       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) 
  3365     hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
  3366       by (simp only: ring_simps diff_def[symmetric])
  3367         hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
  3368 	  by (simp only: add_ac diff_def)
  3369     with nob  have ?case by blast }
  3370   ultimately show ?case by blast
  3371 next
  3372   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+
  3373     let ?e = "Inum (real i # bs) e"
  3374     from prems have "isint e (real i #bs)"  by simp 
  3375     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"]
  3376       by (simp add: isint_iff)
  3377     from prems have id: "j dvd d" by simp
  3378     from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp
  3379     also have "\<dots> = (j dvd c*i + floor ?e)" 
  3380       using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
  3381     also have "\<dots> = (j dvd c*i - c*d + floor ?e)" 
  3382       using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
  3383     also have "\<dots> = (real j rdvd real (c*i - c*d + floor ?e))" 
  3384       using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
  3385       ie by simp
  3386     also have "\<dots> = (real j rdvd real (c*(i - d)) + ?e)" 
  3387       using ie by (simp add:ring_simps)
  3388     finally show ?case 
  3389       using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
  3390       by (simp add: ring_simps)
  3391 next
  3392   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+
  3393     let ?e = "Inum (real i # bs) e"
  3394     from prems have "isint e (real i #bs)"  by simp 
  3395     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"]
  3396       by (simp add: isint_iff)
  3397     from prems have id: "j dvd d" by simp
  3398     from ie[symmetric] have "?p i = (\<not> (real j rdvd real (c*i+ floor ?e)))" by simp
  3399     also have "\<dots> = Not (j dvd c*i + floor ?e)" 
  3400       using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
  3401     also have "\<dots> = Not (j dvd c*i - c*d + floor ?e)" 
  3402       using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
  3403     also have "\<dots> = Not (real j rdvd real (c*i - c*d + floor ?e))" 
  3404       using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
  3405       ie by simp
  3406     also have "\<dots> = Not (real j rdvd real (c*(i - d)) + ?e)" 
  3407       using ie by (simp add:ring_simps)
  3408     finally show ?case 
  3409       using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
  3410       by (simp add: ring_simps)
  3411 qed(auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] nth_pos2)
  3412 
  3413 lemma \<sigma>_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t"
  3414   shows "bound0 (\<sigma> p k t)"
  3415   using \<sigma>\<rho>_nb[OF lp nb] nb by (simp add: \<sigma>_def)
  3416   
  3417 lemma \<rho>':   assumes lp: "iszlfm p (a #bs)"
  3418   and d: "d\<delta> p d"
  3419   and dp: "d > 0"
  3420   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)")
  3421 proof(clarify)
  3422   fix x 
  3423   assume nob1:"?b x" and px: "?P x" 
  3424   from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)".
  3425   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" 
  3426   proof(clarify)
  3427     fix e c j assume ecR: "(e,c) \<in> set (\<rho> p)" and jD: "j\<in> {1 .. c*d}"
  3428       and cx: "real (c*x) = Inum (real x#bs) e + real j"
  3429     let ?e = "Inum (real x#bs) e"
  3430     let ?fe = "floor ?e"
  3431     from \<rho>_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e"
  3432       by auto
  3433     from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" .
  3434     from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp
  3435     hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject)
  3436     hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp)
  3437     hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff)
  3438     hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff])
  3439     from cx have "(c*x) div c = (?fe + j) div c" by simp
  3440     with cp have "x = (?fe + j) div c" by simp
  3441     with px have th: "?P ((?fe + j) div c)" by auto
  3442     from cp have cp': "real c > 0" by simp
  3443     from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp
  3444     from nb have nb': "numbound0 (Add e (C j))" by simp
  3445     have ji: "isint (C j) (real x#bs)" by (simp add: isint_def)
  3446     from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" .
  3447     from th \<sigma>\<rho>[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric]
  3448     have "Ifm (real x#bs) (\<sigma>\<rho> p (Add e (C j), c))" by simp
  3449     with rcdej have th: "Ifm (real x#bs) (\<sigma> p c (Add e (C j)))" by (simp add: \<sigma>_def)
  3450     from th bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"]
  3451     have "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" by blast
  3452       with ecR jD nob1    show "False" by blast
  3453   qed
  3454   from \<rho>[OF lp' px d dp nob] show "?P (x -d )" . 
  3455 qed
  3456 
  3457 
  3458 lemma rl_thm: 
  3459   assumes lp: "iszlfm p (real (i::int)#bs)"
  3460   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)))))"
  3461   (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))" 
  3462     is "?lhs = (?MD \<or> ?RD)"  is "?lhs = ?rhs")
  3463 proof-
  3464   let ?d= "\<delta> p"
  3465   from \<delta>[OF lp] have d:"d\<delta> p ?d" and dp: "?d > 0" by auto
  3466   { assume H:"?MD" hence th:"\<exists> (x::int). ?MP x" by blast
  3467     from H minusinf_ex[OF lp th] have ?thesis  by blast}
  3468   moreover
  3469   { fix e c j assume exR:"(e,c) \<in> ?R" and jD:"j\<in> {1 .. c*?d}" and spx:"?SP c e j"
  3470     from exR \<rho>_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0"
  3471       by auto
  3472     have "isint (C j) (real i#bs)" by (simp add: isint_iff)
  3473     with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]]
  3474     have eji:"isint (Add e (C j)) (real i#bs)" by simp
  3475     from nb have nb': "numbound0 (Add e (C j))" by simp
  3476     from spx bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"]
  3477     have spx': "Ifm (real i # bs) (\<sigma> p c (Add e (C j)))" by blast
  3478     from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))" 
  3479       and sr:"Ifm (real i#bs) (\<sigma>\<rho> p (Add e (C j),c))" by (simp add: \<sigma>_def)+
  3480     from rcdej eji[simplified isint_iff] 
  3481     have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp
  3482     hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff)
  3483     from cp have cp': "real c > 0" by simp
  3484     from \<sigma>\<rho>[OF lp cp' nb' eji cdej] spx' have "?P (\<lfloor>Inum (real i # bs) (Add e (C j))\<rfloor> div c)"
  3485       by (simp add: \<sigma>_def)
  3486     hence ?lhs by blast
  3487     with exR jD spx have ?thesis by blast}
  3488   moreover
  3489   { fix x assume px: "?P x" and nob: "\<not> ?RD"
  3490     from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" .
  3491     from \<rho>'[OF lp' d dp, rule_format, OF nob] have th:"\<forall> x. ?P x \<longrightarrow> ?P (x - ?d)" by blast
  3492     from minusinf_inf[OF lp] obtain z where z:"\<forall> x<z. ?MP x = ?P x" by blast
  3493     have zp: "abs (x - z) + 1 \<ge> 0" by arith
  3494     from decr_lemma[OF dp,where x="x" and z="z"] 
  3495       decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\<exists> x. ?MP x" by auto
  3496     with minusinf_bex[OF lp] px nob have ?thesis by blast}
  3497   ultimately show ?thesis by blast
  3498 qed
  3499 
  3500 lemma mirror_\<alpha>\<rho>:   assumes lp: "iszlfm p (a#bs)"
  3501   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))"
  3502 using lp
  3503 by (induct p rule: mirror.induct, simp_all add: split_def image_Un )
  3504   
  3505 text {* The @{text "\<real>"} part*}
  3506 
  3507 text{* Linearity for fm where Bound 0 ranges over @{text "\<real>"}*}
  3508 consts
  3509   isrlfm :: "fm \<Rightarrow> bool"   (* Linearity test for fm *)
  3510 recdef isrlfm "measure size"
  3511   "isrlfm (And p q) = (isrlfm p \<and> isrlfm q)" 
  3512   "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)" 
  3513   "isrlfm (Eq  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3514   "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3515   "isrlfm (Lt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3516   "isrlfm (Le  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3517   "isrlfm (Gt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3518   "isrlfm (Ge  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
  3519   "isrlfm p = (isatom p \<and> (bound0 p))"
  3520 
  3521 constdefs fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm"
  3522   "fp p n s j \<equiv> (if n > 0 then 
  3523             (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j)))))
  3524                         (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1))))))))
  3525             else 
  3526             (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) 
  3527                         (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))"
  3528 
  3529   (* splits the bounded from the unbounded part*)
  3530 consts rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list" 
  3531 recdef rsplit0 "measure num_size"
  3532   "rsplit0 (Bound 0) = [(T,1,C 0)]"
  3533   "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b 
  3534               in map (\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) [(a,b). a\<leftarrow>acs,b\<leftarrow>bcs])"
  3535   "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))"
  3536   "rsplit0 (Neg a) = map (\<lambda> (p,n,s). (p,-n,Neg s)) (rsplit0 a)"
  3537   "rsplit0 (Floor a) = foldl (op @) [] (map 
  3538       (\<lambda> (p,n,s). if n=0 then [(p,0,Floor s)]
  3539           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))))
  3540        (rsplit0 a))"
  3541   "rsplit0 (CN 0 c a) = map (\<lambda> (p,n,s). (p,n+c,s)) (rsplit0 a)"
  3542   "rsplit0 (CN m c a) = map (\<lambda> (p,n,s). (p,n,CN m c s)) (rsplit0 a)"
  3543   "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)"
  3544   "rsplit0 (Mul c a) = map (\<lambda> (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)"
  3545   "rsplit0 t = [(T,0,t)]"
  3546 
  3547 lemma not_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm (not p)"
  3548   by (induct p rule: isrlfm.induct, auto)
  3549 lemma conj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (conj p q)"
  3550   using conj_def by (cases p, auto)
  3551 lemma disj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (disj p q)"
  3552   using disj_def by (cases p, auto)
  3553 
  3554 
  3555 lemma rsplit0_cs:
  3556   shows "\<forall> (p,n,s) \<in> set (rsplit0 t). 
  3557   (Ifm (x#bs) p \<longrightarrow>  (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p" 
  3558   (is "\<forall> (p,n,s) \<in> ?SS t. (?I p \<longrightarrow> ?N t = ?N (CN 0 n s)) \<and> _ \<and> _ ")
  3559 proof(induct t rule: rsplit0.induct)
  3560   case (5 a) 
  3561   let ?p = "\<lambda> (p,n,s) j. fp p n s j"
  3562   let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))"
  3563   let ?J = "\<lambda> n. if n>0 then iupt (0,n) else iupt (n,0)"
  3564   let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
  3565   have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
  3566   have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
  3567     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto
  3568   have U2': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0}. 
  3569     ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))" by auto
  3570   hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
  3571     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). 
  3572     set (map (?f(p,n,s)) (iupt(0,n)))))"
  3573   proof-
  3574     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3575     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3576     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3577       by (auto simp add: split_def)
  3578   qed
  3579   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))"
  3580     by auto
  3581   hence U3: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
  3582     (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)))))"
  3583       proof-
  3584     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3585     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3586     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3587       by (auto simp add: split_def)
  3588   qed
  3589   from foldl_append_map_Nil_set[where xs="rsplit0 a" and f="?ff"]
  3590   have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" by auto
  3591   also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by auto
  3592   also have "\<dots> = 
  3593     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3594     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3595     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" 
  3596     using int_cases[rule_format] by blast
  3597   also have "\<dots> =  
  3598     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un 
  3599    (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 
  3600    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). 
  3601     set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3)
  3602   also have "\<dots> =  
  3603     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3604     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
  3605     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
  3606     by (simp only: set_map iupt_set set.simps)
  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) j| j. j\<in> {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) j| j. j\<in> {n .. 0}})))" by blast
  3611   finally 
  3612   have FS: "?SS (Floor a) =   
  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   show ?case
  3617     proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -)
  3618       fix p n s
  3619       let ?ths = "(?I p \<longrightarrow> (?N (Floor a) = ?N (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p"
  3620       assume "(\<exists>ba. (p, 0, ba) \<in> set (rsplit0 a) \<and> n = 0 \<and> s = Floor ba) \<or>
  3621        (\<exists>ab ac ba.
  3622            (ab, ac, ba) \<in> set (rsplit0 a) \<and>
  3623            0 < ac \<and>
  3624            (\<exists>j. p = fp ab ac ba j \<and>
  3625                 n = 0 \<and> s = Add (Floor ba) (C j) \<and> 0 \<le> j \<and> j \<le> ac)) \<or>
  3626        (\<exists>ab ac ba.
  3627            (ab, ac, ba) \<in> set (rsplit0 a) \<and>
  3628            ac < 0 \<and>
  3629            (\<exists>j. p = fp ab ac ba j \<and>
  3630                 n = 0 \<and> s = Add (Floor ba) (C j) \<and> ac \<le> j \<and> j \<le> 0))"
  3631       moreover 
  3632       {fix s'
  3633 	assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'"
  3634 	hence ?ths using prems by auto}
  3635       moreover
  3636       {	fix p' n' s' j
  3637 	assume pns: "(p', n', s') \<in> ?SS a" 
  3638 	  and np: "0 < n'" 
  3639 	  and p_def: "p = ?p (p',n',s') j" 
  3640 	  and n0: "n = 0" 
  3641 	  and s_def: "s = (Add (Floor s') (C j))" 
  3642 	  and jp: "0 \<le> j" and jn: "j \<le> n'"
  3643 	from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
  3644           Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
  3645           numbound0 s' \<and> isrlfm p'" by blast
  3646 	hence nb: "numbound0 s'" by simp
  3647 	from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb)
  3648 	let ?nxs = "CN 0 n' s'"
  3649 	let ?l = "floor (?N s') + j"
  3650 	from H 
  3651 	have "?I (?p (p',n',s') j) \<longrightarrow> 
  3652 	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
  3653 	  by (simp add: fp_def np ring_simps numsub numadd numfloor)
  3654 	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
  3655 	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
  3656 	moreover
  3657 	have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
  3658 	ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
  3659 	  by blast
  3660 	with s_def n0 p_def nb nf have ?ths by auto}
  3661       moreover
  3662       {fix p' n' s' j
  3663 	assume pns: "(p', n', s') \<in> ?SS a" 
  3664 	  and np: "n' < 0" 
  3665 	  and p_def: "p = ?p (p',n',s') j" 
  3666 	  and n0: "n = 0" 
  3667 	  and s_def: "s = (Add (Floor s') (C j))" 
  3668 	  and jp: "n' \<le> j" and jn: "j \<le> 0"
  3669 	from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
  3670           Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
  3671           numbound0 s' \<and> isrlfm p'" by blast
  3672 	hence nb: "numbound0 s'" by simp
  3673 	from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb)
  3674 	let ?nxs = "CN 0 n' s'"
  3675 	let ?l = "floor (?N s') + j"
  3676 	from H 
  3677 	have "?I (?p (p',n',s') j) \<longrightarrow> 
  3678 	  (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
  3679 	  by (simp add: np fp_def ring_simps numneg numfloor numadd numsub)
  3680 	also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
  3681 	  using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
  3682 	moreover
  3683 	have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"  by simp
  3684 	ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
  3685 	  by blast
  3686 	with s_def n0 p_def nb nf have ?ths by auto}
  3687       ultimately show ?ths by auto
  3688     qed
  3689 next
  3690   case (3 a b) thus ?case by auto
  3691 qed (auto simp add: Let_def split_def ring_simps conj_rl)
  3692 
  3693 lemma real_in_int_intervals: 
  3694   assumes xb: "real m \<le> x \<and> x < real ((n::int) + 1)"
  3695   shows "\<exists> j\<in> {m.. n}. real j \<le> x \<and> x < real (j+1)" (is "\<exists> j\<in> ?N. ?P j")
  3696 by (rule bexI[where P="?P" and x="floor x" and A="?N"]) 
  3697 (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"]])
  3698 
  3699 lemma rsplit0_complete:
  3700   assumes xp:"0 \<le> x" and x1:"x < 1"
  3701   shows "\<exists> (p,n,s) \<in> set (rsplit0 t). Ifm (x#bs) p" (is "\<exists> (p,n,s) \<in> ?SS t. ?I p")
  3702 proof(induct t rule: rsplit0.induct)
  3703   case (2 a b) 
  3704   from prems have "\<exists> (pa,na,sa) \<in> ?SS a. ?I pa" by auto
  3705   then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\<in> ?SS a \<and> ?I pa" by blast
  3706   from prems have "\<exists> (pb,nb,sb) \<in> ?SS b. ?I pb" by auto
  3707   then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\<in> ?SS b \<and> ?I pb" by blast
  3708   from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \<in> set[(x,y). x\<leftarrow>rsplit0 a, y\<leftarrow>rsplit0 b]"
  3709     by (auto)
  3710   let ?f="(\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))"
  3711   from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \<in> ?SS (Add a b)"
  3712     by (simp add: Let_def)
  3713   hence "(And pa pb, na +nb, Add sa sb) \<in> ?SS (Add a b)" by simp
  3714   moreover from pa pb have "?I (And pa pb)" by simp
  3715   ultimately show ?case by blast
  3716 next
  3717   case (5 a) 
  3718   let ?p = "\<lambda> (p,n,s) j. fp p n s j"
  3719   let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))"
  3720   let ?J = "\<lambda> n. if n>0 then iupt (0,n) else iupt (n,0)"
  3721   let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
  3722   have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
  3723   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
  3724   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))"
  3725     by auto
  3726   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)))))"
  3727   proof-
  3728     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3729     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3730     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3731       by (auto simp add: split_def)
  3732   qed
  3733   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))"
  3734     by auto
  3735   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)))))"
  3736   proof-
  3737     fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
  3738     assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
  3739     thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
  3740       by (auto simp add: split_def)
  3741   qed
  3742   from foldl_append_map_Nil_set[where xs="rsplit0 a" and f="?ff"]
  3743   have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" by auto
  3744   also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by auto
  3745   also have "\<dots> = 
  3746     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3747     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
  3748     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" 
  3749     using int_cases[rule_format] by blast
  3750   also have "\<dots> =  
  3751     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un 
  3752     (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 
  3753     (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)
  3754   also have "\<dots> =  
  3755     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
  3756     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
  3757     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
  3758     by (simp only: set_map iupt_set set.simps)
  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) j| j. j\<in> {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) j| j. j\<in> {n .. 0}})))" by blast
  3763   finally 
  3764   have FS: "?SS (Floor a) =   
  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   from prems have "\<exists> (p,n,s) \<in> ?SS a. ?I p" by auto
  3769   then obtain "p" "n" "s" where pns: "(p,n,s) \<in> ?SS a \<and> ?I p" by blast
  3770   let ?N = "\<lambda> t. Inum (x#bs) t"
  3771   from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \<and> numbound0 s \<and> isrlfm p"
  3772     by auto
  3773   
  3774   have "n=0 \<or> n >0 \<or> n <0" by arith
  3775   moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto }
  3776   moreover
  3777   {
  3778     assume np: "n > 0"
  3779     from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \<le> ?N s" by simp
  3780     also from mult_left_mono[OF xp] np have "?N s \<le> real n * x + ?N s" by simp
  3781     finally have "?N (Floor s) \<le> real n * x + ?N s" .
  3782     moreover
  3783     {from mult_strict_left_mono[OF x1] np 
  3784       have "real n *x + ?N s < real n + ?N s" by simp
  3785       also from real_of_int_floor_add_one_gt[where r="?N s"] 
  3786       have "\<dots> < real n + ?N (Floor s) + 1" by simp
  3787       finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp}
  3788     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
  3789     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
  3790     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
  3791     
  3792     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"
  3793       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)"]) 
  3794     hence "\<exists> j\<in> {0.. n}. ?I (?p (p,n,s) j)"
  3795       using pns by (simp add: fp_def np ring_simps numsub numadd)
  3796     then obtain "j" where j_def: "j\<in> {0 .. n} \<and> ?I (?p (p,n,s) j)" by blast
  3797     hence "\<exists>x \<in> {?p (p,n,s) j |j. 0\<le> j \<and> j \<le> n }. ?I x" by auto
  3798     hence ?case using pns 
  3799       by (simp only: FS,simp add: bex_Un) 
  3800     (rule disjI2, rule disjI1,rule exI [where x="p"],
  3801       rule exI [where x="n"],rule exI [where x="s"],simp_all add: np)
  3802   }
  3803   moreover
  3804   { assume nn: "n < 0" hence np: "-n >0" by simp
  3805     from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp
  3806     moreover from mult_left_mono_neg[OF xp] nn have "?N s \<ge> real n * x + ?N s" by simp
  3807     ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith 
  3808     moreover
  3809     {from mult_strict_left_mono_neg[OF x1, where c="real n"] nn
  3810       have "real n *x + ?N s \<ge> real n + ?N s" by simp 
  3811       moreover from real_of_int_floor_le[where r="?N s"]  have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp
  3812       ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n" 
  3813 	by (simp only: ring_simps)}
  3814     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
  3815     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
  3816     have th1: "\<forall> (a::real). (- a > 0) = (a < 0)" by auto
  3817     have th2: "\<forall> (a::real). (0 \<ge> - a) = (a \<ge> 0)" by auto
  3818     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
  3819     
  3820     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"
  3821       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)"]) 
  3822     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])
  3823     hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
  3824       using pns by (simp add: fp_def nn diff_def add_ac mult_ac numfloor numadd numneg
  3825 	del: diff_less_0_iff_less diff_le_0_iff_le) 
  3826     then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
  3827     hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
  3828     hence ?case using pns 
  3829       by (simp only: FS,simp add: bex_Un)
  3830     (rule disjI2, rule disjI2,rule exI [where x="p"],
  3831       rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn)
  3832   }
  3833   ultimately show ?case by blast
  3834 qed (auto simp add: Let_def split_def)
  3835 
  3836     (* Linearize a formula where Bound 0 ranges over [0,1) *)
  3837 
  3838 constdefs rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm"
  3839   "rsplit f a \<equiv> foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) (rsplit0 a)) F"
  3840 
  3841 lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\<exists> x \<in> set xs. Ifm bs (f x))"
  3842 by(induct xs, simp_all)
  3843 
  3844 lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\<forall> x \<in> set xs. Ifm bs (f x))"
  3845 by(induct xs, simp_all)
  3846 
  3847 lemma foldr_disj_map_rlfm: 
  3848   assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
  3849   and \<phi>: "\<forall> (\<phi>,n,s) \<in> set xs. numbound0 s \<and> isrlfm \<phi>"
  3850   shows "isrlfm (foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) xs) F)"
  3851 using lf \<phi> by (induct xs, auto)
  3852 
  3853 lemma rsplit_ex: "Ifm bs (rsplit f a) = (\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). Ifm bs (conj \<phi> (f n s)))"
  3854 using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def)
  3855 
  3856 lemma rsplit_l: assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
  3857   shows "isrlfm (rsplit f a)"
  3858 proof-
  3859   from rsplit0_cs[where t="a"] have th: "\<forall> (\<phi>,n,s) \<in> set (rsplit0 a). numbound0 s \<and> isrlfm \<phi>" by blast
  3860   from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp
  3861 qed
  3862 
  3863 lemma rsplit: 
  3864   assumes xp: "x \<ge> 0" and x1: "x < 1"
  3865   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))"
  3866   shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)"
  3867 proof(auto)
  3868   let ?I = "\<lambda>x p. Ifm (x#bs) p"
  3869   let ?N = "\<lambda> x t. Inum (x#bs) t"
  3870   assume "?I x (rsplit f a)"
  3871   hence "\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). ?I x (And \<phi> (f n s))" using rsplit_ex by simp
  3872   then obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and "?I x (And \<phi> (f n s))" by blast
  3873   hence \<phi>: "?I x \<phi>" and fns: "?I x (f n s)" by auto
  3874   from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \<phi> 
  3875   have th: "(?N x a = ?N x (CN 0 n s)) \<and> numbound0 s" by auto
  3876   from f[rule_format, OF th] fns show "?I x (g a)" by simp
  3877 next
  3878   let ?I = "\<lambda>x p. Ifm (x#bs) p"
  3879   let ?N = "\<lambda> x t. Inum (x#bs) t"
  3880   assume ga: "?I x (g a)"
  3881   from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] 
  3882   obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and fx: "?I x \<phi>" by blast
  3883   from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx
  3884   have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto
  3885   with ga f have "?I x (f n s)" by auto
  3886   with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto
  3887 qed
  3888 
  3889 definition lt :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3890   lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) 
  3891                         else (Gt (CN 0 (-c) (Neg t))))"
  3892 
  3893 definition  le :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3894   le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) 
  3895                         else (Ge (CN 0 (-c) (Neg t))))"
  3896 
  3897 definition  gt :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3898   gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) 
  3899                         else (Lt (CN 0 (-c) (Neg t))))"
  3900 
  3901 definition  ge :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3902   ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) 
  3903                         else (Le (CN 0 (-c) (Neg t))))"
  3904 
  3905 definition  eq :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3906   eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) 
  3907                         else (Eq (CN 0 (-c) (Neg t))))"
  3908 
  3909 definition neq :: "int \<Rightarrow> num \<Rightarrow> fm" where
  3910   neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) 
  3911                         else (NEq (CN 0 (-c) (Neg t))))"
  3912 
  3913 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)"
  3914   (is "\<forall> a n s . ?N a = ?N (CN 0 n s) \<and> _\<longrightarrow> ?I (lt n s) = ?I (Lt a)")
  3915 proof(clarify)
  3916   fix a n s
  3917   assume H: "?N a = ?N (CN 0 n s)"
  3918   show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def))
  3919   (cases "n > 0", simp_all add: lt_def ring_simps myless[rule_format, where b="0"])
  3920 qed
  3921 
  3922 lemma lt_l: "isrlfm (rsplit lt a)"
  3923   by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def,
  3924     case_tac s, simp_all, case_tac "nat", simp_all)
  3925 
  3926 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)")
  3927 proof(clarify)
  3928   fix a n s
  3929   assume H: "?N a = ?N (CN 0 n s)"
  3930   show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def))
  3931   (cases "n > 0", simp_all add: le_def ring_simps myl[rule_format, where b="0"])
  3932 qed
  3933 
  3934 lemma le_l: "isrlfm (rsplit le a)"
  3935   by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) 
  3936 (case_tac s, simp_all, case_tac "nat",simp_all)
  3937 
  3938 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)")
  3939 proof(clarify)
  3940   fix a n s
  3941   assume H: "?N a = ?N (CN 0 n s)"
  3942   show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def))
  3943   (cases "n > 0", simp_all add: gt_def ring_simps myless[rule_format, where b="0"])
  3944 qed
  3945 lemma gt_l: "isrlfm (rsplit gt a)"
  3946   by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) 
  3947 (case_tac s, simp_all, case_tac "nat", simp_all)
  3948 
  3949 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)")
  3950 proof(clarify)
  3951   fix a n s 
  3952   assume H: "?N a = ?N (CN 0 n s)"
  3953   show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def))
  3954   (cases "n > 0", simp_all add: ge_def ring_simps myl[rule_format, where b="0"])
  3955 qed
  3956 lemma ge_l: "isrlfm (rsplit ge a)"
  3957   by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) 
  3958 (case_tac s, simp_all, case_tac "nat", simp_all)
  3959 
  3960 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)")
  3961 proof(clarify)
  3962   fix a n s 
  3963   assume H: "?N a = ?N (CN 0 n s)"
  3964   show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def ring_simps)
  3965 qed
  3966 lemma eq_l: "isrlfm (rsplit eq a)"
  3967   by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) 
  3968 (case_tac s, simp_all, case_tac"nat", simp_all)
  3969 
  3970 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)")
  3971 proof(clarify)
  3972   fix a n s bs
  3973   assume H: "?N a = ?N (CN 0 n s)"
  3974   show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def ring_simps)
  3975 qed
  3976 
  3977 lemma neq_l: "isrlfm (rsplit neq a)"
  3978   by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) 
  3979 (case_tac s, simp_all, case_tac"nat", simp_all)
  3980 
  3981 lemma small_le: 
  3982   assumes u0:"0 \<le> u" and u1: "u < 1"
  3983   shows "(-u \<le> real (n::int)) = (0 \<le> n)"
  3984 using u0 u1  by auto
  3985 
  3986 lemma small_lt: 
  3987   assumes u0:"0 \<le> u" and u1: "u < 1"
  3988   shows "(real (n::int) < real (m::int) - u) = (n < m)"
  3989 using u0 u1  by auto
  3990 
  3991 lemma rdvd01_cs: 
  3992   assumes up: "u \<ge> 0" and u1: "u<1" and np: "real n > 0"
  3993   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")
  3994 proof-
  3995   let ?ss = "s - real (floor s)"
  3996   from real_of_int_floor_add_one_gt[where r="s", simplified myless[rule_format,where a="s"]] 
  3997     real_of_int_floor_le[where r="s"]  have ss0:"?ss \<ge> 0" and ss1:"?ss < 1" 
  3998     by (auto simp add: myl[rule_format, where b="s", symmetric] myless[rule_format, where a="?ss"])
  3999   from np have n0: "real n \<ge> 0" by simp
  4000   from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] 
  4001   have nu0:"real n * u - s \<ge> -s" and nun:"real n * u -s < real n - s" by auto  
  4002   from int_rdvd_real[where i="i" and x="real (n::int) * u - s"] 
  4003   have "real i rdvd real n * u - s = 
  4004     (i dvd floor (real n * u -s) \<and> (real (floor (real n * u - s)) = real n * u - s ))" 
  4005     (is "_ = (?DE)" is "_ = (?D \<and> ?E)") by simp
  4006   also have "\<dots> = (?DE \<and> real(floor (real n * u - s) + floor s)\<ge> -?ss 
  4007     \<and> real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \<and>real ?a \<ge> _ \<and> real ?a < _)")
  4008     using nu0 nun  by auto
  4009   also have "\<dots> = (?DE \<and> ?a \<ge> 0 \<and> ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1])
  4010   also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. ?a = j))" by simp
  4011   also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. real (\<lfloor>real n * u - s\<rfloor>) = real j - real \<lfloor>s\<rfloor> ))"
  4012     by (simp only: ring_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff)
  4013   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>"]
  4014     by (auto cong: conj_cong)
  4015   also have "\<dots> = ?rhs" by(simp cong: conj_cong) (simp add: ring_simps )
  4016   finally show ?thesis .
  4017 qed
  4018 
  4019 definition
  4020   DVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
  4021 where
  4022   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)"
  4023 
  4024 definition
  4025   NDVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
  4026 where
  4027   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)"
  4028 
  4029 lemma DVDJ_DVD: 
  4030   assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
  4031   shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))"
  4032 proof-
  4033   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))))"
  4034   let ?s= "Inum (x#bs) s"
  4035   from foldr_disj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"]
  4036   have "Ifm (x#bs) (DVDJ i n s) = (\<exists> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
  4037     by (simp add: iupt_set np DVDJ_def del: iupt.simps)
  4038   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])
  4039   also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
  4040   have "\<dots> = (real i rdvd real n * x - (-?s))" by simp
  4041   finally show ?thesis by simp
  4042 qed
  4043 
  4044 lemma NDVDJ_NDVD: 
  4045   assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
  4046   shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))"
  4047 proof-
  4048   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))))"
  4049   let ?s= "Inum (x#bs) s"
  4050   from foldr_conj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"]
  4051   have "Ifm (x#bs) (NDVDJ i n s) = (\<forall> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
  4052     by (simp add: iupt_set np NDVDJ_def del: iupt.simps)
  4053   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])
  4054   also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
  4055   have "\<dots> = (\<not> (real i rdvd real n * x - (-?s)))" by simp
  4056   finally show ?thesis by simp
  4057 qed  
  4058 
  4059 lemma foldr_disj_map_rlfm2: 
  4060   assumes lf: "\<forall> n . isrlfm (f n)"
  4061   shows "isrlfm (foldr disj (map f xs) F)"
  4062 using lf by (induct xs, auto)
  4063 lemma foldr_And_map_rlfm2: 
  4064   assumes lf: "\<forall> n . isrlfm (f n)"
  4065   shows "isrlfm (foldr conj (map f xs) T)"
  4066 using lf by (induct xs, auto)
  4067 
  4068 lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
  4069   shows "isrlfm (DVDJ i n s)"
  4070 proof-
  4071   let ?f="\<lambda>j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
  4072                          (Dvd i (Sub (C j) (Floor (Neg s))))"
  4073   have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
  4074   from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp 
  4075 qed
  4076 
  4077 lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
  4078   shows "isrlfm (NDVDJ i n s)"
  4079 proof-
  4080   let ?f="\<lambda>j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
  4081                       (NDvd i (Sub (C j) (Floor (Neg s))))"
  4082   have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
  4083   from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto
  4084 qed
  4085 
  4086 definition DVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
  4087   DVD_def: "DVD i c t =
  4088   (if i=0 then eq c t else 
  4089   if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))"
  4090 
  4091 definition  NDVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
  4092   "NDVD i c t =
  4093   (if i=0 then neq c t else 
  4094   if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))"
  4095 
  4096 lemma DVD_mono: 
  4097   assumes xp: "0\<le> x" and x1: "x < 1" 
  4098   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)"
  4099   (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (DVD i n s) = ?I (Dvd i a)")
  4100 proof(clarify)
  4101   fix a n s 
  4102   assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
  4103   let ?th = "?I (DVD i n s) = ?I (Dvd i a)"
  4104   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
  4105   moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] 
  4106       by (simp add: DVD_def rdvd_left_0_eq)}
  4107   moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H DVD_def) } 
  4108   moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
  4109       by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 
  4110 	rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
  4111   moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)}
  4112   ultimately show ?th by blast
  4113 qed
  4114 
  4115 lemma NDVD_mono:   assumes xp: "0\<le> x" and x1: "x < 1" 
  4116   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)"
  4117   (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (NDVD i n s) = ?I (NDvd i a)")
  4118 proof(clarify)
  4119   fix a n s 
  4120   assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
  4121   let ?th = "?I (NDVD i n s) = ?I (NDvd i a)"
  4122   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
  4123   moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] 
  4124       by (simp add: NDVD_def rdvd_left_0_eq)}
  4125   moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H NDVD_def) } 
  4126   moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
  4127       by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 
  4128 	rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
  4129   moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th 
  4130       by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)}
  4131   ultimately show ?th by blast
  4132 qed
  4133 
  4134 lemma DVD_l: "isrlfm (rsplit (DVD i) a)"
  4135   by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) 
  4136 (case_tac s, simp_all, case_tac "nat", simp_all)
  4137 
  4138 lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)"
  4139   by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) 
  4140 (case_tac s, simp_all, case_tac "nat", simp_all)
  4141 
  4142 consts rlfm :: "fm \<Rightarrow> fm"
  4143 recdef rlfm "measure fmsize"
  4144   "rlfm (And p q) = conj (rlfm p) (rlfm q)"
  4145   "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
  4146   "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
  4147   "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
  4148   "rlfm (Lt a) = rsplit lt a"
  4149   "rlfm (Le a) = rsplit le a"
  4150   "rlfm (Gt a) = rsplit gt a"
  4151   "rlfm (Ge a) = rsplit ge a"
  4152   "rlfm (Eq a) = rsplit eq a"
  4153   "rlfm (NEq a) = rsplit neq a"
  4154   "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
  4155   "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
  4156   "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
  4157   "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
  4158   "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
  4159   "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
  4160   "rlfm (NOT (NOT p)) = rlfm p"
  4161   "rlfm (NOT T) = F"
  4162   "rlfm (NOT F) = T"
  4163   "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
  4164   "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
  4165   "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
  4166   "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
  4167   "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
  4168   "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
  4169   "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
  4170   "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
  4171   "rlfm p = p" (hints simp add: fmsize_pos)
  4172 
  4173 lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p"
  4174   by (induct p rule: isrlfm.induct, auto)
  4175 lemma igcd_le1: assumes ip: "0 < i" shows "igcd i j \<le> i"
  4176 proof-
  4177   from igcd_dvd1 have th: "igcd i j dvd i" by blast
  4178   from zdvd_imp_le[OF th ip] show ?thesis .
  4179 qed
  4180 
  4181 
  4182 lemma simpfm_rl: "isrlfm p \<Longrightarrow> isrlfm (simpfm p)"
  4183 proof (induct p)
  4184   case (Lt a) 
  4185   hence "bound0 (Lt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4186     by (cases a,simp_all, case_tac "nat", simp_all)
  4187   moreover
  4188   {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))"  
  4189       using simpfm_bound0 by blast
  4190     have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def)
  4191     with bn bound0at_l have ?case by blast}
  4192   moreover 
  4193   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4194     {
  4195       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4196       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4197       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4198       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4199 	by (simp add: numgcd_def igcd_le1)
  4200       from prems have th': "c\<noteq>0" by auto
  4201       from prems have cp: "c \<ge> 0" by simp
  4202       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4203 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4204     }
  4205     with prems have ?case
  4206       by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
  4207   ultimately show ?case by blast
  4208 next
  4209   case (Le a)   
  4210   hence "bound0 (Le a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4211     by (cases a,simp_all, case_tac "nat", simp_all)
  4212   moreover
  4213   {assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))"  
  4214       using simpfm_bound0 by blast
  4215     have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def)
  4216     with bn bound0at_l have ?case by blast}
  4217   moreover 
  4218   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4219     {
  4220       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4221       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4222       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4223       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4224 	by (simp add: numgcd_def igcd_le1)
  4225       from prems have th': "c\<noteq>0" by auto
  4226       from prems have cp: "c \<ge> 0" by simp
  4227       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4228 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4229     }
  4230     with prems have ?case
  4231       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4232   ultimately show ?case by blast
  4233 next
  4234   case (Gt a)   
  4235   hence "bound0 (Gt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4236     by (cases a,simp_all, case_tac "nat", simp_all)
  4237   moreover
  4238   {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))"  
  4239       using simpfm_bound0 by blast
  4240     have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def)
  4241     with bn bound0at_l have ?case by blast}
  4242   moreover 
  4243   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4244     {
  4245       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4246       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4247       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4248       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4249 	by (simp add: numgcd_def igcd_le1)
  4250       from prems have th': "c\<noteq>0" by auto
  4251       from prems have cp: "c \<ge> 0" by simp
  4252       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4253 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4254     }
  4255     with prems have ?case
  4256       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4257   ultimately show ?case by blast
  4258 next
  4259   case (Ge a)   
  4260   hence "bound0 (Ge a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4261     by (cases a,simp_all, case_tac "nat", simp_all)
  4262   moreover
  4263   {assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))"  
  4264       using simpfm_bound0 by blast
  4265     have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def)
  4266     with bn bound0at_l have ?case by blast}
  4267   moreover 
  4268   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4269     {
  4270       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4271       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4272       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4273       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4274 	by (simp add: numgcd_def igcd_le1)
  4275       from prems have th': "c\<noteq>0" by auto
  4276       from prems have cp: "c \<ge> 0" by simp
  4277       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4278 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4279     }
  4280     with prems have ?case
  4281       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4282   ultimately show ?case by blast
  4283 next
  4284   case (Eq a)   
  4285   hence "bound0 (Eq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4286     by (cases a,simp_all, case_tac "nat", simp_all)
  4287   moreover
  4288   {assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))"  
  4289       using simpfm_bound0 by blast
  4290     have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def)
  4291     with bn bound0at_l have ?case by blast}
  4292   moreover 
  4293   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4294     {
  4295       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4296       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4297       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4298       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4299 	by (simp add: numgcd_def igcd_le1)
  4300       from prems have th': "c\<noteq>0" by auto
  4301       from prems have cp: "c \<ge> 0" by simp
  4302       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4303 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4304     }
  4305     with prems have ?case
  4306       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4307   ultimately show ?case by blast
  4308 next
  4309   case (NEq a)  
  4310   hence "bound0 (NEq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
  4311     by (cases a,simp_all, case_tac "nat", simp_all)
  4312   moreover
  4313   {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))"  
  4314       using simpfm_bound0 by blast
  4315     have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def)
  4316     with bn bound0at_l have ?case by blast}
  4317   moreover 
  4318   {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e"
  4319     {
  4320       assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
  4321       with numgcd_pos[where t="CN 0 c (simpnum e)"]
  4322       have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
  4323       from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
  4324 	by (simp add: numgcd_def igcd_le1)
  4325       from prems have th': "c\<noteq>0" by auto
  4326       from prems have cp: "c \<ge> 0" by simp
  4327       from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']]
  4328 	have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
  4329     }
  4330     with prems have ?case
  4331       by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
  4332   ultimately show ?case by blast
  4333 next
  4334   case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))"  
  4335     using simpfm_bound0 by blast
  4336   have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
  4337   with bn bound0at_l show ?case by blast
  4338 next
  4339   case (NDvd i a)  hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))"  
  4340     using simpfm_bound0 by blast
  4341   have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
  4342   with bn bound0at_l show ?case by blast
  4343 qed(auto simp add: conj_def imp_def disj_def iff_def Let_def simpfm_bound0 numadd_nb numneg_nb)
  4344 
  4345 lemma rlfm_I:
  4346   assumes qfp: "qfree p"
  4347   and xp: "0 \<le> x" and x1: "x < 1"
  4348   shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \<and> isrlfm (rlfm p)"
  4349   using qfp 
  4350 by (induct p rule: rlfm.induct) 
  4351 (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
  4352                rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l
  4353                rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl)
  4354 lemma rlfm_l:
  4355   assumes qfp: "qfree p"
  4356   shows "isrlfm (rlfm p)"
  4357   using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l 
  4358 by (induct p rule: rlfm.induct,auto simp add: simpfm_rl)
  4359 
  4360     (* Operations needed for Ferrante and Rackoff *)
  4361 lemma rminusinf_inf:
  4362   assumes lp: "isrlfm p"
  4363   shows "\<exists> z. \<forall> x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
  4364 using lp
  4365 proof (induct p rule: minusinf.induct)
  4366   case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
  4367 next
  4368   case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
  4369 next
  4370   case (3 c e) 
  4371   from prems have nb: "numbound0 e" by simp
  4372   from prems have cp: "real c > 0" by simp
  4373   let ?e="Inum (a#bs) e"
  4374   let ?z = "(- ?e) / real c"
  4375   {fix x
  4376     assume xz: "x < ?z"
  4377     hence "(real c * x < - ?e)" 
  4378       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4379     hence "real c * x + ?e < 0" by arith
  4380     hence "real c * x + ?e \<noteq> 0" by simp
  4381     with xz have "?P ?z x (Eq (CN 0 c e))"
  4382       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp  }
  4383   hence "\<forall> x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp
  4384   thus ?case by blast
  4385 next
  4386   case (4 c e)   
  4387   from prems have nb: "numbound0 e" by simp
  4388   from prems have cp: "real c > 0" by simp
  4389   let ?e="Inum (a#bs) e"
  4390   let ?z = "(- ?e) / real c"
  4391   {fix x
  4392     assume xz: "x < ?z"
  4393     hence "(real c * x < - ?e)" 
  4394       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4395     hence "real c * x + ?e < 0" by arith
  4396     hence "real c * x + ?e \<noteq> 0" by simp
  4397     with xz have "?P ?z x (NEq (CN 0 c e))"
  4398       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4399   hence "\<forall> x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp
  4400   thus ?case by blast
  4401 next
  4402   case (5 c e) 
  4403     from prems have nb: "numbound0 e" by simp
  4404   from prems have cp: "real c > 0" by simp
  4405   let ?e="Inum (a#bs) e"
  4406   let ?z = "(- ?e) / real c"
  4407   {fix x
  4408     assume xz: "x < ?z"
  4409     hence "(real c * x < - ?e)" 
  4410       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4411     hence "real c * x + ?e < 0" by arith
  4412     with xz have "?P ?z x (Lt (CN 0 c e))"
  4413       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"]  by simp }
  4414   hence "\<forall> x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp
  4415   thus ?case by blast
  4416 next
  4417   case (6 c e)  
  4418     from prems have nb: "numbound0 e" by simp
  4419   from prems have cp: "real c > 0" by simp
  4420   let ?e="Inum (a#bs) e"
  4421   let ?z = "(- ?e) / real c"
  4422   {fix x
  4423     assume xz: "x < ?z"
  4424     hence "(real c * x < - ?e)" 
  4425       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4426     hence "real c * x + ?e < 0" by arith
  4427     with xz have "?P ?z x (Le (CN 0 c e))"
  4428       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4429   hence "\<forall> x < ?z. ?P ?z x (Le (CN 0 c e))" by simp
  4430   thus ?case by blast
  4431 next
  4432   case (7 c e)  
  4433     from prems have nb: "numbound0 e" by simp
  4434   from prems have cp: "real c > 0" by simp
  4435   let ?e="Inum (a#bs) e"
  4436   let ?z = "(- ?e) / real c"
  4437   {fix x
  4438     assume xz: "x < ?z"
  4439     hence "(real c * x < - ?e)" 
  4440       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4441     hence "real c * x + ?e < 0" by arith
  4442     with xz have "?P ?z x (Gt (CN 0 c e))"
  4443       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4444   hence "\<forall> x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp
  4445   thus ?case by blast
  4446 next
  4447   case (8 c e)  
  4448     from prems have nb: "numbound0 e" by simp
  4449   from prems have cp: "real c > 0" by simp
  4450   let ?e="Inum (a#bs) e"
  4451   let ?z = "(- ?e) / real c"
  4452   {fix x
  4453     assume xz: "x < ?z"
  4454     hence "(real c * x < - ?e)" 
  4455       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) 
  4456     hence "real c * x + ?e < 0" by arith
  4457     with xz have "?P ?z x (Ge (CN 0 c e))"
  4458       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4459   hence "\<forall> x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp
  4460   thus ?case by blast
  4461 qed simp_all
  4462 
  4463 lemma rplusinf_inf:
  4464   assumes lp: "isrlfm p"
  4465   shows "\<exists> z. \<forall> x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
  4466 using lp
  4467 proof (induct p rule: isrlfm.induct)
  4468   case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
  4469 next
  4470   case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
  4471 next
  4472   case (3 c e) 
  4473   from prems have nb: "numbound0 e" by simp
  4474   from prems have cp: "real c > 0" by simp
  4475   let ?e="Inum (a#bs) e"
  4476   let ?z = "(- ?e) / real c"
  4477   {fix x
  4478     assume xz: "x > ?z"
  4479     with mult_strict_right_mono [OF xz cp] cp
  4480     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4481     hence "real c * x + ?e > 0" by arith
  4482     hence "real c * x + ?e \<noteq> 0" by simp
  4483     with xz have "?P ?z x (Eq (CN 0 c e))"
  4484       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4485   hence "\<forall> x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp
  4486   thus ?case by blast
  4487 next
  4488   case (4 c e) 
  4489   from prems have nb: "numbound0 e" by simp
  4490   from prems have cp: "real c > 0" by simp
  4491   let ?e="Inum (a#bs) e"
  4492   let ?z = "(- ?e) / real c"
  4493   {fix x
  4494     assume xz: "x > ?z"
  4495     with mult_strict_right_mono [OF xz cp] cp
  4496     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4497     hence "real c * x + ?e > 0" by arith
  4498     hence "real c * x + ?e \<noteq> 0" by simp
  4499     with xz have "?P ?z x (NEq (CN 0 c e))"
  4500       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4501   hence "\<forall> x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp
  4502   thus ?case by blast
  4503 next
  4504   case (5 c e) 
  4505   from prems have nb: "numbound0 e" by simp
  4506   from prems have cp: "real c > 0" by simp
  4507   let ?e="Inum (a#bs) e"
  4508   let ?z = "(- ?e) / real c"
  4509   {fix x
  4510     assume xz: "x > ?z"
  4511     with mult_strict_right_mono [OF xz cp] cp
  4512     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4513     hence "real c * x + ?e > 0" by arith
  4514     with xz have "?P ?z x (Lt (CN 0 c e))"
  4515       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4516   hence "\<forall> x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp
  4517   thus ?case by blast
  4518 next
  4519   case (6 c e) 
  4520   from prems have nb: "numbound0 e" by simp
  4521   from prems have cp: "real c > 0" by simp
  4522   let ?e="Inum (a#bs) e"
  4523   let ?z = "(- ?e) / real c"
  4524   {fix x
  4525     assume xz: "x > ?z"
  4526     with mult_strict_right_mono [OF xz cp] cp
  4527     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4528     hence "real c * x + ?e > 0" by arith
  4529     with xz have "?P ?z x (Le (CN 0 c e))"
  4530       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4531   hence "\<forall> x > ?z. ?P ?z x (Le (CN 0 c e))" by simp
  4532   thus ?case by blast
  4533 next
  4534   case (7 c e) 
  4535   from prems have nb: "numbound0 e" by simp
  4536   from prems have cp: "real c > 0" by simp
  4537   let ?e="Inum (a#bs) e"
  4538   let ?z = "(- ?e) / real c"
  4539   {fix x
  4540     assume xz: "x > ?z"
  4541     with mult_strict_right_mono [OF xz cp] cp
  4542     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4543     hence "real c * x + ?e > 0" by arith
  4544     with xz have "?P ?z x (Gt (CN 0 c e))"
  4545       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
  4546   hence "\<forall> x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp
  4547   thus ?case by blast
  4548 next
  4549   case (8 c e) 
  4550   from prems have nb: "numbound0 e" by simp
  4551   from prems have cp: "real c > 0" by simp
  4552   let ?e="Inum (a#bs) e"
  4553   let ?z = "(- ?e) / real c"
  4554   {fix x
  4555     assume xz: "x > ?z"
  4556     with mult_strict_right_mono [OF xz cp] cp
  4557     have "(real c * x > - ?e)" by (simp add: mult_ac)
  4558     hence "real c * x + ?e > 0" by arith
  4559     with xz have "?P ?z x (Ge (CN 0 c e))"
  4560       using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"]   by simp }
  4561   hence "\<forall> x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp
  4562   thus ?case by blast
  4563 qed simp_all
  4564 
  4565 lemma rminusinf_bound0:
  4566   assumes lp: "isrlfm p"
  4567   shows "bound0 (minusinf p)"
  4568   using lp
  4569   by (induct p rule: minusinf.induct) simp_all
  4570 
  4571 lemma rplusinf_bound0:
  4572   assumes lp: "isrlfm p"
  4573   shows "bound0 (plusinf p)"
  4574   using lp
  4575   by (induct p rule: plusinf.induct) simp_all
  4576 
  4577 lemma rminusinf_ex:
  4578   assumes lp: "isrlfm p"
  4579   and ex: "Ifm (a#bs) (minusinf p)"
  4580   shows "\<exists> x. Ifm (x#bs) p"
  4581 proof-
  4582   from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
  4583   have th: "\<forall> x. Ifm (x#bs) (minusinf p)" by auto
  4584   from rminusinf_inf[OF lp, where bs="bs"] 
  4585   obtain z where z_def: "\<forall>x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p" by blast
  4586   from th have "Ifm ((z - 1)#bs) (minusinf p)" by simp
  4587   moreover have "z - 1 < z" by simp
  4588   ultimately show ?thesis using z_def by auto
  4589 qed
  4590 
  4591 lemma rplusinf_ex:
  4592   assumes lp: "isrlfm p"
  4593   and ex: "Ifm (a#bs) (plusinf p)"
  4594   shows "\<exists> x. Ifm (x#bs) p"
  4595 proof-
  4596   from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
  4597   have th: "\<forall> x. Ifm (x#bs) (plusinf p)" by auto
  4598   from rplusinf_inf[OF lp, where bs="bs"] 
  4599   obtain z where z_def: "\<forall>x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast
  4600   from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp
  4601   moreover have "z + 1 > z" by simp
  4602   ultimately show ?thesis using z_def by auto
  4603 qed
  4604 
  4605 consts 
  4606   \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list"
  4607   \<upsilon> :: "fm \<Rightarrow> (num \<times> int) \<Rightarrow> fm "
  4608 recdef \<Upsilon> "measure size"
  4609   "\<Upsilon> (And p q) = (\<Upsilon> p @ \<Upsilon> q)" 
  4610   "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)" 
  4611   "\<Upsilon> (Eq  (CN 0 c e)) = [(Neg e,c)]"
  4612   "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]"
  4613   "\<Upsilon> (Lt  (CN 0 c e)) = [(Neg e,c)]"
  4614   "\<Upsilon> (Le  (CN 0 c e)) = [(Neg e,c)]"
  4615   "\<Upsilon> (Gt  (CN 0 c e)) = [(Neg e,c)]"
  4616   "\<Upsilon> (Ge  (CN 0 c e)) = [(Neg e,c)]"
  4617   "\<Upsilon> p = []"
  4618 
  4619 recdef \<upsilon> "measure size"
  4620   "\<upsilon> (And p q) = (\<lambda> (t,n). And (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
  4621   "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
  4622   "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))"
  4623   "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))"
  4624   "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))"
  4625   "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))"
  4626   "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))"
  4627   "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))"
  4628   "\<upsilon> p = (\<lambda> (t,n). p)"
  4629 
  4630 lemma \<upsilon>_I: assumes lp: "isrlfm p"
  4631   and np: "real n > 0" and nbt: "numbound0 t"
  4632   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> _")
  4633   using lp
  4634 proof(induct p rule: \<upsilon>.induct)
  4635   case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4636   have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)"
  4637     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4638   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)"
  4639     by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4640       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4641   also have "\<dots> = (real c *?t + ?n* (?N x e) < 0)"
  4642     using np by simp 
  4643   finally show ?case using nbt nb by (simp add: ring_simps)
  4644 next
  4645   case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4646   have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<le> 0)"
  4647     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4648   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<le> 0)"
  4649     by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4650       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4651   also have "\<dots> = (real c *?t + ?n* (?N x e) \<le> 0)"
  4652     using np by simp 
  4653   finally show ?case using nbt nb by (simp add: ring_simps)
  4654 next
  4655   case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4656   have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)"
  4657     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4658   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)"
  4659     by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4660       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4661   also have "\<dots> = (real c *?t + ?n* (?N x e) > 0)"
  4662     using np by simp 
  4663   finally show ?case using nbt nb by (simp add: ring_simps)
  4664 next
  4665   case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4666   have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<ge> 0)"
  4667     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4668   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<ge> 0)"
  4669     by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4670       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4671   also have "\<dots> = (real c *?t + ?n* (?N x e) \<ge> 0)"
  4672     using np by simp 
  4673   finally show ?case using nbt nb by (simp add: ring_simps)
  4674 next
  4675   case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4676   from np have np: "real n \<noteq> 0" by simp
  4677   have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)"
  4678     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4679   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)"
  4680     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4681       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4682   also have "\<dots> = (real c *?t + ?n* (?N x e) = 0)"
  4683     using np by simp 
  4684   finally show ?case using nbt nb by (simp add: ring_simps)
  4685 next
  4686   case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+
  4687   from np have np: "real n \<noteq> 0" by simp
  4688   have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<noteq> 0)"
  4689     using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
  4690   also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)"
  4691     by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
  4692       and b="0", simplified divide_zero_left]) (simp only: ring_simps)
  4693   also have "\<dots> = (real c *?t + ?n* (?N x e) \<noteq> 0)"
  4694     using np by simp 
  4695   finally show ?case using nbt nb by (simp add: ring_simps)
  4696 qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2)
  4697 
  4698 lemma \<Upsilon>_l:
  4699   assumes lp: "isrlfm p"
  4700   shows "\<forall> (t,k) \<in> set (\<Upsilon> p). numbound0 t \<and> k >0"
  4701 using lp
  4702 by(induct p rule: \<Upsilon>.induct)  auto
  4703 
  4704 lemma rminusinf_\<Upsilon>:
  4705   assumes lp: "isrlfm p"
  4706   and nmi: "\<not> (Ifm (a#bs) (minusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
  4707   and ex: "Ifm (x#bs) p" (is "?I x p")
  4708   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")
  4709 proof-
  4710   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")
  4711     using lp nmi ex
  4712     by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2)
  4713   then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<ge> ?N a s" by blast
  4714   from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
  4715   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" 
  4716     by (auto simp add: mult_commute)
  4717   thus ?thesis using smU by auto
  4718 qed
  4719 
  4720 lemma rplusinf_\<Upsilon>:
  4721   assumes lp: "isrlfm p"
  4722   and nmi: "\<not> (Ifm (a#bs) (plusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
  4723   and ex: "Ifm (x#bs) p" (is "?I x p")
  4724   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")
  4725 proof-
  4726   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")
  4727     using lp nmi ex
  4728     by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2)
  4729   then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<le> ?N a s" by blast
  4730   from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
  4731   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" 
  4732     by (auto simp add: mult_commute)
  4733   thus ?thesis using smU by auto
  4734 qed
  4735 
  4736 lemma lin_dense: 
  4737   assumes lp: "isrlfm p"
  4738   and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (t,n). Inum (x#bs) t / real n) ` set (\<Upsilon> p)" 
  4739   (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (t,n). ?N x t / real n ) ` (?U p)")
  4740   and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p"
  4741   and ly: "l < y" and yu: "y < u"
  4742   shows "Ifm (y#bs) p"
  4743 using lp px noS
  4744 proof (induct p rule: isrlfm.induct)
  4745   case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4746     from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps)
  4747     hence pxc: "x < (- ?N x e) / real c" 
  4748       by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"])
  4749     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4750     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4751     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4752     moreover {assume y: "y < (-?N x e)/ real c"
  4753       hence "y * real c < - ?N x e"
  4754 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4755       hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
  4756       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4757     moreover {assume y: "y > (- ?N x e) / real c" 
  4758       with yu have eu: "u > (- ?N x e) / real c" by auto
  4759       with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
  4760       with lx pxc have "False" by auto
  4761       hence ?case by simp }
  4762     ultimately show ?case by blast
  4763 next
  4764   case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp +
  4765     from prems have "x * real c + ?N x e \<le> 0" by (simp add: ring_simps)
  4766     hence pxc: "x \<le> (- ?N x e) / real c" 
  4767       by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"])
  4768     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4769     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4770     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4771     moreover {assume y: "y < (-?N x e)/ real c"
  4772       hence "y * real c < - ?N x e"
  4773 	by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4774       hence "real c * y + ?N x e < 0" by (simp add: ring_simps)
  4775       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4776     moreover {assume y: "y > (- ?N x e) / real c" 
  4777       with yu have eu: "u > (- ?N x e) / real c" by auto
  4778       with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
  4779       with lx pxc have "False" by auto
  4780       hence ?case by simp }
  4781     ultimately show ?case by blast
  4782 next
  4783   case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4784     from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps)
  4785     hence pxc: "x > (- ?N x e) / real c" 
  4786       by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"])
  4787     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4788     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4789     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4790     moreover {assume y: "y > (-?N x e)/ real c"
  4791       hence "y * real c > - ?N x e"
  4792 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4793       hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
  4794       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4795     moreover {assume y: "y < (- ?N x e) / real c" 
  4796       with ly have eu: "l < (- ?N x e) / real c" by auto
  4797       with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
  4798       with xu pxc have "False" by auto
  4799       hence ?case by simp }
  4800     ultimately show ?case by blast
  4801 next
  4802   case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4803     from prems have "x * real c + ?N x e \<ge> 0" by (simp add: ring_simps)
  4804     hence pxc: "x \<ge> (- ?N x e) / real c" 
  4805       by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"])
  4806     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4807     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4808     hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
  4809     moreover {assume y: "y > (-?N x e)/ real c"
  4810       hence "y * real c > - ?N x e"
  4811 	by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
  4812       hence "real c * y + ?N x e > 0" by (simp add: ring_simps)
  4813       hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
  4814     moreover {assume y: "y < (- ?N x e) / real c" 
  4815       with ly have eu: "l < (- ?N x e) / real c" by auto
  4816       with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
  4817       with xu pxc have "False" by auto
  4818       hence ?case by simp }
  4819     ultimately show ?case by blast
  4820 next
  4821   case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4822     from cp have cnz: "real c \<noteq> 0" by simp
  4823     from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps)
  4824     hence pxc: "x = (- ?N x e) / real c" 
  4825       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"])
  4826     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4827     with lx xu have yne: "x \<noteq> - ?N x e / real c" by auto
  4828     with pxc show ?case by simp
  4829 next
  4830   case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+
  4831     from cp have cnz: "real c \<noteq> 0" by simp
  4832     from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
  4833     with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
  4834     hence "y* real c \<noteq> -?N x e"      
  4835       by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp
  4836     hence "y* real c + ?N x e \<noteq> 0" by (simp add: ring_simps)
  4837     thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] 
  4838       by (simp add: ring_simps)
  4839 qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"])
  4840 
  4841 lemma finite_set_intervals:
  4842   assumes px: "P (x::real)" 
  4843   and lx: "l \<le> x" and xu: "x \<le> u"
  4844   and linS: "l\<in> S" and uinS: "u \<in> S"
  4845   and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
  4846   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"
  4847 proof-
  4848   let ?Mx = "{y. y\<in> S \<and> y \<le> x}"
  4849   let ?xM = "{y. y\<in> S \<and> x \<le> y}"
  4850   let ?a = "Max ?Mx"
  4851   let ?b = "Min ?xM"
  4852   have MxS: "?Mx \<subseteq> S" by blast
  4853   hence fMx: "finite ?Mx" using fS finite_subset by auto
  4854   from lx linS have linMx: "l \<in> ?Mx" by blast
  4855   hence Mxne: "?Mx \<noteq> {}" by blast
  4856   have xMS: "?xM \<subseteq> S" by blast
  4857   hence fxM: "finite ?xM" using fS finite_subset by auto
  4858   from xu uinS have linxM: "u \<in> ?xM" by blast
  4859   hence xMne: "?xM \<noteq> {}" by blast
  4860   have ax:"?a \<le> x" using Mxne fMx by auto
  4861   have xb:"x \<le> ?b" using xMne fxM by auto
  4862   have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast
  4863   have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast
  4864   have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S"
  4865   proof(clarsimp)
  4866     fix y
  4867     assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S"
  4868     from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by auto
  4869     moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by simp}
  4870     moreover {assume "y \<in> ?xM" hence "y \<ge> ?b" using xMne fxM by auto with yb have "False" by simp}
  4871     ultimately show "False" by blast
  4872   qed
  4873   from ainS binS noy ax xb px show ?thesis by blast
  4874 qed
  4875 
  4876 lemma finite_set_intervals2:
  4877   assumes px: "P (x::real)" 
  4878   and lx: "l \<le> x" and xu: "x \<le> u"
  4879   and linS: "l\<in> S" and uinS: "u \<in> S"
  4880   and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u"
  4881   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)"
  4882 proof-
  4883   from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su]
  4884   obtain a and b where 
  4885     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
  4886   from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by auto
  4887   thus ?thesis using px as bs noS by blast 
  4888 qed
  4889 
  4890 lemma rinf_\<Upsilon>:
  4891   assumes lp: "isrlfm p"
  4892   and nmi: "\<not> (Ifm (x#bs) (minusinf p))" (is "\<not> (Ifm (x#bs) (?M p))")
  4893   and npi: "\<not> (Ifm (x#bs) (plusinf p))" (is "\<not> (Ifm (x#bs) (?P p))")
  4894   and ex: "\<exists> x.  Ifm (x#bs) p" (is "\<exists> x. ?I x p")
  4895   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" 
  4896 proof-
  4897   let ?N = "\<lambda> x t. Inum (x#bs) t"
  4898   let ?U = "set (\<Upsilon> p)"
  4899   from ex obtain a where pa: "?I a p" by blast
  4900   from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi
  4901   have nmi': "\<not> (?I a (?M p))" by simp
  4902   from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi
  4903   have npi': "\<not> (?I a (?P p))" by simp
  4904   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"
  4905   proof-
  4906     let ?M = "(\<lambda> (t,c). ?N a t / real c) ` ?U"
  4907     have fM: "finite ?M" by auto
  4908     from rminusinf_\<Upsilon>[OF lp nmi pa] rplusinf_\<Upsilon>[OF lp npi pa] 
  4909     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
  4910     then obtain "t" "n" "s" "m" where 
  4911       tnU: "(t,n) \<in> ?U" and smU: "(s,m) \<in> ?U" 
  4912       and xs1: "a \<le> ?N x s / real m" and tx1: "a \<ge> ?N x t / real n" by blast
  4913     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
  4914     from tnU have Mne: "?M \<noteq> {}" by auto
  4915     hence Une: "?U \<noteq> {}" by simp
  4916     let ?l = "Min ?M"
  4917     let ?u = "Max ?M"
  4918     have linM: "?l \<in> ?M" using fM Mne by simp
  4919     have uinM: "?u \<in> ?M" using fM Mne by simp
  4920     have tnM: "?N a t / real n \<in> ?M" using tnU by auto
  4921     have smM: "?N a s / real m \<in> ?M" using smU by auto 
  4922     have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto
  4923     have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
  4924     have "?l \<le> ?N a t / real n" using tnM Mne by simp hence lx: "?l \<le> a" using tx by simp
  4925     have "?N a s / real m \<le> ?u" using smM Mne by simp hence xu: "a \<le> ?u" using xs by simp
  4926     from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
  4927     have "(\<exists> s\<in> ?M. ?I s p) \<or> 
  4928       (\<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)" .
  4929     moreover { fix u assume um: "u\<in> ?M" and pu: "?I u p"
  4930       hence "\<exists> (tu,nu) \<in> ?U. u = ?N a tu / real nu" by auto
  4931       then obtain "tu" "nu" where tuU: "(tu,nu) \<in> ?U" and tuu:"u= ?N a tu / real nu" by blast
  4932       have "(u + u) / 2 = u" by auto with pu tuu 
  4933       have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp
  4934       with tuU have ?thesis by blast}
  4935     moreover{
  4936       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"
  4937       then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" 
  4938 	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"
  4939 	by blast
  4940       from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real t1n" by auto
  4941       then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast
  4942       from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real t2n" by auto
  4943       then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \<in> ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast
  4944       from t1x xt2 have t1t2: "t1 < t2" by simp
  4945       let ?u = "(t1 + t2) / 2"
  4946       from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto
  4947       from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" .
  4948       with t1uU t2uU t1u t2u have ?thesis by blast}
  4949     ultimately show ?thesis by blast
  4950   qed
  4951   then obtain "l" "n" "s"  "m" where lnU: "(l,n) \<in> ?U" and smU:"(s,m) \<in> ?U" 
  4952     and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast
  4953   from lnU smU \<Upsilon>_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto
  4954   from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] 
  4955     numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
  4956   have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp
  4957   with lnU smU
  4958   show ?thesis by auto
  4959 qed
  4960     (* The Ferrante - Rackoff Theorem *)
  4961 
  4962 theorem fr_eq: 
  4963   assumes lp: "isrlfm p"
  4964   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))"
  4965   (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
  4966 proof
  4967   assume px: "\<exists> x. ?I x p"
  4968   have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
  4969   moreover {assume "?M \<or> ?P" hence "?D" by blast}
  4970   moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
  4971     from rinf_\<Upsilon>[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast}
  4972   ultimately show "?D" by blast
  4973 next
  4974   assume "?D" 
  4975   moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
  4976   moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
  4977   moreover {assume f:"?F" hence "?E" by blast}
  4978   ultimately show "?E" by blast
  4979 qed
  4980 
  4981 
  4982 lemma fr_eq\<upsilon>: 
  4983   assumes lp: "isrlfm p"
  4984   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))))"
  4985   (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
  4986 proof
  4987   assume px: "\<exists> x. ?I x p"
  4988   have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
  4989   moreover {assume "?M \<or> ?P" hence "?D" by blast}
  4990   moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
  4991     let ?f ="\<lambda> (t,n). Inum (x#bs) t / real n"
  4992     let ?N = "\<lambda> t. Inum (x#bs) t"
  4993     {fix t n s m assume "(t,n)\<in> set (\<Upsilon> p)" and "(s,m) \<in> set (\<Upsilon> p)"
  4994       with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0"
  4995 	by auto
  4996       let ?st = "Add (Mul m t) (Mul n s)"
  4997       from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  4998 	by (simp add: mult_commute)
  4999       from tnb snb have st_nb: "numbound0 ?st" by simp
  5000       have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5001 	using mnp mp np by (simp add: ring_simps add_divide_distrib)
  5002       from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"] 
  5003       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])}
  5004     with rinf_\<Upsilon>[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
  5005   ultimately show "?D" by blast
  5006 next
  5007   assume "?D" 
  5008   moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
  5009   moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
  5010   moreover {fix t k s l assume "(t,k) \<in> set (\<Upsilon> p)" and "(s,l) \<in> set (\<Upsilon> p)" 
  5011     and px:"?I x (\<upsilon> p (Add (Mul l t) (Mul k s), 2*k*l))"
  5012     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
  5013     let ?st = "Add (Mul l t) (Mul k s)"
  5014     from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" 
  5015       by (simp add: mult_commute)
  5016     from tnb snb have st_nb: "numbound0 ?st" by simp
  5017     from \<upsilon>_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto}
  5018   ultimately show "?E" by blast
  5019 qed
  5020 
  5021 text{* The overall Part *}
  5022 
  5023 lemma real_ex_int_real01:
  5024   shows "(\<exists> (x::real). P x) = (\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))"
  5025 proof(auto)
  5026   fix x
  5027   assume Px: "P x"
  5028   let ?i = "floor x"
  5029   let ?u = "x - real ?i"
  5030   have "x = real ?i + ?u" by simp
  5031   hence "P (real ?i + ?u)" using Px by simp
  5032   moreover have "real ?i \<le> x" using real_of_int_floor_le by simp hence "0 \<le> ?u" by arith
  5033   moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith 
  5034   ultimately show "(\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))" by blast
  5035 qed
  5036 
  5037 consts exsplitnum :: "num \<Rightarrow> num"
  5038   exsplit :: "fm \<Rightarrow> fm"
  5039 recdef exsplitnum "measure size"
  5040   "exsplitnum (C c) = (C c)"
  5041   "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)"
  5042   "exsplitnum (Bound n) = Bound (n+1)"
  5043   "exsplitnum (Neg a) = Neg (exsplitnum a)"
  5044   "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) "
  5045   "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) "
  5046   "exsplitnum (Mul c a) = Mul c (exsplitnum a)"
  5047   "exsplitnum (Floor a) = Floor (exsplitnum a)"
  5048   "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))"
  5049   "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)"
  5050   "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)"
  5051 
  5052 recdef exsplit "measure size"
  5053   "exsplit (Lt a) = Lt (exsplitnum a)"
  5054   "exsplit (Le a) = Le (exsplitnum a)"
  5055   "exsplit (Gt a) = Gt (exsplitnum a)"
  5056   "exsplit (Ge a) = Ge (exsplitnum a)"
  5057   "exsplit (Eq a) = Eq (exsplitnum a)"
  5058   "exsplit (NEq a) = NEq (exsplitnum a)"
  5059   "exsplit (Dvd i a) = Dvd i (exsplitnum a)"
  5060   "exsplit (NDvd i a) = NDvd i (exsplitnum a)"
  5061   "exsplit (And p q) = And (exsplit p) (exsplit q)"
  5062   "exsplit (Or p q) = Or (exsplit p) (exsplit q)"
  5063   "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)"
  5064   "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)"
  5065   "exsplit (NOT p) = NOT (exsplit p)"
  5066   "exsplit p = p"
  5067 
  5068 lemma exsplitnum: 
  5069   "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t"
  5070   by(induct t rule: exsplitnum.induct) (simp_all add: ring_simps)
  5071 
  5072 lemma exsplit: 
  5073   assumes qfp: "qfree p"
  5074   shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p"
  5075 using qfp exsplitnum[where x="x" and y="y" and bs="bs"]
  5076 by(induct p rule: exsplit.induct) simp_all
  5077 
  5078 lemma splitex:
  5079   assumes qf: "qfree p"
  5080   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")
  5081 proof-
  5082   have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))"
  5083     by (simp add: myless[rule_format, where b="1"] myless[rule_format, where b="0"] add_ac diff_def)
  5084   also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)"
  5085     by (simp only: exsplit[OF qf] add_ac)
  5086   also have "\<dots> = (\<exists> x. Ifm (x#bs) p)" 
  5087     by (simp only: real_ex_int_real01[where P="\<lambda> x. Ifm (x#bs) p"])
  5088   finally show ?thesis by simp
  5089 qed
  5090 
  5091     (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *)
  5092 
  5093 constdefs ferrack01:: "fm \<Rightarrow> fm"
  5094   "ferrack01 p \<equiv> (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p);
  5095                     U = remdups(map simp_num_pair 
  5096                      (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m))
  5097                            (alluopairs (\<Upsilon> p')))) 
  5098   in decr (evaldjf (\<upsilon> p') U ))"
  5099 
  5100 lemma fr_eq_01: 
  5101   assumes qf: "qfree p"
  5102   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)))"
  5103   (is "(\<exists> x. ?I x ?q) = ?F")
  5104 proof-
  5105   let ?rq = "rlfm ?q"
  5106   let ?M = "?I x (minusinf ?rq)"
  5107   let ?P = "?I x (plusinf ?rq)"
  5108   have MF: "?M = False"
  5109     apply (simp add: Let_def reducecoeff_def numgcd_def igcd_def rsplit_def ge_def lt_def conj_def disj_def)
  5110     by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
  5111   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)
  5112     by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
  5113   have "(\<exists> x. ?I x ?q ) = 
  5114     ((?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))))"
  5115     (is "(\<exists> x. ?I x ?q) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
  5116   proof
  5117     assume "\<exists> x. ?I x ?q"  
  5118     then obtain x where qx: "?I x ?q" by blast
  5119     hence xp: "0\<le> x" and x1: "x< 1" and px: "?I x p" 
  5120       by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf])
  5121     from qx have "?I x ?rq " 
  5122       by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
  5123     hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto
  5124     from qf have qfq:"isrlfm ?rq"  
  5125       by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
  5126     with lqx fr_eq\<upsilon>[OF qfq] show "?M \<or> ?P \<or> ?F" by blast
  5127   next
  5128     assume D: "?D"
  5129     let ?U = "set (\<Upsilon> ?rq )"
  5130     from MF PF D have "?F" by auto
  5131     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
  5132     from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] 
  5133       by (auto simp add: rsplit_def lt_def ge_def)
  5134     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)
  5135     let ?st = "Add (Mul m t) (Mul n s)"
  5136     from tnb snb have stnb: "numbound0 ?st" by simp
  5137     from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  5138       by (simp add: mult_commute)
  5139     from conjunct1[OF \<upsilon>_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx
  5140     have "\<exists> x. ?I x ?rq" by auto
  5141     thus "?E" 
  5142       using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def)
  5143   qed
  5144   with MF PF show ?thesis by blast
  5145 qed
  5146 
  5147 lemma \<Upsilon>_cong_aux:
  5148   assumes Ul: "\<forall> (t,n) \<in> set U. numbound0 t \<and> n >0"
  5149   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))"
  5150   (is "?lhs = ?rhs")
  5151 proof(auto)
  5152   fix t n s m
  5153   assume "((t,n),(s,m)) \<in> set (alluopairs U)"
  5154   hence th: "((t,n),(s,m)) \<in> (set U \<times> set U)"
  5155     using alluopairs_set1[where xs="U"] by blast
  5156   let ?N = "\<lambda> t. Inum (x#bs) t"
  5157   let ?st= "Add (Mul m t) (Mul n s)"
  5158   from Ul th have mnz: "m \<noteq> 0" by auto
  5159   from Ul th have  nnz: "n \<noteq> 0" by auto  
  5160   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5161    using mnz nnz by (simp add: ring_simps add_divide_distrib)
  5162  
  5163   thus "(real m *  Inum (x # bs) t + real n * Inum (x # bs) s) /
  5164        (2 * real n * real m)
  5165        \<in> (\<lambda>((t, n), s, m).
  5166              (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
  5167          (set U \<times> set U)"using mnz nnz th  
  5168     apply (auto simp add: th add_divide_distrib ring_simps split_def image_def)
  5169     by (rule_tac x="(s,m)" in bexI,simp_all) 
  5170   (rule_tac x="(t,n)" in bexI,simp_all)
  5171 next
  5172   fix t n s m
  5173   assume tnU: "(t,n) \<in> set U" and smU:"(s,m) \<in> set U" 
  5174   let ?N = "\<lambda> t. Inum (x#bs) t"
  5175   let ?st= "Add (Mul m t) (Mul n s)"
  5176   from Ul smU have mnz: "m \<noteq> 0" by auto
  5177   from Ul tnU have  nnz: "n \<noteq> 0" by auto  
  5178   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5179    using mnz nnz by (simp add: ring_simps add_divide_distrib)
  5180  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"
  5181  have Pc:"\<forall> a b. ?P a b = ?P b a"
  5182    by auto
  5183  from Ul alluopairs_set1 have Up:"\<forall> ((t,n),(s,m)) \<in> set (alluopairs U). n \<noteq> 0 \<and> m \<noteq> 0" by blast
  5184  from alluopairs_ex[OF Pc, where xs="U"] tnU smU
  5185  have th':"\<exists> ((t',n'),(s',m')) \<in> set (alluopairs U). ?P (t',n') (s',m')"
  5186    by blast
  5187  then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \<in> set (alluopairs U)" 
  5188    and Pts': "?P (t',n') (s',m')" by blast
  5189  from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0" by auto
  5190  let ?st' = "Add (Mul m' t') (Mul n' s')"
  5191    have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')"
  5192    using mnz' nnz' by (simp add: ring_simps add_divide_distrib)
  5193  from Pts' have 
  5194    "(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
  5195  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')
  5196  finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2
  5197           \<in> (\<lambda>(t, n). Inum (x # bs) t / real n) `
  5198             (\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) `
  5199             set (alluopairs U)"
  5200    using ts'_U by blast
  5201 qed
  5202 
  5203 lemma \<Upsilon>_cong:
  5204   assumes lp: "isrlfm p"
  5205   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)")
  5206   and U: "\<forall> (t,n) \<in> U. numbound0 t \<and> n > 0"
  5207   and U': "\<forall> (t,n) \<in> U'. numbound0 t \<and> n > 0"
  5208   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)))"
  5209   (is "?lhs = ?rhs")
  5210 proof
  5211   assume ?lhs
  5212   then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and 
  5213     Pst: "Ifm (x#bs) (\<upsilon> p (Add (Mul m t) (Mul n s),2*n*m))" by blast
  5214   let ?N = "\<lambda> t. Inum (x#bs) t"
  5215   from tnU smU U have tnb: "numbound0 t" and np: "n > 0" 
  5216     and snb: "numbound0 s" and mp:"m > 0"  by auto
  5217   let ?st= "Add (Mul m t) (Mul n s)"
  5218   from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  5219       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
  5220     from tnb snb have stnb: "numbound0 ?st" by simp
  5221   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5222    using mp np by (simp add: ring_simps add_divide_distrib)
  5223   from tnU smU UU' have "?g ((t,n),(s,m)) \<in> ?f ` U'" by blast
  5224   hence "\<exists> (t',n') \<in> U'. ?g ((t,n),(s,m)) = ?f (t',n')"
  5225     by auto (rule_tac x="(a,b)" in bexI, auto)
  5226   then obtain t' n' where tnU': "(t',n') \<in> U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast
  5227   from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
  5228   from \<upsilon>_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst 
  5229   have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
  5230   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]]
  5231   have "Ifm (x # bs) (\<upsilon> p (t', n')) " by (simp only: st) 
  5232   then show ?rhs using tnU' by auto 
  5233 next
  5234   assume ?rhs
  5235   then obtain t' n' where tnU': "(t',n') \<in> U'" and Pt': "Ifm (x # bs) (\<upsilon> p (t', n'))" 
  5236     by blast
  5237   from tnU' UU' have "?f (t',n') \<in> ?g ` (U\<times>U)" by blast
  5238   hence "\<exists> ((t,n),(s,m)) \<in> (U\<times>U). ?f (t',n') = ?g ((t,n),(s,m))" 
  5239     by auto (rule_tac x="(a,b)" in bexI, auto)
  5240   then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and 
  5241     th: "?f (t',n') = ?g((t,n),(s,m)) "by blast
  5242     let ?N = "\<lambda> t. Inum (x#bs) t"
  5243   from tnU smU U have tnb: "numbound0 t" and np: "n > 0" 
  5244     and snb: "numbound0 s" and mp:"m > 0"  by auto
  5245   let ?st= "Add (Mul m t) (Mul n s)"
  5246   from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" 
  5247       by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
  5248     from tnb snb have stnb: "numbound0 ?st" by simp
  5249   have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
  5250    using mp np by (simp add: ring_simps add_divide_distrib)
  5251   from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
  5252   from \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt'
  5253   have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
  5254   with \<upsilon>_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast
  5255 qed
  5256   
  5257 lemma ferrack01: 
  5258   assumes qf: "qfree p"
  5259   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> _")
  5260 proof-
  5261   let ?I = "\<lambda> x p. Ifm (x#bs) p"
  5262   let ?N = "\<lambda> t. Inum (x#bs) t"
  5263   let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)"
  5264   let ?U = "\<Upsilon> ?q"
  5265   let ?Up = "alluopairs ?U"
  5266   let ?g = "\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)"
  5267   let ?S = "map ?g ?Up"
  5268   let ?SS = "map simp_num_pair ?S"
  5269   let ?Y = "remdups ?SS"
  5270   let ?f= "(\<lambda> (t,n). ?N t / real n)"
  5271   let ?h = "\<lambda> ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2"
  5272   let ?F = "\<lambda> p. \<exists> a \<in> set (\<Upsilon> p). \<exists> b \<in> set (\<Upsilon> p). ?I x (\<upsilon> p (?g(a,b)))"
  5273   let ?ep = "evaldjf (\<upsilon> ?q) ?Y"
  5274   from rlfm_l[OF qf] have lq: "isrlfm ?q" 
  5275     by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def igcd_def)
  5276   from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \<le> (set ?U \<times> set ?U)" by simp
  5277   from \<Upsilon>_l[OF lq] have U_l: "\<forall> (t,n) \<in> set ?U. numbound0 t \<and> n > 0" .
  5278   from U_l UpU 
  5279   have Up_: "\<forall> ((t,n),(s,m)) \<in> set ?Up. numbound0 t \<and> n> 0 \<and> numbound0 s \<and> m > 0" by auto
  5280   hence Snb: "\<forall> (t,n) \<in> set ?S. numbound0 t \<and> n > 0 "
  5281     by (auto simp add: mult_pos_pos)
  5282   have Y_l: "\<forall> (t,n) \<in> set ?Y. numbound0 t \<and> n > 0" 
  5283   proof-
  5284     { fix t n assume tnY: "(t,n) \<in> set ?Y" 
  5285       hence "(t,n) \<in> set ?SS" by simp
  5286       hence "\<exists> (t',n') \<in> set ?S. simp_num_pair (t',n') = (t,n)"
  5287 	by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
  5288       then obtain t' n' where tn'S: "(t',n') \<in> set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast
  5289       from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto
  5290       from simp_num_pair_l[OF tnb np tns]
  5291       have "numbound0 t \<and> n > 0" . }
  5292     thus ?thesis by blast
  5293   qed
  5294 
  5295   have YU: "(?f ` set ?Y) = (?h ` (set ?U \<times> set ?U))"
  5296   proof-
  5297      from simp_num_pair_ci[where bs="x#bs"] have 
  5298     "\<forall>x. (?f o simp_num_pair) x = ?f x" by auto
  5299      hence th: "?f o simp_num_pair = ?f" using ext by blast
  5300     have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose)
  5301     also have "\<dots> = (?f ` set ?S)" by (simp add: th)
  5302     also have "\<dots> = ((?f o ?g) ` set ?Up)" 
  5303       by (simp only: set_map o_def image_compose[symmetric])
  5304     also have "\<dots> = (?h ` (set ?U \<times> set ?U))"
  5305       using \<Upsilon>_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast
  5306     finally show ?thesis .
  5307   qed
  5308   have "\<forall> (t,n) \<in> set ?Y. bound0 (\<upsilon> ?q (t,n))"
  5309   proof-
  5310     { fix t n assume tnY: "(t,n) \<in> set ?Y"
  5311       with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto
  5312       from \<upsilon>_I[OF lq np tnb]
  5313     have "bound0 (\<upsilon> ?q (t,n))"  by simp}
  5314     thus ?thesis by blast
  5315   qed
  5316   hence ep_nb: "bound0 ?ep"  using evaldjf_bound0[where xs="?Y" and f="\<upsilon> ?q"]
  5317     by auto
  5318 
  5319   from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q"
  5320     by (simp only: split_def fst_conv snd_conv)
  5321   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]
  5322     by (simp only: split_def fst_conv snd_conv) 
  5323   also have "\<dots> = (Ifm (x#bs) ?ep)" 
  5324     using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\<upsilon> ?q",symmetric]
  5325     by (simp only: split_def pair_collapse)
  5326   also have "\<dots> = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast
  5327   finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def)
  5328   from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def)
  5329   with lr show ?thesis by blast
  5330 qed
  5331 
  5332 lemma cp_thm': 
  5333   assumes lp: "iszlfm p (real (i::int)#bs)"
  5334   and up: "d\<beta> p 1" and dd: "d\<delta> p d" and dp: "d > 0"
  5335   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))"
  5336   using cp_thm[OF lp up dd dp] by auto
  5337 
  5338 constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int"
  5339   "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;
  5340              B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
  5341              in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
  5342 
  5343 lemma unit: assumes qf: "qfree p"
  5344   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)"
  5345 proof-
  5346   fix q B d 
  5347   assume qBd: "unit p = (q,B,d)"
  5348   let ?thes = "((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and>
  5349     Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\<beta> q) \<and>
  5350     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)"
  5351   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5352   let ?p' = "zlfm p"
  5353   let ?l = "\<zeta> ?p'"
  5354   let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a\<beta> ?p' ?l)"
  5355   let ?d = "\<delta> ?q"
  5356   let ?B = "set (\<beta> ?q)"
  5357   let ?B'= "remdups (map simpnum (\<beta> ?q))"
  5358   let ?A = "set (\<alpha> ?q)"
  5359   let ?A'= "remdups (map simpnum (\<alpha> ?q))"
  5360   from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] 
  5361   have pp': "\<forall> i. ?I i ?p' = ?I i p" by auto
  5362   from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]]
  5363   have lp': "\<forall> (i::int). iszlfm ?p' (real i#bs)" by simp 
  5364   hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp
  5365   from lp' \<zeta>[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d\<beta> ?p' ?l" by auto
  5366   from a\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp'
  5367   have pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by (simp add: int_rdvd_iff) 
  5368   from lp'' lp a\<beta>[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d\<beta> ?q 1" 
  5369     by (auto simp add: isint_def)
  5370   from \<delta>[OF lq] have dp:"?d >0" and dd: "d\<delta> ?q ?d" by blast+
  5371   let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
  5372   have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_compose) 
  5373   also have "\<dots> = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto
  5374   finally have BB': "?N ` set ?B' = ?N ` ?B" .
  5375   have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_compose) 
  5376   also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto
  5377   finally have AA': "?N ` set ?A' = ?N ` ?A" .
  5378   from \<beta>_numbound0[OF lq] have B_nb:"\<forall> b\<in> set ?B'. numbound0 b"
  5379     by (simp add: simpnum_numbound0)
  5380   from \<alpha>_l[OF lq] have A_nb: "\<forall> b\<in> set ?A'. numbound0 b"
  5381     by (simp add: simpnum_numbound0)
  5382     {assume "length ?B' \<le> length ?A'"
  5383     hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
  5384       using qBd by (auto simp add: Let_def unit_def)
  5385     with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<beta> q)" 
  5386       and bn: "\<forall>b\<in> set B. numbound0 b" by simp+ 
  5387   with pq_ex dp uq dd lq q d have ?thes by simp}
  5388   moreover 
  5389   {assume "\<not> (length ?B' \<le> length ?A')"
  5390     hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
  5391       using qBd by (auto simp add: Let_def unit_def)
  5392     with AA' mirror\<alpha>\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)" 
  5393       and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
  5394     from mirror_ex[OF lq] pq_ex q 
  5395     have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
  5396     from lq uq q mirror_d\<beta> [where p="?q" and bs="bs" and a="real i"]
  5397     have lq': "iszlfm q (real i#bs)" and uq: "d\<beta> q 1" by auto
  5398     from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq:"d\<delta> q d " by auto
  5399     from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp
  5400   }
  5401   ultimately show ?thes by blast
  5402 qed
  5403     (* Cooper's Algorithm *)
  5404 
  5405 constdefs cooper :: "fm \<Rightarrow> fm"
  5406   "cooper p \<equiv> 
  5407   (let (q,B,d) = unit p; js = iupt (1,d);
  5408        mq = simpfm (minusinf q);
  5409        md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) js
  5410    in if md = T then T else
  5411     (let qd = evaldjf (\<lambda> t. simpfm (subst0 t q)) 
  5412                                (remdups (map (\<lambda> (b,j). simpnum (Add b (C j))) 
  5413                                             [(b,j). b\<leftarrow>B,j\<leftarrow>js]))
  5414      in decr (disj md qd)))"
  5415 lemma cooper: assumes qf: "qfree p"
  5416   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \<and> qfree (cooper p)" 
  5417   (is "(?lhs = ?rhs) \<and> _")
  5418 proof-
  5419 
  5420   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5421   let ?q = "fst (unit p)"
  5422   let ?B = "fst (snd(unit p))"
  5423   let ?d = "snd (snd (unit p))"
  5424   let ?js = "iupt (1,?d)"
  5425   let ?mq = "minusinf ?q"
  5426   let ?smq = "simpfm ?mq"
  5427   let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
  5428   let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
  5429   let ?bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]"
  5430   let ?sbjs = "map (\<lambda> (b,j). simpnum (Add b (C j))) ?bjs"
  5431   let ?qd = "evaldjf (\<lambda> t. simpfm (subst0 t ?q)) (remdups ?sbjs)"
  5432   have qbf:"unit p = (?q,?B,?d)" by simp
  5433   from unit[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and 
  5434     B:"?N ` set ?B = ?N ` set (\<beta> ?q)" and 
  5435     uq:"d\<beta> ?q 1" and dd: "d\<delta> ?q ?d" and dp: "?d > 0" and 
  5436     lq: "iszlfm ?q (real i#bs)" and 
  5437     Bn: "\<forall> b\<in> set ?B. numbound0 b" by auto
  5438   from zlin_qfree[OF lq] have qfq: "qfree ?q" .
  5439   from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
  5440   have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
  5441   hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
  5442     by (auto simp only: subst0_bound0[OF qfmq])
  5443   hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
  5444     by (auto simp add: simpfm_bound0)
  5445   from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
  5446   from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))"
  5447     by simp
  5448   hence "\<forall> (b,j) \<in> set ?bjs. numbound0 (simpnum (Add b (C j)))"
  5449     using simpnum_numbound0 by blast
  5450   hence "\<forall> t \<in> set ?sbjs. numbound0 t" by simp
  5451   hence "\<forall> t \<in> set (remdups ?sbjs). bound0 (subst0 t ?q)"
  5452     using subst0_bound0[OF qfq] by auto 
  5453   hence th': "\<forall> t \<in> set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))"
  5454     using simpfm_bound0 by blast
  5455   from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp
  5456   from mdb qdb 
  5457   have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
  5458   from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B
  5459   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
  5460   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
  5461   also have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> (\<lambda> (b,j). ?N (Add b (C j))) ` set ?bjs. Ifm (t #bs) ?q))" by simp
  5462   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 ?bjs. Ifm (t #bs) ?q))" by (simp only: simpnum_ci)
  5463   also  have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> set ?sbjs. Ifm (?N t #bs) ?q))" 
  5464     by (auto simp add: split_def) 
  5465   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)
  5466   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)
  5467   finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by (simp add: disj) 
  5468   hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp
  5469   {assume mdT: "?md = T"
  5470     hence cT:"cooper p = T" 
  5471       by (simp only: cooper_def unit_def split_def Let_def if_True) simp
  5472     from mdT mdqd have lhs:"?lhs" by (auto simp add: disj)
  5473     from mdT have "?rhs" by (simp add: cooper_def unit_def split_def)
  5474     with lhs cT have ?thesis by simp }
  5475   moreover
  5476   {assume mdT: "?md \<noteq> T" hence "cooper p = decr (disj ?md ?qd)" 
  5477       by (simp only: cooper_def unit_def split_def Let_def if_False) 
  5478     with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
  5479   ultimately show ?thesis by blast
  5480 qed
  5481 
  5482 lemma DJcooper: 
  5483   assumes qf: "qfree p"
  5484   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \<and> qfree (DJ cooper p)"
  5485 proof-
  5486   from cooper have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (cooper p)" by  blast
  5487   from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast
  5488   have "Ifm bs (DJ cooper p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (cooper q))" 
  5489      by (simp add: DJ_def evaldjf_ex)
  5490   also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs)  q)" 
  5491     using cooper disjuncts_qf[OF qf] by blast
  5492   also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
  5493   finally show ?thesis using thqf by blast
  5494 qed
  5495 
  5496     (* Redy and Loveland *)
  5497 
  5498 lemma \<sigma>\<rho>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
  5499   shows "Ifm (a#bs) (\<sigma>\<rho> p (t,c)) = Ifm (a#bs) (\<sigma>\<rho> p (t',c))"
  5500   using lp 
  5501   by (induct p rule: iszlfm.induct, auto simp add: tt')
  5502 
  5503 lemma \<sigma>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
  5504   shows "Ifm (a#bs) (\<sigma> p c t) = Ifm (a#bs) (\<sigma> p c t')"
  5505   by (simp add: \<sigma>_def tt' \<sigma>\<rho>_cong[OF lp tt'])
  5506 
  5507 lemma \<rho>_cong: assumes lp: "iszlfm p (a#bs)" 
  5508   and RR: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R =  (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
  5509   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))))"
  5510   (is "?lhs = ?rhs")
  5511 proof
  5512   let ?d = "\<delta> p"
  5513   assume ?lhs then obtain e c j where ecR: "(e,c) \<in> R" and jD:"j \<in> {1 .. c*?d}" 
  5514     and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
  5515   from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" by auto
  5516   hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" using RR by simp
  5517   hence "\<exists> (e',c') \<in> set (\<rho> p). Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
  5518   then obtain e' c' where ecRo:"(e',c') \<in> set (\<rho> p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
  5519     and cc':"c = c'" by blast
  5520   from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
  5521   
  5522   from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
  5523   from ecRo jD px' cc'  show ?rhs apply auto 
  5524     by (rule_tac x="(e', c')" in bexI,simp_all)
  5525   (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
  5526 next
  5527   let ?d = "\<delta> p"
  5528   assume ?rhs then obtain e c j where ecR: "(e,c) \<in> set (\<rho> p)" and jD:"j \<in> {1 .. c*?d}" 
  5529     and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
  5530   from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" by auto
  5531   hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp
  5532   hence "\<exists> (e',c') \<in> R. Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
  5533   then obtain e' c' where ecRo:"(e',c') \<in> R" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
  5534     and cc':"c = c'" by blast
  5535   from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
  5536   from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
  5537   from ecRo jD px' cc'  show ?lhs apply auto 
  5538     by (rule_tac x="(e', c')" in bexI,simp_all)
  5539   (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
  5540 qed
  5541 
  5542 lemma rl_thm': 
  5543   assumes lp: "iszlfm p (real (i::int)#bs)" 
  5544   and R: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R =  (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
  5545   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)))))"
  5546   using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp 
  5547 
  5548 constdefs chooset:: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int"
  5549   "chooset p \<equiv> (let q = zlfm p ; d = \<delta> q;
  5550              B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ; 
  5551              a = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<alpha>\<rho> q))
  5552              in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
  5553 
  5554 lemma chooset: assumes qf: "qfree p"
  5555   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)"
  5556 proof-
  5557   fix q B d 
  5558   assume qBd: "chooset p = (q,B,d)"
  5559   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)" 
  5560   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5561   let ?q = "zlfm p"
  5562   let ?d = "\<delta> ?q"
  5563   let ?B = "set (\<rho> ?q)"
  5564   let ?f = "\<lambda> (t,k). (simpnum t,k)"
  5565   let ?B'= "remdups (map ?f (\<rho> ?q))"
  5566   let ?A = "set (\<alpha>\<rho> ?q)"
  5567   let ?A'= "remdups (map ?f (\<alpha>\<rho> ?q))"
  5568   from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] 
  5569   have pp': "\<forall> i. ?I i ?q = ?I i p" by auto
  5570   hence pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by simp 
  5571   from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"]
  5572   have lq: "iszlfm ?q (real (i::int)#bs)" . 
  5573   from \<delta>[OF lq] have dp:"?d >0" by blast
  5574   let ?N = "\<lambda> (t,c). (Inum (real (i::int)#bs) t,c)"
  5575   have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_compose)
  5576   also have "\<dots> = ?N ` ?B"
  5577     by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def)
  5578   finally have BB': "?N ` set ?B' = ?N ` ?B" .
  5579   have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_compose) 
  5580   also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"]
  5581     by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) 
  5582   finally have AA': "?N ` set ?A' = ?N ` ?A" .
  5583   from \<rho>_l[OF lq] have B_nb:"\<forall> (e,c)\<in> set ?B'. numbound0 e \<and> c > 0"
  5584     by (simp add: simpnum_numbound0 split_def)
  5585   from \<alpha>\<rho>_l[OF lq] have A_nb: "\<forall> (e,c)\<in> set ?A'. numbound0 e \<and> c > 0"
  5586     by (simp add: simpnum_numbound0 split_def)
  5587     {assume "length ?B' \<le> length ?A'"
  5588     hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
  5589       using qBd by (auto simp add: Let_def chooset_def)
  5590     with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<rho> q)" 
  5591       and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto
  5592   with pq_ex dp lq q d have ?thes by simp}
  5593   moreover 
  5594   {assume "\<not> (length ?B' \<le> length ?A')"
  5595     hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
  5596       using qBd by (auto simp add: Let_def chooset_def)
  5597     with AA' mirror_\<alpha>\<rho>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<rho> q)" 
  5598       and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto 
  5599     from mirror_ex[OF lq] pq_ex q 
  5600     have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
  5601     from lq q mirror_l [where p="?q" and bs="bs" and a="real i"]
  5602     have lq': "iszlfm q (real i#bs)" by auto
  5603     from mirror_\<delta>[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp
  5604   }
  5605   ultimately show ?thes by blast
  5606 qed
  5607 
  5608 constdefs stage:: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm"
  5609   "stage p d \<equiv> (\<lambda> (e,c). evaldjf (\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))) (iupt (1,c*d)))"
  5610 lemma stage:
  5611   shows "Ifm bs (stage p d (e,c)) = (\<exists> j\<in>{1 .. c*d}. Ifm bs (\<sigma> p c (Add e (C j))))"
  5612   by (unfold stage_def split_def ,simp only: evaldjf_ex iupt_set simpfm) simp
  5613 
  5614 lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e"
  5615   shows "bound0 (stage p d (e,c))"
  5616 proof-
  5617   let ?f = "\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))"
  5618   have th: "\<forall> j\<in> set (iupt(1,c*d)). bound0 (?f j)"
  5619   proof
  5620     fix j
  5621     from nb have nb':"numbound0 (Add e (C j))" by simp
  5622     from simpfm_bound0[OF \<sigma>_nb[OF lp nb', where k="c"]]
  5623     show "bound0 (simpfm (\<sigma> p c (Add e (C j))))" .
  5624   qed
  5625   from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp
  5626 qed
  5627 
  5628 constdefs redlove:: "fm \<Rightarrow> fm"
  5629   "redlove p \<equiv> 
  5630   (let (q,B,d) = chooset p;
  5631        mq = simpfm (minusinf q);
  5632        md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) (iupt (1,d))
  5633    in if md = T then T else
  5634     (let qd = evaldjf (stage q d) B
  5635      in decr (disj md qd)))"
  5636 
  5637 lemma redlove: assumes qf: "qfree p"
  5638   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \<and> qfree (redlove p)" 
  5639   (is "(?lhs = ?rhs) \<and> _")
  5640 proof-
  5641 
  5642   let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
  5643   let ?q = "fst (chooset p)"
  5644   let ?B = "fst (snd(chooset p))"
  5645   let ?d = "snd (snd (chooset p))"
  5646   let ?js = "iupt (1,?d)"
  5647   let ?mq = "minusinf ?q"
  5648   let ?smq = "simpfm ?mq"
  5649   let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
  5650   let ?N = "\<lambda> (t,k). (Inum (real (i::int)#bs) t,k)"
  5651   let ?qd = "evaldjf (stage ?q ?d) ?B"
  5652   have qbf:"chooset p = (?q,?B,?d)" by simp
  5653   from chooset[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and 
  5654     B:"?N ` set ?B = ?N ` set (\<rho> ?q)" and dd: "\<delta> ?q = ?d" and dp: "?d > 0" and 
  5655     lq: "iszlfm ?q (real i#bs)" and 
  5656     Bn: "\<forall> (e,c)\<in> set ?B. numbound0 e \<and> c > 0" by auto
  5657   from zlin_qfree[OF lq] have qfq: "qfree ?q" .
  5658   from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
  5659   have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
  5660   hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
  5661     by (auto simp only: subst0_bound0[OF qfmq])
  5662   hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
  5663     by (auto simp add: simpfm_bound0)
  5664   from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
  5665   from Bn stage_nb[OF lq] have th:"\<forall> x \<in> set ?B. bound0 (stage ?q ?d x)" by auto
  5666   from evaldjf_bound0[OF th]  have qdb: "bound0 ?qd" .
  5667   from mdb qdb 
  5668   have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
  5669   from trans [OF pq_ex rl_thm'[OF lq B]] dd
  5670   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
  5671   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) )))" 
  5672     by (simp add: simpfm stage split_def)
  5673   also have "\<dots> = ((\<exists> j\<in> {1 .. ?d}. ?I i (subst0 (C j) ?smq))  \<or> ?I i ?qd)"
  5674     by (simp add: evaldjf_ex subst0_I[OF qfmq])
  5675   finally have mdqd:"?lhs = (?I i ?md \<or> ?I i ?qd)" by (simp only: evaldjf_ex iupt_set simpfm) 
  5676   also have "\<dots> = (?I i (disj ?md ?qd))" by (simp add: disj)
  5677   also have "\<dots> = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) 
  5678   finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . 
  5679   {assume mdT: "?md = T"
  5680     hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def)
  5681     from mdT have lhs:"?lhs" using mdqd by simp 
  5682     from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def)
  5683     with lhs cT have ?thesis by simp }
  5684   moreover
  5685   {assume mdT: "?md \<noteq> T" hence "redlove p = decr (disj ?md ?qd)" 
  5686       by (simp add: redlove_def chooset_def split_def Let_def)
  5687     with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
  5688   ultimately show ?thesis by blast
  5689 qed
  5690 
  5691 lemma DJredlove: 
  5692   assumes qf: "qfree p"
  5693   shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \<and> qfree (DJ redlove p)"
  5694 proof-
  5695   from redlove have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (redlove p)" by  blast
  5696   from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast
  5697   have "Ifm bs (DJ redlove p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (redlove q))" 
  5698      by (simp add: DJ_def evaldjf_ex)
  5699   also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs)  q)" 
  5700     using redlove disjuncts_qf[OF qf] by blast
  5701   also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
  5702   finally show ?thesis using thqf by blast
  5703 qed
  5704 
  5705 
  5706 lemma exsplit_qf: assumes qf: "qfree p"
  5707   shows "qfree (exsplit p)"
  5708 using qf by (induct p rule: exsplit.induct, auto)
  5709 
  5710 constdefs mircfr :: "fm \<Rightarrow> fm"
  5711 "mircfr \<equiv> (DJ cooper) o ferrack01 o simpfm o exsplit"
  5712 
  5713 constdefs mirlfr :: "fm \<Rightarrow> fm"
  5714 "mirlfr \<equiv> (DJ redlove) o ferrack01 o simpfm o exsplit"
  5715 
  5716 
  5717 lemma mircfr: "\<forall> bs p. qfree p \<longrightarrow> qfree (mircfr p) \<and> Ifm bs (mircfr p) = Ifm bs (E p)"
  5718 proof(clarsimp simp del: Ifm.simps)
  5719   fix bs p
  5720   assume qf: "qfree p"
  5721   show "qfree (mircfr p)\<and>(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
  5722   proof-
  5723     let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
  5724     have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" 
  5725       using splitex[OF qf] by simp
  5726     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+
  5727     with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def)
  5728   qed
  5729 qed
  5730   
  5731 lemma mirlfr: "\<forall> bs p. qfree p \<longrightarrow> qfree(mirlfr p) \<and> Ifm bs (mirlfr p) = Ifm bs (E p)"
  5732 proof(clarsimp simp del: Ifm.simps)
  5733   fix bs p
  5734   assume qf: "qfree p"
  5735   show "qfree (mirlfr p)\<and>(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
  5736   proof-
  5737     let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
  5738     have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" 
  5739       using splitex[OF qf] by simp
  5740     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+
  5741     with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def)
  5742   qed
  5743 qed
  5744   
  5745 constdefs mircfrqe:: "fm \<Rightarrow> fm"
  5746   "mircfrqe \<equiv> (\<lambda> p. qelim (prep p) mircfr)"
  5747 
  5748 constdefs mirlfrqe:: "fm \<Rightarrow> fm"
  5749   "mirlfrqe \<equiv> (\<lambda> p. qelim (prep p) mirlfr)"
  5750 
  5751 theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \<and> qfree (mircfrqe p)"
  5752   using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def)
  5753 
  5754 theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \<and> qfree (mirlfrqe p)"
  5755   using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def)
  5756 
  5757 declare zdvd_iff_zmod_eq_0 [code]
  5758 declare max_def [code unfold]
  5759 
  5760 definition
  5761   "test1 (u\<Colon>unit) = mircfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))"
  5762 
  5763 definition
  5764   "test2 (u\<Colon>unit) = mircfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))"
  5765 
  5766 definition
  5767   "test3 (u\<Colon>unit) = mirlfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))"
  5768 
  5769 definition
  5770   "test4 (u\<Colon>unit) = mirlfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))"
  5771 
  5772 definition
  5773   "test5 (u\<Colon>unit) = mircfrqe (A(E(And (Ge(Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg(Bound 0))))))))"
  5774 
  5775 code_gen mircfrqe mirlfrqe test1 test2 test3 test4 test5
  5776   in SML module_name Mir
  5777 
  5778 (*code_gen mircfrqe mirlfrqe
  5779   in SML module_name Mir file "raw_mir.ML"*)
  5780 
  5781 ML "set Toplevel.timing"
  5782 ML "Mir.test1 ()"
  5783 ML "Mir.test2 ()"
  5784 ML "Mir.test3 ()"
  5785 ML "Mir.test4 ()"
  5786 ML "Mir.test5 ()"
  5787 ML "reset Toplevel.timing"
  5788 
  5789 use "mireif.ML"
  5790 oracle mircfr_oracle ("term") = ReflectedMir.mircfr_oracle
  5791 oracle mirlfr_oracle ("term") = ReflectedMir.mirlfr_oracle
  5792 use "mirtac.ML"
  5793 setup "MirTac.setup"
  5794 
  5795 ML "set Toplevel.timing"
  5796 
  5797 lemma "ALL (x::real). (\<lfloor>x\<rfloor> = \<lceil>x\<rceil> = (x = real \<lfloor>x\<rfloor>))"
  5798 apply mir
  5799 done
  5800 
  5801 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))"
  5802 apply mir
  5803 done
  5804 
  5805 lemma "ALL (x::real). 2*\<lfloor>x\<rfloor> \<le> \<lfloor>2*x\<rfloor> \<and> \<lfloor>2*x\<rfloor> \<le> 2*\<lfloor>x+1\<rfloor>"
  5806 apply mir 
  5807 done
  5808 
  5809 lemma "ALL (x::real). \<exists>y \<le> x. (\<lfloor>x\<rfloor> = \<lceil>y\<rceil>)"
  5810 apply mir
  5811 done
  5812 
  5813 ML "reset Toplevel.timing"
  5814 
  5815 end