src/HOL/Decision_Procs/MIR.thy
author haftmann
Sat, 05 Jul 2014 11:01:53 +0200
changeset 57514 bdc2c6b40bf2
parent 57512 cc97b347b301
child 57816 d8bbb97689d3
permissions -rw-r--r--
prefer ac_simps collections over separate name bindings for add and mult
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30439
57c68b3af2ea Updated paths in Decision_Procs comments and NEWS
hoelzl
parents: 30242
diff changeset
     1
(*  Title:      HOL/Decision_Procs/MIR.thy
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
     2
    Author:     Amine Chaieb
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
     3
*)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
     4
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
     5
theory MIR
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
     6
imports Complex_Main Dense_Linear_Order DP_Library
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
     7
  "~~/src/HOL/Library/Code_Target_Numeral" "~~/src/HOL/Library/Old_Recdef"
27368
9f90ac19e32b established Plain theory and image
haftmann
parents: 26935
diff changeset
     8
begin
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
     9
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
    10
section {* Quantifier elimination for @{text "\<real> (0, 1, +, floor, <)"} *}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
    11
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    12
declare real_of_int_floor_cancel [simp del]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    13
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    14
lemma myle:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    15
  fixes a b :: "'a::{ordered_ab_group_add}"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
    16
  shows "(a \<le> b) = (0 \<le> b - a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    17
  by (metis add_0_left add_le_cancel_right diff_add_cancel)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    18
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    19
lemma myless:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    20
  fixes a b :: "'a::{ordered_ab_group_add}"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
    21
  shows "(a < b) = (0 < b - a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    22
  by (metis le_iff_diff_le_0 less_le_not_le myle)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    23
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    24
  (* Maybe should be added to the library \<dots> *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    25
lemma floor_int_eq: "(real n\<le> x \<and> x < real (n+1)) = (floor x = n)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    26
proof( auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    27
  assume lb: "real n \<le> x"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    28
    and ub: "x < real n + 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    29
  have "real (floor x) \<le> x" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    30
  hence "real (floor x) < real (n + 1) " using ub by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    31
  hence "floor x < n+1" by simp
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 30042
diff changeset
    32
  moreover from lb have "n \<le> floor x" using floor_mono[where x="real n" and y="x"] 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    33
    by simp ultimately show "floor x = n" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    34
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    35
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    36
(* Periodicity of dvd *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    37
lemma dvd_period:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    38
  assumes advdd: "(a::int) dvd d"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    39
  shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    40
  using advdd  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    41
proof-
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    42
  { fix x k
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
    43
    from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"]  
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
    44
    have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp}
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
    45
  hence "\<forall>x.\<forall>k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))"  by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    46
  then show ?thesis by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    47
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    48
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
    49
(* The Divisibility relation between reals *)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    50
definition rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    51
  where "x rdvd y \<longleftrightarrow> (\<exists>k\<Colon>int. y = x * real k)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    52
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    53
lemma int_rdvd_real: 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    54
  "real (i::int) rdvd x = (i dvd (floor x) \<and> real (floor x) = x)" (is "?l = ?r")
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    55
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    56
  assume "?l" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    57
  hence th: "\<exists> k. x=real (i*k)" by (simp add: rdvd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    58
  hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    59
  with th have "\<exists> k. real (floor x) = real (i*k)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    60
  hence "\<exists> k. floor x = i*k" by (simp only: real_of_int_inject)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    61
  thus ?r  using th' by (simp add: dvd_def) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    62
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    63
  assume "?r" hence "(i\<Colon>int) dvd \<lfloor>x\<Colon>real\<rfloor>" ..
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    64
  hence "\<exists> k. real (floor x) = real (i*k)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    65
    by (simp only: real_of_int_inject) (simp add: dvd_def)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
    66
  thus ?l using `?r` by (simp add: rdvd_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    67
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    68
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    69
lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    70
  by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only: real_of_int_mult[symmetric])
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    71
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    72
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    73
lemma rdvd_abs1: "(abs (real d) rdvd t) = (real (d ::int) rdvd t)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    74
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    75
  assume d: "real d rdvd t"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    76
  from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    77
    by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    78
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
    79
  from iffD2[OF abs_dvd_iff] d2 have "(abs d) dvd (floor t)" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    80
  with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    81
  thus "abs (real d) rdvd t" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    82
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    83
  assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    84
  with int_rdvd_real[where i="abs d" and x="t"] have d2: "abs d dvd floor t" and ti: "real (floor t) =t"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    85
    by auto
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
    86
  from iffD1[OF abs_dvd_iff] d2 have "d dvd floor t" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    87
  with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    88
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    89
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    90
lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    91
  apply (auto simp add: rdvd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    92
  apply (rule_tac x="-k" in exI, simp) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    93
  apply (rule_tac x="-k" in exI, simp)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    94
  done
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    95
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    96
lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
    97
  by (auto simp add: rdvd_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    98
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    99
lemma rdvd_mult: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   100
  assumes knz: "k\<noteq>0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   101
  shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   102
  using knz by (simp add: rdvd_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   103
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   104
  (*********************************************************************************)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   105
  (****                            SHADOW SYNTAX AND SEMANTICS                  ****)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   106
  (*********************************************************************************)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   107
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   108
datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   109
  | Mul int num | Floor num| CF int num num
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   110
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   111
  (* A size for num to make inductive proofs simpler*)
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   112
primrec num_size :: "num \<Rightarrow> nat" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   113
 "num_size (C c) = 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   114
| "num_size (Bound n) = 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   115
| "num_size (Neg a) = 1 + num_size a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   116
| "num_size (Add a b) = 1 + num_size a + num_size b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   117
| "num_size (Sub a b) = 3 + num_size a + num_size b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   118
| "num_size (CN n c a) = 4 + num_size a "
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   119
| "num_size (CF c a b) = 4 + num_size a + num_size b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   120
| "num_size (Mul c a) = 1 + num_size a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   121
| "num_size (Floor a) = 1 + num_size a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   122
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   123
  (* Semantics of numeral terms (num) *)
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   124
primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   125
  "Inum bs (C c) = (real c)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   126
| "Inum bs (Bound n) = bs!n"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   127
| "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   128
| "Inum bs (Neg a) = -(Inum bs a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   129
| "Inum bs (Add a b) = Inum bs a + Inum bs b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   130
| "Inum bs (Sub a b) = Inum bs a - Inum bs b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   131
| "Inum bs (Mul c a) = (real c) * Inum bs a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   132
| "Inum bs (Floor a) = real (floor (Inum bs a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   133
| "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   134
definition "isint t bs \<equiv> real (floor (Inum bs t)) = Inum bs t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   135
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   136
lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   137
  by (simp add: isint_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   138
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   139
lemma isint_Floor: "isint (Floor n) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   140
  by (simp add: isint_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   141
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   142
lemma isint_Mul: "isint e bs \<Longrightarrow> isint (Mul c e) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   143
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   144
  let ?e = "Inum bs e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   145
  let ?fe = "floor ?e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   146
  assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   147
  have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   148
  also have "\<dots> = real (c* ?fe)" by (simp only: floor_real_of_int) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   149
  also have "\<dots> = real c * ?e" using efe by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   150
  finally show ?thesis using isint_iff by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   151
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   152
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   153
lemma isint_neg: "isint e bs \<Longrightarrow> isint (Neg e) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   154
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   155
  let ?I = "\<lambda> t. Inum bs t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   156
  assume ie: "isint e bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   157
  hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   158
  have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   159
  also have "\<dots> = - real (floor (?I e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   160
  finally show "isint (Neg e) bs" by (simp add: isint_def th)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   161
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   162
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   163
lemma isint_sub: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   164
  assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   165
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   166
  let ?I = "\<lambda> t. Inum bs t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   167
  from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   168
  have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   169
  also have "\<dots> = real (c- floor (?I e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   170
  finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   171
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   172
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   173
lemma isint_add:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   174
  assumes ai: "isint a bs" and bi: "isint b bs"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   175
  shows "isint (Add a b) bs"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   176
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   177
  let ?a = "Inum bs a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   178
  let ?b = "Inum bs b"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   179
  from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   180
    by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   181
  also have "\<dots> = real (floor ?a) + real (floor ?b)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   182
  also have "\<dots> = ?a + ?b" using ai bi isint_iff by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   183
  finally show "isint (Add a b) bs" by (simp add: isint_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   184
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   185
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   186
lemma isint_c: "isint (C j) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   187
  by (simp add: isint_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   188
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   189
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   190
    (* FORMULAE *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   191
datatype fm  = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   192
  T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num|
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   193
  NOT fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   194
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   195
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   196
  (* A size for fm *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   197
fun fmsize :: "fm \<Rightarrow> nat" where
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   198
 "fmsize (NOT p) = 1 + fmsize p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   199
| "fmsize (And p q) = 1 + fmsize p + fmsize q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   200
| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   201
| "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   202
| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   203
| "fmsize (E p) = 1 + fmsize p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   204
| "fmsize (A p) = 4+ fmsize p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   205
| "fmsize (Dvd i t) = 2"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   206
| "fmsize (NDvd i t) = 2"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   207
| "fmsize p = 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   208
  (* several lemmas about fmsize *)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   209
lemma fmsize_pos: "fmsize p > 0"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   210
  by (induct p rule: fmsize.induct) simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   211
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   212
  (* Semantics of formulae (fm) *)
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   213
primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   214
  "Ifm bs T = True"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   215
| "Ifm bs F = False"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   216
| "Ifm bs (Lt a) = (Inum bs a < 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   217
| "Ifm bs (Gt a) = (Inum bs a > 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   218
| "Ifm bs (Le a) = (Inum bs a \<le> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   219
| "Ifm bs (Ge a) = (Inum bs a \<ge> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   220
| "Ifm bs (Eq a) = (Inum bs a = 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   221
| "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   222
| "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   223
| "Ifm bs (NDvd i b) = (\<not>(real i rdvd Inum bs b))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   224
| "Ifm bs (NOT p) = (\<not> (Ifm bs p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   225
| "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   226
| "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   227
| "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   228
| "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   229
| "Ifm bs (E p) = (\<exists> x. Ifm (x#bs) p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   230
| "Ifm bs (A p) = (\<forall> x. Ifm (x#bs) p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   231
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   232
consts prep :: "fm \<Rightarrow> fm"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   233
recdef prep "measure fmsize"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   234
  "prep (E T) = T"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   235
  "prep (E F) = F"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   236
  "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   237
  "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   238
  "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   239
  "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   240
  "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   241
  "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   242
  "prep (E p) = E (prep p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   243
  "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   244
  "prep (A p) = prep (NOT (E (NOT p)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   245
  "prep (NOT (NOT p)) = prep p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   246
  "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   247
  "prep (NOT (A p)) = prep (E (NOT p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   248
  "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   249
  "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   250
  "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   251
  "prep (NOT p) = NOT (prep p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   252
  "prep (Or p q) = Or (prep p) (prep q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   253
  "prep (And p q) = And (prep p) (prep q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   254
  "prep (Imp p q) = prep (Or (NOT p) q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   255
  "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   256
  "prep p = p"
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   257
(hints simp add: fmsize_pos)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   258
lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   259
  by (induct p rule: prep.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   260
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   261
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   262
  (* Quantifier freeness *)
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   263
fun qfree:: "fm \<Rightarrow> bool" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   264
  "qfree (E p) = False"
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   265
  | "qfree (A p) = False"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   266
  | "qfree (NOT p) = qfree p" 
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   267
  | "qfree (And p q) = (qfree p \<and> qfree q)" 
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   268
  | "qfree (Or  p q) = (qfree p \<and> qfree q)" 
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   269
  | "qfree (Imp p q) = (qfree p \<and> qfree q)" 
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   270
  | "qfree (Iff p q) = (qfree p \<and> qfree q)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   271
  | "qfree p = True"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   272
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   273
  (* Boundedness and substitution *)
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   274
primrec numbound0 :: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *) where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   275
  "numbound0 (C c) = True"
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   276
  | "numbound0 (Bound n) = (n>0)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   277
  | "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   278
  | "numbound0 (Neg a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   279
  | "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   280
  | "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)" 
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   281
  | "numbound0 (Mul i a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   282
  | "numbound0 (Floor a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   283
  | "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)" 
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   284
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   285
lemma numbound0_I:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   286
  assumes nb: "numbound0 a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   287
  shows "Inum (b#bs) a = Inum (b'#bs) a"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   288
  using nb by (induct a) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   289
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   290
lemma numbound0_gen: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   291
  assumes nb: "numbound0 t" and ti: "isint t (x#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   292
  shows "\<forall> y. isint t (y#bs)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   293
  using nb ti 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   294
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   295
  fix y
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   296
  from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   297
  show "isint t (y#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   298
    by (simp add: isint_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   299
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   300
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   301
primrec bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *) where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   302
  "bound0 T = True"
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   303
  | "bound0 F = True"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   304
  | "bound0 (Lt a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   305
  | "bound0 (Le a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   306
  | "bound0 (Gt a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   307
  | "bound0 (Ge a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   308
  | "bound0 (Eq a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   309
  | "bound0 (NEq a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   310
  | "bound0 (Dvd i a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   311
  | "bound0 (NDvd i a) = numbound0 a"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   312
  | "bound0 (NOT p) = bound0 p"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   313
  | "bound0 (And p q) = (bound0 p \<and> bound0 q)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   314
  | "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   315
  | "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   316
  | "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   317
  | "bound0 (E p) = False"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   318
  | "bound0 (A p) = False"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   319
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   320
lemma bound0_I:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   321
  assumes bp: "bound0 p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   322
  shows "Ifm (b#bs) p = Ifm (b'#bs) p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   323
  using bp numbound0_I [where b="b" and bs="bs" and b'="b'"]
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   324
  by (induct p) auto
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   325
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   326
primrec numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *) where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   327
  "numsubst0 t (C c) = (C c)"
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   328
  | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   329
  | "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   330
  | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   331
  | "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   332
  | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   333
  | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" 
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   334
  | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   335
  | "numsubst0 t (Floor a) = Floor (numsubst0 t a)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   336
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   337
lemma numsubst0_I:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   338
  shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   339
  by (induct t) simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   340
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   341
primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *) where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   342
  "subst0 t T = T"
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   343
  | "subst0 t F = F"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   344
  | "subst0 t (Lt a) = Lt (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   345
  | "subst0 t (Le a) = Le (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   346
  | "subst0 t (Gt a) = Gt (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   347
  | "subst0 t (Ge a) = Ge (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   348
  | "subst0 t (Eq a) = Eq (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   349
  | "subst0 t (NEq a) = NEq (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   350
  | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   351
  | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   352
  | "subst0 t (NOT p) = NOT (subst0 t p)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   353
  | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   354
  | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   355
  | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   356
  | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   357
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   358
lemma subst0_I: assumes qfp: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   359
  shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   360
  using qfp numsubst0_I[where b="b" and bs="bs" and a="a"]
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
   361
  by (induct p) simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   362
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   363
fun decrnum:: "num \<Rightarrow> num" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   364
  "decrnum (Bound n) = Bound (n - 1)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   365
| "decrnum (Neg a) = Neg (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   366
| "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   367
| "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   368
| "decrnum (Mul c a) = Mul c (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   369
| "decrnum (Floor a) = Floor (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   370
| "decrnum (CN n c a) = CN (n - 1) c (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   371
| "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   372
| "decrnum a = a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   373
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   374
fun decr :: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   375
  "decr (Lt a) = Lt (decrnum a)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   376
| "decr (Le a) = Le (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   377
| "decr (Gt a) = Gt (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   378
| "decr (Ge a) = Ge (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   379
| "decr (Eq a) = Eq (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   380
| "decr (NEq a) = NEq (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   381
| "decr (Dvd i a) = Dvd i (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   382
| "decr (NDvd i a) = NDvd i (decrnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   383
| "decr (NOT p) = NOT (decr p)" 
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   384
| "decr (And p q) = And (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   385
| "decr (Or p q) = Or (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   386
| "decr (Imp p q) = Imp (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   387
| "decr (Iff p q) = Iff (decr p) (decr q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   388
| "decr p = p"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   389
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   390
lemma decrnum: assumes nb: "numbound0 t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   391
  shows "Inum (x#bs) t = Inum bs (decrnum t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   392
  using nb by (induct t rule: decrnum.induct) simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   393
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   394
lemma decr: assumes nb: "bound0 p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   395
  shows "Ifm (x#bs) p = Ifm bs (decr p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   396
  using nb by (induct p rule: decr.induct) (simp_all add: decrnum)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   397
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   398
lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   399
  by (induct p) simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   400
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   401
fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *) where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   402
  "isatom T = True"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   403
| "isatom F = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   404
| "isatom (Lt a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   405
| "isatom (Le a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   406
| "isatom (Gt a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   407
| "isatom (Ge a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   408
| "isatom (Eq a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   409
| "isatom (NEq a) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   410
| "isatom (Dvd i b) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   411
| "isatom (NDvd i b) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   412
| "isatom p = False"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   413
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   414
lemma numsubst0_numbound0:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   415
  assumes nb: "numbound0 t"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   416
  shows "numbound0 (numsubst0 t a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   417
  using nb by (induct a) auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   418
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   419
lemma subst0_bound0:
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   420
  assumes qf: "qfree p" and nb: "numbound0 t"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   421
  shows "bound0 (subst0 t p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   422
  using qf numsubst0_numbound0[OF nb] by (induct p) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   423
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   424
lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   425
  by (induct p) simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   426
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   427
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   428
definition djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   429
  "djf f p q = (if q=T then T else if q=F then f p else 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   430
  (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or fp q))"
25765
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   431
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   432
definition evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
49580bd58a21 some more primrec
haftmann
parents: 25162
diff changeset
   433
  "evaldjf f ps = foldr (djf f) ps F"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   434
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   435
lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   436
  by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) 
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   437
  (cases "f p", simp_all add: Let_def djf_def) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   438
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   439
lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm bs (f p))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   440
  by (induct ps) (simp_all add: evaldjf_def djf_Or)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   441
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   442
lemma evaldjf_bound0: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   443
  assumes nb: "\<forall> x\<in> set xs. bound0 (f x)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   444
  shows "bound0 (evaldjf f xs)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   445
  using nb
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   446
  apply (induct xs)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   447
  apply (auto simp add: evaldjf_def djf_def Let_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   448
  apply (case_tac "f a")
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   449
  apply auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   450
  done
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   451
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   452
lemma evaldjf_qf: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   453
  assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   454
  shows "qfree (evaldjf f xs)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   455
  using nb
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   456
  apply (induct xs)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   457
  apply (auto simp add: evaldjf_def djf_def Let_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   458
  apply (case_tac "f a")
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   459
  apply auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   460
  done
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   461
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   462
fun disjuncts :: "fm \<Rightarrow> fm list" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   463
  "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   464
| "disjuncts F = []"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   465
| "disjuncts p = [p]"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   466
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   467
fun conjuncts :: "fm \<Rightarrow> fm list" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   468
  "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   469
| "conjuncts T = []"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   470
| "conjuncts p = [p]"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   471
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   472
lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm bs q) = Ifm bs p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   473
  by (induct p rule: conjuncts.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   474
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   475
lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   476
proof -
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   477
  assume qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   478
  hence "list_all qfree (disjuncts p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   479
    by (induct p rule: disjuncts.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   480
  thus ?thesis by (simp only: list_all_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   481
qed
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   482
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   483
lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   484
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   485
  assume qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   486
  hence "list_all qfree (conjuncts p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   487
    by (induct p rule: conjuncts.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   488
  thus ?thesis by (simp only: list_all_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   489
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   490
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   491
definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   492
  "DJ f p \<equiv> evaldjf f (disjuncts p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   493
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   494
lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   495
  and fF: "f F = F"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   496
  shows "Ifm bs (DJ f p) = Ifm bs (f p)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   497
proof -
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   498
  have "Ifm bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm bs (f q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   499
    by (simp add: DJ_def evaldjf_ex) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   500
  also have "\<dots> = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   501
  finally show ?thesis .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   502
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   503
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   504
lemma DJ_qf: assumes 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   505
  fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   506
  shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   507
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   508
  fix  p assume qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   509
  have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   510
  from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   511
  with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   512
  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   513
  from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   514
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   515
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   516
lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   517
  shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm bs ((DJ qe p)) = Ifm bs (E p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   518
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   519
  fix p::fm and bs
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   520
  assume qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   521
  from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   522
  from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   523
  have "Ifm bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (qe q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   524
    by (simp add: DJ_def evaldjf_ex)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   525
  also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   526
  also have "\<dots> = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   527
  finally show "qfree (DJ qe p) \<and> Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   528
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   529
  (* Simplification *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   530
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   531
  (* Algebraic simplifications for nums *)
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   532
fun bnds:: "num \<Rightarrow> nat list" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   533
  "bnds (Bound n) = [n]"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   534
| "bnds (CN n c a) = n#(bnds a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   535
| "bnds (Neg a) = bnds a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   536
| "bnds (Add a b) = (bnds a)@(bnds b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   537
| "bnds (Sub a b) = (bnds a)@(bnds b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   538
| "bnds (Mul i a) = bnds a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   539
| "bnds (Floor a) = bnds a"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   540
| "bnds (CF c a b) = (bnds a)@(bnds b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   541
| "bnds a = []"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   542
fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool" where
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   543
  "lex_ns [] ms = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   544
| "lex_ns ns [] = False"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   545
| "lex_ns (n#ns) (m#ms) = (n<m \<or> ((n = m) \<and> lex_ns ns ms)) "
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   546
definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   547
  "lex_bnd t s \<equiv> lex_ns (bnds t) (bnds s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   548
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   549
fun maxcoeff:: "num \<Rightarrow> int" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   550
  "maxcoeff (C i) = abs i"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   551
| "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   552
| "maxcoeff (CF c t s) = max (abs c) (maxcoeff s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   553
| "maxcoeff t = 1"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   554
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   555
lemma maxcoeff_pos: "maxcoeff t \<ge> 0"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   556
  by (induct t rule: maxcoeff.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   557
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   558
fun numgcdh:: "num \<Rightarrow> int \<Rightarrow> int" where
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   559
  "numgcdh (C i) = (\<lambda>g. gcd i g)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   560
| "numgcdh (CN n c t) = (\<lambda>g. gcd c (numgcdh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   561
| "numgcdh (CF c s t) = (\<lambda>g. gcd c (numgcdh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   562
| "numgcdh t = (\<lambda>g. 1)"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
   563
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   564
definition numgcd :: "num \<Rightarrow> int"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   565
  where "numgcd t = numgcdh t (maxcoeff t)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   566
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   567
fun reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   568
  "reducecoeffh (C i) = (\<lambda> g. C (i div g))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   569
| "reducecoeffh (CN n c t) = (\<lambda> g. CN n (c div g) (reducecoeffh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   570
| "reducecoeffh (CF c s t) = (\<lambda> g. CF (c div g)  s (reducecoeffh t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   571
| "reducecoeffh t = (\<lambda>g. t)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   572
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   573
definition reducecoeff :: "num \<Rightarrow> num"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
   574
where
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   575
  "reducecoeff t =
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   576
    (let g = numgcd t in 
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   577
     if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   578
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   579
fun dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   580
  "dvdnumcoeff (C i) = (\<lambda> g. g dvd i)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   581
| "dvdnumcoeff (CN n c t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   582
| "dvdnumcoeff (CF c s t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   583
| "dvdnumcoeff t = (\<lambda>g. False)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   584
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   585
lemma dvdnumcoeff_trans: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   586
  assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   587
  shows "dvdnumcoeff t g"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   588
  using dgt' gdg 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   589
  by (induct t rule: dvdnumcoeff.induct) (simp_all add: gdg dvd_trans[OF gdg])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
   590
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
   591
declare dvd_trans [trans add]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   592
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   593
lemma numgcd0:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   594
  assumes g0: "numgcd t = 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   595
  shows "Inum bs t = 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   596
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   597
  have "\<And>x. numgcdh t x= 0 \<Longrightarrow> Inum bs t = 0"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   598
    by (induct t rule: numgcdh.induct, auto)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   599
  thus ?thesis using g0[simplified numgcd_def] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   600
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   601
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   602
lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   603
  using gp by (induct t rule: numgcdh.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   604
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   605
lemma numgcd_pos: "numgcd t \<ge>0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   606
  by (simp add: numgcd_def numgcdh_pos maxcoeff_pos)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   607
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   608
lemma reducecoeffh:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   609
  assumes gt: "dvdnumcoeff t g" and gp: "g > 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   610
  shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   611
  using gt
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   612
proof(induct t rule: reducecoeffh.induct) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   613
  case (1 i) hence gd: "g dvd i" by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
   614
  from assms 1 show ?case by (simp add: real_of_int_div[OF gd])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   615
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   616
  case (2 n c t)  hence gd: "g dvd c" by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
   617
  from assms 2 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   618
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   619
  case (3 c s t)  hence gd: "g dvd c" by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
   620
  from assms 3 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   621
qed (auto simp add: numgcd_def gp)
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   622
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   623
fun ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   624
  "ismaxcoeff (C i) = (\<lambda> x. abs i \<le> x)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   625
| "ismaxcoeff (CN n c t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   626
| "ismaxcoeff (CF c s t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   627
| "ismaxcoeff t = (\<lambda>x. True)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   628
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   629
lemma ismaxcoeff_mono: "ismaxcoeff t c \<Longrightarrow> c \<le> c' \<Longrightarrow> ismaxcoeff t c'"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   630
  by (induct t rule: ismaxcoeff.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   631
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   632
lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   633
proof (induct t rule: maxcoeff.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   634
  case (2 n c t)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   635
  hence H:"ismaxcoeff t (maxcoeff t)" .
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   636
  have thh: "maxcoeff t \<le> max (abs c) (maxcoeff t)" by simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   637
  from ismaxcoeff_mono[OF H thh] show ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   638
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   639
  case (3 c t s) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   640
  hence H1:"ismaxcoeff s (maxcoeff s)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   641
  have thh1: "maxcoeff s \<le> max \<bar>c\<bar> (maxcoeff s)" by (simp add: max_def)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   642
  from ismaxcoeff_mono[OF H1 thh1] show ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   643
qed simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   644
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   645
lemma zgcd_gt1: "gcd i j > (1::int) \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))"
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   646
  apply (unfold gcd_int_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   647
  apply (cases "i = 0", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   648
  apply (cases "j = 0", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   649
  apply (cases "abs i = 1", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   650
  apply (cases "abs j = 1", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   651
  apply auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   652
  done
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   653
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   654
lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow>  m =0"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   655
  by (induct t rule: numgcdh.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   656
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   657
lemma dvdnumcoeff_aux:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   658
  assumes "ismaxcoeff t m" and mp:"m \<ge> 0" and "numgcdh t m > 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   659
  shows "dvdnumcoeff t (numgcdh t m)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   660
using assms
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   661
proof(induct t rule: numgcdh.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   662
  case (2 n c t) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   663
  let ?g = "numgcdh t m"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   664
  from 2 have th:"gcd c ?g > 1" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27456
diff changeset
   665
  from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   666
  have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   667
  moreover {assume "abs c > 1" and gp: "?g > 1" with 2
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   668
    have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   669
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   670
    from dvdnumcoeff_trans[OF th' th] have ?case by simp }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   671
  moreover {assume "abs c = 0 \<and> ?g > 1"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   672
    with 2 have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   673
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   674
    from dvdnumcoeff_trans[OF th' th] have ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   675
    hence ?case by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   676
  moreover {assume "abs c > 1" and g0:"?g = 0" 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   677
    from numgcdh0[OF g0] have "m=0". with 2 g0 have ?case by simp }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   678
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   679
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   680
  case (3 c s t) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   681
  let ?g = "numgcdh t m"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   682
  from 3 have th:"gcd c ?g > 1" by simp
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 27456
diff changeset
   683
  from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   684
  have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   685
  moreover {assume "abs c > 1" and gp: "?g > 1" with 3
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   686
    have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   687
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   688
    from dvdnumcoeff_trans[OF th' th] have ?case by simp }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   689
  moreover {assume "abs c = 0 \<and> ?g > 1"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   690
    with 3 have th: "dvdnumcoeff t ?g" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   691
    have th': "gcd c ?g dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   692
    from dvdnumcoeff_trans[OF th' th] have ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   693
    hence ?case by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   694
  moreover {assume "abs c > 1" and g0:"?g = 0" 
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   695
    from numgcdh0[OF g0] have "m=0". with 3 g0 have ?case by simp }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   696
  ultimately show ?case by blast
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   697
qed auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   698
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   699
lemma dvdnumcoeff_aux2:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   700
  assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   701
  using assms 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   702
proof (simp add: numgcd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   703
  let ?mc = "maxcoeff t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   704
  let ?g = "numgcdh t ?mc"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   705
  have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   706
  have th2: "?mc \<ge> 0" by (rule maxcoeff_pos)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   707
  assume H: "numgcdh t ?mc > 1"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   708
  from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" .
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   709
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   710
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   711
lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   712
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   713
  let ?g = "numgcd t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   714
  have "?g \<ge> 0"  by (simp add: numgcd_pos)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   715
  hence "?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   716
  moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   717
  moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   718
  moreover { assume g1:"?g > 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   719
    from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   720
    from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   721
      by (simp add: reducecoeff_def Let_def)} 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   722
  ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   723
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   724
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   725
lemma reducecoeffh_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeffh t g)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   726
  by (induct t rule: reducecoeffh.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   727
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   728
lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   729
  using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   730
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   731
consts numadd:: "num \<times> num \<Rightarrow> num"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   732
recdef numadd "measure (\<lambda> (t,s). size t + size s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   733
  "numadd (CN n1 c1 r1,CN n2 c2 r2) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   734
  (if n1=n2 then 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   735
  (let c = c1 + c2
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   736
  in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2))))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   737
  else if n1 \<le> n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   738
  else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   739
  "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   740
  "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   741
  "numadd (CF c1 t1 r1,CF c2 t2 r2) = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   742
   (if t1 = t2 then 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   743
    (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   744
   else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   745
   else CF c2 t2 (numadd(CF c1 t1 r1,r2)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   746
  "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   747
  "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   748
  "numadd (C b1, C b2) = C (b1+b2)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   749
  "numadd (a,b) = Add a b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   750
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   751
lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   752
apply (induct t s rule: numadd.induct, simp_all add: Let_def)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23464
diff changeset
   753
 apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
   754
  apply (case_tac "n1 = n2", simp_all add: algebra_simps)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49069
diff changeset
   755
  apply (simp only: distrib_right[symmetric])
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23464
diff changeset
   756
 apply simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   757
apply (case_tac "lex_bnd t1 t2", simp_all)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23464
diff changeset
   758
 apply (case_tac "c1+c2 = 0")
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   759
  apply (case_tac "t1 = t2")
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   760
   apply (simp_all add: algebra_simps distrib_right[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add distrib_right)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   761
  done
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   762
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   763
lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   764
  by (induct t s rule: numadd.induct) (auto simp add: Let_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   765
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   766
fun nummul:: "num \<Rightarrow> int \<Rightarrow> num" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   767
  "nummul (C j) = (\<lambda> i. C (i*j))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   768
| "nummul (CN n c t) = (\<lambda> i. CN n (c*i) (nummul t i))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   769
| "nummul (CF c t s) = (\<lambda> i. CF (c*i) t (nummul s i))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   770
| "nummul (Mul c t) = (\<lambda> i. nummul t (i*c))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   771
| "nummul t = (\<lambda> i. Mul i t)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   772
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   773
lemma nummul[simp]: "\<And> i. Inum bs (nummul t i) = Inum bs (Mul i t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   774
  by (induct t rule: nummul.induct) (auto simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   775
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   776
lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   777
  by (induct t rule: nummul.induct) auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   778
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   779
definition numneg :: "num \<Rightarrow> num"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   780
  where "numneg t \<equiv> nummul t (- 1)"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   781
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   782
definition numsub :: "num \<Rightarrow> num \<Rightarrow> num"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   783
  where "numsub s t \<equiv> (if s = t then C 0 else numadd (s,numneg t))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   784
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   785
lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   786
  using numneg_def nummul by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   787
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   788
lemma numneg_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numneg t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   789
  using numneg_def by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   790
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   791
lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   792
  using numsub_def by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   793
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   794
lemma numsub_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numsub t s)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   795
  using numsub_def by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   796
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   797
lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   798
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   799
  have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   800
  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   801
  have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   802
  also have "\<dots>" by (simp add: isint_add cti si)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   803
  finally show ?thesis .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   804
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   805
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   806
fun split_int:: "num \<Rightarrow> num \<times> num" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   807
  "split_int (C c) = (C 0, C c)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   808
| "split_int (CN n c b) = 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   809
     (let (bv,bi) = split_int b 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   810
       in (CN n c bv, bi))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   811
| "split_int (CF c a b) = 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   812
     (let (bv,bi) = split_int b 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   813
       in (bv, CF c a bi))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   814
| "split_int a = (a,C 0)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   815
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   816
lemma split_int: "\<And>tv ti. split_int t = (tv,ti) \<Longrightarrow> (Inum bs (Add tv ti) = Inum bs t) \<and> isint ti bs"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   817
proof (induct t rule: split_int.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   818
  case (2 c n b tv ti)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   819
  let ?bv = "fst (split_int b)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   820
  let ?bi = "snd (split_int b)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   821
  have "split_int b = (?bv,?bi)" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   822
  with 2(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   823
  from 2(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   824
  from 2(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   825
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   826
  case (3 c a b tv ti) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   827
  let ?bv = "fst (split_int b)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   828
  let ?bi = "snd (split_int b)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   829
  have "split_int b = (?bv,?bi)" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   830
  with 3(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   831
  from 3(2) have tibi: "ti = CF c a ?bi"
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   832
    by (simp add: Let_def split_def)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   833
  from 3(2) b[symmetric] bii show ?case
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   834
    by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
   835
qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   836
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   837
lemma split_int_nb: "numbound0 t \<Longrightarrow> numbound0 (fst (split_int t)) \<and> numbound0 (snd (split_int t)) "
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   838
  by (induct t rule: split_int.induct) (auto simp add: Let_def split_def)
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   839
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   840
definition numfloor:: "num \<Rightarrow> num"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
   841
where
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   842
  "numfloor t = (let (tv,ti) = split_int t in 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   843
  (case tv of C i \<Rightarrow> numadd (tv,ti) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   844
  | _ \<Rightarrow> numadd(CF 1 tv (C 0),ti)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   845
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   846
lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   847
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   848
  let ?tv = "fst (split_int t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   849
  let ?ti = "snd (split_int t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   850
  have tvti:"split_int t = (?tv,?ti)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   851
  {assume H: "\<forall> v. ?tv \<noteq> C v"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   852
    hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   853
      by (cases ?tv) (auto simp add: numfloor_def Let_def split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   854
    from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   855
    hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   856
    also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   857
      by (simp,subst tii[simplified isint_iff, symmetric]) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   858
    also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   859
    finally have ?thesis using th1 by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   860
  moreover {fix v assume H:"?tv = C v" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   861
    from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   862
    hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   863
    also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   864
      by (simp,subst tii[simplified isint_iff, symmetric]) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   865
    also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   866
    finally have ?thesis by (simp add: H numfloor_def Let_def split_def) }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   867
  ultimately show ?thesis by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   868
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   869
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   870
lemma numfloor_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numfloor t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   871
  using split_int_nb[where t="t"]
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   872
  by (cases "fst (split_int t)") (auto simp add: numfloor_def Let_def split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   873
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   874
function simpnum:: "num \<Rightarrow> num" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   875
  "simpnum (C j) = C j"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   876
| "simpnum (Bound n) = CN n 1 (C 0)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   877
| "simpnum (Neg t) = numneg (simpnum t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   878
| "simpnum (Add t s) = numadd (simpnum t,simpnum s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   879
| "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   880
| "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   881
| "simpnum (Floor t) = numfloor (simpnum t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   882
| "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   883
| "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   884
by pat_completeness auto
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   885
termination by (relation "measure num_size") auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   886
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   887
lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   888
  by (induct t rule: simpnum.induct) auto
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   889
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   890
lemma simpnum_numbound0[simp]: "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   891
  by (induct t rule: simpnum.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   892
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   893
fun nozerocoeff:: "num \<Rightarrow> bool" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   894
  "nozerocoeff (C c) = True"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   895
| "nozerocoeff (CN n c t) = (c\<noteq>0 \<and> nozerocoeff t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   896
| "nozerocoeff (CF c s t) = (c \<noteq> 0 \<and> nozerocoeff t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   897
| "nozerocoeff (Mul c t) = (c\<noteq>0 \<and> nozerocoeff t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
   898
| "nozerocoeff t = True"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   899
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   900
lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd (a,b))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   901
  by (induct a b rule: numadd.induct) (auto simp add: Let_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   902
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   903
lemma nummul_nz : "\<And> i. i\<noteq>0 \<Longrightarrow> nozerocoeff a \<Longrightarrow> nozerocoeff (nummul a i)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   904
  by (induct a rule: nummul.induct) (auto simp add: Let_def numadd_nz)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   905
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   906
lemma numneg_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff (numneg a)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   907
  by (simp add: numneg_def nummul_nz)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   908
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   909
lemma numsub_nz: "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numsub a b)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   910
  by (simp add: numsub_def numneg_nz numadd_nz)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   911
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   912
lemma split_int_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (fst (split_int t)) \<and> nozerocoeff (snd (split_int t))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   913
  by (induct t rule: split_int.induct) (auto simp add: Let_def split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   914
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   915
lemma numfloor_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (numfloor t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   916
  by (simp add: numfloor_def Let_def split_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   917
    (cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   918
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   919
lemma simpnum_nz: "nozerocoeff (simpnum t)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   920
  by (induct t rule: simpnum.induct)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   921
    (auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   922
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   923
lemma maxcoeff_nz: "nozerocoeff t \<Longrightarrow> maxcoeff t = 0 \<Longrightarrow> t = C 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   924
proof (induct t rule: maxcoeff.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   925
  case (2 n c t)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   926
  hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   927
  have "max (abs c) (maxcoeff t) \<ge> abs c" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   928
  with cnz have "max (abs c) (maxcoeff t) > 0" by arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   929
  with 2 show ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   930
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   931
  case (3 c s t) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   932
  hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
   933
  have "max (abs c) (maxcoeff t) \<ge> abs c" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   934
  with cnz have "max (abs c) (maxcoeff t) > 0" by arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   935
  with 3 show ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   936
qed auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   937
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   938
lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   939
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   940
  from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   941
  from numgcdh0[OF th]  have th:"maxcoeff t = 0" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   942
  from maxcoeff_nz[OF nz th] show ?thesis .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   943
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   944
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   945
definition simp_num_pair :: "(num \<times> int) \<Rightarrow> num \<times> int" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   946
  "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   947
   (let t' = simpnum t ; g = numgcd t' in 
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   948
      if g > 1 then (let g' = gcd n g in 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   949
        if g' = 1 then (t',n) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   950
        else (reducecoeffh t' g', n div g')) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   951
      else (t',n))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   952
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   953
lemma simp_num_pair_ci:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   954
  shows "((\<lambda> (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\<lambda> (t,n). Inum bs t / real n) (t,n))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   955
  (is "?lhs = ?rhs")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   956
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   957
  let ?t' = "simpnum t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   958
  let ?g = "numgcd ?t'"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   959
  let ?g' = "gcd n ?g"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   960
  {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   961
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   962
  { assume nnz: "n \<noteq> 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   963
    {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   964
    moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   965
    {assume g1:"?g>1" hence g0: "?g > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   966
      from g1 nnz have gp0: "?g' \<noteq> 0" by simp
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31730
diff changeset
   967
      hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   968
      hence "?g'= 1 \<or> ?g' > 1" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   969
      moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   970
      moreover {assume g'1:"?g'>1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   971
        from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   972
        let ?tt = "reducecoeffh ?t' ?g'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   973
        let ?t = "Inum bs ?tt"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   974
        have gpdg: "?g' dvd ?g" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   975
        have gpdd: "?g' dvd n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   976
        have gpdgp: "?g' dvd ?g'" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   977
        from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   978
        have th2:"real ?g' * ?t = Inum bs ?t'" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   979
        from nnz g1 g'1 have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   980
        also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   981
        also have "\<dots> = (Inum bs ?t' / real n)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
   982
          using real_of_int_div[OF gpdd] th2 gp0 by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
   983
        finally have "?lhs = Inum bs t / real n" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   984
        then have ?thesis using nnz g1 g'1 by (simp add: simp_num_pair_def) }
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   985
      ultimately have ?thesis by blast }
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   986
    ultimately have ?thesis by blast }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   987
  ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   988
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   989
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   990
lemma simp_num_pair_l:
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   991
  assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   992
  shows "numbound0 t' \<and> n' >0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   993
proof-
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   994
  let ?t' = "simpnum t"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   995
  let ?g = "numgcd ?t'"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
   996
  let ?g' = "gcd n ?g"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
   997
  { assume nz: "n = 0" hence ?thesis using assms by (simp add: Let_def simp_num_pair_def) }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   998
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
   999
  { assume nnz: "n \<noteq> 0"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1000
    {assume "\<not> ?g > 1" hence ?thesis using assms by (auto simp add: Let_def simp_num_pair_def) }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1001
    moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1002
    {assume g1:"?g>1" hence g0: "?g > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1003
      from g1 nnz have gp0: "?g' \<noteq> 0" by simp
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31730
diff changeset
  1004
      hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1005
      hence "?g'= 1 \<or> ?g' > 1" by arith
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1006
      moreover {assume "?g'=1" hence ?thesis using assms g1 g0
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1007
          by (auto simp add: Let_def simp_num_pair_def) }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1008
      moreover {assume g'1:"?g'>1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1009
        have gpdg: "?g' dvd ?g" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1010
        have gpdd: "?g' dvd n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1011
        have gpdgp: "?g' dvd ?g'" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1012
        from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  1013
        from zdiv_mono1[OF g'n g'p, simplified div_self[OF gp0]]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1014
        have "n div ?g' >0" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1015
        hence ?thesis using assms g1 g'1
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1016
          by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)}
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1017
      ultimately have ?thesis by blast }
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1018
    ultimately have ?thesis by blast } 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1019
  ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1020
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1021
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1022
fun not:: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1023
  "not (NOT p) = p"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1024
| "not T = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1025
| "not F = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1026
| "not (Lt t) = Ge t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1027
| "not (Le t) = Gt t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1028
| "not (Gt t) = Le t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1029
| "not (Ge t) = Lt t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1030
| "not (Eq t) = NEq t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1031
| "not (NEq t) = Eq t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1032
| "not (Dvd i t) = NDvd i t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1033
| "not (NDvd i t) = Dvd i t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1034
| "not (And p q) = Or (not p) (not q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1035
| "not (Or p q) = And (not p) (not q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1036
| "not p = NOT p"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1037
lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1038
  by (induct p) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1039
lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1040
  by (induct p) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1041
lemma not_nb[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1042
  by (induct p) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1043
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1044
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1045
  "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 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1046
   if p = q then p else And p q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1047
lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1048
  by (cases "p=F \<or> q=F", simp_all add: conj_def) (cases p, simp_all)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1049
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1050
lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1051
  using conj_def by auto 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1052
lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1053
  using conj_def by auto 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1054
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1055
definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1056
  "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1057
       else if p=q then p else Or p q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1058
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1059
lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1060
  by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1061
lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1062
  using disj_def by auto 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1063
lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1064
  using disj_def by auto 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1065
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1066
definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1067
  "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 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1068
    else Imp p q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1069
lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1070
  by (cases "p=F \<or> q=T",simp_all add: imp_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1071
lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1072
  using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1073
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1074
definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1075
  "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1076
       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 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1077
  Iff p q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1078
lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1079
  by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1080
(cases "not p= q", auto simp add:not)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1081
lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1082
  by (unfold iff_def,cases "p=q", auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1083
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1084
fun check_int:: "num \<Rightarrow> bool" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1085
  "check_int (C i) = True"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1086
| "check_int (Floor t) = True"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1087
| "check_int (Mul i t) = check_int t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1088
| "check_int (Add t s) = (check_int t \<and> check_int s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1089
| "check_int (Neg t) = check_int t"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1090
| "check_int (CF c t s) = check_int s"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1091
| "check_int t = False"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1092
lemma check_int: "check_int t \<Longrightarrow> isint t bs"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1093
  by (induct t) (auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1094
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1095
lemma rdvd_left1_int: "real \<lfloor>t\<rfloor> = t \<Longrightarrow> 1 rdvd t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1096
  by (simp add: rdvd_def,rule_tac x="\<lfloor>t\<rfloor>" in exI) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1097
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1098
lemma rdvd_reduce: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1099
  assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1100
  shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1101
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1102
  assume d: "real d rdvd real c * t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1103
  from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1104
  from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1105
  from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1106
  from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1107
  hence "real kc * t = real kd * real k" using gp by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1108
  hence th:"real kd rdvd real kc * t" using rdvd_def by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1109
  from kd_def gp have th':"kd = d div g" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1110
  from kc_def gp have "kc = c div g" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1111
  with th th' show "real (d div g) rdvd real (c div g) * t" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1112
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1113
  assume d: "real (d div g) rdvd real (c div g) * t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1114
  from gp have gnz: "g \<noteq> 0" by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  1115
  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 gd] real_of_int_div[OF gc] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1116
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1117
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1118
definition simpdvd :: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1119
  "simpdvd d t \<equiv> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1120
   (let g = numgcd t in 
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1121
      if g > 1 then (let g' = gcd d g in 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1122
        if g' = 1 then (d, t) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1123
        else (d div g',reducecoeffh t g')) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1124
      else (d, t))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1125
lemma simpdvd: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1126
  assumes tnz: "nozerocoeff t" and dnz: "d \<noteq> 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1127
  shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1128
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1129
  let ?g = "numgcd t"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1130
  let ?g' = "gcd d ?g"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1131
  {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1132
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1133
  {assume g1:"?g>1" hence g0: "?g > 0" by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1134
    from g1 dnz have gp0: "?g' \<noteq> 0" by simp
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31730
diff changeset
  1135
    hence g'p: "?g' > 0" using gcd_ge_0_int[where x="d" and y="numgcd t"] by arith
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1136
    hence "?g'= 1 \<or> ?g' > 1" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1137
    moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1138
    moreover {assume g'1:"?g'>1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1139
      from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" ..
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1140
      let ?tt = "reducecoeffh t ?g'"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1141
      let ?t = "Inum bs ?tt"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1142
      have gpdg: "?g' dvd ?g" by simp
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1143
      have gpdd: "?g' dvd d" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1144
      have gpdgp: "?g' dvd ?g'" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1145
      from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1146
      have th2:"real ?g' * ?t = Inum bs t" by simp
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1147
      from assms g1 g0 g'1
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1148
      have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1149
        by (simp add: simpdvd_def Let_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1150
      also have "\<dots> = (real d rdvd (Inum bs t))"
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  1151
        using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified div_self[OF gp0]] 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1152
          th2[symmetric] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1153
      finally have ?thesis by simp  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1154
    ultimately have ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1155
  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1156
  ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1157
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1158
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1159
function (sequential) simpfm :: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1160
  "simpfm (And p q) = conj (simpfm p) (simpfm q)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1161
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1162
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1163
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1164
| "simpfm (NOT p) = not (simpfm p)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1165
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1166
  | _ \<Rightarrow> Lt (reducecoeff a'))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1167
| "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'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1168
| "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'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1169
| "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'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1170
| "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'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1171
| "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'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1172
| "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1173
             else if (abs i = 1) \<and> check_int a then T
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1174
             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))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1175
| "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1176
             else if (abs i = 1) \<and> check_int a then F
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1177
             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))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1178
| "simpfm p = p"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1179
by pat_completeness auto
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1180
termination by (relation "measure fmsize") auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1181
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1182
lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1183
proof(induct p rule: simpfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1184
  case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1185
  {fix v assume "?sa = C v" hence ?case using sa by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1186
  moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1187
    let ?g = "numgcd ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1188
    let ?rsa = "reducecoeff ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1189
    let ?r = "Inum bs ?rsa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1190
    have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1191
    {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1192
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1193
    hence gp: "real ?g > 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1194
    have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1195
    with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1196
    also have "\<dots> = (?r < 0)" using gp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1197
      by (simp only: mult_less_cancel_left) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1198
    finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1199
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1200
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1201
  case (7 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1202
  {fix v assume "?sa = C v" hence ?case using sa by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1203
  moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1204
    let ?g = "numgcd ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1205
    let ?rsa = "reducecoeff ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1206
    let ?r = "Inum bs ?rsa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1207
    have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1208
    {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1209
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1210
    hence gp: "real ?g > 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1211
    have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1212
    with sa have "Inum bs a \<le> 0 = (real ?g * ?r \<le> real ?g * 0)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1213
    also have "\<dots> = (?r \<le> 0)" using gp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1214
      by (simp only: mult_le_cancel_left) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1215
    finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1216
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1217
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1218
  case (8 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1219
  {fix v assume "?sa = C v" hence ?case using sa by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1220
  moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1221
    let ?g = "numgcd ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1222
    let ?rsa = "reducecoeff ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1223
    let ?r = "Inum bs ?rsa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1224
    have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1225
    {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1226
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1227
    hence gp: "real ?g > 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1228
    have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1229
    with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1230
    also have "\<dots> = (?r > 0)" using gp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1231
      by (simp only: mult_less_cancel_left) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1232
    finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1233
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1234
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1235
  case (9 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1236
  {fix v assume "?sa = C v" hence ?case using sa by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1237
  moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1238
    let ?g = "numgcd ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1239
    let ?rsa = "reducecoeff ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1240
    let ?r = "Inum bs ?rsa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1241
    have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1242
    {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1243
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1244
    hence gp: "real ?g > 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1245
    have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1246
    with sa have "Inum bs a \<ge> 0 = (real ?g * ?r \<ge> real ?g * 0)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1247
    also have "\<dots> = (?r \<ge> 0)" using gp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1248
      by (simp only: mult_le_cancel_left) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1249
    finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1250
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1251
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1252
  case (10 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1253
  {fix v assume "?sa = C v" hence ?case using sa by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1254
  moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1255
    let ?g = "numgcd ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1256
    let ?rsa = "reducecoeff ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1257
    let ?r = "Inum bs ?rsa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1258
    have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1259
    {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1260
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1261
    hence gp: "real ?g > 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1262
    have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1263
    with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1264
    also have "\<dots> = (?r = 0)" using gp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1265
      by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1266
    finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1267
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1268
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1269
  case (11 a)  let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1270
  {fix v assume "?sa = C v" hence ?case using sa by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1271
  moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1272
    let ?g = "numgcd ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1273
    let ?rsa = "reducecoeff ?sa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1274
    let ?r = "Inum bs ?rsa"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1275
    have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1276
    {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1277
    with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1278
    hence gp: "real ?g > 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1279
    have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1280
    with sa have "Inum bs a \<noteq> 0 = (real ?g * ?r \<noteq> 0)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1281
    also have "\<dots> = (?r \<noteq> 0)" using gp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1282
      by simp
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1283
    finally have ?case using H by (cases "?sa") (simp_all add: Let_def) }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1284
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1285
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1286
  case (12 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1287
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1288
  {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1289
  moreover 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1290
  {assume ai1: "abs i = 1" and ai: "check_int a" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1291
    hence "i=1 \<or> i= - 1" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1292
    moreover {assume i1: "i = 1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1293
      from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1294
      have ?case using i1 ai by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1295
    moreover {assume i1: "i = - 1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1296
      from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1297
        rdvd_abs1[where d="- 1" and t="Inum bs a"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1298
      have ?case using i1 ai by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1299
    ultimately have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1300
  moreover   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1301
  {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1302
    {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1303
        by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1304
    moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1305
      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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1306
      from simpnum_nz have nz:"nozerocoeff ?sa" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1307
      from simpdvd [OF nz inz] th have ?case using sa by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1308
    ultimately have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1309
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1310
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1311
  case (13 i a)  let ?sa = "simpnum a"   have sa: "Inum bs ?sa = Inum bs a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1312
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1313
  {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1314
  moreover 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1315
  {assume ai1: "abs i = 1" and ai: "check_int a" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1316
    hence "i=1 \<or> i= - 1" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1317
    moreover {assume i1: "i = 1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1318
      from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1319
      have ?case using i1 ai by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1320
    moreover {assume i1: "i = - 1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1321
      from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1322
        rdvd_abs1[where d="- 1" and t="Inum bs a"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1323
      have ?case using i1 ai by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1324
    ultimately have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1325
  moreover   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1326
  {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1327
    {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1328
        by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1329
    moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1330
      hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  1331
        by (cases ?sa, auto simp add: Let_def split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1332
      from simpnum_nz have nz:"nozerocoeff ?sa" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1333
      from simpdvd [OF nz inz] th have ?case using sa by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1334
    ultimately have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1335
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1336
qed (induct p rule: simpfm.induct, simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1337
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1338
lemma simpdvd_numbound0: "numbound0 t \<Longrightarrow> numbound0 (snd (simpdvd d t))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1339
  by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1340
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1341
lemma simpfm_bound0[simp]: "bound0 p \<Longrightarrow> bound0 (simpfm p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1342
proof(induct p rule: simpfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1343
  case (6 a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1344
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1345
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1346
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1347
  case (7 a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1348
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1349
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1350
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1351
  case (8 a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1352
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1353
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1354
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1355
  case (9 a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1356
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1357
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1358
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1359
  case (10 a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1360
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1361
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1362
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1363
  case (11 a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1364
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1365
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1366
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1367
  case (12 i a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1368
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1369
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1370
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1371
  case (13 i a) hence nb: "numbound0 a" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1372
  hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1373
  thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1374
qed(auto simp add: disj_def imp_def iff_def conj_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1375
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1376
lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1377
by (induct p rule: simpfm.induct, auto simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1378
(case_tac "simpnum a",auto simp add: split_def Let_def)+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1379
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1380
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1381
  (* Generic quantifier elimination *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1382
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1383
definition list_conj :: "fm list \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1384
  "list_conj ps \<equiv> foldr conj ps T"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1385
lemma list_conj: "Ifm bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm bs p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1386
  by (induct ps, auto simp add: list_conj_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1387
lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1388
  by (induct ps, auto simp add: list_conj_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1389
lemma list_conj_nb: " \<forall>p\<in> set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1390
  by (induct ps, auto simp add: list_conj_def)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  1391
definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1392
  "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = List.partition bound0 cjs
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1393
                   in conj (decr (list_conj yes)) (f (list_conj no)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1394
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1395
lemma CJNB_qe: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1396
  assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1397
  shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm bs ((CJNB qe p)) = Ifm bs (E p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1398
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1399
  fix bs p
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1400
  assume qfp: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1401
  let ?cjs = "conjuncts p"
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1402
  let ?yes = "fst (List.partition bound0 ?cjs)"
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1403
  let ?no = "snd (List.partition bound0 ?cjs)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1404
  let ?cno = "list_conj ?no"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1405
  let ?cyes = "list_conj ?yes"
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29667
diff changeset
  1406
  have part: "List.partition bound0 ?cjs = (?yes,?no)" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1407
  from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1408
  hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1409
  hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1410
  from conjuncts_qf[OF qfp] partition_set[OF part] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1411
  have " \<forall>q\<in> set ?no. qfree q" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1412
  hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1413
  with qe have cno_qf:"qfree (qe ?cno )" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1414
    and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1415
  from cno_qf yes_qf have qf: "qfree (CJNB qe p)" 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1416
    by (simp add: CJNB_def Let_def split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1417
  {fix bs
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1418
    from conjuncts have "Ifm bs p = (\<forall>q\<in> set ?cjs. Ifm bs q)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1419
    also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm bs q) \<and> (\<forall>q\<in> set ?no. Ifm bs q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1420
      using partition_set[OF part] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1421
    finally have "Ifm bs p = ((Ifm bs ?cyes) \<and> (Ifm bs ?cno))" using list_conj by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1422
  hence "Ifm bs (E p) = (\<exists>x. (Ifm (x#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  1423
  also fix y have "\<dots> = (\<exists>x. (Ifm (y#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1424
    using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1425
  also have "\<dots> = (Ifm bs (decr ?cyes) \<and> Ifm bs (E ?cno))"
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33063
diff changeset
  1426
    by (auto simp add: decr[OF yes_nb] simp del: partition_filter_conv)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1427
  also have "\<dots> = (Ifm bs (conj (decr ?cyes) (qe ?cno)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1428
    using qe[rule_format, OF no_qf] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1429
  finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1430
    by (simp add: Let_def CJNB_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1431
  with qf show "qfree (CJNB qe p) \<and> Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1432
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1433
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1434
function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1435
  "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1436
| "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1437
| "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1438
| "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))" 
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1439
| "qelim (Or  p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))" 
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1440
| "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1441
| "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1442
| "qelim p = (\<lambda> y. simpfm p)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1443
by pat_completeness auto
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1444
termination by (relation "measure fmsize") auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1445
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1446
lemma qelim_ci:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1447
  assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1448
  shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1449
  using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] 
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1450
  by (induct p rule: qelim.induct) (auto simp del: simpfm.simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1451
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1452
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  1453
text {* The @{text "\<int>"} Part *}
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  1454
text{* Linearity for fm where Bound 0 ranges over @{text "\<int>"} *}
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1455
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1456
function zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*) where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1457
  "zsplit0 (C c) = (0,C c)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1458
| "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1459
| "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1460
| "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1461
| "zsplit0 (Neg a) = (let (i',a') =  zsplit0 a in (-i', Neg a'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1462
| "zsplit0 (Add a b) = (let (ia,a') =  zsplit0 a ; 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1463
                            (ib,b') =  zsplit0 b 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1464
                            in (ia+ib, Add a' b'))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1465
| "zsplit0 (Sub a b) = (let (ia,a') =  zsplit0 a ; 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1466
                            (ib,b') =  zsplit0 b 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1467
                            in (ia-ib, Sub a' b'))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1468
| "zsplit0 (Mul i a) = (let (i',a') =  zsplit0 a in (i*i', Mul i a'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1469
| "zsplit0 (Floor a) = (let (i',a') =  zsplit0 a in (i',Floor a'))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1470
by pat_completeness auto
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1471
termination by (relation "measure num_size") auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1472
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1473
lemma zsplit0_I:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1474
  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"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1475
  (is "\<And> n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1476
proof(induct t rule: zsplit0.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1477
  case (1 c n a) thus ?case by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1478
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1479
  case (2 m n a) thus ?case by (cases "m=0") auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1480
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1481
  case (3 n i a n a') thus ?case by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1482
next 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1483
  case (4 c a b n a') thus ?case by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1484
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1485
  case (5 t n a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1486
  let ?nt = "fst (zsplit0 t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1487
  let ?at = "snd (zsplit0 t)"
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 41464
diff changeset
  1488
  have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \<and> n=-?nt" using 5 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1489
    by (simp add: Let_def split_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1490
  from abj 5 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1491
  from th2[simplified] th[simplified] show ?case by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1492
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1493
  case (6 s t n a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1494
  let ?ns = "fst (zsplit0 s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1495
  let ?as = "snd (zsplit0 s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1496
  let ?nt = "fst (zsplit0 t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1497
  let ?at = "snd (zsplit0 t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1498
  have abjs: "zsplit0 s = (?ns,?as)" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1499
  moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1500
  ultimately have th: "a=Add ?as ?at \<and> n=?ns + ?nt" using 6
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1501
    by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1502
  from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1503
  from 6 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 blast (*FIXME*)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1504
  with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1505
  from abjs 6  have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1506
  from th3[simplified] th2[simplified] th[simplified] show ?case 
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49069
diff changeset
  1507
    by (simp add: distrib_right)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1508
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1509
  case (7 s t n a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1510
  let ?ns = "fst (zsplit0 s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1511
  let ?as = "snd (zsplit0 s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1512
  let ?nt = "fst (zsplit0 t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1513
  let ?at = "snd (zsplit0 t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1514
  have abjs: "zsplit0 s = (?ns,?as)" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1515
  moreover have abjt:  "zsplit0 t = (?nt,?at)" by simp 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1516
  ultimately have th: "a=Sub ?as ?at \<and> n=?ns - ?nt" using 7
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1517
    by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1518
  from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1519
  from 7 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 blast (*FIXME*)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1520
  with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1521
  from abjs 7 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1522
  from th3[simplified] th2[simplified] th[simplified] show ?case 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1523
    by (simp add: left_diff_distrib)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1524
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1525
  case (8 i t n a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1526
  let ?nt = "fst (zsplit0 t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1527
  let ?at = "snd (zsplit0 t)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1528
  have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \<and> n=i*?nt" using 8
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1529
    by (simp add: Let_def split_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1530
  from abj 8 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1531
  hence "?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49069
diff changeset
  1532
  also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: distrib_left)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1533
  finally show ?case using th th2 by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1534
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1535
  case (9 t n a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1536
  let ?nt = "fst (zsplit0 t)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1537
  let ?at = "snd (zsplit0 t)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1538
  have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \<and> n=?nt" using 9
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1539
    by (simp add: Let_def split_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1540
  from abj 9 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1541
  hence na: "?N a" using th by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1542
  have th': "(real ?nt)*(real x) = real (?nt * x)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1543
  have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1544
  also have "\<dots> = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1545
  also have "\<dots> = real (floor (?I x ?at + real (?nt* x)))" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1546
  also have "\<dots> = real (floor (?I x ?at) + (?nt* x))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1547
    using floor_add[where x="?I x ?at" and a="?nt* x"] by simp 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1548
  also have "\<dots> = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1549
  finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1550
  with na show ?case by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1551
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1552
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1553
consts
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1554
  iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool"   (* Linearity test for fm *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1555
  zlfm :: "fm \<Rightarrow> fm"       (* Linearity transformation for fm *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1556
recdef iszlfm "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1557
  "iszlfm (And p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1558
  "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1559
  "iszlfm (Eq  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1560
  "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1561
  "iszlfm (Lt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1562
  "iszlfm (Le  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1563
  "iszlfm (Gt  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1564
  "iszlfm (Ge  (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1565
  "iszlfm (Dvd i (CN 0 c e)) = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1566
                 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1567
  "iszlfm (NDvd i (CN 0 c e))= 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1568
                 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1569
  "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1570
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1571
lemma zlin_qfree: "iszlfm p bs \<Longrightarrow> qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1572
  by (induct p rule: iszlfm.induct) auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1573
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1574
lemma iszlfm_gen:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1575
  assumes lp: "iszlfm p (x#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1576
  shows "\<forall> y. iszlfm p (y#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1577
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1578
  fix y
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1579
  show "iszlfm p (y#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1580
    using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1581
  by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1582
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1583
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1584
lemma conj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (conj p q) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1585
  using conj_def by (cases p,auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1586
lemma disj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (disj p q) bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1587
  using disj_def by (cases p,auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1588
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1589
recdef zlfm "measure fmsize"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1590
  "zlfm (And p q) = conj (zlfm p) (zlfm q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1591
  "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1592
  "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1593
  "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1594
  "zlfm (Lt a) = (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1595
     if c=0 then Lt r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1596
     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))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1597
     else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1598
  "zlfm (Le a) = (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1599
     if c=0 then Le r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1600
     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))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1601
     else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1602
  "zlfm (Gt a) = (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1603
     if c=0 then Gt r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1604
     if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1605
     else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1606
  "zlfm (Ge a) = (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1607
     if c=0 then Ge r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1608
     if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1609
     else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1610
  "zlfm (Eq a) = (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1611
              if c=0 then Eq r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1612
      if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r)))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1613
      else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1614
  "zlfm (NEq a) = (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1615
              if c=0 then NEq r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1616
      if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r)))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1617
      else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1618
  "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1619
  else (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1620
              if c=0 then Dvd (abs i) r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1621
      if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1622
      else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1623
  "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1624
  else (let (c,r) = zsplit0 a in 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1625
              if c=0 then NDvd (abs i) r else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1626
      if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1627
      else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1628
  "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1629
  "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1630
  "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1631
  "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1632
  "zlfm (NOT (NOT p)) = zlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1633
  "zlfm (NOT T) = F"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1634
  "zlfm (NOT F) = T"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1635
  "zlfm (NOT (Lt a)) = zlfm (Ge a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1636
  "zlfm (NOT (Le a)) = zlfm (Gt a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1637
  "zlfm (NOT (Gt a)) = zlfm (Le a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1638
  "zlfm (NOT (Ge a)) = zlfm (Lt a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1639
  "zlfm (NOT (Eq a)) = zlfm (NEq a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1640
  "zlfm (NOT (NEq a)) = zlfm (Eq a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1641
  "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1642
  "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1643
  "zlfm p = p" (hints simp add: fmsize_pos)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1644
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1645
lemma split_int_less_real: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1646
  "(real (a::int) < b) = (a < floor b \<or> (a = floor b \<and> real (floor b) < b))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1647
proof( auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1648
  assume alb: "real a < b" and agb: "\<not> a < floor b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1649
  from agb have "floor b \<le> a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1650
  from floor_eq[OF alb th] show "a= floor b" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1651
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1652
  assume alb: "a < floor b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1653
  hence "real a < real (floor b)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1654
  moreover have "real (floor b) \<le> b" by simp ultimately show  "real a < b" by arith 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1655
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1656
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1657
lemma split_int_less_real': 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1658
  "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \<or> (real a - real (floor (-b)) = 0 \<and> real (floor (-b)) + b < 0))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1659
proof- 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1660
  have "(real a + b <0) = (real a < -b)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1661
  with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1662
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1663
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1664
lemma split_int_gt_real': 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1665
  "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \<or> (real a + real (floor b) = 0 \<and> real (floor b) - b < 0))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1666
proof- 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1667
  have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  1668
  show ?thesis using myless[of _ "real (floor b)"] 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1669
    by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1670
    (simp add: algebra_simps,arith)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1671
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1672
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1673
lemma split_int_le_real: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1674
  "(real (a::int) \<le> b) = (a \<le> floor b \<or> (a = floor b \<and> real (floor b) < b))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1675
proof( auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1676
  assume alb: "real a \<le> b" and agb: "\<not> a \<le> floor b"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 30042
diff changeset
  1677
  from alb have "floor (real a) \<le> floor b " by (simp only: floor_mono) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1678
  hence "a \<le> floor b" by simp with agb show "False" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1679
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1680
  assume alb: "a \<le> floor b"
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 30042
diff changeset
  1681
  hence "real a \<le> real (floor b)" by (simp only: floor_mono)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1682
  also have "\<dots>\<le> b" by simp  finally show  "real a \<le> b" . 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1683
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1684
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1685
lemma split_int_le_real': 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1686
  "(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))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1687
proof- 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1688
  have "(real a + b \<le>0) = (real a \<le> -b)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1689
  with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1690
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1691
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1692
lemma split_int_ge_real': 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1693
  "(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))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1694
proof- 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1695
  have th: "(real a + b \<ge>0) = (real (-a) + (-b) \<le> 0)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1696
  show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"])
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  1697
    (simp add: algebra_simps ,arith)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1698
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1699
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1700
lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \<and> b = real (floor b))" (is "?l = ?r")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1701
by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1702
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1703
lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \<and> real (floor (-b)) + b = 0)" (is "?l = ?r")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1704
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1705
  have "?l = (real a = -b)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1706
  with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1707
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1708
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1709
lemma zlfm_I:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1710
  assumes qfp: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1711
  shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \<and> iszlfm (zlfm p) (real (i::int) #bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1712
  (is "(?I (?l p) = ?I p) \<and> ?L (?l p)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1713
using qfp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1714
proof(induct p rule: zlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1715
  case (5 a) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1716
  let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1717
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1718
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1719
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1720
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1721
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1722
  have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1723
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1724
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1725
      by (cases "?r", simp_all add: Let_def split_def,case_tac "nat", simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1726
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1727
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1728
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1729
    have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  1730
    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)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1731
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1732
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1733
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1734
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1735
    have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1736
    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 ac_simps, arith)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1737
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1738
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1739
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1740
  case (6 a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1741
  let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1742
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1743
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1744
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1745
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1746
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1747
  have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1748
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1749
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1750
      by (cases "?r", simp_all add: Let_def split_def, case_tac "nat",simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1751
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1752
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1753
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1754
    have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  1755
    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)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1756
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1757
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1758
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1759
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1760
    have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1761
    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 ac_simps, arith)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1762
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1763
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1764
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1765
  case (7 a) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1766
  let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1767
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1768
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1769
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1770
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1771
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1772
  have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1773
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1774
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1775
      by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1776
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1777
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1778
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1779
    have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  1780
    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)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1781
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1782
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1783
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1784
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1785
    have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1786
    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 ac_simps, arith)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1787
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1788
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1789
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1790
  case (8 a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1791
   let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1792
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1793
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1794
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1795
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1796
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1797
  have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1798
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1799
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1800
      by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1801
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1802
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1803
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1804
    have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  1805
    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)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1806
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1807
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1808
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1809
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1810
    have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1811
    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 ac_simps, arith)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1812
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1813
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1814
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1815
  case (9 a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1816
  let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1817
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1818
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1819
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1820
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1821
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1822
  have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1823
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1824
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1825
      by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1826
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1827
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1828
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1829
    have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1830
    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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1831
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1832
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1833
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1834
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1835
    have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1836
    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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1837
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1838
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1839
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1840
  case (10 a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1841
  let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1842
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1843
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1844
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1845
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1846
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1847
  have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1848
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1849
  {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1850
      by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1851
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1852
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1853
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1854
    have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1855
    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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1856
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1857
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1858
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1859
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1860
    have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1861
    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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1862
    finally have ?case using l by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1863
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1864
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1865
  case (11 j a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1866
  let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1867
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1868
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1869
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1870
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1871
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1872
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1873
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1874
  { assume j: "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1875
    hence ?case using 11 j by (simp del: zlfm.simps add: rdvd_left_0_eq)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1876
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1877
  {assume "?c=0" and "j\<noteq>0" hence ?case 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1878
      using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1879
      by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1880
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1881
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1882
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1883
    have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1884
      using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1885
    also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1886
      by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1887
    also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1888
       (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1889
      by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1890
    also have "\<dots> = (?I (?l (Dvd j a)))" using cp cnz jnz  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1891
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1892
        del: real_of_int_mult) (auto simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1893
    finally have ?case using l jnz  by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1894
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1895
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1896
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1897
    have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1898
      using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1899
    also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1900
      by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1901
    also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1902
       (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1903
      by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1904
    also have "\<dots> = (?I (?l (Dvd j a)))" using cn cnz jnz
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1905
      using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1906
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1907
        del: real_of_int_mult) (auto simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1908
    finally have ?case using l jnz by blast }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1909
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1910
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1911
  case (12 j a)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1912
  let ?c = "fst (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1913
  let ?r = "snd (zsplit0 a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1914
  have spl: "zsplit0 a = (?c,?r)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1915
  from zsplit0_I[OF spl, where x="i" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1916
  have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1917
  let ?N = "\<lambda> t. Inum (real i#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1918
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1919
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1920
  {assume j: "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  1921
    hence ?case using 12 j by (simp del: zlfm.simps add: rdvd_left_0_eq)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1922
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1923
  {assume "?c=0" and "j\<noteq>0" hence ?case 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1924
      using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1925
      by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1926
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1927
  {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1928
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1929
    have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1930
      using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1931
    also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1932
      by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1933
    also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1934
       (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1935
      by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1936
    also have "\<dots> = (?I (?l (NDvd j a)))" using cp cnz jnz  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1937
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1938
        del: real_of_int_mult) (auto simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1939
    finally have ?case using l jnz  by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1940
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1941
  {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1942
      by (simp add: nb Let_def split_def isint_Floor isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1943
    have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1944
      using Ia by (simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1945
    also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1946
      by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1947
    also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1948
       (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1949
      by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1950
    also have "\<dots> = (?I (?l (NDvd j a)))" using cn cnz jnz
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1951
      using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1952
      by (simp add: Let_def split_def int_rdvd_iff[symmetric]  
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1953
        del: real_of_int_mult) (auto simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1954
    finally have ?case using l jnz by blast }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1955
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1956
qed auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1957
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  1958
text{* plusinf : Virtual substitution of @{text "+\<infinity>"}
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  1959
       minusinf: Virtual substitution of @{text "-\<infinity>"}
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  1960
       @{text "\<delta>"} Compute lcm @{text "d| Dvd d  c*x+t \<in> p"}
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  1961
       @{text "d_\<delta>"} checks if a given l divides all the ds above*}
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  1962
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1963
fun minusinf:: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1964
  "minusinf (And p q) = conj (minusinf p) (minusinf q)" 
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1965
| "minusinf (Or p q) = disj (minusinf p) (minusinf q)" 
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1966
| "minusinf (Eq  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1967
| "minusinf (NEq (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1968
| "minusinf (Lt  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1969
| "minusinf (Le  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1970
| "minusinf (Gt  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1971
| "minusinf (Ge  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1972
| "minusinf p = p"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1973
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1974
lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1975
  by (induct p rule: minusinf.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1976
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1977
fun plusinf:: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  1978
  "plusinf (And p q) = conj (plusinf p) (plusinf q)" 
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1979
| "plusinf (Or p q) = disj (plusinf p) (plusinf q)" 
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1980
| "plusinf (Eq  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1981
| "plusinf (NEq (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1982
| "plusinf (Lt  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1983
| "plusinf (Le  (CN 0 c e)) = F"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1984
| "plusinf (Gt  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1985
| "plusinf (Ge  (CN 0 c e)) = T"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1986
| "plusinf p = p"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1987
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1988
fun \<delta> :: "fm \<Rightarrow> int" where
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  1989
  "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)" 
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1990
| "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)" 
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1991
| "\<delta> (Dvd i (CN 0 c e)) = i"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1992
| "\<delta> (NDvd i (CN 0 c e)) = i"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1993
| "\<delta> p = 1"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  1994
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  1995
fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" where
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  1996
  "d_\<delta> (And p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  1997
| "d_\<delta> (Or p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  1998
| "d_\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  1999
| "d_\<delta> (NDvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2000
| "d_\<delta> p = (\<lambda> d. True)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2001
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2002
lemma delta_mono: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2003
  assumes lin: "iszlfm p bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2004
  and d: "d dvd d'"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2005
  and ad: "d_\<delta> p d"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2006
  shows "d_\<delta> p d'"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2007
  using lin ad d
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2008
proof(induct p rule: iszlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2009
  case (9 i c e)  thus ?case using d
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2010
    by (simp add: dvd_trans[of "i" "d" "d'"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2011
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2012
  case (10 i c e) thus ?case using d
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2013
    by (simp add: dvd_trans[of "i" "d" "d'"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2014
qed simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2015
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2016
lemma \<delta> : assumes lin:"iszlfm p bs"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2017
  shows "d_\<delta> p (\<delta> p) \<and> \<delta> p >0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2018
using lin
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2019
proof (induct p rule: iszlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2020
  case (1 p q) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2021
  let ?d = "\<delta> (And p q)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2022
  from 1 lcm_pos_int have dp: "?d >0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2023
  have d1: "\<delta> p dvd \<delta> (And p q)" using 1 by simp 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2024
  hence th: "d_\<delta> p ?d" 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2025
    using delta_mono 1 by (simp only: iszlfm.simps) blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2026
  have "\<delta> q dvd \<delta> (And p q)" using 1 by simp 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2027
  hence th': "d_\<delta> q ?d" using delta_mono 1 by (simp only: iszlfm.simps) blast
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  2028
  from th th' dp show ?case by simp 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2029
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2030
  case (2 p q)  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2031
  let ?d = "\<delta> (And p q)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2032
  from 2 lcm_pos_int have dp: "?d >0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2033
  have "\<delta> p dvd \<delta> (And p q)" using 2 by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2034
  hence th: "d_\<delta> p ?d" using delta_mono 2 by (simp only: iszlfm.simps) blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2035
  have "\<delta> q dvd \<delta> (And p q)" using 2 by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2036
  hence th': "d_\<delta> q ?d" using delta_mono 2 by (simp only: iszlfm.simps) blast
31730
d74830dc3e4a added lemmas; tuned
nipkow
parents: 31706
diff changeset
  2037
  from th th' dp show ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2038
qed simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2039
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2040
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2041
lemma minusinf_inf:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2042
  assumes linp: "iszlfm p (a # bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2043
  shows "\<exists> (z::int). \<forall> x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2044
  (is "?P p" is "\<exists> (z::int). \<forall> x < z. ?I x (?M p) = ?I x p")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2045
using linp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2046
proof (induct p rule: minusinf.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2047
  case (1 f g)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2048
  then have "?P f" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2049
  then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2050
  with 1 have "?P g" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2051
  then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2052
  let ?z = "min z1 z2"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2053
  from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2054
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2055
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2056
  case (2 f g)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2057
  then have "?P f" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2058
  then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2059
  with 2 have "?P g" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2060
  then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2061
  let ?z = "min z1 z2"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2062
  from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2063
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2064
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2065
  case (3 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2066
  then have "c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2067
  hence rcpos: "real c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2068
  from 3 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2069
  fix y
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2070
  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))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2071
  proof (simp add: less_floor_eq , rule allI, rule impI) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2072
    fix x :: int
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2073
    assume A: "real x + 1 \<le> - (Inum (y # bs) e / real c)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2074
    hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2075
    with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2076
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2077
    hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2078
    thus "real c * real x + Inum (real x # bs) e \<noteq> 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2079
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"]  by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2080
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2081
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2082
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2083
  case (4 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2084
  then have "c > 0" by simp hence rcpos: "real c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2085
  from 4 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2086
  fix y
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2087
  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))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2088
  proof (simp add: less_floor_eq , rule allI, rule impI) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2089
    fix x :: int
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2090
    assume A: "real x + 1 \<le> - (Inum (y # bs) e / real c)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2091
    hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2092
    with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2093
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2094
    hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos  by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2095
    thus "real c * real x + Inum (real x # bs) e \<noteq> 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2096
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"]  by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2097
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2098
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2099
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2100
  case (5 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2101
  then have "c > 0" by simp hence rcpos: "real c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2102
  from 5 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2103
  fix y
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2104
  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))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2105
  proof (simp add: less_floor_eq , rule allI, rule impI) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2106
    fix x :: int
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2107
    assume A: "real x + 1 \<le> - (Inum (y # bs) e / real c)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2108
    hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2109
    with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2110
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2111
    thus "real c * real x + Inum (real x # bs) e < 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2112
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2113
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2114
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2115
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2116
  case (6 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2117
  then have "c > 0" by simp hence rcpos: "real c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2118
  from 6 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2119
  fix y
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2120
  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))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2121
  proof (simp add: less_floor_eq , rule allI, rule impI) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2122
    fix x :: int
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2123
    assume A: "real x + 1 \<le> - (Inum (y # bs) e / real c)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2124
    hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2125
    with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2126
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2127
    thus "real c * real x + Inum (real x # bs) e \<le> 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2128
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2129
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2130
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2131
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2132
  case (7 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2133
  then have "c > 0" by simp hence rcpos: "real c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2134
  from 7 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2135
  fix y
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2136
  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))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2137
  proof (simp add: less_floor_eq , rule allI, rule impI) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2138
    fix x :: int
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2139
    assume A: "real x + 1 \<le> - (Inum (y # bs) e / real c)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2140
    hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2141
    with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2142
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2143
    thus "\<not> (real c * real x + Inum (real x # bs) e>0)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2144
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2145
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2146
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2147
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2148
  case (8 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2149
  then have "c > 0" by simp hence rcpos: "real c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2150
  from 8 have nbe: "numbound0 e" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2151
  fix y
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2152
  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))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  2153
  proof (simp add: less_floor_eq , rule allI, rule impI) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2154
    fix x :: int
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2155
    assume A: "real x + 1 \<le> - (Inum (y # bs) e / real c)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2156
    hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2157
    with rcpos  have "(real c)*(real  x) < (real c)*(- (Inum (y # bs) e / real c))"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36531
diff changeset
  2158
      by (simp only: mult_strict_left_mono [OF th1 rcpos])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2159
    thus "\<not> real c * real x + Inum (real x # bs) e \<ge> 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2160
      using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2161
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2162
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2163
qed simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2164
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2165
lemma minusinf_repeats:
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2166
  assumes d: "d_\<delta> p d" and linp: "iszlfm p (a # bs)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2167
  shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2168
using linp d
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2169
proof(induct p rule: iszlfm.induct) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2170
  case (9 i c e) hence nbe: "numbound0 e"  and id: "i dvd d" by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2171
    hence "\<exists> k. d=i*k" by (simp add: dvd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2172
    then obtain "di" where di_def: "d=i*di" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2173
    show ?case 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2174
    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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2175
      assume 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2176
        "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2177
      (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2178
      hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2179
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2180
        by (simp add: algebra_simps di_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2181
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2182
        by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2183
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2184
      thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2185
    next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2186
      assume 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2187
        "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2188
      hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2189
      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2190
      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)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2191
      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2192
      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2193
        by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2194
      thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2195
    qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2196
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2197
  case (10 i c e) hence nbe: "numbound0 e"  and id: "i dvd d" by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2198
    hence "\<exists> k. d=i*k" by (simp add: dvd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2199
    then obtain "di" where di_def: "d=i*di" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2200
    show ?case 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2201
    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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2202
      assume 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2203
        "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2204
      (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2205
      hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2206
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2207
        by (simp add: algebra_simps di_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2208
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2209
        by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2210
      hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2211
      thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2212
    next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2213
      assume 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2214
        "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2215
      hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2216
        by (simp add: rdvd_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2217
      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2218
        by simp
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2219
      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2220
        by (simp add: di_def)
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2221
      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2222
        by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2223
      hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2224
        by blast
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2225
      thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2226
        using rdvd_def by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2227
    qed
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  2228
qed (auto simp add: 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)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2229
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2230
lemma minusinf_ex:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2231
  assumes lin: "iszlfm p (real (a::int) #bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2232
  and exmi: "\<exists> (x::int). Ifm (real x#bs) (minusinf p)" (is "\<exists> x. ?P1 x")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2233
  shows "\<exists> (x::int). Ifm (real x#bs) p" (is "\<exists> x. ?P x")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2234
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2235
  let ?d = "\<delta> p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2236
  from \<delta> [OF lin] have dpos: "?d >0" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2237
  from \<delta> [OF lin] have alld: "d_\<delta> p ?d" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2238
  from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P1 x = ?P1 (x - (k * ?d))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2239
  from minusinf_inf[OF lin] have th2:"\<exists> z. \<forall> x. x<z \<longrightarrow> (?P x = ?P1 x)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2240
  from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2241
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2242
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2243
lemma minusinf_bex:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2244
  assumes lin: "iszlfm p (real (a::int) #bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2245
  shows "(\<exists> (x::int). Ifm (real x#bs) (minusinf p)) = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2246
         (\<exists> (x::int)\<in> {1..\<delta> p}. Ifm (real x#bs) (minusinf p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2247
  (is "(\<exists> x. ?P x) = _")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2248
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2249
  let ?d = "\<delta> p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2250
  from \<delta> [OF lin] have dpos: "?d >0" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2251
  from \<delta> [OF lin] have alld: "d_\<delta> p ?d" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2252
  from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P x = ?P (x - (k * ?d))" by simp
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2253
  from periodic_finite_ex[OF dpos th1] show ?thesis by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2254
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2255
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2256
lemma dvd1_eq1: "x >0 \<Longrightarrow> (x::int) dvd 1 = (x = 1)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2257
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2258
consts 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2259
  a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coeffitients of a formula *)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2260
  d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2261
  \<zeta>  :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2262
  \<beta> :: "fm \<Rightarrow> num list"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2263
  \<alpha> :: "fm \<Rightarrow> num list"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2264
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2265
recdef a_\<beta> "measure size"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2266
  "a_\<beta> (And p q) = (\<lambda> k. And (a_\<beta> p k) (a_\<beta> q k))" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2267
  "a_\<beta> (Or p q) = (\<lambda> k. Or (a_\<beta> p k) (a_\<beta> q k))" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2268
  "a_\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2269
  "a_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2270
  "a_\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2271
  "a_\<beta> (Le  (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2272
  "a_\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2273
  "a_\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2274
  "a_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2275
  "a_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2276
  "a_\<beta> p = (\<lambda> k. p)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2277
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2278
recdef d_\<beta> "measure size"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2279
  "d_\<beta> (And p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2280
  "d_\<beta> (Or p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2281
  "d_\<beta> (Eq  (CN 0 c e)) = (\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2282
  "d_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2283
  "d_\<beta> (Lt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2284
  "d_\<beta> (Le  (CN 0 c e)) = (\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2285
  "d_\<beta> (Gt  (CN 0 c e)) = (\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2286
  "d_\<beta> (Ge  (CN 0 c e)) = (\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2287
  "d_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2288
  "d_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2289
  "d_\<beta> p = (\<lambda> k. True)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2290
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2291
recdef \<zeta> "measure size"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  2292
  "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)" 
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  2293
  "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)" 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2294
  "\<zeta> (Eq  (CN 0 c e)) = c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2295
  "\<zeta> (NEq (CN 0 c e)) = c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2296
  "\<zeta> (Lt  (CN 0 c e)) = c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2297
  "\<zeta> (Le  (CN 0 c e)) = c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2298
  "\<zeta> (Gt  (CN 0 c e)) = c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2299
  "\<zeta> (Ge  (CN 0 c e)) = c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2300
  "\<zeta> (Dvd i (CN 0 c e)) = c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2301
  "\<zeta> (NDvd i (CN 0 c e))= c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2302
  "\<zeta> p = 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2303
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2304
recdef \<beta> "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2305
  "\<beta> (And p q) = (\<beta> p @ \<beta> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2306
  "\<beta> (Or p q) = (\<beta> p @ \<beta> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2307
  "\<beta> (Eq  (CN 0 c e)) = [Sub (C -1) e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2308
  "\<beta> (NEq (CN 0 c e)) = [Neg e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2309
  "\<beta> (Lt  (CN 0 c e)) = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2310
  "\<beta> (Le  (CN 0 c e)) = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2311
  "\<beta> (Gt  (CN 0 c e)) = [Neg e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2312
  "\<beta> (Ge  (CN 0 c e)) = [Sub (C -1) e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2313
  "\<beta> p = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2314
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2315
recdef \<alpha> "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2316
  "\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2317
  "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2318
  "\<alpha> (Eq  (CN 0 c e)) = [Add (C -1) e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2319
  "\<alpha> (NEq (CN 0 c e)) = [e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2320
  "\<alpha> (Lt  (CN 0 c e)) = [e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2321
  "\<alpha> (Le  (CN 0 c e)) = [Add (C -1) e]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2322
  "\<alpha> (Gt  (CN 0 c e)) = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2323
  "\<alpha> (Ge  (CN 0 c e)) = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2324
  "\<alpha> p = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2325
consts mirror :: "fm \<Rightarrow> fm"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2326
recdef mirror "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2327
  "mirror (And p q) = And (mirror p) (mirror q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2328
  "mirror (Or p q) = Or (mirror p) (mirror q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2329
  "mirror (Eq  (CN 0 c e)) = Eq (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2330
  "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2331
  "mirror (Lt  (CN 0 c e)) = Gt (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2332
  "mirror (Le  (CN 0 c e)) = Ge (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2333
  "mirror (Gt  (CN 0 c e)) = Lt (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2334
  "mirror (Ge  (CN 0 c e)) = Le (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2335
  "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2336
  "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2337
  "mirror p = p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2338
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2339
lemma mirror_\<alpha>_\<beta>:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2340
  assumes lp: "iszlfm p (a#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2341
  shows "(Inum (real (i::int)#bs)) ` set (\<alpha> p) = (Inum (real i#bs)) ` set (\<beta> (mirror p))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2342
  using lp by (induct p rule: mirror.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2343
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2344
lemma mirror: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2345
  assumes lp: "iszlfm p (a#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2346
  shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p" 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2347
  using lp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2348
proof(induct p rule: iszlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2349
  case (9 j c e)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2350
  have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2351
       (real j rdvd - (real c * real x - Inum (real x # bs) e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2352
    by (simp only: rdvd_minus[symmetric])
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2353
  from 9 th show ?case
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2354
    by (simp add: algebra_simps
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2355
      numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2356
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2357
  case (10 j c e)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2358
  have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2359
       (real j rdvd - (real c * real x - Inum (real x # bs) e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2360
    by (simp only: rdvd_minus[symmetric])
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2361
  from 10 th show  ?case
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2362
    by (simp add: algebra_simps
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2363
      numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  2364
qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2365
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2366
lemma mirror_l: "iszlfm p (a#bs) \<Longrightarrow> iszlfm (mirror p) (a#bs)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2367
  by (induct p rule: mirror.induct) (auto simp add: isint_neg)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2368
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2369
lemma mirror_d_\<beta>: "iszlfm p (a#bs) \<and> d_\<beta> p 1 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2370
  \<Longrightarrow> iszlfm (mirror p) (a#bs) \<and> d_\<beta> (mirror p) 1"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2371
  by (induct p rule: mirror.induct) (auto simp add: isint_neg)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2372
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2373
lemma mirror_\<delta>: "iszlfm p (a#bs) \<Longrightarrow> \<delta> (mirror p) = \<delta> p"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  2374
  by (induct p rule: mirror.induct) auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2375
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2376
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2377
lemma mirror_ex: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2378
  assumes lp: "iszlfm p (real (i::int)#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2379
  shows "(\<exists> (x::int). Ifm (real x#bs) (mirror p)) = (\<exists> (x::int). Ifm (real x#bs) p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2380
  (is "(\<exists> x. ?I x ?mp) = (\<exists> x. ?I x p)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2381
proof(auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2382
  fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2383
  thus "\<exists> x. ?I x p" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2384
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2385
  fix x assume "?I x p" hence "?I (- x) ?mp" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2386
    using mirror[OF lp, where x="- x", symmetric] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2387
  thus "\<exists> x. ?I x ?mp" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2388
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2389
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2390
lemma \<beta>_numbound0: assumes lp: "iszlfm p bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2391
  shows "\<forall> b\<in> set (\<beta> p). numbound0 b"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2392
  using lp by (induct p rule: \<beta>.induct,auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2393
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2394
lemma d_\<beta>_mono: 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2395
  assumes linp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2396
  and dr: "d_\<beta> p l"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2397
  and d: "l dvd l'"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2398
  shows "d_\<beta> p l'"
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2399
using dr linp dvd_trans[of _ "l" "l'", simplified d]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2400
by (induct p rule: iszlfm.induct) simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2401
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2402
lemma \<alpha>_l: assumes lp: "iszlfm p (a#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2403
  shows "\<forall> b\<in> set (\<alpha> p). numbound0 b \<and> isint b (a#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2404
using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2405
by(induct p rule: \<alpha>.induct, auto simp add: isint_add isint_c)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2406
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2407
lemma \<zeta>: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2408
  assumes linp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2409
  shows "\<zeta> p > 0 \<and> d_\<beta> p (\<zeta> p)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2410
using linp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2411
proof(induct p rule: iszlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2412
  case (1 p q)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2413
  then  have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2414
  from 1 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2415
  from 1 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"] 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2416
    d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"] 
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31730
diff changeset
  2417
    dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2418
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2419
  case (2 p q)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2420
  then have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2421
  from 2 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2422
  from 2 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"] 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2423
    d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"] 
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31730
diff changeset
  2424
    dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31730
diff changeset
  2425
qed (auto simp add: lcm_pos_int)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2426
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2427
lemma a_\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d_\<beta> p l" and lp: "l > 0"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2428
  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)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2429
using linp d
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2430
proof (induct p rule: iszlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2431
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2432
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2433
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2434
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2435
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2436
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2437
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2438
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2439
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2440
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2441
    hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\<Colon>real)) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2442
          (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2443
      by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2444
    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2445
    also have "\<dots> = (real c * real x + Inum (real x # bs) e < 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2446
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2447
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2448
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2449
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2450
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2451
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2452
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2453
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2454
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2455
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2456
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2457
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2458
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2459
    hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0\<Colon>real)) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2460
          (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<le> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2461
      by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2462
    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: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2463
    also have "\<dots> = (real c * real x + Inum (real x # bs) e \<le> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2464
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2465
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2466
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2467
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2468
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2469
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2470
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2471
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2472
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2473
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2474
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2475
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2476
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2477
    hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\<Colon>real)) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2478
          (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2479
      by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2480
    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2481
    also have "\<dots> = (real c * real x + Inum (real x # bs) e > 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2482
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2483
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2484
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2485
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2486
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2487
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2488
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2489
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2490
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2491
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2492
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2493
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2494
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2495
    hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0\<Colon>real)) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2496
          (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<ge> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2497
      by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2498
    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: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2499
    also have "\<dots> = (real c * real x + Inum (real x # bs) e \<ge> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2500
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2501
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2502
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2503
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2504
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2505
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2506
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2507
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2508
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2509
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2510
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2511
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2512
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2513
    hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\<Colon>real)) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2514
          (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2515
      by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2516
    also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2517
    also have "\<dots> = (real c * real x + Inum (real x # bs) e = 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2518
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2519
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2520
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2521
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2522
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2523
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2524
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2525
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2526
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2527
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2528
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2529
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2530
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2531
    hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0\<Colon>real)) =
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2532
          (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<noteq> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2533
      by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2534
    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: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2535
    also have "\<dots> = (real c * real x + Inum (real x # bs) e \<noteq> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2536
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2537
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2538
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2539
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2540
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2541
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2542
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2543
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2544
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2545
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2546
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2547
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2548
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2549
    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
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2550
    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: algebra_simps)
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2551
    also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2552
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2553
  also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2554
  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 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2555
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2556
  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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2557
    from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2558
    from cp have cnz: "c \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2559
    have "c div c\<le> l div c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2560
      by (simp add: zdiv_mono1[OF clel cp])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2561
    then have ldcp:"0 < l div c" 
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  2562
      by (simp add: div_self[OF cnz])
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29823
diff changeset
  2563
    have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2564
    hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2565
      by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2566
    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
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2567
    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: algebra_simps)
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  2568
    also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2569
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2570
  also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2571
  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
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  2572
qed (simp_all add: numbound0_I[where bs="bs" and b="real (l * x)" and b'="real x"] isint_Mul del: real_of_int_mult)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2573
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2574
lemma a_\<beta>_ex: assumes linp: "iszlfm p (a#bs)" and d: "d_\<beta> p l" and lp: "l>0"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2575
  shows "(\<exists> x. l dvd x \<and> Ifm (real x #bs) (a_\<beta> p l)) = (\<exists> (x::int). Ifm (real x#bs) p)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2576
  (is "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> x. ?P' x)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2577
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2578
  have "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> (x::int). ?P (l*x))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2579
    using unity_coeff_ex[where l="l" and P="?P", simplified] by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2580
  also have "\<dots> = (\<exists> (x::int). ?P' x)" using a_\<beta>[OF linp d lp] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2581
  finally show ?thesis  . 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2582
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2583
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2584
lemma \<beta>:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2585
  assumes lp: "iszlfm p (a#bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2586
  and u: "d_\<beta> p 1"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2587
  and d: "d_\<delta> p d"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2588
  and dp: "d > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2589
  and nob: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2590
  and p: "Ifm (real x#bs) p" (is "?P x")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2591
  shows "?P (x - d)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2592
using lp u d dp nob p
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2593
proof(induct p rule: iszlfm.induct)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2594
  case (5 c e) hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2595
  with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] 5
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2596
  show ?case by (simp del: real_of_int_minus)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2597
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2598
  case (6 c e)  hence c1: "c=1" and  bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2599
  with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] 6
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2600
  show ?case by (simp del: real_of_int_minus)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2601
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2602
  case (7 c e) hence p: "Ifm (real x #bs) (Gt (CN 0 c e))" and c1: "c=1"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2603
    and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2604
  let ?e = "Inum (real x # bs) e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2605
  from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2606
      numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2607
    by (simp add: isint_iff)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2608
    {assume "real (x-d) +?e > 0" hence ?case using c1 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2609
      numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2610
        by (simp del: real_of_int_minus)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2611
    moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2612
    {assume H: "\<not> real (x-d) + ?e > 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2613
      let ?v="Neg e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2614
      have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp
55584
a879f14b6f95 merged 'List.set' with BNF-generated 'set'
blanchet
parents: 54489
diff changeset
  2615
      from 7(5)[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"]] 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2616
      have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e + real j)" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2617
      from H p have "real x + ?e > 0 \<and> real x + ?e \<le> real d" by (simp add: c1)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2618
      hence "real (x + floor ?e) > real (0::int) \<and> real (x + floor ?e) \<le> real d"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2619
        using ie by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2620
      hence "x + floor ?e \<ge> 1 \<and> x + floor ?e \<le> d"  by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2621
      hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2622
      hence "\<exists> (j::int) \<in> {1 .. d}. real x = real (- floor ?e + j)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2623
        by (simp only: real_of_int_inject) (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2624
      hence "\<exists> (j::int) \<in> {1 .. d}. real x = - ?e + real j" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2625
        by (simp add: ie[simplified isint_iff])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2626
      with nob have ?case by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2627
    ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2628
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2629
  case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2630
    and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2631
    let ?e = "Inum (real x # bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2632
    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"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2633
      by (simp add: isint_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2634
    {assume "real (x-d) +?e \<ge> 0" hence ?case using  c1 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2635
      numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2636
        by (simp del: real_of_int_minus)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2637
    moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2638
    {assume H: "\<not> real (x-d) + ?e \<ge> 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2639
      let ?v="Sub (C -1) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2640
      have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp
55584
a879f14b6f95 merged 'List.set' with BNF-generated 'set'
blanchet
parents: 54489
diff changeset
  2641
      from 8(5)[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"]] 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2642
      have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e - 1 + real j)" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2643
      from H p have "real x + ?e \<ge> 0 \<and> real x + ?e < real d" by (simp add: c1)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2644
      hence "real (x + floor ?e) \<ge> real (0::int) \<and> real (x + floor ?e) < real d"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2645
        using ie by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2646
      hence "x + floor ?e +1 \<ge> 1 \<and> x + floor ?e + 1 \<le> d"  by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2647
      hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e + 1" by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  2648
      hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2649
      hence "\<exists> (j::int) \<in> {1 .. d}. real x= real (- floor ?e - 1 + j)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2650
        by (simp only: real_of_int_inject)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2651
      hence "\<exists> (j::int) \<in> {1 .. d}. real x= - ?e - 1 + real j" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2652
        by (simp add: ie[simplified isint_iff])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2653
      with nob have ?case by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2654
    ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2655
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2656
  case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2657
    and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2658
    let ?e = "Inum (real x # bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2659
    let ?v="(Sub (C -1) e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2660
    have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2661
    from p have "real x= - ?e" by (simp add: c1) with 3(5) show ?case using dp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2662
      by simp (erule ballE[where x="1"],
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2663
        simp_all add:algebra_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2664
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2665
  case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2666
    and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2667
    let ?e = "Inum (real x # bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2668
    let ?v="Neg e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2669
    have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2670
    {assume "real x - real d + Inum ((real (x -d)) # bs) e \<noteq> 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2671
      hence ?case by (simp add: c1)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2672
    moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2673
    {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2674
      hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2675
      hence "real x = - Inum (a # bs) e + real d"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2676
        by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"])
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2677
       with 4(5) have ?case using dp by simp}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2678
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2679
next 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2680
  case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2681
    and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2682
  let ?e = "Inum (real x # bs) e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2683
  from 9 have "isint e (a #bs)"  by simp 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2684
  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"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2685
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2686
  from 9 have id: "j dvd d" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2687
  from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2688
  also have "\<dots> = (j dvd x + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2689
    using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2690
  also have "\<dots> = (j dvd x - d + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2691
    using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2692
  also have "\<dots> = (real j rdvd real (x - d + floor ?e))" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2693
    using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2694
      ie by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2695
  also have "\<dots> = (real j rdvd real x - real d + ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2696
    using ie by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2697
  finally show ?case 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2698
    using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2699
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2700
  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+
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2701
  let ?e = "Inum (real x # bs) e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2702
  from 10 have "isint e (a#bs)"  by simp 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2703
  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"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2704
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2705
  from 10 have id: "j dvd d" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2706
  from c1 ie[symmetric] have "?p x = (\<not> real j rdvd real (x+ floor ?e))" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2707
  also have "\<dots> = (\<not> j dvd x + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2708
    using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2709
  also have "\<dots> = (\<not> j dvd x - d + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2710
    using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2711
  also have "\<dots> = (\<not> real j rdvd real (x - d + floor ?e))" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2712
    using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2713
      ie by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2714
  also have "\<dots> = (\<not> real j rdvd real x - real d + ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2715
    using ie by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2716
  finally show ?case
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2717
    using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2718
qed (auto simp add: numbound0_I[where bs="bs" and b="real (x - d)" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2719
  simp del: real_of_int_diff)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2720
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2721
lemma \<beta>':   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2722
  assumes lp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2723
  and u: "d_\<beta> p 1"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2724
  and d: "d_\<delta> p d"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2725
  and dp: "d > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2726
  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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2727
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2728
  fix x 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2729
  assume nb:"?b" and px: "?P x" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2730
  hence nb2: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2731
    by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2732
  from  \<beta>[OF lp u d dp nb2 px] show "?P (x -d )" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2733
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2734
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2735
lemma \<beta>_int: assumes lp: "iszlfm p bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2736
  shows "\<forall> b\<in> set (\<beta> p). isint b bs"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2737
using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2738
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2739
lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x))
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2740
==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) 
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2741
==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D))))
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2742
==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2743
apply(rule iffI)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2744
prefer 2
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2745
apply(drule minusinfinity)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2746
apply assumption+
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44121
diff changeset
  2747
apply(fastforce)
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2748
apply clarsimp
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2749
apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)")
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2750
apply(frule_tac x = x and z=z in decr_lemma)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2751
apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2752
prefer 2
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2753
apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2754
prefer 2 apply arith
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44121
diff changeset
  2755
 apply fastforce
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2756
apply(drule (1)  periodic_finite_ex)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2757
apply blast
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2758
apply(blast dest:decr_mult_lemma)
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2759
done
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2760
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2761
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2762
theorem cp_thm:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2763
  assumes lp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2764
  and u: "d_\<beta> p 1"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2765
  and d: "d_\<delta> p d"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2766
  and dp: "d > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2767
  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))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2768
  (is "(\<exists> (x::int). ?P (real x)) = (\<exists> j\<in> ?D. ?M j \<or> (\<exists> b\<in> ?B. ?P (?I b + real j)))")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2769
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2770
  from minusinf_inf[OF lp] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2771
  have th: "\<exists>(z::int). \<forall>x<z. ?P (real x) = ?M x" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2772
  let ?B' = "{floor (?I b) | b. b\<in> ?B}"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2773
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2774
  from B[rule_format] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2775
  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))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2776
    by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2777
  also have "\<dots> = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real (floor (?I b) + j)))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2778
  also have"\<dots> = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j)))"  by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2779
  finally have BB': 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2780
    "(\<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)))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2781
    by blast 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2782
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2783
  from minusinf_repeats[OF d lp]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2784
  have th3: "\<forall> x k. ?M x = ?M (x-k*d)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2785
  from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2786
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2787
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2788
    (* Reddy and Loveland *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2789
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2790
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2791
consts 
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  2792
  \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2793
  \<sigma>_\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*)
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2794
  \<alpha>_\<rho> :: "fm \<Rightarrow> (num\<times>int) list"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2795
  a_\<rho> :: "fm \<Rightarrow> int \<Rightarrow> fm"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2796
recdef \<rho> "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2797
  "\<rho> (And p q) = (\<rho> p @ \<rho> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2798
  "\<rho> (Or p q) = (\<rho> p @ \<rho> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2799
  "\<rho> (Eq  (CN 0 c e)) = [(Sub (C -1) e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2800
  "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2801
  "\<rho> (Lt  (CN 0 c e)) = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2802
  "\<rho> (Le  (CN 0 c e)) = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2803
  "\<rho> (Gt  (CN 0 c e)) = [(Neg e, c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2804
  "\<rho> (Ge  (CN 0 c e)) = [(Sub (C (-1)) e, c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2805
  "\<rho> p = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2806
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2807
recdef \<sigma>_\<rho> "measure size"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2808
  "\<sigma>_\<rho> (And p q) = (\<lambda> (t,k). And (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2809
  "\<sigma>_\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2810
  "\<sigma>_\<rho> (Eq  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2811
                                            else (Eq (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2812
  "\<sigma>_\<rho> (NEq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2813
                                            else (NEq (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2814
  "\<sigma>_\<rho> (Lt  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2815
                                            else (Lt (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2816
  "\<sigma>_\<rho> (Le  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2817
                                            else (Le (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2818
  "\<sigma>_\<rho> (Gt  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2819
                                            else (Gt (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2820
  "\<sigma>_\<rho> (Ge  (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2821
                                            else (Ge (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2822
  "\<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)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2823
                                            else (Dvd (i*k) (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2824
  "\<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)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2825
                                            else (NDvd (i*k) (Add (Mul c t) (Mul k e))))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2826
  "\<sigma>_\<rho> p = (\<lambda> (t,k). p)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2827
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2828
recdef \<alpha>_\<rho> "measure size"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2829
  "\<alpha>_\<rho> (And p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2830
  "\<alpha>_\<rho> (Or p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)" 
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2831
  "\<alpha>_\<rho> (Eq  (CN 0 c e)) = [(Add (C -1) e,c)]"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2832
  "\<alpha>_\<rho> (NEq (CN 0 c e)) = [(e,c)]"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2833
  "\<alpha>_\<rho> (Lt  (CN 0 c e)) = [(e,c)]"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2834
  "\<alpha>_\<rho> (Le  (CN 0 c e)) = [(Add (C -1) e,c)]"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2835
  "\<alpha>_\<rho> p = []"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2836
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2837
    (* Simulates normal substituion by modifying the formula see correctness theorem *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2838
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  2839
definition \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2840
  "\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>_\<rho> p (t,k))"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2841
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2842
lemma \<sigma>_\<rho>:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2843
  assumes linp: "iszlfm p (real (x::int)#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2844
  and kpos: "real k > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2845
  and tnb: "numbound0 t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2846
  and tint: "isint t (real x#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2847
  and kdt: "k dvd floor (Inum (b'#bs) t)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2848
  shows "Ifm (real x#bs) (\<sigma>_\<rho> p (t,k)) = 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2849
  (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2850
  (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2851
using linp kpos tnb
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  2852
proof(induct p rule: \<sigma>_\<rho>.induct)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2853
  case (3 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2854
  from 3 have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2855
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2856
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2857
    from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2858
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2859
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2860
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2861
  { assume *: "\<not> k dvd c"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2862
    from kpos have knz': "real k \<noteq> 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2863
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2864
      using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2865
    from assms * 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)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2866
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2867
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2868
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2869
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2870
      also have "\<dots> = (?I ?tk (Eq (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2871
        using nonzero_eq_divide_eq[OF knz',
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2872
            where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2873
          real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2874
          numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  2875
        by (simp add: ti)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2876
      finally have ?case . }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2877
    ultimately show ?case by blast 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2878
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2879
  case (4 c e)  
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2880
  then have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2881
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2882
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2883
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2884
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2885
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2886
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2887
  { assume *: "\<not> k dvd c"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2888
    from kpos have knz': "real k \<noteq> 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2889
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2890
    from assms * 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)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2891
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2892
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2893
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2894
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2895
    also have "\<dots> = (?I ?tk (NEq (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2896
      using nonzero_eq_divide_eq[OF knz',
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2897
          where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2898
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2899
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2900
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2901
    finally have ?case . }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2902
  ultimately show ?case by blast 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2903
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2904
  case (5 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2905
  then have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2906
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2907
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2908
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2909
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2910
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2911
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2912
  { assume *: "\<not> k dvd c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2913
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2914
    from assms * 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)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2915
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2916
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2917
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2918
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2919
    also have "\<dots> = (?I ?tk (Lt (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2920
      using pos_less_divide_eq[OF kpos,
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2921
          where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2922
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2923
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2924
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2925
    finally have ?case . }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2926
  ultimately show ?case by blast 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2927
next
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2928
  case (6 c e)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2929
  then have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2930
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2931
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2932
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2933
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2934
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2935
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2936
  { assume *: "\<not> k dvd c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2937
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2938
    from assms * 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)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2939
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2940
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2941
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2942
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2943
    also have "\<dots> = (?I ?tk (Le (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2944
      using pos_le_divide_eq[OF kpos,
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2945
          where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2946
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2947
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2948
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2949
    finally have ?case . }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2950
  ultimately show ?case by blast 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2951
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2952
  case (7 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2953
  then have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2954
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2955
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2956
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2957
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2958
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2959
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2960
  { assume *: "\<not> k dvd c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2961
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2962
    from assms * 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)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2963
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2964
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2965
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2966
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2967
    also have "\<dots> = (?I ?tk (Gt (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2968
      using pos_divide_less_eq[OF kpos,
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2969
          where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2970
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2971
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2972
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2973
    finally have ?case . }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2974
  ultimately show ?case by blast 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2975
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2976
  case (8 c e)  
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2977
  then have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2978
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2979
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2980
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2981
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2982
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2983
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2984
  { assume *: "\<not> k dvd c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2985
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2986
    from assms * 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)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2987
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2988
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2989
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2990
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2991
    also have "\<dots> = (?I ?tk (Ge (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2992
      using pos_divide_le_eq[OF kpos,
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2993
          where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  2994
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2995
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2996
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2997
    finally have ?case . }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  2998
  ultimately show ?case by blast 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  2999
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3000
  case (9 i c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3001
  then have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3002
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3003
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3004
    from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3005
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3006
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3007
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3008
  { assume *: "\<not> k dvd c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3009
    from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3010
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3011
    from assms * 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)"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3012
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3013
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3014
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3015
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3016
    also have "\<dots> = (?I ?tk (Dvd i (CN 0 c e)))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3017
      using rdvd_mult[OF knz, where n="i"]
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3018
        real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3019
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3020
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3021
    finally have ?case . }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3022
  ultimately show ?case by blast 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3023
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3024
  case (10 i c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3025
  then have cp: "c > 0" and nb: "numbound0 e" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3026
  { assume kdc: "k dvd c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3027
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3028
    from kdc have  ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3029
      numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3030
      numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3031
  moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3032
  { assume *: "\<not> k dvd c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3033
    from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3034
    from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3035
    from assms * 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))"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3036
      using real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3037
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3038
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3039
      by (simp add: ti algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3040
    also have "\<dots> = (?I ?tk (NDvd i (CN 0 c e)))"
46670
e9aa6d151329 removing unnecessary assumptions in RealDef;
bulwahn
parents: 46130
diff changeset
  3041
      using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF kdt]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3042
        numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3043
        numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3044
      by (simp add: ti)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3045
    finally have ?case . }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3046
  ultimately show ?case by blast 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3047
qed (simp_all add: bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3048
  numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"])
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3049
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3050
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3051
lemma \<sigma>_\<rho>_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3052
  shows "bound0 (\<sigma>_\<rho> p (t,k))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3053
  using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3054
  by (induct p rule: iszlfm.induct, auto simp add: nb)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3055
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3056
lemma \<rho>_l:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3057
  assumes lp: "iszlfm p (real (i::int)#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3058
  shows "\<forall> (b,k) \<in> set (\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3059
using lp by (induct p rule: \<rho>.induct, auto simp add: isint_sub isint_neg)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3060
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3061
lemma \<alpha>_\<rho>_l:
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3062
  assumes lp: "iszlfm p (real (i::int)#bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3063
  shows "\<forall> (b,k) \<in> set (\<alpha>_\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3064
using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"]
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3065
 by (induct p rule: \<alpha>_\<rho>.induct, auto)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3066
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3067
lemma \<rho>: assumes lp: "iszlfm p (real (i::int) #bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3068
  and pi: "Ifm (real i#bs) p"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3069
  and d: "d_\<delta> p d"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3070
  and dp: "d > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3071
  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"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3072
  (is "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. _ \<noteq> ?N i e + _")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3073
  shows "Ifm (real(i - d)#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3074
  using lp pi d nob
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3075
proof(induct p rule: iszlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3076
  case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3077
    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"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3078
    by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3079
  from mult_strict_left_mono[OF dp cp]  have one:"1 \<in> {1 .. c*d}" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3080
  from nob[rule_format, where j="1", OF one] pi show ?case by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3081
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3082
  case (4 c e)  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3083
  hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3084
    and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3085
    by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3086
  {assume "real (c*i) \<noteq> - ?N i e + real (c*d)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3087
    with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"]
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3088
    have ?case by (simp add: algebra_simps)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3089
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3090
  {assume pi: "real (c*i) = - ?N i e + real (c*d)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3091
    from mult_strict_left_mono[OF dp cp] have d: "(c*d) \<in> {1 .. c*d}" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3092
    from nob[rule_format, where j="c*d", OF d] pi have ?case by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3093
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3094
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3095
  case (5 c e) hence cp: "c > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3096
  from 5 mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3097
    real_of_int_mult]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3098
  show ?case using 5 dp 
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  3099
    by (simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  3100
      algebra_simps del: mult_pos_pos)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3101
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3102
  case (6 c e) hence cp: "c > 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3103
  from 6 mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3104
    real_of_int_mult]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3105
  show ?case using 6 dp 
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  3106
    by (simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] 
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  3107
      algebra_simps del: mult_pos_pos)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3108
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3109
  case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3110
    and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3111
    and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3112
    by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3113
  let ?fe = "floor (?N i e)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3114
  from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3115
  from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3116
  hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3117
  have "real (c*i) + ?N i e > real (c*d) \<or> real (c*i) + ?N i e \<le> real (c*d)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3118
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3119
  {assume "real (c*i) + ?N i e > real (c*d)" hence ?case
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3120
      by (simp add: algebra_simps 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3121
        numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3122
  moreover 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3123
  {assume H:"real (c*i) + ?N i e \<le> real (c*d)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3124
    with ei[simplified isint_iff] have "real (c*i + ?fe) \<le> real (c*d)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3125
    hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3126
    with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3127
    hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1" 
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  3128
      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff])
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  3129
        (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3130
    with nob  have ?case by blast }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3131
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3132
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3133
  case (8 c e)  hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3134
    and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - 1 - ?N i e + real j"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3135
    and pi: "real (c*i) + ?N i e \<ge> 0" and cp': "real c >0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3136
    by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3137
  let ?fe = "floor (?N i e)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3138
  from pi cp have th:"(real i +?N i e / real c)*real c \<ge> 0" by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3139
  from pi ei[simplified isint_iff] have "real (c*i + ?fe) \<ge> real (0::int)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3140
  hence pi': "c*i + 1 + ?fe \<ge> 1" by (simp only: real_of_int_le_iff[symmetric])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3141
  have "real (c*i) + ?N i e \<ge> real (c*d) \<or> real (c*i) + ?N i e < real (c*d)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3142
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3143
  {assume "real (c*i) + ?N i e \<ge> real (c*d)" hence ?case
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3144
      by (simp add: algebra_simps 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3145
        numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3146
  moreover 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3147
  {assume H:"real (c*i) + ?N i e < real (c*d)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3148
    with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3149
    hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3150
    with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3151
    hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  3152
      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] real_of_one) 
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  3153
        (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3154
    hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3155
      by (simp only: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3156
        hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
  3157
          by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3158
    with nob  have ?case by blast }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3159
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3160
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3161
  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+
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3162
  let ?e = "Inum (real i # bs) e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3163
  from 9 have "isint e (real i #bs)"  by simp 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3164
  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"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3165
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3166
  from 9 have id: "j dvd d" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3167
  from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3168
  also have "\<dots> = (j dvd c*i + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3169
    using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3170
  also have "\<dots> = (j dvd c*i - c*d + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3171
    using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3172
  also have "\<dots> = (real j rdvd real (c*i - c*d + floor ?e))" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3173
    using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3174
      ie by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3175
  also have "\<dots> = (real j rdvd real (c*(i - d)) + ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3176
    using ie by (simp add:algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3177
  finally show ?case 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3178
    using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3179
    by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3180
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3181
  case (10 j c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3182
  hence p: "\<not> (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3183
    by simp+
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3184
  let ?e = "Inum (real i # bs) e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3185
  from 10 have "isint e (real i #bs)"  by simp 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3186
  hence ie: "real (floor ?e) = ?e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3187
    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"]
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3188
    by (simp add: isint_iff)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3189
  from 10 have id: "j dvd d" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3190
  from ie[symmetric] have "?p i = (\<not> (real j rdvd real (c*i+ floor ?e)))" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3191
  also have "\<dots> = Not (j dvd c*i + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3192
    using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3193
  also have "\<dots> = Not (j dvd c*i - c*d + floor ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3194
    using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3195
  also have "\<dots> = Not (real j rdvd real (c*i - c*d + floor ?e))" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3196
    using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3197
      ie by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3198
  also have "\<dots> = Not (real j rdvd real (c*(i - d)) + ?e)" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3199
    using ie by (simp add:algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3200
  finally show ?case 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3201
    using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3202
    by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3203
qed (auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3204
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3205
lemma \<sigma>_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3206
  shows "bound0 (\<sigma> p k t)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3207
  using \<sigma>_\<rho>_nb[OF lp nb] nb by (simp add: \<sigma>_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3208
  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3209
lemma \<rho>':   assumes lp: "iszlfm p (a #bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3210
  and d: "d_\<delta> p d"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3211
  and dp: "d > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3212
  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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3213
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3214
  fix x 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3215
  assume nob1:"?b x" and px: "?P x" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3216
  from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)".
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3217
  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" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3218
  proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3219
    fix e c j assume ecR: "(e,c) \<in> set (\<rho> p)" and jD: "j\<in> {1 .. c*d}"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3220
      and cx: "real (c*x) = Inum (real x#bs) e + real j"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3221
    let ?e = "Inum (real x#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3222
    let ?fe = "floor ?e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3223
    from \<rho>_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3224
      by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3225
    from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3226
    from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3227
    hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3228
    hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3229
    hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3230
    hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3231
    from cx have "(c*x) div c = (?fe + j) div c" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3232
    with cp have "x = (?fe + j) div c" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3233
    with px have th: "?P ((?fe + j) div c)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3234
    from cp have cp': "real c > 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3235
    from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3236
    from nb have nb': "numbound0 (Add e (C j))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3237
    have ji: "isint (C j) (real x#bs)" by (simp add: isint_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3238
    from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" .
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3239
    from th \<sigma>_\<rho>[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric]
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3240
    have "Ifm (real x#bs) (\<sigma>_\<rho> p (Add e (C j), c))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3241
    with rcdej have th: "Ifm (real x#bs) (\<sigma> p c (Add e (C j)))" by (simp add: \<sigma>_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3242
    from th bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3243
    have "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3244
      with ecR jD nob1    show "False" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3245
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3246
  from \<rho>[OF lp' px d dp nob] show "?P (x -d )" . 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3247
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3248
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3249
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3250
lemma rl_thm: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3251
  assumes lp: "iszlfm p (real (i::int)#bs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3252
  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)))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3253
  (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))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3254
    is "?lhs = (?MD \<or> ?RD)"  is "?lhs = ?rhs")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3255
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3256
  let ?d= "\<delta> p"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3257
  from \<delta>[OF lp] have d:"d_\<delta> p ?d" and dp: "?d > 0" by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3258
  { assume H:"?MD" hence th:"\<exists> (x::int). ?MP x" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3259
    from H minusinf_ex[OF lp th] have ?thesis  by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3260
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3261
  { fix e c j assume exR:"(e,c) \<in> ?R" and jD:"j\<in> {1 .. c*?d}" and spx:"?SP c e j"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3262
    from exR \<rho>_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3263
      by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3264
    have "isint (C j) (real i#bs)" by (simp add: isint_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3265
    with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3266
    have eji:"isint (Add e (C j)) (real i#bs)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3267
    from nb have nb': "numbound0 (Add e (C j))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3268
    from spx bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3269
    have spx': "Ifm (real i # bs) (\<sigma> p c (Add e (C j)))" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3270
    from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))" 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3271
      and sr:"Ifm (real i#bs) (\<sigma>_\<rho> p (Add e (C j),c))" by (simp add: \<sigma>_def)+
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3272
    from rcdej eji[simplified isint_iff] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3273
    have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3274
    hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3275
    from cp have cp': "real c > 0" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3276
    from \<sigma>_\<rho>[OF lp cp' nb' eji cdej] spx' have "?P (\<lfloor>Inum (real i # bs) (Add e (C j))\<rfloor> div c)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3277
      by (simp add: \<sigma>_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3278
    hence ?lhs by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3279
    with exR jD spx have ?thesis by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3280
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3281
  { fix x assume px: "?P x" and nob: "\<not> ?RD"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3282
    from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3283
    from \<rho>'[OF lp' d dp, rule_format, OF nob] have th:"\<forall> x. ?P x \<longrightarrow> ?P (x - ?d)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3284
    from minusinf_inf[OF lp] obtain z where z:"\<forall> x<z. ?MP x = ?P x" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3285
    have zp: "abs (x - z) + 1 \<ge> 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3286
    from decr_lemma[OF dp,where x="x" and z="z"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3287
      decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\<exists> x. ?MP x" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3288
    with minusinf_bex[OF lp] px nob have ?thesis by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3289
  ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3290
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3291
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3292
lemma mirror_\<alpha>_\<rho>:   assumes lp: "iszlfm p (a#bs)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  3293
  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))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3294
  using lp
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3295
  by (induct p rule: mirror.induct) (simp_all add: split_def image_Un)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3296
  
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  3297
text {* The @{text "\<real>"} part*}
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  3298
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  3299
text{* Linearity for fm where Bound 0 ranges over @{text "\<real>"}*}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3300
consts
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3301
  isrlfm :: "fm \<Rightarrow> bool"   (* Linearity test for fm *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3302
recdef isrlfm "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3303
  "isrlfm (And p q) = (isrlfm p \<and> isrlfm q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3304
  "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3305
  "isrlfm (Eq  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3306
  "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3307
  "isrlfm (Lt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3308
  "isrlfm (Le  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3309
  "isrlfm (Gt  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3310
  "isrlfm (Ge  (CN 0 c e)) = (c>0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3311
  "isrlfm p = (isatom p \<and> (bound0 p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3312
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  3313
definition fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3314
  "fp p n s j \<equiv> (if n > 0 then 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3315
            (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j)))))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3316
                        (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1))))))))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3317
            else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3318
            (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3319
                        (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3320
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3321
  (* splits the bounded from the unbounded part*)
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3322
function (sequential) rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3323
  "rsplit0 (Bound 0) = [(T,1,C 0)]"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3324
| "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b 
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  3325
              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])"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3326
| "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3327
| "rsplit0 (Neg a) = map (\<lambda> (p,n,s). (p,-n,Neg s)) (rsplit0 a)"
46130
4821af078cd6 prefer concat over foldl append []
haftmann
parents: 45740
diff changeset
  3328
| "rsplit0 (Floor a) = concat (map 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3329
      (\<lambda> (p,n,s). if n=0 then [(p,0,Floor s)]
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3330
          else (map (\<lambda> j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then [0 .. n] else [n .. 0])))
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3331
       (rsplit0 a))"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3332
| "rsplit0 (CN 0 c a) = map (\<lambda> (p,n,s). (p,n+c,s)) (rsplit0 a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3333
| "rsplit0 (CN m c a) = map (\<lambda> (p,n,s). (p,n,CN m c s)) (rsplit0 a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3334
| "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3335
| "rsplit0 (Mul c a) = map (\<lambda> (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3336
| "rsplit0 t = [(T,0,t)]"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3337
by pat_completeness auto
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  3338
termination by (relation "measure num_size") auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3339
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3340
lemma conj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (conj p q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3341
  using conj_def by (cases p, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3342
lemma disj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (disj p q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3343
  using disj_def by (cases p, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3344
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3345
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3346
lemma rsplit0_cs:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3347
  shows "\<forall> (p,n,s) \<in> set (rsplit0 t). 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3348
  (Ifm (x#bs) p \<longrightarrow>  (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3349
  (is "\<forall> (p,n,s) \<in> ?SS t. (?I p \<longrightarrow> ?N t = ?N (CN 0 n s)) \<and> _ \<and> _ ")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3350
proof(induct t rule: rsplit0.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3351
  case (5 a) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3352
  let ?p = "\<lambda> (p,n,s) j. fp p n s j"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3353
  let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3354
  let ?J = "\<lambda> n. if n>0 then [0..n] else [n..0]"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3355
  let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3356
  have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3357
  have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3358
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3359
  have U2': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0}. 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3360
    ?ff (p,n,s) = map (?f(p,n,s)) [0..n]" by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3361
  hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3362
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3363
    set (map (?f(p,n,s)) [0..n])))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3364
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3365
    fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3366
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3367
    thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3368
      by (auto simp add: split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3369
  qed
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3370
  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)) [n..0]"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3371
    by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3372
  hence U3: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3373
    (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n<0} (\<lambda>(p,n,s). set (map (?f(p,n,s)) [n..0])))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3374
      proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3375
    fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3376
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3377
    thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3378
      by (auto simp add: split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3379
  qed
41464
cb2e3e651893 adopting proofs due to new list comprehension to set comprehension simproc
bulwahn
parents: 41413
diff changeset
  3380
  have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))"
46130
4821af078cd6 prefer concat over foldl append []
haftmann
parents: 45740
diff changeset
  3381
    by auto
41464
cb2e3e651893 adopting proofs due to new list comprehension to set comprehension simproc
bulwahn
parents: 41413
diff changeset
  3382
  also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3383
  also have "\<dots> = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3384
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3385
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3386
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3387
    using int_cases[rule_format] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3388
  also have "\<dots> =  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3389
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3390
   (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n>0} (\<lambda>(p,n,s). set(map(?f(p,n,s)) [0..n]))) Un 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3391
   (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3392
    set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3393
  also have "\<dots> =  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3394
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3395
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3396
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
55584
a879f14b6f95 merged 'List.set' with BNF-generated 'set'
blanchet
parents: 54489
diff changeset
  3397
    by (simp only: set_map set_upto set_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3398
  also have "\<dots> =   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3399
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3400
    (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 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3401
    (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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3402
  finally 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3403
  have FS: "?SS (Floor a) =   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3404
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3405
    (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 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3406
    (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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3407
  show ?case
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3408
  proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3409
    fix p n s
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3410
    let ?ths = "(?I p \<longrightarrow> (?N (Floor a) = ?N (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3411
    assume "(\<exists>ba. (p, 0, ba) \<in> set (rsplit0 a) \<and> n = 0 \<and> s = Floor ba) \<or>
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3412
       (\<exists>ab ac ba.
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3413
           (ab, ac, ba) \<in> set (rsplit0 a) \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3414
           0 < ac \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3415
           (\<exists>j. p = fp ab ac ba j \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3416
                n = 0 \<and> s = Add (Floor ba) (C j) \<and> 0 \<le> j \<and> j \<le> ac)) \<or>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3417
       (\<exists>ab ac ba.
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3418
           (ab, ac, ba) \<in> set (rsplit0 a) \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3419
           ac < 0 \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3420
           (\<exists>j. p = fp ab ac ba j \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3421
                n = 0 \<and> s = Add (Floor ba) (C j) \<and> ac \<le> j \<and> j \<le> 0))"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3422
    moreover 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3423
    { fix s'
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3424
      assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3425
      hence ?ths using 5(1) by auto }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3426
    moreover
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3427
    { fix p' n' s' j
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3428
      assume pns: "(p', n', s') \<in> ?SS a" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3429
        and np: "0 < n'" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3430
        and p_def: "p = ?p (p',n',s') j" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3431
        and n0: "n = 0" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3432
        and s_def: "s = (Add (Floor s') (C j))" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3433
        and jp: "0 \<le> j" and jn: "j \<le> n'"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3434
      from 5 pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3435
          Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3436
          numbound0 s' \<and> isrlfm p'" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3437
      hence nb: "numbound0 s'" by simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3438
      from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3439
      let ?nxs = "CN 0 n' s'"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3440
      let ?l = "floor (?N s') + j"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3441
      from H 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3442
      have "?I (?p (p',n',s') j) \<longrightarrow> 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3443
          (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3444
        by (simp add: fp_def np algebra_simps)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3445
      also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3446
        using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3447
      moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3448
      have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3449
      ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3450
        by blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3451
      with s_def n0 p_def nb nf have ?ths by auto}
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3452
    moreover
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3453
    { fix p' n' s' j
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3454
      assume pns: "(p', n', s') \<in> ?SS a" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3455
        and np: "n' < 0" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3456
        and p_def: "p = ?p (p',n',s') j" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3457
        and n0: "n = 0" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3458
        and s_def: "s = (Add (Floor s') (C j))" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3459
        and jp: "n' \<le> j" and jn: "j \<le> 0"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3460
      from 5 pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3461
          Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3462
          numbound0 s' \<and> isrlfm p'" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3463
      hence nb: "numbound0 s'" by simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3464
      from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3465
      let ?nxs = "CN 0 n' s'"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3466
      let ?l = "floor (?N s') + j"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3467
      from H 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3468
      have "?I (?p (p',n',s') j) \<longrightarrow> 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3469
          (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3470
        by (simp add: np fp_def algebra_simps)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3471
      also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3472
        using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3473
      moreover
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3474
      have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"  by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3475
      ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3476
        by blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3477
      with s_def n0 p_def nb nf have ?ths by auto}
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3478
    ultimately show ?ths by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3479
  qed
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3480
next
28741
1b257449f804 simproc for let
haftmann
parents: 28290
diff changeset
  3481
  case (3 a b) then show ?case
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  3482
    by auto
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3483
qed (auto simp add: Let_def split_def algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3484
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3485
lemma real_in_int_intervals: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3486
  assumes xb: "real m \<le> x \<and> x < real ((n::int) + 1)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3487
  shows "\<exists> j\<in> {m.. n}. real j \<le> x \<and> x < real (j+1)" (is "\<exists> j\<in> ?N. ?P j")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3488
by (rule bexI[where P="?P" and x="floor x" and A="?N"]) 
30097
57df8626c23b generalize floor/ceiling to work with real and rat; rename floor_mono2 to floor_mono
huffman
parents: 30042
diff changeset
  3489
(auto simp add: floor_less_eq[where x="x" and a="n+1", simplified] xb[simplified] floor_mono[where x="real m" and y="x", OF conjunct1[OF xb], simplified floor_real_of_int[where n="m"]])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3490
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3491
lemma rsplit0_complete:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3492
  assumes xp:"0 \<le> x" and x1:"x < 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3493
  shows "\<exists> (p,n,s) \<in> set (rsplit0 t). Ifm (x#bs) p" (is "\<exists> (p,n,s) \<in> ?SS t. ?I p")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3494
proof(induct t rule: rsplit0.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3495
  case (2 a b) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3496
  then have "\<exists> (pa,na,sa) \<in> ?SS a. ?I pa" by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3497
  then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\<in> ?SS a \<and> ?I pa" by blast
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3498
  with 2 have "\<exists> (pb,nb,sb) \<in> ?SS b. ?I pb" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3499
  then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\<in> ?SS b \<and> ?I pb" by blast
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  3500
  from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \<in> set[(x,y). x\<leftarrow>rsplit0 a, y\<leftarrow>rsplit0 b]"
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  3501
    by (auto)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3502
  let ?f="(\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3503
  from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \<in> ?SS (Add a b)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3504
    by (simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3505
  hence "(And pa pb, na +nb, Add sa sb) \<in> ?SS (Add a b)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3506
  moreover from pa pb have "?I (And pa pb)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3507
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3508
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3509
  case (5 a) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3510
  let ?p = "\<lambda> (p,n,s) j. fp p n s j"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3511
  let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3512
  let ?J = "\<lambda> n. if n>0 then [0..n] else [n..0]"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3513
  let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3514
  have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3515
  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
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3516
  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)) [0..n]"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3517
    by auto
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3518
  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)) [0..n])))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3519
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3520
    fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3521
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3522
    thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3523
      by (auto simp add: split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3524
  qed
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3525
  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)) [n..0]"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3526
    by auto
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3527
  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)) [n..0])))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3528
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3529
    fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3530
    assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3531
    thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3532
      by (auto simp add: split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3533
  qed
24473
acd19ea21fbb fixed Proofs
chaieb
parents: 24348
diff changeset
  3534
46130
4821af078cd6 prefer concat over foldl append []
haftmann
parents: 45740
diff changeset
  3535
  have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" by auto
41464
cb2e3e651893 adopting proofs due to new list comprehension to set comprehension simproc
bulwahn
parents: 41413
diff changeset
  3536
  also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3537
  also have "\<dots> = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3538
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3539
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3540
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3541
    using int_cases[rule_format] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3542
  also have "\<dots> =  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3543
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3544
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) [0..n]))) Un 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3545
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) [n..0]))))"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3546
    by (simp only: U1 U2 U3)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3547
  also have "\<dots> =  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3548
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3549
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3550
    (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
55584
a879f14b6f95 merged 'List.set' with BNF-generated 'set'
blanchet
parents: 54489
diff changeset
  3551
    by (simp only: set_map set_upto set_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3552
  also have "\<dots> =   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3553
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3554
    (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 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3555
    (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}})))"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3556
    by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3557
  finally 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3558
  have FS: "?SS (Floor a) =   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3559
    ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3560
    (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 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3561
    (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}})))"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3562
    by blast
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3563
  from 5 have "\<exists> (p,n,s) \<in> ?SS a. ?I p" by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3564
  then obtain "p" "n" "s" where pns: "(p,n,s) \<in> ?SS a \<and> ?I p" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3565
  let ?N = "\<lambda> t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3566
  from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \<and> numbound0 s \<and> isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3567
    by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3568
  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3569
  have "n=0 \<or> n >0 \<or> n <0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3570
  moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3571
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3572
  {
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3573
    assume np: "n > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3574
    from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \<le> ?N s" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3575
    also from mult_left_mono[OF xp] np have "?N s \<le> real n * x + ?N s" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3576
    finally have "?N (Floor s) \<le> real n * x + ?N s" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3577
    moreover
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30439
diff changeset
  3578
    {from x1 np have "real n *x + ?N s < real n + ?N s" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3579
      also from real_of_int_floor_add_one_gt[where r="?N s"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3580
      have "\<dots> < real n + ?N (Floor s) + 1" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3581
      finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3582
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3583
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3584
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3585
    
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3586
    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"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3587
      by(simp only: myle[of _ "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)"]) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3588
    hence "\<exists> j\<in> {0.. n}. ?I (?p (p,n,s) j)"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3589
      using pns by (simp add: fp_def np algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3590
    then obtain "j" where j_def: "j\<in> {0 .. n} \<and> ?I (?p (p,n,s) j)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3591
    hence "\<exists>x \<in> {?p (p,n,s) j |j. 0\<le> j \<and> j \<le> n }. ?I x" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3592
    hence ?case using pns 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3593
      by (simp only: FS,simp add: bex_Un) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3594
    (rule disjI2, rule disjI1,rule exI [where x="p"],
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3595
      rule exI [where x="n"],rule exI [where x="s"],simp_all add: np)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3596
  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3597
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3598
  { assume nn: "n < 0" hence np: "-n >0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3599
    from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3600
    moreover from mult_left_mono_neg[OF xp] nn have "?N s \<ge> real n * x + ?N s" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3601
    ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3602
    moreover
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30439
diff changeset
  3603
    {from x1 nn have "real n *x + ?N s \<ge> real n + ?N s" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3604
      moreover from real_of_int_floor_le[where r="?N s"]  have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3605
      ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3606
        by (simp only: algebra_simps)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3607
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3608
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3609
    have th1: "\<forall> (a::real). (- a > 0) = (a < 0)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3610
    have th2: "\<forall> (a::real). (0 \<ge> - a) = (a \<ge> 0)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3611
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3612
    
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3613
    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"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3614
      by(simp only: myle[of _ "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)"]) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3615
    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])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3616
    hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  3617
      using pns by (simp add: fp_def nn algebra_simps
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3618
        del: diff_less_0_iff_less diff_le_0_iff_le) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3619
    then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3620
    hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3621
    hence ?case using pns 
23464
bc2563c37b1a tuned proofs -- avoid implicit prems;
wenzelm
parents: 23316
diff changeset
  3622
      by (simp only: FS,simp add: bex_Un)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3623
    (rule disjI2, rule disjI2,rule exI [where x="p"],
23464
bc2563c37b1a tuned proofs -- avoid implicit prems;
wenzelm
parents: 23316
diff changeset
  3624
      rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3625
  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3626
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3627
qed (auto simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3628
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3629
    (* Linearize a formula where Bound 0 ranges over [0,1) *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3630
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  3631
definition rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3632
  "rsplit f a \<equiv> foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) (rsplit0 a)) F"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3633
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3634
lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\<exists> x \<in> set xs. Ifm bs (f x))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3635
by(induct xs, simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3636
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3637
lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\<forall> x \<in> set xs. Ifm bs (f x))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3638
by(induct xs, simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3639
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3640
lemma foldr_disj_map_rlfm: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3641
  assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3642
  and \<phi>: "\<forall> (\<phi>,n,s) \<in> set xs. numbound0 s \<and> isrlfm \<phi>"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3643
  shows "isrlfm (foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) xs) F)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3644
using lf \<phi> by (induct xs, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3645
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3646
lemma rsplit_ex: "Ifm bs (rsplit f a) = (\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). Ifm bs (conj \<phi> (f n s)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3647
using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3648
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3649
lemma rsplit_l: assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3650
  shows "isrlfm (rsplit f a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3651
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3652
  from rsplit0_cs[where t="a"] have th: "\<forall> (\<phi>,n,s) \<in> set (rsplit0 a). numbound0 s \<and> isrlfm \<phi>" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3653
  from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3654
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3655
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3656
lemma rsplit: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3657
  assumes xp: "x \<ge> 0" and x1: "x < 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3658
  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))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3659
  shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3660
proof(auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3661
  let ?I = "\<lambda>x p. Ifm (x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3662
  let ?N = "\<lambda> x t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3663
  assume "?I x (rsplit f a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3664
  hence "\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). ?I x (And \<phi> (f n s))" using rsplit_ex by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3665
  then obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and "?I x (And \<phi> (f n s))" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3666
  hence \<phi>: "?I x \<phi>" and fns: "?I x (f n s)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3667
  from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \<phi> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3668
  have th: "(?N x a = ?N x (CN 0 n s)) \<and> numbound0 s" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3669
  from f[rule_format, OF th] fns show "?I x (g a)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3670
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3671
  let ?I = "\<lambda>x p. Ifm (x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3672
  let ?N = "\<lambda> x t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3673
  assume ga: "?I x (g a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3674
  from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3675
  obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and fx: "?I x \<phi>" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3676
  from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3677
  have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3678
  with ga f have "?I x (f n s)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3679
  with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3680
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3681
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3682
definition lt :: "int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3683
  lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3684
                        else (Gt (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3685
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3686
definition  le :: "int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3687
  le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3688
                        else (Ge (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3689
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3690
definition  gt :: "int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3691
  gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3692
                        else (Lt (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3693
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3694
definition  ge :: "int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3695
  ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3696
                        else (Le (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3697
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3698
definition  eq :: "int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3699
  eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3700
                        else (Eq (CN 0 (-c) (Neg t))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3701
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3702
definition neq :: "int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3703
  neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3704
                        else (NEq (CN 0 (-c) (Neg t))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3705
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3706
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)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3707
  (is "\<forall> a n s . ?N a = ?N (CN 0 n s) \<and> _\<longrightarrow> ?I (lt n s) = ?I (Lt a)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3708
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3709
  fix a n s
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3710
  assume H: "?N a = ?N (CN 0 n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3711
  show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def))
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3712
  (cases "n > 0", simp_all add: lt_def algebra_simps myless[of _ "0"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3713
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3714
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3715
lemma lt_l: "isrlfm (rsplit lt a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3716
  by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def,
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3717
    case_tac s, simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3718
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3719
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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3720
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3721
  fix a n s
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3722
  assume H: "?N a = ?N (CN 0 n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3723
  show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def))
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3724
  (cases "n > 0", simp_all add: le_def algebra_simps myle[of _ "0"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3725
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3726
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3727
lemma le_l: "isrlfm (rsplit le a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3728
  by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3729
(case_tac s, simp_all, case_tac "nat",simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3730
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3731
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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3732
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3733
  fix a n s
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3734
  assume H: "?N a = ?N (CN 0 n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3735
  show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def))
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3736
  (cases "n > 0", simp_all add: gt_def algebra_simps myless[of _ "0"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3737
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3738
lemma gt_l: "isrlfm (rsplit gt a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3739
  by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3740
(case_tac s, simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3741
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3742
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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3743
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3744
  fix a n s 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3745
  assume H: "?N a = ?N (CN 0 n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3746
  show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def))
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3747
  (cases "n > 0", simp_all add: ge_def algebra_simps myle[of _ "0"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3748
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3749
lemma ge_l: "isrlfm (rsplit ge a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3750
  by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3751
(case_tac s, simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3752
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3753
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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3754
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3755
  fix a n s 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3756
  assume H: "?N a = ?N (CN 0 n s)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3757
  show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3758
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3759
lemma eq_l: "isrlfm (rsplit eq a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3760
  by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3761
(case_tac s, simp_all, case_tac"nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3762
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3763
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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3764
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3765
  fix a n s bs
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3766
  assume H: "?N a = ?N (CN 0 n s)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3767
  show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3768
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3769
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3770
lemma neq_l: "isrlfm (rsplit neq a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3771
  by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3772
(case_tac s, simp_all, case_tac"nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3773
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3774
lemma small_le: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3775
  assumes u0:"0 \<le> u" and u1: "u < 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3776
  shows "(-u \<le> real (n::int)) = (0 \<le> n)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3777
using u0 u1  by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3778
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3779
lemma small_lt: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3780
  assumes u0:"0 \<le> u" and u1: "u < 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3781
  shows "(real (n::int) < real (m::int) - u) = (n < m)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3782
using u0 u1  by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3783
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3784
lemma rdvd01_cs: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3785
  assumes up: "u \<ge> 0" and u1: "u<1" and np: "real n > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3786
  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")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3787
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3788
  let ?ss = "s - real (floor s)"
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3789
  from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]] 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3790
    real_of_int_floor_le[where r="s"]  have ss0:"?ss \<ge> 0" and ss1:"?ss < 1" 
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3791
    by (auto simp add: myle[of _ "s", symmetric] myless[of "?ss"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3792
  from np have n0: "real n \<ge> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3793
  from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3794
  have nu0:"real n * u - s \<ge> -s" and nun:"real n * u -s < real n - s" by auto  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3795
  from int_rdvd_real[where i="i" and x="real (n::int) * u - s"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3796
  have "real i rdvd real n * u - s = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3797
    (i dvd floor (real n * u -s) \<and> (real (floor (real n * u - s)) = real n * u - s ))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3798
    (is "_ = (?DE)" is "_ = (?D \<and> ?E)") by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3799
  also have "\<dots> = (?DE \<and> real(floor (real n * u - s) + floor s)\<ge> -?ss 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3800
    \<and> real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \<and>real ?a \<ge> _ \<and> real ?a < _)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3801
    using nu0 nun  by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3802
  also have "\<dots> = (?DE \<and> ?a \<ge> 0 \<and> ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3803
  also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. ?a = j))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3804
  also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. real (\<lfloor>real n * u - s\<rfloor>) = real j - real \<lfloor>s\<rfloor> ))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3805
    by (simp only: algebra_simps real_of_int_diff[symmetric] real_of_int_inject)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3806
  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>"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3807
    by (auto cong: conj_cong)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  3808
  also have "\<dots> = ?rhs" by(simp cong: conj_cong) (simp add: algebra_simps )
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3809
  finally show ?thesis .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3810
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3811
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3812
definition
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3813
  DVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3814
where
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3815
  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))))) [0..n - 1]) F)"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3816
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3817
definition
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3818
  NDVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3819
where
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3820
  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))))) [0..n - 1]) T)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3821
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3822
lemma DVDJ_DVD: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3823
  assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3824
  shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3825
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3826
  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))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3827
  let ?s= "Inum (x#bs) s"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3828
  from foldr_disj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3829
  have "Ifm (x#bs) (DVDJ i n s) = (\<exists> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3830
    by (simp add: np DVDJ_def)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3831
  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)))"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3832
    by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3833
  also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3834
  have "\<dots> = (real i rdvd real n * x - (-?s))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3835
  finally show ?thesis by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3836
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3837
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3838
lemma NDVDJ_NDVD: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3839
  assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3840
  shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3841
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3842
  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))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3843
  let ?s= "Inum (x#bs) s"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3844
  from foldr_conj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3845
  have "Ifm (x#bs) (NDVDJ i n s) = (\<forall> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  3846
    by (simp add: np NDVDJ_def)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3847
  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))))"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  3848
    by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3849
  also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3850
  have "\<dots> = (\<not> (real i rdvd real n * x - (-?s)))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3851
  finally show ?thesis by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3852
qed  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3853
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3854
lemma foldr_disj_map_rlfm2: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3855
  assumes lf: "\<forall> n . isrlfm (f n)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3856
  shows "isrlfm (foldr disj (map f xs) F)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3857
using lf by (induct xs, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3858
lemma foldr_And_map_rlfm2: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3859
  assumes lf: "\<forall> n . isrlfm (f n)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3860
  shows "isrlfm (foldr conj (map f xs) T)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3861
using lf by (induct xs, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3862
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3863
lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3864
  shows "isrlfm (DVDJ i n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3865
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3866
  let ?f="\<lambda>j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3867
                         (Dvd i (Sub (C j) (Floor (Neg s))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3868
  have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3869
  from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3870
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3871
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3872
lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3873
  shows "isrlfm (NDVDJ i n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3874
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3875
  let ?f="\<lambda>j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3876
                      (NDvd i (Sub (C j) (Floor (Neg s))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3877
  have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3878
  from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3879
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3880
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3881
definition DVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3882
  DVD_def: "DVD i c t =
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3883
  (if i=0 then eq c t else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3884
  if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3885
23997
a23d0b4b1c1f Updated proofs;
chaieb
parents: 23993
diff changeset
  3886
definition  NDVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  3887
  "NDVD i c t =
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3888
  (if i=0 then neq c t else 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3889
  if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3890
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3891
lemma DVD_mono: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3892
  assumes xp: "0\<le> x" and x1: "x < 1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3893
  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)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3894
  (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (DVD i n s) = ?I (Dvd i a)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3895
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3896
  fix a n s 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3897
  assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3898
  let ?th = "?I (DVD i n s) = ?I (Dvd i a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3899
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3900
  moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3901
      by (simp add: DVD_def rdvd_left_0_eq)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3902
  moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H DVD_def) } 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3903
  moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3904
      by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3905
        rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3906
  moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3907
  ultimately show ?th by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3908
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3909
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3910
lemma NDVD_mono:   assumes xp: "0\<le> x" and x1: "x < 1" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3911
  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)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3912
  (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (NDVD i n s) = ?I (NDvd i a)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3913
proof(clarify)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3914
  fix a n s 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3915
  assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3916
  let ?th = "?I (NDVD i n s) = ?I (NDvd i a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3917
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3918
  moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3919
      by (simp add: NDVD_def rdvd_left_0_eq)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3920
  moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H NDVD_def) } 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3921
  moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3922
      by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  3923
        rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3924
  moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3925
      by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3926
  ultimately show ?th by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3927
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3928
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3929
lemma DVD_l: "isrlfm (rsplit (DVD i) a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3930
  by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3931
(case_tac s, simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3932
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3933
lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3934
  by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3935
(case_tac s, simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3936
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3937
consts rlfm :: "fm \<Rightarrow> fm"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3938
recdef rlfm "measure fmsize"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3939
  "rlfm (And p q) = conj (rlfm p) (rlfm q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3940
  "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3941
  "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3942
  "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3943
  "rlfm (Lt a) = rsplit lt a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3944
  "rlfm (Le a) = rsplit le a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3945
  "rlfm (Gt a) = rsplit gt a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3946
  "rlfm (Ge a) = rsplit ge a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3947
  "rlfm (Eq a) = rsplit eq a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3948
  "rlfm (NEq a) = rsplit neq a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3949
  "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3950
  "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3951
  "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3952
  "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3953
  "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3954
  "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3955
  "rlfm (NOT (NOT p)) = rlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3956
  "rlfm (NOT T) = F"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3957
  "rlfm (NOT F) = T"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3958
  "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3959
  "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3960
  "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3961
  "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3962
  "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3963
  "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3964
  "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3965
  "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3966
  "rlfm p = p" (hints simp add: fmsize_pos)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3967
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3968
lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3969
  by (induct p rule: isrlfm.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3970
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3971
lemma simpfm_rl: "isrlfm p \<Longrightarrow> isrlfm (simpfm p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3972
proof (induct p)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3973
  case (Lt a) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3974
  hence "bound0 (Lt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3975
    by (cases a,simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3976
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3977
  {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3978
      using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3979
    have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3980
    with bn bound0at_l have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3981
  moreover 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3982
  { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3983
    { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3984
      with numgcd_pos[where t="CN 0 c (simpnum e)"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3985
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3986
      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  3987
        by (simp add: numgcd_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3988
      from `c > 0` have th': "c\<noteq>0" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3989
      from `c > 0` have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  3990
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3991
      have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3992
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  3993
    with Lt a have ?case
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3994
      by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3995
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3996
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3997
  case (Le a)   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3998
  hence "bound0 (Le a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  3999
    by (cases a,simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4000
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4001
  { assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))"  
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4002
      using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4003
    have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4004
    with bn bound0at_l have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4005
  moreover 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4006
  { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4007
    { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4008
      with numgcd_pos[where t="CN 0 c (simpnum e)"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4009
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4010
      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4011
        by (simp add: numgcd_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4012
      from `c > 0` have th': "c\<noteq>0" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4013
      from `c > 0` have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4014
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4015
      have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4016
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4017
    with Le a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4018
      by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4019
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4020
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4021
  case (Gt a)   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4022
  hence "bound0 (Gt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4023
    by (cases a,simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4024
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4025
  {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4026
      using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4027
    have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4028
    with bn bound0at_l have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4029
  moreover 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4030
  { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4031
    { assume cn1: "numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4032
      with numgcd_pos[where t="CN 0 c (simpnum e)"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4033
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4034
      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4035
        by (simp add: numgcd_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4036
      from `c > 0` have th': "c\<noteq>0" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4037
      from `c > 0` have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4038
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4039
      have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4040
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4041
    with Gt a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4042
      by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4043
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4044
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4045
  case (Ge a)   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4046
  hence "bound0 (Ge a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4047
    by (cases a,simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4048
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4049
  { assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))"  
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4050
      using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4051
    have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4052
    with bn bound0at_l have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4053
  moreover 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4054
  { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4055
    { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4056
      with numgcd_pos[where t="CN 0 c (simpnum e)"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4057
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4058
      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4059
        by (simp add: numgcd_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4060
      from `c > 0` have th': "c\<noteq>0" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4061
      from `c > 0` have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4062
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4063
      have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4064
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4065
    with Ge a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4066
      by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4067
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4068
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4069
  case (Eq a)   
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4070
  hence "bound0 (Eq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4071
    by (cases a,simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4072
  moreover
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4073
  { assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))"  
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4074
      using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4075
    have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4076
    with bn bound0at_l have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4077
  moreover 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4078
  { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4079
    { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4080
      with numgcd_pos[where t="CN 0 c (simpnum e)"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4081
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4082
      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4083
        by (simp add: numgcd_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4084
      from `c > 0` have th': "c\<noteq>0" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4085
      from `c > 0` have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4086
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4087
      have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4088
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4089
    with Eq a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4090
      by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4091
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4092
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4093
  case (NEq a)  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4094
  hence "bound0 (NEq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4095
    by (cases a,simp_all, case_tac "nat", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4096
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4097
  {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4098
      using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4099
    have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4100
    with bn bound0at_l have ?case by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4101
  moreover 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4102
  { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4103
    { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4104
      with numgcd_pos[where t="CN 0 c (simpnum e)"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4105
      have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4106
      from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c" 
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4107
        by (simp add: numgcd_def)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4108
      from `c > 0` have th': "c\<noteq>0" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4109
      from `c > 0` have cp: "c \<ge> 0" by simp
47142
d64fa2ca54b8 remove redundant lemmas
huffman
parents: 47108
diff changeset
  4110
      from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4111
      have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4112
    }
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4113
    with NEq a have ?case
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4114
      by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4115
  ultimately show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4116
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4117
  case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4118
    using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4119
  have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4120
  with bn bound0at_l show ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4121
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4122
  case (NDvd i a)  hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4123
    using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4124
  have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4125
  with bn bound0at_l show ?case by blast
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4126
qed(auto simp add: conj_def imp_def disj_def iff_def Let_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4127
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4128
lemma rlfm_I:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4129
  assumes qfp: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4130
  and xp: "0 \<le> x" and x1: "x < 1"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4131
  shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \<and> isrlfm (rlfm p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4132
  using qfp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4133
by (induct p rule: rlfm.induct) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4134
(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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4135
               rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4136
               rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4137
lemma rlfm_l:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4138
  assumes qfp: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4139
  shows "isrlfm (rlfm p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4140
  using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  4141
by (induct p rule: rlfm.induct) (auto simp add: simpfm_rl)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4142
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4143
    (* Operations needed for Ferrante and Rackoff *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4144
lemma rminusinf_inf:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4145
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4146
  shows "\<exists> z. \<forall> x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4147
using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4148
proof (induct p rule: minusinf.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4149
  case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4150
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4151
  case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4152
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4153
  case (3 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4154
  from 3 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4155
  from 3 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4156
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4157
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4158
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4159
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4160
    assume xz: "x < ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4161
    hence "(real c * x < - ?e)" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4162
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4163
    hence "real c * x + ?e < 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4164
    hence "real c * x + ?e \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4165
    with xz have "?P ?z x (Eq (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4166
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4167
  hence "\<forall> x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4168
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4169
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4170
  case (4 c e)   
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4171
  from 4 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4172
  from 4 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4173
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4174
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4175
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4176
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4177
    assume xz: "x < ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4178
    hence "(real c * x < - ?e)" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4179
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4180
    hence "real c * x + ?e < 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4181
    hence "real c * x + ?e \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4182
    with xz have "?P ?z x (NEq (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4183
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4184
  hence "\<forall> x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4185
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4186
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4187
  case (5 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4188
  from 5 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4189
  from 5 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4190
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4191
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4192
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4193
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4194
    assume xz: "x < ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4195
    hence "(real c * x < - ?e)" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4196
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4197
    hence "real c * x + ?e < 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4198
    with xz have "?P ?z x (Lt (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4199
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"]  by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4200
  hence "\<forall> x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4201
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4202
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4203
  case (6 c e)  
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4204
  from 6 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4205
  from 6 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4206
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4207
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4208
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4209
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4210
    assume xz: "x < ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4211
    hence "(real c * x < - ?e)" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4212
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4213
    hence "real c * x + ?e < 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4214
    with xz have "?P ?z x (Le (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4215
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4216
  hence "\<forall> x < ?z. ?P ?z x (Le (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4217
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4218
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4219
  case (7 c e)  
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4220
  from 7 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4221
  from 7 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4222
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4223
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4224
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4225
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4226
    assume xz: "x < ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4227
    hence "(real c * x < - ?e)" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4228
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4229
    hence "real c * x + ?e < 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4230
    with xz have "?P ?z x (Gt (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4231
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4232
  hence "\<forall> x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4233
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4234
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4235
  case (8 c e)  
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4236
  from 8 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4237
  from 8 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4238
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4239
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4240
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4241
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4242
    assume xz: "x < ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4243
    hence "(real c * x < - ?e)" 
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4244
      by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4245
    hence "real c * x + ?e < 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4246
    with xz have "?P ?z x (Ge (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4247
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4248
  hence "\<forall> x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4249
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4250
qed simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4251
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4252
lemma rplusinf_inf:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4253
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4254
  shows "\<exists> z. \<forall> x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4255
using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4256
proof (induct p rule: isrlfm.induct)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4257
  case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4258
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4259
  case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4260
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4261
  case (3 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4262
  from 3 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4263
  from 3 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4264
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4265
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4266
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4267
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4268
    assume xz: "x > ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4269
    with mult_strict_right_mono [OF xz cp] cp
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4270
    have "(real c * x > - ?e)" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4271
    hence "real c * x + ?e > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4272
    hence "real c * x + ?e \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4273
    with xz have "?P ?z x (Eq (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4274
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4275
  hence "\<forall> x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4276
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4277
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4278
  case (4 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4279
  from 4 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4280
  from 4 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4281
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4282
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4283
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4284
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4285
    assume xz: "x > ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4286
    with mult_strict_right_mono [OF xz cp] cp
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4287
    have "(real c * x > - ?e)" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4288
    hence "real c * x + ?e > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4289
    hence "real c * x + ?e \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4290
    with xz have "?P ?z x (NEq (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4291
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4292
  hence "\<forall> x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4293
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4294
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4295
  case (5 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4296
  from 5 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4297
  from 5 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4298
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4299
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4300
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4301
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4302
    assume xz: "x > ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4303
    with mult_strict_right_mono [OF xz cp] cp
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4304
    have "(real c * x > - ?e)" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4305
    hence "real c * x + ?e > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4306
    with xz have "?P ?z x (Lt (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4307
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4308
  hence "\<forall> x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4309
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4310
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4311
  case (6 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4312
  from 6 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4313
  from 6 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4314
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4315
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4316
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4317
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4318
    assume xz: "x > ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4319
    with mult_strict_right_mono [OF xz cp] cp
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4320
    have "(real c * x > - ?e)" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4321
    hence "real c * x + ?e > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4322
    with xz have "?P ?z x (Le (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4323
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4324
  hence "\<forall> x > ?z. ?P ?z x (Le (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4325
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4326
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4327
  case (7 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4328
  from 7 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4329
  from 7 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4330
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4331
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4332
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4333
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4334
    assume xz: "x > ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4335
    with mult_strict_right_mono [OF xz cp] cp
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4336
    have "(real c * x > - ?e)" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4337
    hence "real c * x + ?e > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4338
    with xz have "?P ?z x (Gt (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4339
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4340
  hence "\<forall> x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4341
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4342
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4343
  case (8 c e) 
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4344
  from 8 have nb: "numbound0 e" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4345
  from 8 have cp: "real c > 0" by simp
26932
c398a3866082 avoid undeclared variables within proofs;
wenzelm
parents: 25765
diff changeset
  4346
  fix a
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4347
  let ?e="Inum (a#bs) e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4348
  let ?z = "(- ?e) / real c"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4349
  {fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4350
    assume xz: "x > ?z"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4351
    with mult_strict_right_mono [OF xz cp] cp
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4352
    have "(real c * x > - ?e)" by (simp add: ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4353
    hence "real c * x + ?e > 0" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4354
    with xz have "?P ?z x (Ge (CN 0 c e))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4355
      using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"]   by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4356
  hence "\<forall> x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4357
  thus ?case by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4358
qed simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4359
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4360
lemma rminusinf_bound0:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4361
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4362
  shows "bound0 (minusinf p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4363
  using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4364
  by (induct p rule: minusinf.induct) simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4365
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4366
lemma rplusinf_bound0:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4367
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4368
  shows "bound0 (plusinf p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4369
  using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4370
  by (induct p rule: plusinf.induct) simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4371
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4372
lemma rminusinf_ex:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4373
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4374
  and ex: "Ifm (a#bs) (minusinf p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4375
  shows "\<exists> x. Ifm (x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4376
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4377
  from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4378
  have th: "\<forall> x. Ifm (x#bs) (minusinf p)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4379
  from rminusinf_inf[OF lp, where bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4380
  obtain z where z_def: "\<forall>x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4381
  from th have "Ifm ((z - 1)#bs) (minusinf p)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4382
  moreover have "z - 1 < z" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4383
  ultimately show ?thesis using z_def by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4384
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4385
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4386
lemma rplusinf_ex:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4387
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4388
  and ex: "Ifm (a#bs) (plusinf p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4389
  shows "\<exists> x. Ifm (x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4390
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4391
  from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4392
  have th: "\<forall> x. Ifm (x#bs) (plusinf p)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4393
  from rplusinf_inf[OF lp, where bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4394
  obtain z where z_def: "\<forall>x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4395
  from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4396
  moreover have "z + 1 > z" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4397
  ultimately show ?thesis using z_def by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4398
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4399
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4400
consts 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4401
  \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4402
  \<upsilon> :: "fm \<Rightarrow> (num \<times> int) \<Rightarrow> fm "
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4403
recdef \<Upsilon> "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4404
  "\<Upsilon> (And p q) = (\<Upsilon> p @ \<Upsilon> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4405
  "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4406
  "\<Upsilon> (Eq  (CN 0 c e)) = [(Neg e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4407
  "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4408
  "\<Upsilon> (Lt  (CN 0 c e)) = [(Neg e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4409
  "\<Upsilon> (Le  (CN 0 c e)) = [(Neg e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4410
  "\<Upsilon> (Gt  (CN 0 c e)) = [(Neg e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4411
  "\<Upsilon> (Ge  (CN 0 c e)) = [(Neg e,c)]"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4412
  "\<Upsilon> p = []"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4413
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4414
recdef \<upsilon> "measure size"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4415
  "\<upsilon> (And p q) = (\<lambda> (t,n). And (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4416
  "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4417
  "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4418
  "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4419
  "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4420
  "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4421
  "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4422
  "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4423
  "\<upsilon> p = (\<lambda> (t,n). p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4424
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4425
lemma \<upsilon>_I: assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4426
  and np: "real n > 0" and nbt: "numbound0 t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4427
  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> _")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4428
  using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4429
proof(induct p rule: \<upsilon>.induct)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4430
  case (5 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4431
  from 5 have cp: "c >0" and nb: "numbound0 e" by simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4432
  have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4433
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4434
  also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4435
    by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4436
      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4437
  also have "\<dots> = (real c *?t + ?n* (?N x e) < 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4438
    using np by simp 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4439
  finally show ?case using nbt nb by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4440
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4441
  case (6 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4442
  from 6 have cp: "c >0" and nb: "numbound0 e" by simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4443
  have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<le> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4444
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4445
  also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<le> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4446
    by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4447
      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4448
  also have "\<dots> = (real c *?t + ?n* (?N x e) \<le> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4449
    using np by simp 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4450
  finally show ?case using nbt nb by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4451
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4452
  case (7 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4453
  from 7 have cp: "c >0" and nb: "numbound0 e" by simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4454
  have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4455
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4456
  also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4457
    by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4458
      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4459
  also have "\<dots> = (real c *?t + ?n* (?N x e) > 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4460
    using np by simp 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4461
  finally show ?case using nbt nb by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4462
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4463
  case (8 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4464
  from 8 have cp: "c >0" and nb: "numbound0 e" by simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4465
  have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<ge> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4466
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4467
  also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<ge> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4468
    by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4469
      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4470
  also have "\<dots> = (real c *?t + ?n* (?N x e) \<ge> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4471
    using np by simp 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4472
  finally show ?case using nbt nb by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4473
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4474
  case (3 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4475
  from 3 have cp: "c >0" and nb: "numbound0 e" by simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4476
  from np have np: "real n \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4477
  have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4478
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4479
  also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4480
    by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4481
      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4482
  also have "\<dots> = (real c *?t + ?n* (?N x e) = 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4483
    using np by simp 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4484
  finally show ?case using nbt nb by (simp add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4485
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4486
  case (4 c e)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4487
  from 4 have cp: "c >0" and nb: "numbound0 e" by simp_all
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4488
  from np have np: "real n \<noteq> 0" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4489
  have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<noteq> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4490
    using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4491
  also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4492
    by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4493
      and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4494
  also have "\<dots> = (real c *?t + ?n* (?N x e) \<noteq> 0)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4495
    using np by simp 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4496
  finally show ?case using nbt nb by (simp add: algebra_simps)
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4497
qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4498
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4499
lemma \<Upsilon>_l:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4500
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4501
  shows "\<forall> (t,k) \<in> set (\<Upsilon> p). numbound0 t \<and> k >0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4502
using lp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4503
by(induct p rule: \<Upsilon>.induct)  auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4504
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4505
lemma rminusinf_\<Upsilon>:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4506
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4507
  and nmi: "\<not> (Ifm (a#bs) (minusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4508
  and ex: "Ifm (x#bs) p" (is "?I x p")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4509
  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")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4510
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4511
  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")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4512
    using lp nmi ex
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4513
    by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4514
  then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<ge> ?N a s" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4515
  from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4516
  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" 
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4517
    by (auto simp add: mult.commute)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4518
  thus ?thesis using smU by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4519
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4520
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4521
lemma rplusinf_\<Upsilon>:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4522
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4523
  and nmi: "\<not> (Ifm (a#bs) (plusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4524
  and ex: "Ifm (x#bs) p" (is "?I x p")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4525
  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")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4526
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4527
  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")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4528
    using lp nmi ex
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4529
    by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4530
  then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<le> ?N a s" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4531
  from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4532
  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" 
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4533
    by (auto simp add: mult.commute)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4534
  thus ?thesis using smU by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4535
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4536
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4537
lemma lin_dense: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4538
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4539
  and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (t,n). Inum (x#bs) t / real n) ` set (\<Upsilon> p)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4540
  (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (t,n). ?N x t / real n ) ` (?U p)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4541
  and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4542
  and ly: "l < y" and yu: "y < u"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4543
  shows "Ifm (y#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4544
using lp px noS
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4545
proof (induct p rule: isrlfm.induct)
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4546
  case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4547
  from 5 have "x * real c + ?N x e < 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4548
  hence pxc: "x < (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4549
    by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4550
  from 5 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4551
  with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4552
  hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4553
  moreover {assume y: "y < (-?N x e)/ real c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4554
    hence "y * real c < - ?N x e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4555
      by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4556
    hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4557
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4558
  moreover {assume y: "y > (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4559
    with yu have eu: "u > (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4560
    with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4561
    with lx pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4562
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4563
  ultimately show ?case by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4564
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4565
  case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4566
  from 6 have "x * real c + ?N x e \<le> 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4567
  hence pxc: "x \<le> (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4568
    by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4569
  from 6 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4570
  with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4571
  hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4572
  moreover {assume y: "y < (-?N x e)/ real c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4573
    hence "y * real c < - ?N x e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4574
      by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4575
    hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4576
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4577
  moreover {assume y: "y > (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4578
    with yu have eu: "u > (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4579
    with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4580
    with lx pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4581
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4582
  ultimately show ?case by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4583
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4584
  case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4585
  from 7 have "x * real c + ?N x e > 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4586
  hence pxc: "x > (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4587
    by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4588
  from 7 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4589
  with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4590
  hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4591
  moreover {assume y: "y > (-?N x e)/ real c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4592
    hence "y * real c > - ?N x e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4593
      by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4594
    hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4595
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4596
  moreover {assume y: "y < (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4597
    with ly have eu: "l < (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4598
    with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4599
    with xu pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4600
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4601
  ultimately show ?case by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4602
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4603
  case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4604
  from 8 have "x * real c + ?N x e \<ge> 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4605
  hence pxc: "x \<ge> (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4606
    by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4607
  from 8 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4608
  with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4609
  hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4610
  moreover {assume y: "y > (-?N x e)/ real c"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4611
    hence "y * real c > - ?N x e"
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4612
      by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4613
    hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4614
    hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4615
  moreover {assume y: "y < (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4616
    with ly have eu: "l < (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4617
    with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4618
    with xu pxc have "False" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4619
    hence ?case by simp }
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4620
  ultimately show ?case by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4621
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4622
  case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4623
  from cp have cnz: "real c \<noteq> 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4624
  from 3 have "x * real c + ?N x e = 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4625
  hence pxc: "x = (- ?N x e) / real c" 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4626
    by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"])
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4627
  from 3 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4628
  with lx xu have yne: "x \<noteq> - ?N x e / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4629
  with pxc show ?case by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4630
next
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4631
  case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4632
  from cp have cnz: "real c \<noteq> 0" by simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4633
  from 4 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4634
  with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4635
  hence "y* real c \<noteq> -?N x e"      
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4636
    by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4637
  hence "y* real c + ?N x e \<noteq> 0" by (simp add: algebra_simps)
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4638
  thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] 
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4639
    by (simp add: algebra_simps)
41849
1a65b780bd56 Some cleaning up
nipkow
parents: 41839
diff changeset
  4640
qed (auto simp add: numbound0_I[where bs="bs" and b="y" and b'="x"])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4641
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4642
lemma rinf_\<Upsilon>:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4643
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4644
  and nmi: "\<not> (Ifm (x#bs) (minusinf p))" (is "\<not> (Ifm (x#bs) (?M p))")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4645
  and npi: "\<not> (Ifm (x#bs) (plusinf p))" (is "\<not> (Ifm (x#bs) (?P p))")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4646
  and ex: "\<exists> x.  Ifm (x#bs) p" (is "\<exists> x. ?I x p")
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4647
  shows "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p).
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  4648
    ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4649
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4650
  let ?N = "\<lambda> x t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4651
  let ?U = "set (\<Upsilon> p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4652
  from ex obtain a where pa: "?I a p" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4653
  from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4654
  have nmi': "\<not> (?I a (?M p))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4655
  from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4656
  have npi': "\<not> (?I a (?P p))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4657
  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"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4658
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4659
    let ?M = "(\<lambda> (t,c). ?N a t / real c) ` ?U"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4660
    have fM: "finite ?M" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4661
    from rminusinf_\<Upsilon>[OF lp nmi pa] rplusinf_\<Upsilon>[OF lp npi pa] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4662
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4663
    then obtain "t" "n" "s" "m" where 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4664
      tnU: "(t,n) \<in> ?U" and smU: "(s,m) \<in> ?U" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4665
      and xs1: "a \<le> ?N x s / real m" and tx1: "a \<ge> ?N x t / real n" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4666
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4667
    from tnU have Mne: "?M \<noteq> {}" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4668
    hence Une: "?U \<noteq> {}" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4669
    let ?l = "Min ?M"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4670
    let ?u = "Max ?M"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4671
    have linM: "?l \<in> ?M" using fM Mne by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4672
    have uinM: "?u \<in> ?M" using fM Mne by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4673
    have tnM: "?N a t / real n \<in> ?M" using tnU by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4674
    have smM: "?N a s / real m \<in> ?M" using smU by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4675
    have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4676
    have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4677
    have "?l \<le> ?N a t / real n" using tnM Mne by simp hence lx: "?l \<le> a" using tx by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4678
    have "?N a s / real m \<le> ?u" using smM Mne by simp hence xu: "a \<le> ?u" using xs by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4679
    from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4680
    have "(\<exists> s\<in> ?M. ?I s p) \<or> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4681
      (\<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)" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4682
    moreover { fix u assume um: "u\<in> ?M" and pu: "?I u p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4683
      hence "\<exists> (tu,nu) \<in> ?U. u = ?N a tu / real nu" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4684
      then obtain "tu" "nu" where tuU: "(tu,nu) \<in> ?U" and tuu:"u= ?N a tu / real nu" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4685
      have "(u + u) / 2 = u" by auto with pu tuu 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4686
      have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4687
      with tuU have ?thesis by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4688
    moreover{
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4689
      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"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4690
      then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  4691
        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"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  4692
        by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4693
      from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real t1n" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4694
      then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4695
      from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real t2n" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4696
      then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \<in> ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4697
      from t1x xt2 have t1t2: "t1 < t2" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4698
      let ?u = "(t1 + t2) / 2"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4699
      from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4700
      from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4701
      with t1uU t2uU t1u t2u have ?thesis by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4702
    ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4703
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4704
  then obtain "l" "n" "s"  "m" where lnU: "(l,n) \<in> ?U" and smU:"(s,m) \<in> ?U" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4705
    and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4706
  from lnU smU \<Upsilon>_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4707
  from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4708
    numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4709
  have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4710
  with lnU smU
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4711
  show ?thesis by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4712
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4713
    (* The Ferrante - Rackoff Theorem *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4714
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4715
theorem fr_eq: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4716
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4717
  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))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4718
  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4719
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4720
  assume px: "\<exists> x. ?I x p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4721
  have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4722
  moreover {assume "?M \<or> ?P" hence "?D" by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4723
  moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4724
    from rinf_\<Upsilon>[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4725
  ultimately show "?D" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4726
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4727
  assume "?D" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4728
  moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4729
  moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4730
  moreover {assume f:"?F" hence "?E" by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4731
  ultimately show "?E" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4732
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4733
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4734
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  4735
lemma fr_eq_\<upsilon>: 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4736
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4737
  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))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4738
  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4739
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4740
  assume px: "\<exists> x. ?I x p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4741
  have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4742
  moreover {assume "?M \<or> ?P" hence "?D" by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4743
  moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4744
    let ?f ="\<lambda> (t,n). Inum (x#bs) t / real n"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4745
    let ?N = "\<lambda> t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4746
    {fix t n s m assume "(t,n)\<in> set (\<Upsilon> p)" and "(s,m) \<in> set (\<Upsilon> p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4747
      with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  4748
        by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4749
      let ?st = "Add (Mul m t) (Mul n s)"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4750
      from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult.commute)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4751
      from tnb snb have st_nb: "numbound0 ?st" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4752
      have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31952
diff changeset
  4753
        using mnp mp np by (simp add: algebra_simps add_divide_distrib)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4754
      from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4755
      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])}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4756
    with rinf_\<Upsilon>[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4757
  ultimately show "?D" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4758
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4759
  assume "?D" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4760
  moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4761
  moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4762
  moreover {fix t k s l assume "(t,k) \<in> set (\<Upsilon> p)" and "(s,l) \<in> set (\<Upsilon> p)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4763
    and px:"?I x (\<upsilon> p (Add (Mul l t) (Mul k s), 2*k*l))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4764
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4765
    let ?st = "Add (Mul l t) (Mul k s)"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4766
    from np mp have mnp: "real (2*k*l) > 0" by (simp add: mult.commute)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4767
    from tnb snb have st_nb: "numbound0 ?st" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4768
    from \<upsilon>_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4769
  ultimately show "?E" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4770
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4771
23316
26c978a475de tuned Proof and Document
chaieb
parents: 23264
diff changeset
  4772
text{* The overall Part *}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4773
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4774
lemma real_ex_int_real01:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4775
  shows "(\<exists> (x::real). P x) = (\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4776
proof(auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4777
  fix x
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4778
  assume Px: "P x"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4779
  let ?i = "floor x"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4780
  let ?u = "x - real ?i"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4781
  have "x = real ?i + ?u" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4782
  hence "P (real ?i + ?u)" using Px by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4783
  moreover have "real ?i \<le> x" using real_of_int_floor_le by simp hence "0 \<le> ?u" by arith
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4784
  moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4785
  ultimately show "(\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4786
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4787
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4788
fun exsplitnum :: "num \<Rightarrow> num" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4789
  "exsplitnum (C c) = (C c)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4790
| "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4791
| "exsplitnum (Bound n) = Bound (n+1)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4792
| "exsplitnum (Neg a) = Neg (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4793
| "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) "
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4794
| "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) "
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4795
| "exsplitnum (Mul c a) = Mul c (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4796
| "exsplitnum (Floor a) = Floor (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4797
| "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4798
| "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4799
| "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4800
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4801
fun exsplit :: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4802
  "exsplit (Lt a) = Lt (exsplitnum a)"
41839
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4803
| "exsplit (Le a) = Le (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4804
| "exsplit (Gt a) = Gt (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4805
| "exsplit (Ge a) = Ge (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4806
| "exsplit (Eq a) = Eq (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4807
| "exsplit (NEq a) = NEq (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4808
| "exsplit (Dvd i a) = Dvd i (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4809
| "exsplit (NDvd i a) = NDvd i (exsplitnum a)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4810
| "exsplit (And p q) = And (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4811
| "exsplit (Or p q) = Or (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4812
| "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4813
| "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4814
| "exsplit (NOT p) = NOT (exsplit p)"
421a795cee05 recdef -> fun(ction)
krauss
parents: 41836
diff changeset
  4815
| "exsplit p = p"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4816
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4817
lemma exsplitnum: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4818
  "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4819
  by(induct t rule: exsplitnum.induct) (simp_all add: algebra_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4820
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4821
lemma exsplit: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4822
  assumes qfp: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4823
  shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4824
using qfp exsplitnum[where x="x" and y="y" and bs="bs"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4825
by(induct p rule: exsplit.induct) simp_all
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4826
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4827
lemma splitex:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4828
  assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4829
  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")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4830
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4831
  have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4832
    by (simp add: myless[of _ "1"] myless[of _ "0"] ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4833
  also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4834
    by (simp only: exsplit[OF qf] ac_simps)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4835
  also have "\<dots> = (\<exists> x. Ifm (x#bs) p)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4836
    by (simp only: real_ex_int_real01[where P="\<lambda> x. Ifm (x#bs) p"])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4837
  finally show ?thesis by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4838
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4839
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4840
    (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4841
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  4842
definition ferrack01 :: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4843
  "ferrack01 p \<equiv> (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p);
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4844
                    U = remdups(map simp_num_pair 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4845
                     (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4846
                           (alluopairs (\<Upsilon> p')))) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4847
  in decr (evaldjf (\<upsilon> p') U ))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4848
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4849
lemma fr_eq_01: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4850
  assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4851
  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)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4852
  (is "(\<exists> x. ?I x ?q) = ?F")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4853
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4854
  let ?rq = "rlfm ?q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4855
  let ?M = "?I x (minusinf ?rq)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4856
  let ?P = "?I x (plusinf ?rq)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4857
  have MF: "?M = False"
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  4858
    apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4859
    by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  4860
  have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4861
    by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4862
  have "(\<exists> x. ?I x ?q ) = 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4863
    ((?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))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4864
    (is "(\<exists> x. ?I x ?q) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4865
  proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4866
    assume "\<exists> x. ?I x ?q"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4867
    then obtain x where qx: "?I x ?q" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4868
    hence xp: "0\<le> x" and x1: "x< 1" and px: "?I x p" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4869
      by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4870
    from qx have "?I x ?rq " 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4871
      by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4872
    hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4873
    from qf have qfq:"isrlfm ?rq"  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4874
      by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  4875
    with lqx fr_eq_\<upsilon>[OF qfq] show "?M \<or> ?P \<or> ?F" by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4876
  next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4877
    assume D: "?D"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4878
    let ?U = "set (\<Upsilon> ?rq )"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4879
    from MF PF D have "?F" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4880
    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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4881
    from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4882
      by (auto simp add: rsplit_def lt_def ge_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4883
    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)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4884
    let ?st = "Add (Mul m t) (Mul n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4885
    from tnb snb have stnb: "numbound0 ?st" by simp
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4886
    from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult.commute)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4887
    from conjunct1[OF \<upsilon>_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4888
    have "\<exists> x. ?I x ?rq" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4889
    thus "?E" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4890
      using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4891
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4892
  with MF PF show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4893
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4894
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4895
lemma \<Upsilon>_cong_aux:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4896
  assumes Ul: "\<forall> (t,n) \<in> set U. numbound0 t \<and> n >0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4897
  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))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4898
  (is "?lhs = ?rhs")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4899
proof(auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4900
  fix t n s m
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4901
  assume "((t,n),(s,m)) \<in> set (alluopairs U)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4902
  hence th: "((t,n),(s,m)) \<in> (set U \<times> set U)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4903
    using alluopairs_set1[where xs="U"] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4904
  let ?N = "\<lambda> t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4905
  let ?st= "Add (Mul m t) (Mul n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4906
  from Ul th have mnz: "m \<noteq> 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4907
  from Ul th have  nnz: "n \<noteq> 0" by auto  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4908
  have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4909
   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4910
 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4911
  thus "(real m *  Inum (x # bs) t + real n * Inum (x # bs) s) /
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4912
       (2 * real n * real m)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4913
       \<in> (\<lambda>((t, n), s, m).
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4914
             (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4915
         (set U \<times> set U)"using mnz nnz th  
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4916
    apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4917
    by (rule_tac x="(s,m)" in bexI,simp_all) 
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4918
  (rule_tac x="(t,n)" in bexI,simp_all add: mult.commute)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4919
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4920
  fix t n s m
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4921
  assume tnU: "(t,n) \<in> set U" and smU:"(s,m) \<in> set U" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4922
  let ?N = "\<lambda> t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4923
  let ?st= "Add (Mul m t) (Mul n s)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4924
  from Ul smU have mnz: "m \<noteq> 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4925
  from Ul tnU have  nnz: "n \<noteq> 0" by auto  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4926
  have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4927
   using mnz nnz by (simp add: algebra_simps add_divide_distrib)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4928
 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"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4929
 have Pc:"\<forall> a b. ?P a b = ?P b a"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4930
   by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4931
 from Ul alluopairs_set1 have Up:"\<forall> ((t,n),(s,m)) \<in> set (alluopairs U). n \<noteq> 0 \<and> m \<noteq> 0" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4932
 from alluopairs_ex[OF Pc, where xs="U"] tnU smU
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4933
 have th':"\<exists> ((t',n'),(s',m')) \<in> set (alluopairs U). ?P (t',n') (s',m')"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4934
   by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4935
 then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \<in> set (alluopairs U)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4936
   and Pts': "?P (t',n') (s',m')" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4937
 from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4938
 let ?st' = "Add (Mul m' t') (Mul n' s')"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4939
   have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4940
   using mnz' nnz' by (simp add: algebra_simps add_divide_distrib)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4941
 from Pts' have 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4942
   "(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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4943
 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')
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4944
 finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4945
          \<in> (\<lambda>(t, n). Inum (x # bs) t / real n) `
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4946
            (\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) `
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4947
            set (alluopairs U)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4948
   using ts'_U by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4949
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4950
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4951
lemma \<Upsilon>_cong:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4952
  assumes lp: "isrlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4953
  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)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4954
  and U: "\<forall> (t,n) \<in> U. numbound0 t \<and> n > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4955
  and U': "\<forall> (t,n) \<in> U'. numbound0 t \<and> n > 0"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4956
  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)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4957
  (is "?lhs = ?rhs")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4958
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4959
  assume ?lhs
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4960
  then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4961
    Pst: "Ifm (x#bs) (\<upsilon> p (Add (Mul m t) (Mul n s),2*n*m))" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4962
  let ?N = "\<lambda> t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4963
  from tnU smU U have tnb: "numbound0 t" and np: "n > 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4964
    and snb: "numbound0 s" and mp:"m > 0"  by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4965
  let ?st= "Add (Mul m t) (Mul n s)"
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  4966
  from np mp have mnp: "real (2*n*m) > 0" 
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4967
      by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4968
    from tnb snb have stnb: "numbound0 ?st" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4969
  have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4970
   using mp np by (simp add: algebra_simps add_divide_distrib)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4971
  from tnU smU UU' have "?g ((t,n),(s,m)) \<in> ?f ` U'" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4972
  hence "\<exists> (t',n') \<in> U'. ?g ((t,n),(s,m)) = ?f (t',n')"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4973
    by auto (rule_tac x="(a,b)" in bexI, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4974
  then obtain t' n' where tnU': "(t',n') \<in> U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4975
  from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4976
  from \<upsilon>_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4977
  have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4978
  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]]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4979
  have "Ifm (x # bs) (\<upsilon> p (t', n')) " by (simp only: st) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4980
  then show ?rhs using tnU' by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4981
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4982
  assume ?rhs
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4983
  then obtain t' n' where tnU': "(t',n') \<in> U'" and Pt': "Ifm (x # bs) (\<upsilon> p (t', n'))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4984
    by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4985
  from tnU' UU' have "?f (t',n') \<in> ?g ` (U\<times>U)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4986
  hence "\<exists> ((t,n),(s,m)) \<in> (U\<times>U). ?f (t',n') = ?g ((t,n),(s,m))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4987
    by auto (rule_tac x="(a,b)" in bexI, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4988
  then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4989
    th: "?f (t',n') = ?g((t,n),(s,m)) "by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4990
    let ?N = "\<lambda> t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4991
  from tnU smU U have tnb: "numbound0 t" and np: "n > 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4992
    and snb: "numbound0 s" and mp:"m > 0"  by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4993
  let ?st= "Add (Mul m t) (Mul n s)"
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  4994
  from np mp have mnp: "real (2*n*m) > 0" 
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
  4995
      by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4996
    from tnb snb have stnb: "numbound0 ?st" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4997
  have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29265
diff changeset
  4998
   using mp np by (simp add: algebra_simps add_divide_distrib)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  4999
  from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5000
  from \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt'
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5001
  have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5002
  with \<upsilon>_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5003
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5004
  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5005
lemma ferrack01: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5006
  assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5007
  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> _")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5008
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5009
  let ?I = "\<lambda> x p. Ifm (x#bs) p"
26935
ee6bcb1b8953 avoid undeclared variables within proofs;
wenzelm
parents: 26932
diff changeset
  5010
  fix x
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5011
  let ?N = "\<lambda> t. Inum (x#bs) t"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5012
  let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5013
  let ?U = "\<Upsilon> ?q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5014
  let ?Up = "alluopairs ?U"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5015
  let ?g = "\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5016
  let ?S = "map ?g ?Up"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5017
  let ?SS = "map simp_num_pair ?S"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5018
  let ?Y = "remdups ?SS"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5019
  let ?f= "(\<lambda> (t,n). ?N t / real n)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5020
  let ?h = "\<lambda> ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5021
  let ?F = "\<lambda> p. \<exists> a \<in> set (\<Upsilon> p). \<exists> b \<in> set (\<Upsilon> p). ?I x (\<upsilon> p (?g(a,b)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5022
  let ?ep = "evaldjf (\<upsilon> ?q) ?Y"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5023
  from rlfm_l[OF qf] have lq: "isrlfm ?q" 
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30649
diff changeset
  5024
    by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5025
  from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \<le> (set ?U \<times> set ?U)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5026
  from \<Upsilon>_l[OF lq] have U_l: "\<forall> (t,n) \<in> set ?U. numbound0 t \<and> n > 0" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5027
  from U_l UpU 
50241
8b0fdeeefef7 eliminated some improper identifiers;
wenzelm
parents: 49962
diff changeset
  5028
  have "\<forall> ((t,n),(s,m)) \<in> set ?Up. numbound0 t \<and> n> 0 \<and> numbound0 s \<and> m > 0" by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5029
  hence Snb: "\<forall> (t,n) \<in> set ?S. numbound0 t \<and> n > 0 "
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56479
diff changeset
  5030
    by (auto)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5031
  have Y_l: "\<forall> (t,n) \<in> set ?Y. numbound0 t \<and> n > 0" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5032
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5033
    { fix t n assume tnY: "(t,n) \<in> set ?Y" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5034
      hence "(t,n) \<in> set ?SS" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5035
      hence "\<exists> (t',n') \<in> set ?S. simp_num_pair (t',n') = (t,n)"
33639
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33063
diff changeset
  5036
        by (auto simp add: split_def simp del: map_map)
603320b93668 New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
hoelzl
parents: 33063
diff changeset
  5037
           (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5038
      then obtain t' n' where tn'S: "(t',n') \<in> set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5039
      from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5040
      from simp_num_pair_l[OF tnb np tns]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5041
      have "numbound0 t \<and> n > 0" . }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5042
    thus ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5043
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5044
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5045
  have YU: "(?f ` set ?Y) = (?h ` (set ?U \<times> set ?U))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5046
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5047
     from simp_num_pair_ci[where bs="x#bs"] have 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5048
    "\<forall>x. (?f o simp_num_pair) x = ?f x" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5049
     hence th: "?f o simp_num_pair = ?f" using ext by blast
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5050
    have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_comp comp_assoc)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5051
    also have "\<dots> = (?f ` set ?S)" by (simp add: th)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5052
    also have "\<dots> = ((?f o ?g) ` set ?Up)" 
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5053
      by (simp only: set_map o_def image_comp)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5054
    also have "\<dots> = (?h ` (set ?U \<times> set ?U))"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5055
      using \<Upsilon>_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_comp] by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5056
    finally show ?thesis .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5057
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5058
  have "\<forall> (t,n) \<in> set ?Y. bound0 (\<upsilon> ?q (t,n))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5059
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5060
    { fix t n assume tnY: "(t,n) \<in> set ?Y"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5061
      with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5062
      from \<upsilon>_I[OF lq np tnb]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5063
    have "bound0 (\<upsilon> ?q (t,n))"  by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5064
    thus ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5065
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5066
  hence ep_nb: "bound0 ?ep"  using evaldjf_bound0[where xs="?Y" and f="\<upsilon> ?q"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5067
    by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5068
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5069
  from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5070
    by (simp only: split_def fst_conv snd_conv)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5071
  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]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5072
    by (simp only: split_def fst_conv snd_conv) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5073
  also have "\<dots> = (Ifm (x#bs) ?ep)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5074
    using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\<upsilon> ?q",symmetric]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5075
    by (simp only: split_def pair_collapse)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5076
  also have "\<dots> = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5077
  finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5078
  from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5079
  with lr show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5080
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5081
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5082
lemma cp_thm': 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5083
  assumes lp: "iszlfm p (real (i::int)#bs)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5084
  and up: "d_\<beta> p 1" and dd: "d_\<delta> p d" and dp: "d > 0"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5085
  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))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5086
  using cp_thm[OF lp up dd dp] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5087
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5088
definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" where
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5089
  "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;
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5090
             B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5091
             in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5092
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5093
lemma unit: assumes qf: "qfree p"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5094
  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)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5095
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5096
  fix q B d 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5097
  assume qBd: "unit p = (q,B,d)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5098
  let ?thes = "((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and>
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5099
    Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\<beta> q) \<and>
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5100
    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)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5101
  let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5102
  let ?p' = "zlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5103
  let ?l = "\<zeta> ?p'"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5104
  let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a_\<beta> ?p' ?l)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5105
  let ?d = "\<delta> ?q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5106
  let ?B = "set (\<beta> ?q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5107
  let ?B'= "remdups (map simpnum (\<beta> ?q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5108
  let ?A = "set (\<alpha> ?q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5109
  let ?A'= "remdups (map simpnum (\<alpha> ?q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5110
  from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5111
  have pp': "\<forall> i. ?I i ?p' = ?I i p" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5112
  from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5113
  have lp': "\<forall> (i::int). iszlfm ?p' (real i#bs)" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5114
  hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5115
  from lp' \<zeta>[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d_\<beta> ?p' ?l" by auto
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5116
  from a_\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp'
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5117
  have pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by (simp add: int_rdvd_iff) 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5118
  from lp'' lp a_\<beta>[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d_\<beta> ?q 1" 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5119
    by (auto simp add: isint_def)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5120
  from \<delta>[OF lq] have dp:"?d >0" and dd: "d_\<delta> ?q ?d" by blast+
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5121
  let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5122
  have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_comp) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5123
  also have "\<dots> = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5124
  finally have BB': "?N ` set ?B' = ?N ` ?B" .
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5125
  have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_comp) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5126
  also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5127
  finally have AA': "?N ` set ?A' = ?N ` ?A" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5128
  from \<beta>_numbound0[OF lq] have B_nb:"\<forall> b\<in> set ?B'. numbound0 b"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5129
    by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5130
  from \<alpha>_l[OF lq] have A_nb: "\<forall> b\<in> set ?A'. numbound0 b"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5131
    by simp
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5132
  { assume "length ?B' \<le> length ?A'"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5133
    hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5134
      using qBd by (auto simp add: Let_def unit_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5135
    with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<beta> q)" 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5136
      and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5137
    with pq_ex dp uq dd lq q d have ?thes by simp }
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5138
  moreover 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5139
  { assume "\<not> (length ?B' \<le> length ?A')"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5140
    hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5141
      using qBd by (auto simp add: Let_def unit_def)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5142
    with AA' mirror_\<alpha>_\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)" 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5143
      and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5144
    from mirror_ex[OF lq] pq_ex q 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5145
    have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5146
    from lq uq q mirror_d_\<beta> [where p="?q" and bs="bs" and a="real i"]
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5147
    have lq': "iszlfm q (real i#bs)" and uq: "d_\<beta> q 1" by auto
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5148
    from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq:"d_\<delta> q d " by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5149
    from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5150
  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5151
  ultimately show ?thes by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5152
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5153
    (* Cooper's Algorithm *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5154
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5155
definition cooper :: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5156
  "cooper p \<equiv> 
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5157
  (let (q,B,d) = unit p; js = [1..d];
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5158
       mq = simpfm (minusinf q);
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5159
       md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) js
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5160
   in if md = T then T else
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5161
    (let qd = evaldjf (\<lambda> t. simpfm (subst0 t q)) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5162
                               (remdups (map (\<lambda> (b,j). simpnum (Add b (C j))) 
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5163
                                            [(b,j). b\<leftarrow>B,j\<leftarrow>js]))
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5164
     in decr (disj md qd)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5165
lemma cooper: assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5166
  shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \<and> qfree (cooper p)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5167
  (is "(?lhs = ?rhs) \<and> _")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5168
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5169
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5170
  let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5171
  let ?q = "fst (unit p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5172
  let ?B = "fst (snd(unit p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5173
  let ?d = "snd (snd (unit p))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5174
  let ?js = "[1..?d]"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5175
  let ?mq = "minusinf ?q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5176
  let ?smq = "simpfm ?mq"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5177
  let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
26935
ee6bcb1b8953 avoid undeclared variables within proofs;
wenzelm
parents: 26932
diff changeset
  5178
  fix i
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5179
  let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5180
  let ?bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]"
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5181
  let ?sbjs = "map (\<lambda> (b,j). simpnum (Add b (C j))) ?bjs"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5182
  let ?qd = "evaldjf (\<lambda> t. simpfm (subst0 t ?q)) (remdups ?sbjs)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5183
  have qbf:"unit p = (?q,?B,?d)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5184
  from unit[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5185
    B:"?N ` set ?B = ?N ` set (\<beta> ?q)" and 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5186
    uq:"d_\<beta> ?q 1" and dd: "d_\<delta> ?q ?d" and dp: "?d > 0" and 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5187
    lq: "iszlfm ?q (real i#bs)" and 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5188
    Bn: "\<forall> b\<in> set ?B. numbound0 b" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5189
  from zlin_qfree[OF lq] have qfq: "qfree ?q" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5190
  from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5191
  have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5192
  hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5193
    by (auto simp only: subst0_bound0[OF qfmq])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5194
  hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53168
diff changeset
  5195
    by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5196
  from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5197
  from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))"
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5198
    by simp
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5199
  hence "\<forall> (b,j) \<in> set ?bjs. numbound0 (simpnum (Add b (C j)))"
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5200
    using simpnum_numbound0 by blast
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5201
  hence "\<forall> t \<in> set ?sbjs. numbound0 t" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5202
  hence "\<forall> t \<in> set (remdups ?sbjs). bound0 (subst0 t ?q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5203
    using subst0_bound0[OF qfq] by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5204
  hence th': "\<forall> t \<in> set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5205
    using simpfm_bound0 by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5206
  from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5207
  from mdb qdb 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5208
  have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5209
  from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5210
  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
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5211
  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))" by auto
24336
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5212
  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
fff40259f336 removed allpairs
nipkow
parents: 24249
diff changeset
  5213
  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)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5214
  also  have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> set ?sbjs. Ifm (?N t #bs) ?q))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5215
    by (auto simp add: split_def) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5216
  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))"
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5217
    by (simp only: simpfm subst0_I[OF qfq] Inum.simps subst0_I[OF qfmq] set_remdups)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5218
  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)
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5219
  finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5220
  hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5221
  {assume mdT: "?md = T"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5222
    hence cT:"cooper p = T" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5223
      by (simp only: cooper_def unit_def split_def Let_def if_True) simp
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5224
    from mdT mdqd have lhs:"?lhs" by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5225
    from mdT have "?rhs" by (simp add: cooper_def unit_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5226
    with lhs cT have ?thesis by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5227
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5228
  {assume mdT: "?md \<noteq> T" hence "cooper p = decr (disj ?md ?qd)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5229
      by (simp only: cooper_def unit_def split_def Let_def if_False) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5230
    with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5231
  ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5232
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5233
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5234
lemma DJcooper: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5235
  assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5236
  shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \<and> qfree (DJ cooper p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5237
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5238
  from cooper have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (cooper p)" by  blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5239
  from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5240
  have "Ifm bs (DJ cooper p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (cooper q))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5241
     by (simp add: DJ_def evaldjf_ex)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5242
  also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs)  q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5243
    using cooper disjuncts_qf[OF qf] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5244
  also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5245
  finally show ?thesis using thqf by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5246
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5247
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5248
    (* Redy and Loveland *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5249
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5250
lemma \<sigma>_\<rho>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5251
  shows "Ifm (a#bs) (\<sigma>_\<rho> p (t,c)) = Ifm (a#bs) (\<sigma>_\<rho> p (t',c))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5252
  using lp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5253
  by (induct p rule: iszlfm.induct, auto simp add: tt')
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5254
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5255
lemma \<sigma>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5256
  shows "Ifm (a#bs) (\<sigma> p c t) = Ifm (a#bs) (\<sigma> p c t')"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5257
  by (simp add: \<sigma>_def tt' \<sigma>_\<rho>_cong[OF lp tt'])
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5258
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5259
lemma \<rho>_cong: assumes lp: "iszlfm p (a#bs)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5260
  and RR: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R =  (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5261
  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))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5262
  (is "?lhs = ?rhs")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5263
proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5264
  let ?d = "\<delta> p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5265
  assume ?lhs then obtain e c j where ecR: "(e,c) \<in> R" and jD:"j \<in> {1 .. c*?d}" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5266
    and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5267
  from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5268
  hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" using RR by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5269
  hence "\<exists> (e',c') \<in> set (\<rho> p). Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5270
  then obtain e' c' where ecRo:"(e',c') \<in> set (\<rho> p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5271
    and cc':"c = c'" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5272
  from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5273
  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5274
  from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56544
diff changeset
  5275
  from ecRo jD px' show ?rhs apply (auto simp: cc')
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5276
    by (rule_tac x="(e', c')" in bexI,simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5277
  (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5278
next
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5279
  let ?d = "\<delta> p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5280
  assume ?rhs then obtain e c j where ecR: "(e,c) \<in> set (\<rho> p)" and jD:"j \<in> {1 .. c*?d}" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5281
    and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5282
  from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5283
  hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5284
  hence "\<exists> (e',c') \<in> R. Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5285
  then obtain e' c' where ecRo:"(e',c') \<in> R" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5286
    and cc':"c = c'" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5287
  from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5288
  from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56544
diff changeset
  5289
  from ecRo jD px' show ?lhs apply (auto simp: cc')
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5290
    by (rule_tac x="(e', c')" in bexI,simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5291
  (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5292
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5293
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5294
lemma rl_thm': 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5295
  assumes lp: "iszlfm p (real (i::int)#bs)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5296
  and R: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R =  (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5297
  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)))))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5298
  using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5299
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5300
definition chooset :: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5301
  "chooset p \<equiv> (let q = zlfm p ; d = \<delta> q;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5302
             B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ; 
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5303
             a = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<alpha>_\<rho> q))
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5304
             in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5305
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5306
lemma chooset: assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5307
  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)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5308
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5309
  fix q B d 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5310
  assume qBd: "chooset p = (q,B,d)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5311
  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)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5312
  let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5313
  let ?q = "zlfm p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5314
  let ?d = "\<delta> ?q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5315
  let ?B = "set (\<rho> ?q)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5316
  let ?f = "\<lambda> (t,k). (simpnum t,k)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5317
  let ?B'= "remdups (map ?f (\<rho> ?q))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5318
  let ?A = "set (\<alpha>_\<rho> ?q)"
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5319
  let ?A'= "remdups (map ?f (\<alpha>_\<rho> ?q))"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5320
  from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5321
  have pp': "\<forall> i. ?I i ?q = ?I i p" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5322
  hence pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5323
  from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5324
  have lq: "iszlfm ?q (real (i::int)#bs)" . 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5325
  from \<delta>[OF lq] have dp:"?d >0" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5326
  let ?N = "\<lambda> (t,c). (Inum (real (i::int)#bs) t,c)"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5327
  have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_comp)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5328
  also have "\<dots> = ?N ` ?B"
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5329
    by(simp add: split_def image_comp simpnum_ci[where bs="real i #bs"] image_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5330
  finally have BB': "?N ` set ?B' = ?N ` ?B" .
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5331
  have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_comp) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5332
  also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"]
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55584
diff changeset
  5333
    by(simp add: split_def image_comp simpnum_ci[where bs="real i #bs"] image_def) 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5334
  finally have AA': "?N ` set ?A' = ?N ` ?A" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5335
  from \<rho>_l[OF lq] have B_nb:"\<forall> (e,c)\<in> set ?B'. numbound0 e \<and> c > 0"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5336
    by (simp add: split_def)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5337
  from \<alpha>_\<rho>_l[OF lq] have A_nb: "\<forall> (e,c)\<in> set ?A'. numbound0 e \<and> c > 0"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5338
    by (simp add: split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5339
    {assume "length ?B' \<le> length ?A'"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5340
    hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5341
      using qBd by (auto simp add: Let_def chooset_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5342
    with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<rho> q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5343
      and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5344
  with pq_ex dp lq q d have ?thes by simp}
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5345
  moreover 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5346
  {assume "\<not> (length ?B' \<le> length ?A')"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5347
    hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5348
      using qBd by (auto simp add: Let_def chooset_def)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50241
diff changeset
  5349
    with AA' mirror_\<alpha>_\<rho>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<rho> q)" 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5350
      and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5351
    from mirror_ex[OF lq] pq_ex q 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5352
    have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5353
    from lq q mirror_l [where p="?q" and bs="bs" and a="real i"]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5354
    have lq': "iszlfm q (real i#bs)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5355
    from mirror_\<delta>[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5356
  }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5357
  ultimately show ?thes by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5358
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5359
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5360
definition stage :: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm" where
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5361
  "stage p d \<equiv> (\<lambda> (e,c). evaldjf (\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))) [1..c*d])"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5362
lemma stage:
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5363
  shows "Ifm bs (stage p d (e,c)) = (\<exists> j\<in>{1 .. c*d}. Ifm bs (\<sigma> p c (Add e (C j))))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5364
  by (unfold stage_def split_def ,simp only: evaldjf_ex simpfm) simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5365
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5366
lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5367
  shows "bound0 (stage p d (e,c))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5368
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5369
  let ?f = "\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5370
  have th: "\<forall> j\<in> set [1..c*d]. bound0 (?f j)"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5371
  proof
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5372
    fix j
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5373
    from nb have nb':"numbound0 (Add e (C j))" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5374
    from simpfm_bound0[OF \<sigma>_nb[OF lp nb', where k="c"]]
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5375
    show "bound0 (simpfm (\<sigma> p c (Add e (C j))))" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5376
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5377
  from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5378
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5379
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
  5380
definition redlove :: "fm \<Rightarrow> fm" where
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5381
  "redlove p \<equiv> 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5382
  (let (q,B,d) = chooset p;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5383
       mq = simpfm (minusinf q);
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5384
       md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) [1..d]
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5385
   in if md = T then T else
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5386
    (let qd = evaldjf (stage q d) B
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5387
     in decr (disj md qd)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5388
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5389
lemma redlove: assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5390
  shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \<and> qfree (redlove p)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5391
  (is "(?lhs = ?rhs) \<and> _")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5392
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5393
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5394
  let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5395
  let ?q = "fst (chooset p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5396
  let ?B = "fst (snd(chooset p))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5397
  let ?d = "snd (snd (chooset p))"
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5398
  let ?js = "[1..?d]"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5399
  let ?mq = "minusinf ?q"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5400
  let ?smq = "simpfm ?mq"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5401
  let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
26935
ee6bcb1b8953 avoid undeclared variables within proofs;
wenzelm
parents: 26932
diff changeset
  5402
  fix i
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5403
  let ?N = "\<lambda> (t,k). (Inum (real (i::int)#bs) t,k)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5404
  let ?qd = "evaldjf (stage ?q ?d) ?B"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5405
  have qbf:"chooset p = (?q,?B,?d)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5406
  from chooset[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5407
    B:"?N ` set ?B = ?N ` set (\<rho> ?q)" and dd: "\<delta> ?q = ?d" and dp: "?d > 0" and 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5408
    lq: "iszlfm ?q (real i#bs)" and 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5409
    Bn: "\<forall> (e,c)\<in> set ?B. numbound0 e \<and> c > 0" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5410
  from zlin_qfree[OF lq] have qfq: "qfree ?q" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5411
  from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5412
  have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5413
  hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5414
    by (auto simp only: subst0_bound0[OF qfmq])
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5415
  hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5416
    by auto
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5417
  from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5418
  from Bn stage_nb[OF lq] have th:"\<forall> x \<in> set ?B. bound0 (stage ?q ?d x)" by auto
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5419
  from evaldjf_bound0[OF th]  have qdb: "bound0 ?qd" .
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5420
  from mdb qdb 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5421
  have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5422
  from trans [OF pq_ex rl_thm'[OF lq B]] dd
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5423
  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
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5424
  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) )))" 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5425
    by (simp add: stage split_def)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5426
  also have "\<dots> = ((\<exists> j\<in> {1 .. ?d}. ?I i (subst0 (C j) ?smq))  \<or> ?I i ?qd)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5427
    by (simp add: evaldjf_ex subst0_I[OF qfmq])
41836
c9d788ff7940 eliminated clones of List.upto
krauss
parents: 41807
diff changeset
  5428
  finally have mdqd:"?lhs = (?I i ?md \<or> ?I i ?qd)" by (simp only: evaldjf_ex set_upto simpfm) 
51369
960b0ca9ae5d tuned proofs -- more structure, less warnings;
wenzelm
parents: 51272
diff changeset
  5429
  also have "\<dots> = (?I i (disj ?md ?qd))" by simp
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5430
  also have "\<dots> = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5431
  finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5432
  {assume mdT: "?md = T"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5433
    hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5434
    from mdT have lhs:"?lhs" using mdqd by simp 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5435
    from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5436
    with lhs cT have ?thesis by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5437
  moreover
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5438
  {assume mdT: "?md \<noteq> T" hence "redlove p = decr (disj ?md ?qd)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5439
      by (simp add: redlove_def chooset_def split_def Let_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5440
    with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5441
  ultimately show ?thesis by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5442
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5443
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5444
lemma DJredlove: 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5445
  assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5446
  shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \<and> qfree (DJ redlove p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5447
proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5448
  from redlove have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (redlove p)" by  blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5449
  from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5450
  have "Ifm bs (DJ redlove p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (redlove q))" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5451
     by (simp add: DJ_def evaldjf_ex)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5452
  also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs)  q)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5453
    using redlove disjuncts_qf[OF qf] by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5454
  also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5455
  finally show ?thesis using thqf by blast
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5456
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5457
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5458
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5459
lemma exsplit_qf: assumes qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5460
  shows "qfree (exsplit p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5461
using qf by (induct p rule: exsplit.induct, auto)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5462
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5463
definition mircfr :: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5464
  "mircfr = DJ cooper o ferrack01 o simpfm o exsplit"
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5465
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5466
definition mirlfr :: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5467
  "mirlfr = DJ redlove o ferrack01 o simpfm o exsplit"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5468
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5469
lemma mircfr: "\<forall> bs p. qfree p \<longrightarrow> qfree (mircfr p) \<and> Ifm bs (mircfr p) = Ifm bs (E p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5470
proof(clarsimp simp del: Ifm.simps)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5471
  fix bs p
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5472
  assume qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5473
  show "qfree (mircfr p)\<and>(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5474
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5475
    let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5476
    have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5477
      using splitex[OF qf] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5478
    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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5479
    with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5480
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5481
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5482
  
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5483
lemma mirlfr: "\<forall> bs p. qfree p \<longrightarrow> qfree(mirlfr p) \<and> Ifm bs (mirlfr p) = Ifm bs (E p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5484
proof(clarsimp simp del: Ifm.simps)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5485
  fix bs p
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5486
  assume qf: "qfree p"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5487
  show "qfree (mirlfr p)\<and>(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5488
  proof-
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5489
    let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5490
    have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" 
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5491
      using splitex[OF qf] by simp
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5492
    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+
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5493
    with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5494
  qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5495
qed
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5496
  
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5497
definition mircfrqe:: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5498
  "mircfrqe p = qelim (prep p) mircfr"
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5499
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5500
definition mirlfrqe:: "fm \<Rightarrow> fm" where
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5501
  "mirlfrqe p = qelim (prep p) mirlfr"
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5502
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5503
theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \<and> qfree (mircfrqe p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5504
  using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5505
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5506
theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \<and> qfree (mirlfrqe p)"
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5507
  using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5508
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5509
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5510
  "problem1 = A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0))))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5511
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5512
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5513
  "problem2 = A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5514
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5515
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5516
  "problem3 = A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0))))))"
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5517
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5518
definition
36870
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5519
  "problem4 = E (And (Ge (Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg (Bound 0))))))"
b897bd9ca71b tuned test problems
haftmann
parents: 36778
diff changeset
  5520
51272
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5521
ML_val {* @{code mircfrqe} @{code problem1} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5522
ML_val {* @{code mirlfrqe} @{code problem1} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5523
ML_val {* @{code mircfrqe} @{code problem2} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5524
ML_val {* @{code mirlfrqe} @{code problem2} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5525
ML_val {* @{code mircfrqe} @{code problem3} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5526
ML_val {* @{code mirlfrqe} @{code problem3} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5527
ML_val {* @{code mircfrqe} @{code problem4} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5528
ML_val {* @{code mirlfrqe} @{code problem4} *}
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 51143
diff changeset
  5529
24249
1f60b45c5f97 renamed keyword "to" to "module_name"
haftmann
parents: 23997
diff changeset
  5530
36531
19f6e3b0d9b6 code_reflect: specify module name directly after keyword
haftmann
parents: 36526
diff changeset
  5531
(*code_reflect Mir
36526
353041483b9b use code_reflect
haftmann
parents: 35416
diff changeset
  5532
  functions mircfrqe mirlfrqe
353041483b9b use code_reflect
haftmann
parents: 35416
diff changeset
  5533
  file "mir.ML"*)
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5534
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28264
diff changeset
  5535
oracle mirfr_oracle = {* fn (proofs, ct) =>
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5536
let
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5537
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5538
val mk_C = @{code C} o @{code int_of_integer};
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5539
val mk_Dvd = @{code Dvd} o apfst @{code int_of_integer};
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5540
val mk_Bound = @{code Bound} o @{code nat_of_integer};
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5541
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5542
fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5543
     of NONE => error "Variable not found in the list!"
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5544
      | SOME n => mk_Bound n)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5545
  | num_of_term vs @{term "real (0::int)"} = mk_C 0
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5546
  | num_of_term vs @{term "real (1::int)"} = mk_C 1
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5547
  | num_of_term vs @{term "0::real"} = mk_C 0
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5548
  | num_of_term vs @{term "1::real"} = mk_C 1
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
  5549
  | num_of_term vs @{term "- 1::real"} = mk_C (~ 1)
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5550
  | num_of_term vs (Bound i) = mk_Bound i
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5551
  | num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t')
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5552
  | num_of_term vs (@{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5553
      @{code Add} (num_of_term vs t1, num_of_term vs t2)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5554
  | num_of_term vs (@{term "op - :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5555
      @{code Sub} (num_of_term vs t1, num_of_term vs t2)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5556
  | num_of_term vs (@{term "op * :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5557
      (case (num_of_term vs t1)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5558
       of @{code C} i => @{code Mul} (i, num_of_term vs t2)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5559
        | _ => error "num_of_term: unsupported Multiplication")
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  5560
  | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t')) =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5561
      mk_C (HOLogic.dest_num t')
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
  5562
  | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> int"} $ t')) =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5563
      mk_C (~ (HOLogic.dest_num t'))
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5564
  | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ t')) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5565
      @{code Floor} (num_of_term vs t')
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5566
  | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ t')) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5567
      @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t')))
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  5568
  | num_of_term vs (@{term "numeral :: _ \<Rightarrow> real"} $ t') =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5569
      mk_C (HOLogic.dest_num t')
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
  5570
  | num_of_term vs (@{term "- numeral :: _ \<Rightarrow> real"} $ t') =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5571
      mk_C (~ (HOLogic.dest_num t'))
28264
e1dae766c108 local @{context};
wenzelm
parents: 27567
diff changeset
  5572
  | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t);
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5573
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5574
fun fm_of_term vs @{term True} = @{code T}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5575
  | fm_of_term vs @{term False} = @{code F}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5576
  | fm_of_term vs (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5577
      @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5578
  | fm_of_term vs (@{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5579
      @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5580
  | fm_of_term vs (@{term "op = :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5581
      @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) 
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
  5582
  | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5583
      mk_Dvd (HOLogic.dest_num t1, num_of_term vs t2)
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
  5584
  | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5585
      mk_Dvd (~ (HOLogic.dest_num t1), num_of_term vs t2)
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5586
  | fm_of_term vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5587
      @{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
  5588
  | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5589
      @{code And} (fm_of_term vs t1, fm_of_term vs t2)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
  5590
  | fm_of_term vs (@{term HOL.disj} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5591
      @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38558
diff changeset
  5592
  | fm_of_term vs (@{term HOL.implies} $ t1 $ t2) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5593
      @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5594
  | fm_of_term vs (@{term "Not"} $ t') =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5595
      @{code NOT} (fm_of_term vs t')
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
  5596
  | fm_of_term vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5597
      @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
  5598
  | fm_of_term vs (Const (@{const_name All}, _) $ Abs (xn, xT, p)) =
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5599
      @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
28264
e1dae766c108 local @{context};
wenzelm
parents: 27567
diff changeset
  5600
  | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5601
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5602
fun term_of_num vs (@{code C} i) = @{term "real :: int \<Rightarrow> real"} $
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5603
      HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i)
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5604
  | term_of_num vs (@{code Bound} n) =
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5605
      let
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5606
        val m = @{code integer_of_nat} n;
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5607
      in fst (the (find_first (fn (_, q) => m = q) vs)) end
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5608
  | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5609
      @{term "real :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ term_of_num vs t')
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5610
  | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \<Rightarrow> real"} $ term_of_num vs t'
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5611
  | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5612
      term_of_num vs t1 $ term_of_num vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5613
  | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: real \<Rightarrow> real \<Rightarrow> real"} $
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5614
      term_of_num vs t1 $ term_of_num vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5615
  | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \<Rightarrow> real \<Rightarrow> real"} $
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5616
      term_of_num vs (@{code C} i) $ term_of_num vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5617
  | term_of_num vs (@{code Floor} t) = @{term "real :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ term_of_num vs t)
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5618
  | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t))
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5619
  | term_of_num vs (@{code CF} (c, t, s)) = term_of_num vs (@{code Add} (@{code Mul} (c, @{code Floor} t), s));
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5620
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 44890
diff changeset
  5621
fun term_of_fm vs @{code T} = @{term True} 
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 44890
diff changeset
  5622
  | term_of_fm vs @{code F} = @{term False}
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5623
  | term_of_fm vs (@{code Lt} t) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5624
      @{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5625
  | term_of_fm vs (@{code Le} t) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5626
      @{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5627
  | term_of_fm vs (@{code Gt} t) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5628
      @{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ @{term "0::real"} $ term_of_num vs t
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5629
  | term_of_fm vs (@{code Ge} t) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5630
      @{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ @{term "0::real"} $ term_of_num vs t
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5631
  | term_of_fm vs (@{code Eq} t) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5632
      @{term "op = :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5633
  | term_of_fm vs (@{code NEq} t) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5634
      term_of_fm vs (@{code NOT} (@{code Eq} t))
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5635
  | term_of_fm vs (@{code Dvd} (i, t)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5636
      @{term "op rdvd"} $ term_of_num vs (@{code C} i) $ term_of_num vs t
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5637
  | term_of_fm vs (@{code NDvd} (i, t)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5638
      term_of_fm vs (@{code NOT} (@{code Dvd} (i, t)))
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5639
  | term_of_fm vs (@{code NOT} t') =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5640
      HOLogic.Not $ term_of_fm vs t'
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5641
  | term_of_fm vs (@{code And} (t1, t2)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5642
      HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5643
  | term_of_fm vs (@{code Or} (t1, t2)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5644
      HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5645
  | term_of_fm vs (@{code Imp}  (t1, t2)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5646
      HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5647
  | term_of_fm vs (@{code Iff} (t1, t2)) =
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5648
      @{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ term_of_fm vs t1 $ term_of_fm vs t2;
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5649
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28264
diff changeset
  5650
in
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5651
  let 
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28264
diff changeset
  5652
    val thy = Thm.theory_of_cterm ct;
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28264
diff changeset
  5653
    val t = Thm.term_of ct;
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 44013
diff changeset
  5654
    val fs = Misc_Legacy.term_frees t;
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32960
diff changeset
  5655
    val vs = map_index swap fs;
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5656
    val qe = if proofs then @{code mirlfrqe} else @{code mircfrqe};
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5657
    val t' = (term_of_fm vs o qe o fm_of_term vs) t;
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28264
diff changeset
  5658
  in (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end
27456
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5659
end;
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5660
*}
52c7c42e7e27 code antiquotation roaring ahead
haftmann
parents: 27368
diff changeset
  5661
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47432
diff changeset
  5662
ML_file "mir_tac.ML"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  5663
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  5664
method_setup mir = {*
53168
d998de7f0efc tuned signature;
wenzelm
parents: 51369
diff changeset
  5665
  Scan.lift (Args.mode "no_quantify") >>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  5666
    (fn q => fn ctxt => SIMPLE_METHOD' (Mir_Tac.mir_tac ctxt (not q)))
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  5667
*} "decision procedure for MIR arithmetic"
e1576d13e933 more standard method setup;
wenzelm
parents: 47142
diff changeset
  5668
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5669
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5670
lemma "ALL (x::real). (\<lfloor>x\<rfloor> = \<lceil>x\<rceil> = (x = real \<lfloor>x\<rfloor>))"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5671
  by mir
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5672
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5673
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))"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5674
  by mir
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5675
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5676
lemma "ALL (x::real). 2*\<lfloor>x\<rfloor> \<le> \<lfloor>2*x\<rfloor> \<and> \<lfloor>2*x\<rfloor> \<le> 2*\<lfloor>x+1\<rfloor>"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5677
  by mir 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5678
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5679
lemma "ALL (x::real). \<exists>y \<le> x. (\<lfloor>x\<rfloor> = \<lceil>y\<rceil>)"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5680
  by mir
23858
5500610fe1e5 adapted to new code generator framework
haftmann
parents: 23477
diff changeset
  5681
30103
32effb2a8168 add type annotation
huffman
parents: 30097
diff changeset
  5682
lemma "ALL (x::real) (y::real). \<lfloor>x\<rfloor> = \<lfloor>y\<rfloor> \<longrightarrow> 0 \<le> abs (y - x) \<and> abs (y - x) \<le> 1"
41891
d37babdf5cae tuned proofs -- eliminated prems;
wenzelm
parents: 41882
diff changeset
  5683
  by mir
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5684
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
  5685
end
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 50252
diff changeset
  5686